coq-8.11.0/0000755000175000017500000000000013612664533012266 5ustar treinentreinencoq-8.11.0/plugins/0000755000175000017500000000000013612664533013747 5ustar treinentreinencoq-8.11.0/plugins/funind/0000755000175000017500000000000013612664533015232 5ustar treinentreinencoq-8.11.0/plugins/funind/recdef_plugin.mlpack0000644000175000017500000000022613612664533021231 0ustar treinentreinenIndfun_common Glob_termops Recdef Glob_term_to_relation Functional_principles_proofs Functional_principles_types Invfun Indfun Gen_principle G_indfun coq-8.11.0/plugins/funind/functional_principles_proofs.ml0000644000175000017500000017756413612664533023572 0ustar treinentreinenopen Printer open CErrors open Util open Constr open Context open EConstr open Vars open Namegen open Names open Pp open Tacmach open Termops open Tacticals open Tactics open Indfun_common open Libnames open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration let list_chop ?(msg="") n l = try List.chop n l with Failure (msg') -> failwith (msg ^ msg') let pop t = Vars.lift (-1) t let make_refl_eq constructor type_of_t t = (* let refl_equal_term = Lazy.force refl_equal in *) mkApp(constructor,[|type_of_t;t|]) type pte_info = { proving_tac : (Id.t list -> Tacmach.tactic); is_valid : constr -> bool } type ptes_info = pte_info Id.Map.t type 'a dynamic_info = { nb_rec_hyps : int; rec_hyps : Id.t list ; eq_hyps : Id.t list; info : 'a } type body_info = constr dynamic_info let observe_tac s = observe_tac (fun _ _ -> Pp.str s) let finish_proof dynamic_infos g = observe_tac "finish" (Proofview.V82.of_tactic assumption) g let refine c = Refiner.refiner ~check:true EConstr.Unsafe.(to_constr c) let thin l = Proofview.V82.of_tactic (Tactics.clear l) let eq_constr sigma u v = EConstr.eq_constr_nounivs sigma u v let is_trivial_eq sigma t = let res = try begin match EConstr.kind sigma t with | App(f,[|_;t1;t2|]) when eq_constr sigma f (Lazy.force eq) -> eq_constr sigma t1 t2 | App(f,[|t1;a1;t2;a2|]) when eq_constr sigma f (jmeq ()) -> eq_constr sigma t1 t2 && eq_constr sigma a1 a2 | _ -> false end with e when CErrors.noncritical e -> false in (* observe (str "is_trivial_eq " ++ Printer.pr_lconstr t ++ (if res then str " true" else str " false")); *) res let rec incompatible_constructor_terms sigma t1 t2 = let c1,arg1 = decompose_app sigma t1 and c2,arg2 = decompose_app sigma t2 in (not (eq_constr sigma t1 t2)) && isConstruct sigma c1 && isConstruct sigma c2 && ( not (eq_constr sigma c1 c2) || List.exists2 (incompatible_constructor_terms sigma) arg1 arg2 ) let is_incompatible_eq env sigma t = let res = try match EConstr.kind sigma t with | App(f,[|_;t1;t2|]) when eq_constr sigma f (Lazy.force eq) -> incompatible_constructor_terms sigma t1 t2 | App(f,[|u1;t1;u2;t2|]) when eq_constr sigma f (jmeq ()) -> (eq_constr sigma u1 u2 && incompatible_constructor_terms sigma t1 t2) | _ -> false with e when CErrors.noncritical e -> false in if res then observe (str "is_incompatible_eq " ++ pr_leconstr_env env sigma t); res let change_hyp_with_using msg hyp_id t tac : tactic = fun g -> let prov_id = pf_get_new_id hyp_id g in tclTHENS ((* observe_tac msg *) Proofview.V82.of_tactic (assert_by (Name prov_id) t (Proofview.V82.tactic (tclCOMPLETE tac)))) [tclTHENLIST [ (* observe_tac "change_hyp_with_using thin" *) (thin [hyp_id]); (* observe_tac "change_hyp_with_using rename " *) (Proofview.V82.of_tactic (rename_hyp [prov_id,hyp_id])) ]] g exception TOREMOVE let prove_trivial_eq h_id context (constructor,type_of_term,term) = let nb_intros = List.length context in tclTHENLIST [ tclDO nb_intros (Proofview.V82.of_tactic intro); (* introducing context *) (fun g -> let context_hyps = fst (list_chop ~msg:"prove_trivial_eq : " nb_intros (pf_ids_of_hyps g)) in let context_hyps' = (mkApp(constructor,[|type_of_term;term|])):: (List.map mkVar context_hyps) in let to_refine = applist(mkVar h_id,List.rev context_hyps') in refine to_refine g ) ] let find_rectype env sigma c = let (t, l) = decompose_app sigma (Reductionops.whd_betaiotazeta sigma c) in match EConstr.kind sigma t with | Ind ind -> (t, l) | Construct _ -> (t,l) | _ -> raise Not_found let isAppConstruct ?(env=Global.env ()) sigma t = try let t',l = find_rectype env sigma t in observe (str "isAppConstruct : " ++ Printer.pr_leconstr_env env sigma t ++ str " -> " ++ Printer.pr_leconstr_env env sigma (applist (t',l))); true with Not_found -> false exception NoChange let change_eq env sigma hyp_id (context:rel_context) x t end_of_type = let nochange ?t' msg = begin observe (str ("Not treating ( "^msg^" )") ++ pr_leconstr_env env sigma t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_leconstr_env env sigma t ); raise NoChange; end in let eq_constr c1 c2 = try ignore(Evarconv.unify_delay env sigma c1 c2); true with Evarconv.UnableToUnify _ -> false in if not (noccurn sigma 1 end_of_type) then nochange "dependent"; (* if end_of_type depends on this term we don't touch it *) if not (isApp sigma t) then nochange "not an equality"; let f_eq,args = destApp sigma t in let constructor,t1,t2,t1_typ = try if (eq_constr f_eq (Lazy.force eq)) then let t1 = (args.(1),args.(0)) and t2 = (args.(2),args.(0)) and t1_typ = args.(0) in (Lazy.force refl_equal,t1,t2,t1_typ) else if (eq_constr f_eq (jmeq ())) then (jmeq_refl (),(args.(1),args.(0)),(args.(3),args.(2)),args.(0)) else nochange "not an equality" with e when CErrors.noncritical e -> nochange "not an equality" in if not ((closed0 sigma (fst t1)) && (closed0 sigma (snd t1)))then nochange "not a closed lhs"; let rec compute_substitution sub t1 t2 = (* observe (str "compute_substitution : " ++ pr_lconstr t1 ++ str " === " ++ pr_lconstr t2); *) if isRel sigma t2 then let t2 = destRel sigma t2 in begin try let t1' = Int.Map.find t2 sub in if not (eq_constr t1 t1') then nochange "twice bound variable"; sub with Not_found -> assert (closed0 sigma t1); Int.Map.add t2 t1 sub end else if isAppConstruct sigma t1 && isAppConstruct sigma t2 then begin let c1,args1 = find_rectype env sigma t1 and c2,args2 = find_rectype env sigma t2 in if not (eq_constr c1 c2) then nochange "cannot solve (diff)"; List.fold_left2 compute_substitution sub args1 args2 end else if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reductionops.whd_all env sigma t1) t2) "cannot solve (diff)" in let sub = compute_substitution Int.Map.empty (snd t1) (snd t2) in let sub = compute_substitution sub (fst t1) (fst t2) in let end_of_type_with_pop = pop end_of_type in (*the equation will be removed *) let new_end_of_type = (* Ugly hack to prevent Map.fold order change between ocaml-3.08.3 and ocaml-3.08.4 Can be safely replaced by the next comment for Ocaml >= 3.08.4 *) let sub = Int.Map.bindings sub in List.fold_left (fun end_of_type (i,t) -> liftn 1 i (substnl [t] (i-1) end_of_type)) end_of_type_with_pop sub in let old_context_length = List.length context + 1 in let witness_fun = mkLetIn(make_annot Anonymous Sorts.Relevant,make_refl_eq constructor t1_typ (fst t1),t, mkApp(mkVar hyp_id,Array.init old_context_length (fun i -> mkRel (old_context_length - i))) ) in let new_type_of_hyp,ctxt_size,witness_fun = List.fold_left_i (fun i (end_of_type,ctxt_size,witness_fun) decl -> try let witness = Int.Map.find i sub in if is_local_def decl then anomaly (Pp.str "can not redefine a rel!"); (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_annot decl, witness, RelDecl.get_type decl, witness_fun)) with Not_found -> (mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) ) 1 (new_end_of_type,0,witness_fun) context in let new_type_of_hyp = Reductionops.nf_betaiota env sigma new_type_of_hyp in let new_ctxt,new_end_of_type = decompose_prod_n_assum sigma ctxt_size new_type_of_hyp in let prove_new_hyp : tactic = tclTHEN (tclDO ctxt_size (Proofview.V82.of_tactic intro)) (fun g -> let all_ids = pf_ids_of_hyps g in let new_ids,_ = list_chop ctxt_size all_ids in let to_refine = applist(witness_fun,List.rev_map mkVar new_ids) in let evm, _ = pf_apply Typing.type_of g to_refine in tclTHEN (Refiner.tclEVARS evm) (refine to_refine) g ) in let simpl_eq_tac = change_hyp_with_using "prove_pattern_simplification" hyp_id new_type_of_hyp prove_new_hyp in (* observe (str "In " ++ Ppconstr.pr_id hyp_id ++ *) (* str "removing an equation " ++ fnl ()++ *) (* str "old_typ_of_hyp :=" ++ *) (* Printer.pr_lconstr_env *) (* env *) (* (it_mkProd_or_LetIn ~init:end_of_type ((x,None,t)::context)) *) (* ++ fnl () ++ *) (* str "new_typ_of_hyp := "++ *) (* Printer.pr_lconstr_env env new_type_of_hyp ++ fnl () *) (* ++ str "old context := " ++ pr_rel_context env context ++ fnl () *) (* ++ str "new context := " ++ pr_rel_context env new_ctxt ++ fnl () *) (* ++ str "old type := " ++ pr_lconstr end_of_type ++ fnl () *) (* ++ str "new type := " ++ pr_lconstr new_end_of_type ++ fnl () *) (* ); *) new_ctxt,new_end_of_type,simpl_eq_tac let is_property sigma (ptes_info:ptes_info) t_x full_type_of_hyp = if isApp sigma t_x then let pte,args = destApp sigma t_x in if isVar sigma pte && Array.for_all (closed0 sigma) args then try let info = Id.Map.find (destVar sigma pte) ptes_info in info.is_valid full_type_of_hyp with Not_found -> false else false else false let isLetIn sigma t = match EConstr.kind sigma t with | LetIn _ -> true | _ -> false let h_reduce_with_zeta cl = Proofview.V82.of_tactic (reduce (Genredexpr.Cbv {Redops.all_flags with Genredexpr.rDelta = false; }) cl) let rewrite_until_var arg_num eq_ids : tactic = (* tests if the declares recursive argument is neither a Constructor nor an applied Constructor since such a form for the recursive argument will break the Guard when trying to save the Lemma. *) let test_var g = let sigma = project g in let _,args = destApp sigma (pf_concl g) in not ((isConstruct sigma args.(arg_num)) || isAppConstruct sigma args.(arg_num)) in let rec do_rewrite eq_ids g = if test_var g then tclIDTAC g else match eq_ids with | [] -> anomaly (Pp.str "Cannot find a way to prove recursive property."); | eq_id::eq_ids -> tclTHEN (tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar eq_id)))) (do_rewrite eq_ids) g in do_rewrite eq_ids let rec_pte_id = Id.of_string "Hrec" let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = let coq_False = EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.False.type") in let coq_True = EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.True.type") in let coq_I = EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.True.I") in let rec scan_type context type_of_hyp : tactic = if isLetIn sigma type_of_hyp then let real_type_of_hyp = it_mkProd_or_LetIn type_of_hyp context in let reduced_type_of_hyp = Reductionops.nf_betaiotazeta env sigma real_type_of_hyp in (* length of context didn't change ? *) let new_context,new_typ_of_hyp = decompose_prod_n_assum sigma (List.length context) reduced_type_of_hyp in tclTHENLIST [ h_reduce_with_zeta (Locusops.onHyp hyp_id); scan_type new_context new_typ_of_hyp ] else if isProd sigma type_of_hyp then begin let (x,t_x,t') = destProd sigma type_of_hyp in let actual_real_type_of_hyp = it_mkProd_or_LetIn t' context in if is_property sigma ptes_infos t_x actual_real_type_of_hyp then begin let pte,pte_args = (destApp sigma t_x) in let (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar sigma pte) ptes_infos).proving_tac in let popped_t' = pop t' in let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in let prove_new_type_of_hyp = let context_length = List.length context in tclTHENLIST [ tclDO context_length (Proofview.V82.of_tactic intro); (fun g -> let context_hyps_ids = fst (list_chop ~msg:"rec hyp : context_hyps" context_length (pf_ids_of_hyps g)) in let rec_pte_id = pf_get_new_id rec_pte_id g in let to_refine = applist(mkVar hyp_id, List.rev_map mkVar (rec_pte_id::context_hyps_ids) ) in (* observe_tac "rec hyp " *) (tclTHENS (Proofview.V82.of_tactic (assert_before (Name rec_pte_id) t_x)) [ (* observe_tac "prove rec hyp" *) (prove_rec_hyp eq_hyps); (* observe_tac "prove rec hyp" *) (refine to_refine) ]) g ) ] in tclTHENLIST [ (* observe_tac "hyp rec" *) (change_hyp_with_using "rec_hyp_tac" hyp_id real_type_of_hyp prove_new_type_of_hyp); scan_type context popped_t' ] end else if eq_constr sigma t_x coq_False then begin (* observe (str "Removing : "++ Ppconstr.pr_id hyp_id++ *) (* str " since it has False in its preconds " *) (* ); *) raise TOREMOVE; (* False -> .. useless *) end else if is_incompatible_eq env sigma t_x then raise TOREMOVE (* t_x := C1 ... = C2 ... *) else if eq_constr sigma t_x coq_True (* Trivial => we remove this precons *) then (* observe (str "In "++Ppconstr.pr_id hyp_id++ *) (* str " removing useless precond True" *) (* ); *) let popped_t' = pop t' in let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in let prove_trivial = let nb_intro = List.length context in tclTHENLIST [ tclDO nb_intro (Proofview.V82.of_tactic intro); (fun g -> let context_hyps = fst (list_chop ~msg:"removing True : context_hyps "nb_intro (pf_ids_of_hyps g)) in let to_refine = applist (mkVar hyp_id, List.rev (coq_I::List.map mkVar context_hyps) ) in refine to_refine g ) ] in tclTHENLIST[ change_hyp_with_using "prove_trivial" hyp_id real_type_of_hyp ((* observe_tac "prove_trivial" *) prove_trivial); scan_type context popped_t' ] else if is_trivial_eq sigma t_x then (* t_x := t = t => we remove this precond *) let popped_t' = pop t' in let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in let hd,args = destApp sigma t_x in let get_args hd args = if eq_constr sigma hd (Lazy.force eq) then (Lazy.force refl_equal,args.(0),args.(1)) else (jmeq_refl (),args.(0),args.(1)) in tclTHENLIST [ change_hyp_with_using "prove_trivial_eq" hyp_id real_type_of_hyp ((* observe_tac "prove_trivial_eq" *) (prove_trivial_eq hyp_id context (get_args hd args))); scan_type context popped_t' ] else begin try let new_context,new_t',tac = change_eq env sigma hyp_id context x t_x t' in tclTHEN tac (scan_type new_context new_t') with NoChange -> (* Last thing todo : push the rel in the context and continue *) scan_type (LocalAssum (x,t_x) :: context) t' end end else tclIDTAC in try scan_type [] (Typing.unsafe_type_of env sigma (mkVar hyp_id)), [hyp_id] with TOREMOVE -> thin [hyp_id],[] let clean_goal_with_heq ptes_infos continue_tac (dyn_infos:body_info) = fun g -> let env = pf_env g and sigma = project g in let tac,new_hyps = List.fold_left ( fun (hyps_tac,new_hyps) hyp_id -> let hyp_tac,new_hyp = clean_hyp_with_heq ptes_infos dyn_infos.eq_hyps hyp_id env sigma in (tclTHEN hyp_tac hyps_tac),new_hyp@new_hyps ) (tclIDTAC,[]) dyn_infos.rec_hyps in let new_infos = { dyn_infos with rec_hyps = new_hyps; nb_rec_hyps = List.length new_hyps } in tclTHENLIST [ tac ; (* observe_tac "clean_hyp_with_heq continue" *) (continue_tac new_infos) ] g let heq_id = Id.of_string "Heq" let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = fun g -> let nb_first_intro = nb_prod - 1 - dyn_infos.nb_rec_hyps in tclTHENLIST [ (* We first introduce the variables *) tclDO nb_first_intro (Proofview.V82.of_tactic (intro_avoiding (Id.Set.of_list dyn_infos.rec_hyps))); (* Then the equation itself *) Proofview.V82.of_tactic (intro_using heq_id); onLastHypId (fun heq_id -> tclTHENLIST [ (* Then the new hypothesis *) tclMAP (fun id -> Proofview.V82.of_tactic (introduction id)) dyn_infos.rec_hyps; observe_tac "after_introduction" (fun g' -> (* We get infos on the equations introduced*) let new_term_value_eq = pf_unsafe_type_of g' (mkVar heq_id) in (* compute the new value of the body *) let new_term_value = match EConstr.kind (project g') new_term_value_eq with | App(f,[| _;_;args2 |]) -> args2 | _ -> observe (str "cannot compute new term value : " ++ pr_gls g' ++ fnl () ++ str "last hyp is" ++ pr_leconstr_env (pf_env g') (project g') new_term_value_eq ); anomaly (Pp.str "cannot compute new term value.") in let fun_body = mkLambda(make_annot Anonymous Sorts.Relevant, pf_unsafe_type_of g' term, Termops.replace_term (project g') term (mkRel 1) dyn_infos.info ) in let new_body = pf_nf_betaiota g' (mkApp(fun_body,[| new_term_value |])) in let new_infos = {dyn_infos with info = new_body; eq_hyps = heq_id::dyn_infos.eq_hyps } in clean_goal_with_heq ptes_infos continue_tac new_infos g' )]) ] g let my_orelse tac1 tac2 g = try tac1 g with e when CErrors.noncritical e -> (* observe (str "using snd tac since : " ++ CErrors.print e); *) tac2 g let instantiate_hyps_with_args (do_prove:Id.t list -> tactic) hyps args_id = let args = Array.of_list (List.map mkVar args_id) in let instantiate_one_hyp hid = my_orelse ( (* we instantiate the hyp if possible *) fun g -> let prov_hid = pf_get_new_id hid g in let c = mkApp(mkVar hid,args) in let evm, _ = pf_apply Typing.type_of g c in tclTHENLIST[ Refiner.tclEVARS evm; Proofview.V82.of_tactic (pose_proof (Name prov_hid) c); thin [hid]; Proofview.V82.of_tactic (rename_hyp [prov_hid,hid]) ] g ) ( (* if not then we are in a mutual function block and this hyp is a recursive hyp on an other function. We are not supposed to use it while proving this principle so that we can trash it *) (fun g -> (* observe (str "Instantiation: removing hyp " ++ Ppconstr.pr_id hid); *) thin [hid] g ) ) in if List.is_empty args_id then tclTHENLIST [ tclMAP (fun hyp_id -> h_reduce_with_zeta (Locusops.onHyp hyp_id)) hyps; do_prove hyps ] else tclTHENLIST [ tclMAP (fun hyp_id -> h_reduce_with_zeta (Locusops.onHyp hyp_id)) hyps; tclMAP instantiate_one_hyp hyps; (fun g -> let all_g_hyps_id = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in let remaining_hyps = List.filter (fun id -> Id.Set.mem id all_g_hyps_id) hyps in do_prove remaining_hyps g ) ] let build_proof (interactive_proof:bool) (fnames:Constant.t list) ptes_infos dyn_infos : tactic = let rec build_proof_aux do_finalize dyn_infos : tactic = fun g -> let env = pf_env g in let sigma = project g in (* observe (str "proving on " ++ Printer.pr_lconstr_env (pf_env g) term);*) match EConstr.kind sigma dyn_infos.info with | Case(ci,ct,t,cb) -> let do_finalize_t dyn_info' = fun g -> let t = dyn_info'.info in let dyn_infos = {dyn_info' with info = mkCase(ci,ct,t,cb)} in let g_nb_prod = nb_prod (project g) (pf_concl g) in let type_of_term = pf_unsafe_type_of g t in let term_eq = make_refl_eq (Lazy.force refl_equal) type_of_term t in tclTHENLIST [ Proofview.V82.of_tactic (generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps))); thin dyn_infos.rec_hyps; Proofview.V82.of_tactic (pattern_option [Locus.AllOccurrencesBut [1],t] None); (fun g -> observe_tac "toto" ( tclTHENLIST [Proofview.V82.of_tactic (Simple.case t); (fun g' -> let g'_nb_prod = nb_prod (project g') (pf_concl g') in let nb_instantiate_partial = g'_nb_prod - g_nb_prod in observe_tac "treat_new_case" (treat_new_case ptes_infos nb_instantiate_partial (build_proof do_finalize) t dyn_infos) g' ) ]) g ) ] g in build_proof do_finalize_t {dyn_infos with info = t} g | Lambda(n,t,b) -> begin match EConstr.kind sigma (pf_concl g) with | Prod _ -> tclTHEN (Proofview.V82.of_tactic intro) (fun g' -> let open Context.Named.Declaration in let id = pf_last_hyp g' |> get_id in let new_term = pf_nf_betaiota g' (mkApp(dyn_infos.info,[|mkVar id|])) in let new_infos = {dyn_infos with info = new_term} in let do_prove new_hyps = build_proof do_finalize {new_infos with rec_hyps = new_hyps; nb_rec_hyps = List.length new_hyps } in (* observe_tac "Lambda" *) (instantiate_hyps_with_args do_prove new_infos.rec_hyps [id]) g' (* build_proof do_finalize new_infos g' *) ) g | _ -> do_finalize dyn_infos g end | Cast(t,_,_) -> build_proof do_finalize {dyn_infos with info = t} g | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ | Float _ -> do_finalize dyn_infos g | App(_,_) -> let f,args = decompose_app sigma dyn_infos.info in begin match EConstr.kind sigma f with | Int _ -> user_err Pp.(str "integer cannot be applied") | Float _ -> user_err Pp.(str "float cannot be applied") | App _ -> assert false (* we have collected all the app in decompose_app *) | Proj _ -> assert false (*FIXME*) | Var _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ -> let new_infos = { dyn_infos with info = (f,args) } in build_proof_args env sigma do_finalize new_infos g | Const (c,_) when not (List.mem_f Constant.equal c fnames) -> let new_infos = { dyn_infos with info = (f,args) } in (* Pp.msgnl (str "proving in " ++ pr_lconstr_env (pf_env g) dyn_infos.info); *) build_proof_args env sigma do_finalize new_infos g | Const _ -> do_finalize dyn_infos g | Lambda _ -> let new_term = Reductionops.nf_beta env sigma dyn_infos.info in build_proof do_finalize {dyn_infos with info = new_term} g | LetIn _ -> let new_infos = { dyn_infos with info = Reductionops.nf_betaiotazeta env sigma dyn_infos.info } in tclTHENLIST [tclMAP (fun hyp_id -> h_reduce_with_zeta (Locusops.onHyp hyp_id)) dyn_infos.rec_hyps; h_reduce_with_zeta Locusops.onConcl; build_proof do_finalize new_infos ] g | Cast(b,_,_) -> build_proof do_finalize {dyn_infos with info = b } g | Case _ | Fix _ | CoFix _ -> let new_finalize dyn_infos = let new_infos = { dyn_infos with info = dyn_infos.info,args } in build_proof_args env sigma do_finalize new_infos in build_proof new_finalize {dyn_infos with info = f } g end | Fix _ | CoFix _ -> user_err Pp.(str ( "Anonymous local (co)fixpoints are not handled yet")) | Proj _ -> user_err Pp.(str "Prod") | Prod _ -> do_finalize dyn_infos g | LetIn _ -> let new_infos = { dyn_infos with info = Reductionops.nf_betaiotazeta env sigma dyn_infos.info } in tclTHENLIST [tclMAP (fun hyp_id -> h_reduce_with_zeta (Locusops.onHyp hyp_id)) dyn_infos.rec_hyps; h_reduce_with_zeta Locusops.onConcl; build_proof do_finalize new_infos ] g | Rel _ -> anomaly (Pp.str "Free var in goal conclusion!") and build_proof do_finalize dyn_infos g = (* observe (str "proving with "++Printer.pr_lconstr dyn_infos.info++ str " on goal " ++ pr_gls g); *) Indfun_common.observe_tac (fun env sigma -> str "build_proof with " ++ pr_leconstr_env env sigma dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g and build_proof_args env sigma do_finalize dyn_infos (* f_args' args *) :tactic = fun g -> let (f_args',args) = dyn_infos.info in let tac : tactic = fun g -> match args with | [] -> do_finalize {dyn_infos with info = f_args'} g | arg::args -> (* observe (str "build_proof_args with arg := "++ pr_lconstr_env (pf_env g) arg++ *) (* fnl () ++ *) (* pr_goal (Tacmach.sig_it g) *) (* ); *) let do_finalize dyn_infos = let new_arg = dyn_infos.info in (* tclTRYD *) (build_proof_args env sigma do_finalize {dyn_infos with info = (mkApp(f_args',[|new_arg|])), args} ) in build_proof do_finalize {dyn_infos with info = arg } g in (* observe_tac "build_proof_args" *) (tac ) g in let do_finish_proof dyn_infos = (* tclTRYD *) (clean_goal_with_heq ptes_infos finish_proof dyn_infos) in (* observe_tac "build_proof" *) fun g -> build_proof (clean_goal_with_heq ptes_infos do_finish_proof) dyn_infos g (* Proof of principles from structural functions *) type static_fix_info = { idx : int; name : Id.t; types : types; offset : int; nb_realargs : int; body_with_param : constr; num_in_block : int } let prove_rec_hyp_for_struct fix_info = (fun eq_hyps -> tclTHEN (rewrite_until_var (fix_info.idx) eq_hyps) (fun g -> let _,pte_args = destApp (project g) (pf_concl g) in let rec_hyp_proof = mkApp(mkVar fix_info.name,array_get_start pte_args) in refine rec_hyp_proof g )) let prove_rec_hyp fix_info = { proving_tac = prove_rec_hyp_for_struct fix_info ; is_valid = fun _ -> true } let generalize_non_dep hyp g = (* observe (str "rec id := " ++ Ppconstr.pr_id hyp); *) let hyps = [hyp] in let env = Global.env () in let hyp_typ = pf_unsafe_type_of g (mkVar hyp) in let to_revert,_ = let open Context.Named.Declaration in Environ.fold_named_context_reverse (fun (clear,keep) decl -> let decl = map_named_decl EConstr.of_constr decl in let hyp = get_id decl in if Id.List.mem hyp hyps || List.exists (Termops.occur_var_in_decl env (project g) hyp) keep || Termops.occur_var env (project g) hyp hyp_typ || Termops.is_section_variable hyp (* should be dangerous *) then (clear,decl::keep) else (hyp::clear,keep)) ~init:([],[]) (pf_env g) in (* observe (str "to_revert := " ++ prlist_with_sep spc Ppconstr.pr_id to_revert); *) tclTHEN ((* observe_tac "h_generalize" *) (Proofview.V82.of_tactic (generalize (List.map mkVar to_revert) ))) ((* observe_tac "thin" *) (thin to_revert)) g let id_of_decl = RelDecl.get_name %> Nameops.Name.get_id let var_of_decl = id_of_decl %> mkVar let revert idl = tclTHEN (Proofview.V82.of_tactic (generalize (List.map mkVar idl))) (thin idl) let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num = (* observe (str "nb_args := " ++ str (string_of_int nb_args)); *) (* observe (str "nb_params := " ++ str (string_of_int nb_params)); *) (* observe (str "rec_args_num := " ++ str (string_of_int (rec_args_num + 1) )); *) let f_def = Global.lookup_constant (fst (destConst evd f)) in let eq_lhs = mkApp(f,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i))) in let (f_body, _, _) = Option.get (Global.body_of_constant_body Library.indirect_accessor f_def) in let f_body = EConstr.of_constr f_body in let params,f_body_with_params = decompose_lam_n evd nb_params f_body in let (_,num),(_,_,bodies) = destFix evd f_body_with_params in let fnames_with_params = let params = Array.init nb_params (fun i -> mkRel(nb_params - i)) in let fnames = List.rev (Array.to_list (Array.map (fun f -> mkApp(f,params)) fnames)) in fnames in (* observe (str "fnames_with_params " ++ prlist_with_sep fnl pr_lconstr fnames_with_params); *) (* observe (str "body " ++ pr_lconstr bodies.(num)); *) let f_body_with_params_and_other_fun = substl fnames_with_params bodies.(num) in (* observe (str "f_body_with_params_and_other_fun " ++ pr_lconstr f_body_with_params_and_other_fun); *) let eq_rhs = Reductionops.nf_betaiotazeta (Global.env ()) evd (mkApp(compose_lam params f_body_with_params_and_other_fun,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i)))) in (* observe (str "eq_rhs " ++ pr_lconstr eq_rhs); *) let (type_ctxt,type_of_f),evd = let evd,t = Typing.type_of ~refresh:true (Global.env ()) evd f in decompose_prod_n_assum evd (nb_params + nb_args) t,evd in let eqn = mkApp(Lazy.force eq,[|type_of_f;eq_lhs;eq_rhs|]) in let lemma_type = it_mkProd_or_LetIn eqn type_ctxt in (* Pp.msgnl (str "lemma type " ++ Printer.pr_lconstr lemma_type ++ fnl () ++ str "f_body " ++ Printer.pr_lconstr f_body); *) let f_id = Label.to_id (Constant.label (fst (destConst evd f))) in let prove_replacement = tclTHENLIST [ tclDO (nb_params + rec_args_num + 1) (Proofview.V82.of_tactic intro); observe_tac "" (fun g -> let rec_id = pf_nth_hyp_id g 1 in tclTHENLIST [observe_tac "generalize_non_dep in generate_equation_lemma" (generalize_non_dep rec_id); observe_tac "h_case" (Proofview.V82.of_tactic (simplest_case (mkVar rec_id))); (Proofview.V82.of_tactic intros_reflexivity)] g ) ] in (* Pp.msgnl (str "lemma type (2) " ++ Printer.pr_lconstr_env (Global.env ()) evd lemma_type); *) (*i The next call to mk_equation_id is valid since we are constructing the lemma Ensures by: obvious i*) let lemma = Lemmas.start_lemma ~name:(mk_equation_id f_id) ~poly:false evd lemma_type in let lemma,_ = Lemmas.by (Proofview.V82.tactic prove_replacement) lemma in let () = Lemmas.save_lemma_proved ~lemma ~opaque:Proof_global.Transparent ~idopt:None in evd let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num all_funs g = let equation_lemma = try let finfos = match find_Function_infos (fst (destConst !evd f)) (*FIXME*) with | None -> raise Not_found | Some finfos -> finfos in mkConst (Option.get finfos.equation_lemma) with (Not_found | Option.IsNone as e) -> let f_id = Label.to_id (Constant.label (fst (destConst !evd f))) in (*i The next call to mk_equation_id is valid since we will construct the lemma Ensures by: obvious i*) let equation_lemma_id = (mk_equation_id f_id) in evd := generate_equation_lemma !evd all_funs f fun_num (List.length params) (List.length rev_args_id) rec_arg_num; let _ = match e with | Option.IsNone -> let finfos = match find_Function_infos (fst (destConst !evd f)) with | None -> raise Not_found | Some finfos -> finfos in update_Function {finfos with equation_lemma = Some ( match Nametab.locate (qualid_of_ident equation_lemma_id) with | GlobRef.ConstRef c -> c | _ -> CErrors.anomaly (Pp.str "Not a constant.") ) } | _ -> () in (* let res = Constrintern.construct_reference (pf_hyps g) equation_lemma_id in *) let evd',res = Evd.fresh_global (Global.env ()) !evd (Constrintern.locate_reference (qualid_of_ident equation_lemma_id)) in evd:=evd'; let sigma, _ = Typing.type_of ~refresh:true (Global.env ()) !evd res in evd := sigma; res in let nb_intro_to_do = nb_prod (project g) (pf_concl g) in tclTHEN (tclDO nb_intro_to_do (Proofview.V82.of_tactic intro)) ( fun g' -> let just_introduced = nLastDecls nb_intro_to_do g' in let open Context.Named.Declaration in let just_introduced_id = List.map get_id just_introduced in tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR equation_lemma)) (revert just_introduced_id) g' ) g let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnames all_funs _nparams : tactic = fun g -> let princ_type = pf_concl g in (* Pp.msgnl (str "princ_type " ++ Printer.pr_lconstr princ_type); *) (* Pp.msgnl (str "all_funs "); *) (* Array.iter (fun c -> Pp.msgnl (Printer.pr_lconstr c)) all_funs; *) let princ_info = compute_elim_sig (project g) princ_type in let fresh_id = let avoid = ref (pf_ids_of_hyps g) in (fun na -> let new_id = match na with Name id -> fresh_id !avoid (Id.to_string id) | Anonymous -> fresh_id !avoid "H" in avoid := new_id :: !avoid; (Name new_id) ) in let fresh_decl = RelDecl.map_name fresh_id in let princ_info : elim_scheme = { princ_info with params = List.map fresh_decl princ_info.params; predicates = List.map fresh_decl princ_info.predicates; branches = List.map fresh_decl princ_info.branches; args = List.map fresh_decl princ_info.args } in let get_body const = match Global.body_of_constant Library.indirect_accessor const with | Some (body, _, _) -> let env = Global.env () in let sigma = Evd.from_env env in Tacred.cbv_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) env sigma (EConstr.of_constr body) | None -> user_err Pp.(str "Cannot define a principle over an axiom ") in let fbody = get_body fnames.(fun_num) in let f_ctxt,f_body = decompose_lam (project g) fbody in let f_ctxt_length = List.length f_ctxt in let diff_params = princ_info.nparams - f_ctxt_length in let full_params,princ_params,fbody_with_full_params = if diff_params > 0 then let princ_params,full_params = list_chop diff_params princ_info.params in (full_params, (* real params *) princ_params, (* the params of the principle which are not params of the function *) substl (* function instantiated with real params *) (List.map var_of_decl full_params) f_body ) else let f_ctxt_other,f_ctxt_params = list_chop (- diff_params) f_ctxt in let f_body = compose_lam f_ctxt_other f_body in (princ_info.params, (* real params *) [],(* all params are full params *) substl (* function instantiated with real params *) (List.map var_of_decl princ_info.params) f_body ) in observe (str "full_params := " ++ prlist_with_sep spc (RelDecl.get_name %> Nameops.Name.get_id %> Ppconstr.pr_id) full_params ); observe (str "princ_params := " ++ prlist_with_sep spc (RelDecl.get_name %> Nameops.Name.get_id %> Ppconstr.pr_id) princ_params ); observe (str "fbody_with_full_params := " ++ pr_leconstr_env (Global.env ()) !evd fbody_with_full_params ); let all_funs_with_full_params = Array.map (fun f -> applist(f, List.rev_map var_of_decl full_params)) all_funs in let fix_offset = List.length princ_params in let ptes_to_fix,infos = match EConstr.kind (project g) fbody_with_full_params with | Fix((idxs,i),(names,typess,bodies)) -> let bodies_with_all_params = Array.map (fun body -> Reductionops.nf_betaiota (pf_env g) (project g) (applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body, List.rev_map var_of_decl princ_params)) ) bodies in let info_array = Array.mapi (fun i types -> let types = prod_applist (project g) types (List.rev_map var_of_decl princ_params) in { idx = idxs.(i) - fix_offset; name = Nameops.Name.get_id (fresh_id names.(i).binder_name); types = types; offset = fix_offset; nb_realargs = List.length (fst (decompose_lam (project g) bodies.(i))) - fix_offset; body_with_param = bodies_with_all_params.(i); num_in_block = i } ) typess in let pte_to_fix,rev_info = List.fold_left_i (fun i (acc_map,acc_info) decl -> let pte = RelDecl.get_name decl in let infos = info_array.(i) in let type_args,_ = decompose_prod (project g) infos.types in let nargs = List.length type_args in let f = applist(mkConst fnames.(i), List.rev_map var_of_decl princ_info.params) in let first_args = Array.init nargs (fun i -> mkRel (nargs -i)) in let app_f = mkApp(f,first_args) in let pte_args = (Array.to_list first_args)@[app_f] in let app_pte = applist(mkVar (Nameops.Name.get_id pte),pte_args) in let body_with_param,num = let body = get_body fnames.(i) in let body_with_full_params = Reductionops.nf_betaiota (pf_env g) (project g) ( applist(body,List.rev_map var_of_decl full_params)) in match EConstr.kind (project g) body_with_full_params with | Fix((_,num),(_,_,bs)) -> Reductionops.nf_betaiota (pf_env g) (project g) ( (applist (substl (List.rev (Array.to_list all_funs_with_full_params)) bs.(num), List.rev_map var_of_decl princ_params)) ),num | _ -> user_err Pp.(str "Not a mutual block") in let info = {infos with types = compose_prod type_args app_pte; body_with_param = body_with_param; num_in_block = num } in (* observe (str "binding " ++ Ppconstr.pr_id (Nameops.Name.get_id pte) ++ *) (* str " to " ++ Ppconstr.pr_id info.name); *) (Id.Map.add (Nameops.Name.get_id pte) info acc_map,info::acc_info) ) 0 (Id.Map.empty,[]) (List.rev princ_info.predicates) in pte_to_fix,List.rev rev_info | _ -> Id.Map.empty,[] in let mk_fixes : tactic = let pre_info,infos = list_chop fun_num infos in match pre_info,infos with | _,[] -> tclIDTAC | _, this_fix_info::others_infos -> let other_fix_infos = List.map (fun fi -> fi.name,fi.idx + 1 ,fi.types) (pre_info@others_infos) in if List.is_empty other_fix_infos then if this_fix_info.idx + 1 = 0 then tclIDTAC (* Someone tries to defined a principle on a fully parametric definition declared as a fixpoint (strange but ....) *) else Indfun_common.observe_tac (fun _ _ -> str "h_fix " ++ int (this_fix_info.idx +1)) (Proofview.V82.of_tactic (fix this_fix_info.name (this_fix_info.idx +1))) else Proofview.V82.of_tactic (Tactics.mutual_fix this_fix_info.name (this_fix_info.idx + 1) other_fix_infos 0) in let first_tac : tactic = (* every operations until fix creations *) tclTHENLIST [ observe_tac "introducing params" (Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.params))); observe_tac "introducing predictes" (Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.predicates))); observe_tac "introducing branches" (Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.branches))); observe_tac "building fixes" mk_fixes; ] in let intros_after_fixes : tactic = fun gl -> let ctxt,pte_app = (decompose_prod_assum (project gl) (pf_concl gl)) in let pte,pte_args = (decompose_app (project gl) pte_app) in try let pte = try destVar (project gl) pte with DestKO -> anomaly (Pp.str "Property is not a variable.") in let fix_info = Id.Map.find pte ptes_to_fix in let nb_args = fix_info.nb_realargs in tclTHENLIST [ (* observe_tac ("introducing args") *) (tclDO nb_args (Proofview.V82.of_tactic intro)); (fun g -> (* replacement of the function by its body *) let args = nLastDecls nb_args g in let fix_body = fix_info.body_with_param in (* observe (str "fix_body := "++ pr_lconstr_env (pf_env gl) fix_body); *) let open Context.Named.Declaration in let args_id = List.map get_id args in let dyn_infos = { nb_rec_hyps = -100; rec_hyps = []; info = Reductionops.nf_betaiota (pf_env g) (project g) (applist(fix_body,List.rev_map mkVar args_id)); eq_hyps = [] } in tclTHENLIST [ observe_tac "do_replace" (do_replace evd full_params (fix_info.idx + List.length princ_params) (args_id@(List.map (RelDecl.get_name %> Nameops.Name.get_id) princ_params)) (all_funs.(fix_info.num_in_block)) fix_info.num_in_block all_funs ); let do_prove = build_proof interactive_proof (Array.to_list fnames) (Id.Map.map prove_rec_hyp ptes_to_fix) in let prove_tac branches = let dyn_infos = {dyn_infos with rec_hyps = branches; nb_rec_hyps = List.length branches } in observe_tac "cleaning" (clean_goal_with_heq (Id.Map.map prove_rec_hyp ptes_to_fix) do_prove dyn_infos) in (* observe (str "branches := " ++ *) (* prlist_with_sep spc (fun decl -> Ppconstr.pr_id (id_of_decl decl)) princ_info.branches ++ fnl () ++ *) (* str "args := " ++ prlist_with_sep spc Ppconstr.pr_id args_id *) (* ); *) (* observe_tac "instancing" *) (instantiate_hyps_with_args prove_tac (List.rev_map id_of_decl princ_info.branches) (List.rev args_id)) ] g ); ] gl with Not_found -> let nb_args = min (princ_info.nargs) (List.length ctxt) in tclTHENLIST [ tclDO nb_args (Proofview.V82.of_tactic intro); (fun g -> (* replacement of the function by its body *) let args = nLastDecls nb_args g in let open Context.Named.Declaration in let args_id = List.map get_id args in let dyn_infos = { nb_rec_hyps = -100; rec_hyps = []; info = Reductionops.nf_betaiota (pf_env g) (project g) (applist(fbody_with_full_params, (List.rev_map var_of_decl princ_params)@ (List.rev_map mkVar args_id) )); eq_hyps = [] } in let fname = destConst (project g) (fst (decompose_app (project g) (List.hd (List.rev pte_args)))) in tclTHENLIST [Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst fname))]); let do_prove = build_proof interactive_proof (Array.to_list fnames) (Id.Map.map prove_rec_hyp ptes_to_fix) in let prove_tac branches = let dyn_infos = {dyn_infos with rec_hyps = branches; nb_rec_hyps = List.length branches } in clean_goal_with_heq (Id.Map.map prove_rec_hyp ptes_to_fix) do_prove dyn_infos in instantiate_hyps_with_args prove_tac (List.rev_map id_of_decl princ_info.branches) (List.rev args_id) ] g ) ] gl in tclTHEN first_tac intros_after_fixes g (* Proof of principles of general functions *) (* let hrec_id = Recdef.hrec_id *) (* and acc_inv_id = Recdef.acc_inv_id *) (* and ltof_ref = Recdef.ltof_ref *) (* and acc_rel = Recdef.acc_rel *) (* and well_founded = Recdef.well_founded *) (* and list_rewrite = Recdef.list_rewrite *) (* and evaluable_of_global_reference = Recdef.evaluable_of_global_reference *) let prove_with_tcc tcc_lemma_constr eqs : tactic = match !tcc_lemma_constr with | Undefined -> anomaly (Pp.str "No tcc proof !!") | Value lemma -> fun gls -> (* let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in *) (* let ids = hid::pf_ids_of_hyps gls in *) tclTHENLIST [ (* generalize [lemma]; *) (* h_intro hid; *) (* Elim.h_decompose_and (mkVar hid); *) tclTRY(list_rewrite true eqs); (* (fun g -> *) (* let ids' = pf_ids_of_hyps g in *) (* let ids = List.filter (fun id -> not (List.mem id ids)) ids' in *) (* rewrite *) (* ) *) Proofview.V82.of_tactic (Eauto.gen_eauto (false,5) [] (Some [])) ] gls | Not_needed -> tclIDTAC let backtrack_eqs_until_hrec hrec eqs : tactic = fun gls -> let eqs = List.map mkVar eqs in let rewrite = tclFIRST (List.map (fun x -> Proofview.V82.of_tactic (Equality.rewriteRL x)) eqs ) in let _,hrec_concl = decompose_prod (project gls) (pf_unsafe_type_of gls (mkVar hrec)) in let f_app = Array.last (snd (destApp (project gls) hrec_concl)) in let f = (fst (destApp (project gls) f_app)) in let rec backtrack : tactic = fun g -> let f_app = Array.last (snd (destApp (project g) (pf_concl g))) in match EConstr.kind (project g) f_app with | App(f',_) when eq_constr (project g) f' f -> tclIDTAC g | _ -> tclTHEN rewrite backtrack g in backtrack gls let rec rewrite_eqs_in_eqs eqs = match eqs with | [] -> tclIDTAC | eq::eqs -> tclTHEN (tclMAP (fun id gl -> observe_tac (Format.sprintf "rewrite %s in %s " (Id.to_string eq) (Id.to_string id)) (tclTRY (Proofview.V82.of_tactic (Equality.general_rewrite_in true Locus.AllOccurrences true (* dep proofs also: *) true id (mkVar eq) false))) gl ) eqs ) (rewrite_eqs_in_eqs eqs) let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = fun gls -> (tclTHENLIST [ backtrack_eqs_until_hrec hrec eqs; (* observe_tac ("new_prove_with_tcc ( applying "^(Id.to_string hrec)^" )" ) *) (tclTHENS (* We must have exactly ONE subgoal !*) (Proofview.V82.of_tactic (apply (mkVar hrec))) [ tclTHENLIST [ (Proofview.V82.of_tactic (keep (tcc_hyps@eqs))); (Proofview.V82.of_tactic (apply (Lazy.force acc_inv))); (fun g -> if is_mes then Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference (delayed_force ltof_ref))]) g else tclIDTAC g ); observe_tac "rew_and_finish" (tclTHENLIST [tclTRY(list_rewrite false (List.map (fun v -> (mkVar v,true)) eqs)); observe_tac "rewrite_eqs_in_eqs" (rewrite_eqs_in_eqs eqs); (observe_tac "finishing using" ( tclCOMPLETE( Proofview.V82.of_tactic @@ Eauto.eauto_with_bases (true,5) [(fun _ sigma -> (sigma, Lazy.force refl_equal))] [Hints.Hint_db.empty TransparentState.empty false] ) ) ) ] ) ] ]) ]) gls let is_valid_hypothesis sigma predicates_name = let predicates_name = List.fold_right Id.Set.add predicates_name Id.Set.empty in let is_pte typ = if isApp sigma typ then let pte,_ = destApp sigma typ in if isVar sigma pte then Id.Set.mem (destVar sigma pte) predicates_name else false else false in let rec is_valid_hypothesis typ = is_pte typ || match EConstr.kind sigma typ with | Prod(_,pte,typ') -> is_pte pte && is_valid_hypothesis typ' | _ -> false in is_valid_hypothesis let prove_principle_for_gen (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation gl = let princ_type = pf_concl gl in let princ_info = compute_elim_sig (project gl) princ_type in let fresh_id = let avoid = ref (pf_ids_of_hyps gl) in fun na -> let new_id = match na with | Name id -> fresh_id !avoid (Id.to_string id) | Anonymous -> fresh_id !avoid "H" in avoid := new_id :: !avoid; Name new_id in let fresh_decl = map_name fresh_id in let princ_info : elim_scheme = { princ_info with params = List.map fresh_decl princ_info.params; predicates = List.map fresh_decl princ_info.predicates; branches = List.map fresh_decl princ_info.branches; args = List.map fresh_decl princ_info.args } in let wf_tac = if is_mes then (fun b -> Proofview.V82.of_tactic @@ Recdef.tclUSER_if_not_mes Tacticals.New.tclIDTAC b None) else fun _ -> prove_with_tcc tcc_lemma_ref [] in let real_rec_arg_num = rec_arg_num - princ_info.nparams in let npost_rec_arg = princ_info.nargs - real_rec_arg_num + 1 in (* observe ( *) (* str "princ_type := " ++ pr_lconstr princ_type ++ fnl () ++ *) (* str "princ_info.nparams := " ++ int princ_info.nparams ++ fnl () ++ *) (* str "princ_info.nargs := " ++ int princ_info.nargs ++ fnl () ++ *) (* str "rec_arg_num := " ++ int rec_arg_num ++ fnl() ++ *) (* str "real_rec_arg_num := " ++ int real_rec_arg_num ++ fnl () ++ *) (* str "npost_rec_arg := " ++ int npost_rec_arg ); *) let (post_rec_arg,pre_rec_arg) = Util.List.chop npost_rec_arg princ_info.args in let rec_arg_id = match List.rev post_rec_arg with | (LocalAssum ({binder_name=Name id},_) | LocalDef ({binder_name=Name id},_,_)) :: _ -> id | _ -> assert false in (* observe (str "rec_arg_id := " ++ pr_lconstr (mkVar rec_arg_id)); *) let subst_constrs = List.map (get_name %> Nameops.Name.get_id %> mkVar) (pre_rec_arg@princ_info.params) in let relation = substl subst_constrs relation in let input_type = substl subst_constrs rec_arg_type in let wf_thm_id = Nameops.Name.get_id (fresh_id (Name (Id.of_string "wf_R"))) in let acc_rec_arg_id = Nameops.Name.get_id (fresh_id (Name (Id.of_string ("Acc_"^(Id.to_string rec_arg_id))))) in let revert l = tclTHEN (Proofview.V82.of_tactic (Tactics.generalize (List.map mkVar l))) (Proofview.V82.of_tactic (clear l)) in let fix_id = Nameops.Name.get_id (fresh_id (Name hrec_id)) in let prove_rec_arg_acc g = ((* observe_tac "prove_rec_arg_acc" *) (tclCOMPLETE (tclTHEN (Proofview.V82.of_tactic (assert_by (Name wf_thm_id) (mkApp (delayed_force well_founded,[|input_type;relation|])) (Proofview.V82.tactic (fun g -> (* observe_tac "prove wf" *) (tclCOMPLETE (wf_tac is_mes)) g)))) ( (* observe_tac *) (* "apply wf_thm" *) Proofview.V82.of_tactic (Tactics.Simple.apply (mkApp(mkVar wf_thm_id,[|mkVar rec_arg_id|]))) ) ) ) ) g in let args_ids = List.map (get_name %> Nameops.Name.get_id) princ_info.args in let lemma = match !tcc_lemma_ref with | Undefined -> user_err Pp.(str "No tcc proof !!") | Value lemma -> EConstr.of_constr lemma | Not_needed -> EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.True.I") in (* let rec list_diff del_list check_list = *) (* match del_list with *) (* [] -> *) (* [] *) (* | f::r -> *) (* if List.mem f check_list then *) (* list_diff r check_list *) (* else *) (* f::(list_diff r check_list) *) (* in *) let tcc_list = ref [] in let start_tac gls = let hyps = pf_ids_of_hyps gls in let hid = next_ident_away_in_goal (Id.of_string "prov") (Id.Set.of_list hyps) in tclTHENLIST [ Proofview.V82.of_tactic (generalize [lemma]); Proofview.V82.of_tactic (Simple.intro hid); Proofview.V82.of_tactic (Elim.h_decompose_and (mkVar hid)); (fun g -> let new_hyps = pf_ids_of_hyps g in tcc_list := List.rev (List.subtract Id.equal new_hyps (hid::hyps)); if List.is_empty !tcc_list then begin tcc_list := [hid]; tclIDTAC g end else thin [hid] g ) ] gls in tclTHENLIST [ observe_tac "start_tac" start_tac; h_intros (List.rev_map (get_name %> Nameops.Name.get_id) (princ_info.args@princ_info.branches@princ_info.predicates@princ_info.params) ); (* observe_tac "" *) Proofview.V82.of_tactic (assert_by (Name acc_rec_arg_id) (mkApp (delayed_force acc_rel,[|input_type;relation;mkVar rec_arg_id|])) (Proofview.V82.tactic prove_rec_arg_acc) ); (* observe_tac "reverting" *) (revert (List.rev (acc_rec_arg_id::args_ids))); (* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl () ++ *) (* str "fix arg num" ++ int (List.length args_ids + 1) ); tclIDTAC g); *) (* observe_tac "h_fix " *) (Proofview.V82.of_tactic (fix fix_id (List.length args_ids + 1))); (* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl() ++ pr_lconstr_env (pf_env g ) (pf_unsafe_type_of g (mkVar fix_id) )); tclIDTAC g); *) h_intros (List.rev (acc_rec_arg_id::args_ids)); Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_ref)); (* observe_tac "finish" *) (fun gl' -> let body = let _,args = destApp (project gl') (pf_concl gl') in Array.last args in let body_info rec_hyps = { nb_rec_hyps = List.length rec_hyps; rec_hyps = rec_hyps; eq_hyps = []; info = body } in let acc_inv = lazy ( mkApp ( delayed_force acc_inv_id, [|input_type;relation;mkVar rec_arg_id|] ) ) in let acc_inv = lazy (mkApp(Lazy.force acc_inv, [|mkVar acc_rec_arg_id|])) in let predicates_names = List.map (get_name %> Nameops.Name.get_id) princ_info.predicates in let pte_info = { proving_tac = (fun eqs -> (* msgnl (str "tcc_list := "++ prlist_with_sep spc Ppconstr.pr_id !tcc_list); *) (* msgnl (str "princ_info.args := "++ prlist_with_sep spc Ppconstr.pr_id (List.map (fun (na,_,_) -> (Nameops.Name.get_id na)) princ_info.args)); *) (* msgnl (str "princ_info.params := "++ prlist_with_sep spc Ppconstr.pr_id (List.map (fun (na,_,_) -> (Nameops.Name.get_id na)) princ_info.params)); *) (* msgnl (str "acc_rec_arg_id := "++ Ppconstr.pr_id acc_rec_arg_id); *) (* msgnl (str "eqs := "++ prlist_with_sep spc Ppconstr.pr_id eqs); *) (* observe_tac "new_prove_with_tcc" *) (new_prove_with_tcc is_mes acc_inv fix_id (!tcc_list@(List.map (get_name %> Nameops.Name.get_id) (princ_info.args@princ_info.params) )@ ([acc_rec_arg_id])) eqs ) ); is_valid = is_valid_hypothesis (project gl') predicates_names } in let ptes_info : pte_info Id.Map.t = List.fold_left (fun map pte_id -> Id.Map.add pte_id pte_info map ) Id.Map.empty predicates_names in let make_proof rec_hyps = build_proof false [f_ref] ptes_info (body_info rec_hyps) in (* observe_tac "instantiate_hyps_with_args" *) (instantiate_hyps_with_args make_proof (List.map (get_name %> Nameops.Name.get_id) princ_info.branches) (List.rev args_ids) ) gl' ) ] gl coq-8.11.0/plugins/funind/invfun.mli0000644000175000017500000000141213612664533017240 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Names.GlobRef.t option -> unit Proofview.tactic coq-8.11.0/plugins/funind/recdef.ml0000644000175000017500000017125113612664533017023 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (try (match constant_opt_value_in (Global.env ()) sp with | Some c -> c | _ -> raise Not_found) with Not_found -> anomaly (str "Cannot find definition of constant " ++ (Id.print (Label.to_id (Constant.label (fst sp)))) ++ str ".") ) |_ -> assert false let type_of_const sigma t = match (EConstr.kind sigma t) with | Const (sp, u) -> let u = EInstance.kind sigma u in (* FIXME discarding universe constraints *) Typeops.type_of_constant_in (Global.env()) (sp, u) |_ -> assert false let constant sl s = UnivGen.constr_of_monomorphic_global (find_reference sl s) let const_of_ref = function GlobRef.ConstRef kn -> kn | _ -> anomaly (Pp.str "ConstRef expected.") (* Generic values *) let pf_get_new_ids idl g = let ids = pf_ids_of_hyps g in let ids = Id.Set.of_list ids in List.fold_right (fun id acc -> next_global_ident_away id (Id.Set.union (Id.Set.of_list acc) ids)::acc) idl [] let next_ident_away_in_goal ids avoid = next_ident_away_in_goal ids (Id.Set.of_list avoid) let compute_renamed_type gls c = rename_bound_vars_as_displayed (project gls) (*no avoid*) Id.Set.empty (*no rels*) [] (pf_unsafe_type_of gls c) let h'_id = Id.of_string "h'" let teq_id = Id.of_string "teq" let ano_id = Id.of_string "anonymous" let x_id = Id.of_string "x" let k_id = Id.of_string "k" let v_id = Id.of_string "v" let def_id = Id.of_string "def" let p_id = Id.of_string "p" let rec_res_id = Id.of_string "rec_res";; let lt = function () -> (coq_init_constant "lt") [@@@ocaml.warning "-3"] let le = function () -> (Coqlib.gen_reference_in_modules "RecursiveDefinition" Coqlib.init_modules "le") let ex = function () -> (coq_init_constant "ex") let nat = function () -> (coq_init_constant "nat") let iter_ref () = try find_reference ["Recdef"] "iter" with Not_found -> user_err Pp.(str "module Recdef not loaded") let iter_rd = function () -> (constr_of_monomorphic_global (delayed_force iter_ref)) let eq = function () -> (coq_init_constant "eq") let le_lt_SS = function () -> (constant ["Recdef"] "le_lt_SS") let le_lt_n_Sm = function () -> (coq_constant arith_Lt "le_lt_n_Sm") let le_trans = function () -> (coq_constant arith_Nat "le_trans") let le_lt_trans = function () -> (coq_constant arith_Nat "le_lt_trans") let lt_S_n = function () -> (coq_constant arith_Lt "lt_S_n") let le_n = function () -> (coq_init_constant "le_n") let coq_sig_ref = function () -> (find_reference ["Coq";"Init";"Specif"] "sig") let coq_O = function () -> (coq_init_constant "O") let coq_S = function () -> (coq_init_constant "S") let lt_n_O = function () -> (coq_constant arith_Nat "nlt_0_r") let max_ref = function () -> (find_reference ["Recdef"] "max") let max_constr = function () -> EConstr.of_constr (constr_of_monomorphic_global (delayed_force max_ref)) let f_S t = mkApp(delayed_force coq_S, [|t|]);; let rec n_x_id ids n = if Int.equal n 0 then [] else let x = next_ident_away_in_goal x_id ids in x::n_x_id (x::ids) (n-1);; let simpl_iter clause = reduce (Lazy {rBeta=true;rMatch=true;rFix=true;rCofix=true;rZeta=true;rDelta=false; rConst = [ EvalConstRef (const_of_ref (delayed_force iter_ref))]}) clause (* Others ugly things ... *) let (value_f: Constr.t list -> GlobRef.t -> Constr.t) = let open Term in let open Constr in fun al fterm -> let rev_x_id_l = ( List.fold_left (fun x_id_l _ -> let x_id = next_ident_away_in_goal x_id x_id_l in x_id::x_id_l ) [] al ) in let context = List.map (fun (x, c) -> LocalAssum (make_annot (Name x) Sorts.Relevant, c)) (List.combine rev_x_id_l (List.rev al)) in let env = Environ.push_rel_context context (Global.env ()) in let glob_body = DAst.make @@ GCases (RegularStyle,None, [DAst.make @@ GApp(DAst.make @@ GRef(fterm,None), List.rev_map (fun x_id -> DAst.make @@ GVar x_id) rev_x_id_l), (Anonymous,None)], [CAst.make ([v_id], [DAst.make @@ PatCstr ((destIndRef (delayed_force coq_sig_ref),1), [DAst.make @@ PatVar(Name v_id); DAst.make @@ PatVar Anonymous], Anonymous)], DAst.make @@ GVar v_id)]) in let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in let body = EConstr.Unsafe.to_constr body in it_mkLambda_or_LetIn body context let (declare_f : Id.t -> Decls.logical_kind -> Constr.t list -> GlobRef.t -> GlobRef.t) = fun f_id kind input_type fterm_ref -> declare_fun f_id kind (value_f input_type fterm_ref);; let observe_tclTHENLIST s tacl = if do_observe () then let rec aux n = function | [] -> tclIDTAC | [tac] -> observe_tac (fun env sigma -> s env sigma ++ spc () ++ int n) tac | tac::tacl -> observe_tac (fun env sigma -> s env sigma ++ spc () ++ int n) (tclTHEN tac (aux (succ n) tacl)) in aux 0 tacl else tclTHENLIST tacl module New = struct open Tacticals.New let observe_tac = New.observe_tac ~header:(Pp.mt()) let observe_tclTHENLIST s tacl = if do_observe () then let rec aux n = function | [] -> tclIDTAC | [tac] -> observe_tac (fun env sigma -> s env sigma ++ spc () ++ int n) tac | tac::tacl -> observe_tac (fun env sigma -> s env sigma ++ spc () ++ int n) (tclTHEN tac (aux (succ n) tacl)) in aux 0 tacl else tclTHENLIST tacl end (* Conclusion tactics *) (* The boolean value is_mes expresses that the termination is expressed using a measure function instead of a well-founded relation. *) let tclUSER tac is_mes l = let open Tacticals.New in let clear_tac = match l with | None -> tclIDTAC | Some l -> tclMAP (fun id -> tclTRY (clear [id])) (List.rev l) in New.observe_tclTHENLIST (fun _ _ -> str "tclUSER1") [ clear_tac; if is_mes then New.observe_tclTHENLIST (fun _ _ -> str "tclUSER2") [ unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference (delayed_force Indfun_common.ltof_ref))] ; tac ] else tac ] let tclUSER_if_not_mes concl_tac is_mes names_to_suppress = if is_mes then Tacticals.New.tclCOMPLETE (Simple.apply (delayed_force well_founded_ltof)) else (* tclTHEN (Simple.apply (delayed_force acc_intro_generator_function) ) *) (tclUSER concl_tac is_mes names_to_suppress) (* Traveling term. Both definitions of [f_terminate] and [f_equation] use the same generic traveling mechanism. *) (* [check_not_nested forbidden e] checks that [e] does not contains any variable of [forbidden] *) let check_not_nested env sigma forbidden e = let rec check_not_nested e = match EConstr.kind sigma e with | Rel _ -> () | Int _ | Float _ -> () | Var x -> if Id.List.mem x forbidden then user_err ~hdr:"Recdef.check_not_nested" (str "check_not_nested: failure " ++ Id.print x) | Meta _ | Evar _ | Sort _ -> () | Cast(e,_,t) -> check_not_nested e;check_not_nested t | Prod(_,t,b) -> check_not_nested t;check_not_nested b | Lambda(_,t,b) -> check_not_nested t;check_not_nested b | LetIn(_,v,t,b) -> check_not_nested t;check_not_nested b;check_not_nested v | App(f,l) -> check_not_nested f;Array.iter check_not_nested l | Proj (p,c) -> check_not_nested c | Const _ -> () | Ind _ -> () | Construct _ -> () | Case(_,t,e,a) -> check_not_nested t;check_not_nested e;Array.iter check_not_nested a | Fix _ -> user_err Pp.(str "check_not_nested : Fix") | CoFix _ -> user_err Pp.(str "check_not_nested : Fix") in try check_not_nested e with UserError(_,p) -> user_err ~hdr:"_" (str "on expr : " ++ Printer.pr_leconstr_env env sigma e ++ str " " ++ p) (* ['a info] contains the local information for traveling *) type 'a infos = { nb_arg : int; (* function number of arguments *) concl_tac : unit Proofview.tactic; (* final tactic to finish proofs *) rec_arg_id : Id.t; (*name of the declared recursive argument *) is_mes : bool; (* type of recursion *) ih : Id.t; (* induction hypothesis name *) f_id : Id.t; (* function name *) f_constr : constr; (* function term *) f_terminate : constr; (* termination proof term *) func : GlobRef.t; (* functional reference *) info : 'a; is_main_branch : bool; (* on the main branch or on a matched expression *) is_final : bool; (* final first order term or not *) values_and_bounds : (Id.t*Id.t) list; eqs : Id.t list; forbidden_ids : Id.t list; acc_inv : constr lazy_t; acc_id : Id.t; args_assoc : ((constr list)*constr) list; } type ('a,'b) journey_info_tac = 'a -> (* the arguments of the constructor *) 'b infos -> (* infos of the caller *) ('b infos -> tactic) -> (* the continuation tactic of the caller *) 'b infos -> (* argument of the tactic *) tactic (* journey_info : specifies the actions to do on the different term constructors during the traveling of the term *) type journey_info = { letiN : ((Name.t*constr*types*constr),constr) journey_info_tac; lambdA : ((Name.t*types*constr),constr) journey_info_tac; casE : ((constr infos -> tactic) -> constr infos -> tactic) -> ((case_info * constr * constr * constr array),constr) journey_info_tac; otherS : (unit,constr) journey_info_tac; apP : (constr*(constr list),constr) journey_info_tac; app_reC : (constr*(constr list),constr) journey_info_tac; message : string } let add_vars sigma forbidden e = let rec aux forbidden e = match EConstr.kind sigma e with | Var x -> x::forbidden | _ -> EConstr.fold sigma aux forbidden e in aux forbidden e let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = fun g -> let rev_context,b = decompose_lam_n (project g) nb_lam e in let ids = List.fold_left (fun acc (na,_) -> let pre_id = match na.binder_name with | Name x -> x | Anonymous -> ano_id in pre_id::acc ) [] rev_context in let rev_ids = pf_get_new_ids (List.rev ids) g in let new_b = substl (List.map mkVar rev_ids) b in observe_tclTHENLIST (fun _ _ -> str "treat_case1") [ h_intros (List.rev rev_ids); Proofview.V82.of_tactic (intro_using teq_id); onLastHypId (fun heq -> observe_tclTHENLIST (fun _ _ -> str "treat_case2")[ Proofview.V82.of_tactic (clear to_intros); h_intros to_intros; (fun g' -> let ty_teq = pf_unsafe_type_of g' (mkVar heq) in let teq_lhs,teq_rhs = let _,args = try destApp (project g') ty_teq with DestKO -> assert false in args.(1),args.(2) in let new_b' = Termops.replace_term (project g') teq_lhs teq_rhs new_b in let new_infos = { infos with info = new_b'; eqs = heq::infos.eqs; forbidden_ids = if forbid_new_ids then add_vars (project g') infos.forbidden_ids new_b' else infos.forbidden_ids } in finalize_tac new_infos g' ) ] ) ] g let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g = let sigma = project g in let env = pf_env g in match EConstr.kind sigma expr_info.info with | CoFix _ | Fix _ -> user_err Pp.(str "Function cannot treat local fixpoint or cofixpoint") | Proj _ -> user_err Pp.(str "Function cannot treat projections") | LetIn(na,b,t,e) -> begin let new_continuation_tac = jinfo.letiN (na.binder_name,b,t,e) expr_info continuation_tac in travel jinfo new_continuation_tac {expr_info with info = b; is_final=false} g end | Rel _ -> anomaly (Pp.str "Free var in goal conclusion!") | Prod _ -> begin try check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; jinfo.otherS () expr_info continuation_tac expr_info g with e when CErrors.noncritical e -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env env sigma expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id) end | Lambda(n,t,b) -> begin try check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; jinfo.otherS () expr_info continuation_tac expr_info g with e when CErrors.noncritical e -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env env sigma expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id) end | Case(ci,t,a,l) -> begin let continuation_tac_a = jinfo.casE (travel jinfo) (ci,t,a,l) expr_info continuation_tac in travel jinfo continuation_tac_a {expr_info with info = a; is_main_branch = false; is_final = false} g end | App _ -> let f,args = decompose_app sigma expr_info.info in if EConstr.eq_constr sigma f (expr_info.f_constr) then jinfo.app_reC (f,args) expr_info continuation_tac expr_info g else begin match EConstr.kind sigma f with | App _ -> assert false (* f is coming from a decompose_app *) | Const _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ | Var _ -> let new_infos = {expr_info with info=(f,args)} in let new_continuation_tac = jinfo.apP (f,args) expr_info continuation_tac in travel_args jinfo expr_info.is_main_branch new_continuation_tac new_infos g | Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env env sigma expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)") | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr_env env sigma expr_info.info ++ Pp.str ".") end | Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t} g | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ | Float _ -> let new_continuation_tac = jinfo.otherS () expr_info continuation_tac in new_continuation_tac expr_info g and travel_args jinfo is_final continuation_tac infos = let (f_args',args) = infos.info in match args with | [] -> continuation_tac {infos with info = f_args'; is_final = is_final} | arg::args' -> let new_continuation_tac new_infos = let new_arg = new_infos.info in travel_args jinfo is_final continuation_tac {new_infos with info = (mkApp(f_args',[|new_arg|]),args')} in travel jinfo new_continuation_tac {infos with info=arg;is_final=false} and travel jinfo continuation_tac expr_info = observe_tac (fun env sigma -> str jinfo.message ++ Printer.pr_leconstr_env env sigma expr_info.info) (travel_aux jinfo continuation_tac expr_info) (* Termination proof *) let rec prove_lt hyple g = let sigma = project g in begin try let (varx,varz) = match decompose_app sigma (pf_concl g) with | _, x::z::_ when isVar sigma x && isVar sigma z -> x, z | _ -> assert false in let h = List.find (fun id -> match decompose_app sigma (pf_unsafe_type_of g (mkVar id)) with | _, t::_ -> EConstr.eq_constr sigma t varx | _ -> false ) hyple in let y = List.hd (List.tl (snd (decompose_app sigma (pf_unsafe_type_of g (mkVar h))))) in observe_tclTHENLIST (fun _ _ -> str "prove_lt1")[ Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|]))); observe_tac (fun _ _ -> str "prove_lt") (prove_lt hyple) ] with Not_found -> ( ( observe_tclTHENLIST (fun _ _ -> str "prove_lt2")[ Proofview.V82.of_tactic (apply (delayed_force lt_S_n)); (observe_tac (fun _ _ -> str "assumption: " ++ Printer.pr_goal g) (Proofview.V82.of_tactic assumption)) ]) ) end g let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g = match lbounds with | [] -> let ids = pf_ids_of_hyps g in let s_max = mkApp(delayed_force coq_S, [|bound|]) in let k = next_ident_away_in_goal k_id ids in let ids = k::ids in let h' = next_ident_away_in_goal (h'_id) ids in let ids = h'::ids in let def = next_ident_away_in_goal def_id ids in observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux1")[ Proofview.V82.of_tactic (split (ImplicitBindings [s_max])); Proofview.V82.of_tactic (intro_then (fun id -> Proofview.V82.tactic begin observe_tac (fun _ _ -> str "destruct_bounds_aux") (tclTHENS (Proofview.V82.of_tactic (simplest_case (mkVar id))) [ observe_tclTHENLIST (fun _ _ -> str "")[Proofview.V82.of_tactic (intro_using h_id); Proofview.V82.of_tactic (simplest_elim(mkApp(delayed_force lt_n_O,[|s_max|]))); Proofview.V82.of_tactic default_full_auto]; observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux2")[ observe_tac (fun _ _ -> str "clearing k ") (Proofview.V82.of_tactic (clear [id])); h_intros [k;h';def]; observe_tac (fun _ _ -> str "simple_iter") (Proofview.V82.of_tactic (simpl_iter Locusops.onConcl)); observe_tac (fun _ _ -> str "unfold functional") (Proofview.V82.of_tactic (unfold_in_concl[(Locus.OnlyOccurrences [1], evaluable_of_global_reference infos.func)])); ( observe_tclTHENLIST (fun _ _ -> str "test")[ list_rewrite true (List.fold_right (fun e acc -> (mkVar e,true)::acc) infos.eqs (List.map (fun e -> (e,true)) rechyps) ); (* list_rewrite true *) (* (List.map (fun e -> (mkVar e,true)) infos.eqs) *) (* ; *) (observe_tac (fun _ _ -> str "finishing") (tclORELSE (Proofview.V82.of_tactic intros_reflexivity) (observe_tac (fun _ _ -> str "calling prove_lt") (prove_lt hyple))))]) ] ] )end)) ] g | (_,v_bound)::l -> observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux3")[ Proofview.V82.of_tactic (simplest_elim (mkVar v_bound)); Proofview.V82.of_tactic (clear [v_bound]); tclDO 2 (Proofview.V82.of_tactic intro); onNthHypId 1 (fun p_hyp -> (onNthHypId 2 (fun p -> observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux4")[ Proofview.V82.of_tactic (simplest_elim (mkApp(delayed_force max_constr, [| bound; mkVar p|]))); tclDO 3 (Proofview.V82.of_tactic intro); onNLastHypsId 3 (fun lids -> match lids with [hle2;hle1;pmax] -> destruct_bounds_aux infos ((mkVar pmax), hle1::hle2::hyple,(mkVar p_hyp)::rechyps) l | _ -> assert false) ; ] ) ) ) ] g let destruct_bounds infos = destruct_bounds_aux infos (delayed_force coq_O,[],[]) infos.values_and_bounds let terminate_app f_and_args expr_info continuation_tac infos = if expr_info.is_final && expr_info.is_main_branch then observe_tclTHENLIST (fun _ _ -> str "terminate_app1")[ continuation_tac infos; observe_tac (fun _ _ -> str "first split") (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info]))); observe_tac (fun _ _ -> str "destruct_bounds (1)") (destruct_bounds infos) ] else continuation_tac infos let terminate_others _ expr_info continuation_tac infos = if expr_info.is_final && expr_info.is_main_branch then observe_tclTHENLIST (fun _ _ -> str "terminate_others")[ continuation_tac infos; observe_tac (fun _ _ -> str "first split") (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info]))); observe_tac (fun _ _ -> str "destruct_bounds") (destruct_bounds infos) ] else continuation_tac infos let terminate_letin (na,b,t,e) expr_info continuation_tac info g = let sigma = project g in let env = pf_env g in let new_e = subst1 info.info e in let new_forbidden = let forbid = try check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) b; true with e when CErrors.noncritical e -> false in if forbid then match na with | Anonymous -> info.forbidden_ids | Name id -> id::info.forbidden_ids else info.forbidden_ids in continuation_tac {info with info = new_e; forbidden_ids = new_forbidden} g let pf_type c tac gl = let evars, ty = Typing.type_of (pf_env gl) (project gl) c in tclTHEN (Refiner.tclEVARS evars) (tac ty) gl let pf_typel l tac = let rec aux tys l = match l with | [] -> tac (List.rev tys) | hd :: tl -> pf_type hd (fun ty -> aux (ty::tys) tl) in aux [] l (* This is like the previous one except that it also rewrite on all hypotheses except the ones given in the first argument. All the modified hypotheses are generalized in the process and should be introduced back later; the result is the pair of the tactic and the list of hypotheses that have been generalized and cleared. *) let mkDestructEq : Id.t list -> constr -> goal Evd.sigma -> tactic * Id.t list = fun not_on_hyp expr g -> let hyps = pf_hyps g in let to_revert = Util.List.map_filter (fun decl -> let open Context.Named.Declaration in let id = get_id decl in if Id.List.mem id not_on_hyp || not (Termops.dependent (project g) expr (get_type decl)) then None else Some id) hyps in let to_revert_constr = List.rev_map mkVar to_revert in let type_of_expr = pf_unsafe_type_of g expr in let new_hyps = mkApp(Lazy.force refl_equal, [|type_of_expr; expr|]):: to_revert_constr in pf_typel new_hyps (fun _ -> observe_tclTHENLIST (fun _ _ -> str "mkDestructEq") [Proofview.V82.of_tactic (generalize new_hyps); (fun g2 -> let changefun patvars env sigma = pattern_occs [Locus.AllOccurrencesBut [1], expr] (pf_env g2) sigma (pf_concl g2) in Proofview.V82.of_tactic (change_in_concl ~check:true None changefun) g2); Proofview.V82.of_tactic (simplest_case expr)]), to_revert let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g = let sigma = project g in let env = pf_env g in let f_is_present = try check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) a; false with e when CErrors.noncritical e -> true in let a' = infos.info in let new_info = {infos with info = mkCase(ci,t,a',l); is_main_branch = expr_info.is_main_branch; is_final = expr_info.is_final} in let destruct_tac,rev_to_thin_intro = mkDestructEq [expr_info.rec_arg_id] a' g in let to_thin_intro = List.rev rev_to_thin_intro in observe_tac (fun _ _ -> str "treating cases (" ++ int (Array.length l) ++ str")" ++ spc () ++ Printer.pr_leconstr_env (pf_env g) sigma a') (try (tclTHENS destruct_tac (List.map_i (fun i e -> observe_tac (fun _ _ -> str "do treat case") (treat_case f_is_present to_thin_intro (next_step continuation_tac) ci.ci_cstr_ndecls.(i) e new_info)) 0 (Array.to_list l) )) with | UserError(Some "Refiner.thensn_tac3",_) | UserError(Some "Refiner.tclFAIL_s",_) -> (observe_tac (fun _ _ -> str "is computable " ++ Printer.pr_leconstr_env env sigma new_info.info) (next_step continuation_tac {new_info with info = Reductionops.nf_betaiotazeta (pf_env g) sigma new_info.info} ) )) g let terminate_app_rec (f,args) expr_info continuation_tac _ g = let sigma = project g in let env = pf_env g in List.iter (check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids)) args; begin try let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in let new_infos = {expr_info with info = v} in observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec")[ continuation_tac new_infos; if expr_info.is_final && expr_info.is_main_branch then observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec1")[ observe_tac (fun _ _ -> str "first split") (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info]))); observe_tac (fun _ _ -> str "destruct_bounds (3)") (destruct_bounds new_infos) ] else tclIDTAC ] g with Not_found -> observe_tac (fun _ _ -> str "terminate_app_rec not found") (tclTHENS (Proofview.V82.of_tactic (simplest_elim (mkApp(mkVar expr_info.ih,Array.of_list args)))) [ observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec2")[ Proofview.V82.of_tactic (intro_using rec_res_id); Proofview.V82.of_tactic intro; onNthHypId 1 (fun v_bound -> (onNthHypId 2 (fun v -> let new_infos = { expr_info with info = (mkVar v); values_and_bounds = (v,v_bound)::expr_info.values_and_bounds; args_assoc=(args,mkVar v)::expr_info.args_assoc } in observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec3")[ continuation_tac new_infos; if expr_info.is_final && expr_info.is_main_branch then observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec4")[ observe_tac (fun _ _ -> str "first split") (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info]))); observe_tac (fun _ _ -> str "destruct_bounds (2)") (destruct_bounds new_infos) ] else tclIDTAC ] ) ) ) ]; observe_tac (fun _ _ -> str "proving decreasing") ( tclTHENS (* proof of args < formal args *) (Proofview.V82.of_tactic (apply (Lazy.force expr_info.acc_inv))) [ observe_tac (fun _ _ -> str "assumption") (Proofview.V82.of_tactic assumption); observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec5") [ tclTRY(list_rewrite true (List.map (fun e -> mkVar e,true) expr_info.eqs ) ); Proofview.V82.of_tactic @@ tclUSER expr_info.concl_tac true (Some ( expr_info.ih::expr_info.acc_id:: (fun (x,y) -> y) (List.split expr_info.values_and_bounds) ) ); ] ]) ]) g end let terminate_info = { message = "prove_terminate with term "; letiN = terminate_letin; lambdA = (fun _ _ _ _ -> assert false); casE = terminate_case; otherS = terminate_others; apP = terminate_app; app_reC = terminate_app_rec; } let prove_terminate = travel terminate_info (* Equation proof *) let equation_case next_step (ci,a,t,l) expr_info continuation_tac infos = observe_tac (fun _ _ -> str "equation case") (terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos) let rec prove_le g = let sigma = project g in let x,z = let _,args = decompose_app sigma (pf_concl g) in (List.hd args,List.hd (List.tl args)) in tclFIRST[ Proofview.V82.of_tactic assumption; Proofview.V82.of_tactic (apply (delayed_force le_n)); begin try let matching_fun c = match EConstr.kind sigma c with | App (c, [| x0 ; _ |]) -> EConstr.isVar sigma x0 && Id.equal (destVar sigma x0) (destVar sigma x) && EConstr.is_global sigma (le ()) c | _ -> false in let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g) in let h = h.binder_name in let y = let _,args = decompose_app sigma t in List.hd (List.tl args) in observe_tclTHENLIST (fun _ _ -> str "prove_le")[ Proofview.V82.of_tactic (apply(mkApp(le_trans (),[|x;y;z;mkVar h|]))); observe_tac (fun _ _ -> str "prove_le (rec)") (prove_le) ] with Not_found -> tclFAIL 0 (mt()) end; ] g let rec make_rewrite_list expr_info max = function | [] -> tclIDTAC | (_,p,hp)::l -> observe_tac (fun _ _ -> str "make_rewrite_list") (tclTHENS (observe_tac (fun _ _ -> str "rewrite heq on " ++ Id.print p ) ( (fun g -> let sigma = project g in let t_eq = compute_renamed_type g (mkVar hp) in let k,def = let k_na,_,t = destProd sigma t_eq in let _,_,t = destProd sigma t in let def_na,_,_ = destProd sigma t in Nameops.Name.get_id k_na.binder_name,Nameops.Name.get_id def_na.binder_name in Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true (mkVar hp, ExplicitBindings[CAst.make @@ (NamedHyp def, expr_info.f_constr); CAst.make @@ (NamedHyp k, f_S max)]) false) g) ) ) [make_rewrite_list expr_info max l; observe_tclTHENLIST (fun _ _ -> str "make_rewrite_list")[ (* x < S max proof *) Proofview.V82.of_tactic (apply (delayed_force le_lt_n_Sm)); observe_tac (fun _ _ -> str "prove_le(2)") prove_le ] ] ) let make_rewrite expr_info l hp max = tclTHENFIRST (observe_tac (fun _ _ -> str "make_rewrite") (make_rewrite_list expr_info max l)) (observe_tac (fun _ _ -> str "make_rewrite") (tclTHENS (fun g -> let sigma = project g in let t_eq = compute_renamed_type g (mkVar hp) in let k,def = let k_na,_,t = destProd sigma t_eq in let _,_,t = destProd sigma t in let def_na,_,_ = destProd sigma t in Nameops.Name.get_id k_na.binder_name,Nameops.Name.get_id def_na.binder_name in observe_tac (fun _ _ -> str "general_rewrite_bindings") (Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true (mkVar hp, ExplicitBindings[CAst.make @@ (NamedHyp def, expr_info.f_constr); CAst.make @@ (NamedHyp k, f_S (f_S max))]) false)) g) [observe_tac(fun _ _ -> str "make_rewrite finalize") ( (* tclORELSE( h_reflexivity) *) (observe_tclTHENLIST (fun _ _ -> str "make_rewrite")[ Proofview.V82.of_tactic (simpl_iter Locusops.onConcl); observe_tac (fun _ _ -> str "unfold functional") (Proofview.V82.of_tactic (unfold_in_concl[(Locus.OnlyOccurrences [1], evaluable_of_global_reference expr_info.func)])); (list_rewrite true (List.map (fun e -> mkVar e,true) expr_info.eqs)); (observe_tac (fun _ _ -> str "h_reflexivity") (Proofview.V82.of_tactic intros_reflexivity) ) ])) ; observe_tclTHENLIST (fun _ _ -> str "make_rewrite1")[ (* x < S (S max) proof *) Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_lt_SS))); observe_tac (fun _ _ -> str "prove_le (3)") prove_le ] ]) ) let rec compute_max rew_tac max l = match l with | [] -> rew_tac max | (_,p,_)::l -> observe_tclTHENLIST (fun _ _ -> str "compute_max")[ Proofview.V82.of_tactic (simplest_elim (mkApp(delayed_force max_constr, [| max; mkVar p|]))); tclDO 3 (Proofview.V82.of_tactic intro); onNLastHypsId 3 (fun lids -> match lids with | [hle2;hle1;pmax] -> compute_max rew_tac (mkVar pmax) l | _ -> assert false )] let rec destruct_hex expr_info acc l = match l with | [] -> begin match List.rev acc with | [] -> tclIDTAC | (_,p,hp)::tl -> observe_tac (fun _ _ -> str "compute max ") (compute_max (make_rewrite expr_info tl hp) (mkVar p) tl) end | (v,hex)::l -> observe_tclTHENLIST (fun _ _ -> str "destruct_hex")[ Proofview.V82.of_tactic (simplest_case (mkVar hex)); Proofview.V82.of_tactic (clear [hex]); tclDO 2 (Proofview.V82.of_tactic intro); onNthHypId 1 (fun hp -> onNthHypId 2 (fun p -> observe_tac (fun _ _ -> str "destruct_hex after " ++ Id.print hp ++ spc () ++ Id.print p) (destruct_hex expr_info ((v,p,hp)::acc) l) ) ) ] let rec intros_values_eq expr_info acc = tclORELSE( observe_tclTHENLIST (fun _ _ -> str "intros_values_eq")[ tclDO 2 (Proofview.V82.of_tactic intro); onNthHypId 1 (fun hex -> (onNthHypId 2 (fun v -> intros_values_eq expr_info ((v,hex)::acc))) ) ]) (tclCOMPLETE ( destruct_hex expr_info [] acc )) let equation_others _ expr_info continuation_tac infos = if expr_info.is_final && expr_info.is_main_branch then observe_tac (fun env sigma -> str "equation_others (cont_tac +intros) " ++ Printer.pr_leconstr_env env sigma expr_info.info) (tclTHEN (continuation_tac infos) (observe_tac (fun env sigma -> str "intros_values_eq equation_others " ++ Printer.pr_leconstr_env env sigma expr_info.info) (intros_values_eq expr_info []))) else observe_tac (fun env sigma -> str "equation_others (cont_tac) " ++ Printer.pr_leconstr_env env sigma expr_info.info) (continuation_tac infos) let equation_app f_and_args expr_info continuation_tac infos = if expr_info.is_final && expr_info.is_main_branch then ((observe_tac (fun _ _ -> str "intros_values_eq equation_app") (intros_values_eq expr_info []))) else continuation_tac infos let equation_app_rec (f,args) expr_info continuation_tac info g = let sigma = project g in begin try let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in let new_infos = {expr_info with info = v} in observe_tac (fun _ _ -> str "app_rec found") (continuation_tac new_infos) g with Not_found -> if expr_info.is_final && expr_info.is_main_branch then observe_tclTHENLIST (fun _ _ -> str "equation_app_rec") [ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args))); continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc}; observe_tac (fun _ _ -> str "app_rec intros_values_eq") (intros_values_eq expr_info []) ] g else observe_tclTHENLIST (fun _ _ -> str "equation_app_rec1")[ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args))); observe_tac (fun _ _ -> str "app_rec not_found") (continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc}) ] g end let equation_info = {message = "prove_equation with term "; letiN = (fun _ -> assert false); lambdA = (fun _ _ _ _ -> assert false); casE = equation_case; otherS = equation_others; apP = equation_app; app_reC = equation_app_rec } let prove_eq = travel equation_info (* wrappers *) (* [compute_terminate_type] computes the type of the Definition f_terminate from the type of f_F *) let compute_terminate_type nb_args func = let open Term in let open Constr in let open CVars in let _,a_arrow_b,_ = destLambda(def_of_const (constr_of_monomorphic_global func)) in let rev_args,b = decompose_prod_n nb_args a_arrow_b in let left = mkApp(delayed_force iter_rd, Array.of_list (lift 5 a_arrow_b:: mkRel 3:: constr_of_monomorphic_global func::mkRel 1:: List.rev (List.map_i (fun i _ -> mkRel (6+i)) 0 rev_args) ) ) in let right = mkRel 5 in let delayed_force c = EConstr.Unsafe.to_constr (delayed_force c) in let equality = mkApp(delayed_force eq, [|lift 5 b; left; right|]) in let result = (mkProd (make_annot (Name def_id) Sorts.Relevant, lift 4 a_arrow_b, equality)) in let cond = mkApp(delayed_force lt, [|(mkRel 2); (mkRel 1)|]) in let nb_iter = mkApp(delayed_force ex, [|delayed_force nat; (mkLambda (make_annot (Name p_id) Sorts.Relevant, delayed_force nat, (mkProd (make_annot (Name k_id) Sorts.Relevant, delayed_force nat, mkArrow cond Sorts.Relevant result))))|])in let value = mkApp(constr_of_monomorphic_global (Util.delayed_force coq_sig_ref), [|b; (mkLambda (make_annot (Name v_id) Sorts.Relevant, b, nb_iter))|]) in compose_prod rev_args value let termination_proof_header is_mes input_type ids args_id relation rec_arg_num rec_arg_id tac wf_tac : tactic = begin fun g -> let nargs = List.length args_id in let pre_rec_args = List.rev_map mkVar (fst (List.chop (rec_arg_num - 1) args_id)) in let relation = substl pre_rec_args relation in let input_type = substl pre_rec_args input_type in let wf_thm = next_ident_away_in_goal (Id.of_string ("wf_R")) ids in let wf_rec_arg = next_ident_away_in_goal (Id.of_string ("Acc_"^(Id.to_string rec_arg_id))) (wf_thm::ids) in let hrec = next_ident_away_in_goal hrec_id (wf_rec_arg::wf_thm::ids) in let acc_inv = lazy ( mkApp ( delayed_force acc_inv_id, [|input_type;relation;mkVar rec_arg_id|] ) ) in tclTHEN (h_intros args_id) (tclTHENS (observe_tac (fun _ _ -> str "first assert") (Proofview.V82.of_tactic (assert_before (Name wf_rec_arg) (mkApp (delayed_force acc_rel, [|input_type;relation;mkVar rec_arg_id|]) ) )) ) [ (* accesibility proof *) tclTHENS (observe_tac (fun _ _ -> str "second assert") (Proofview.V82.of_tactic (assert_before (Name wf_thm) (mkApp (delayed_force well_founded,[|input_type;relation|])) )) ) [ (* interactive proof that the relation is well_founded *) observe_tac (fun _ _ -> str "wf_tac") (wf_tac is_mes (Some args_id)); (* this gives the accessibility argument *) observe_tac (fun _ _ -> str "apply wf_thm") (Proofview.V82.of_tactic (Simple.apply (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|]))) ) ] ; (* rest of the proof *) observe_tclTHENLIST (fun _ _ -> str "rest of proof") [observe_tac (fun _ _ -> str "generalize") (onNLastHypsId (nargs+1) (tclMAP (fun id -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (Proofview.V82.of_tactic (clear [id]))) )) ; observe_tac (fun _ _ -> str "fix") (Proofview.V82.of_tactic (fix hrec (nargs+1))); h_intros args_id; Proofview.V82.of_tactic (Simple.intro wf_rec_arg); observe_tac (fun _ _ -> str "tac") (tac wf_rec_arg hrec wf_rec_arg acc_inv) ] ] ) g end let rec instantiate_lambda sigma t l = match l with | [] -> t | a::l -> let (_, _, body) = destLambda sigma t in instantiate_lambda sigma (subst1 a body) l let whole_start concl_tac nb_args is_mes func input_type relation rec_arg_num : tactic = begin fun g -> let sigma = project g in let ids = Termops.ids_of_named_context (pf_hyps g) in let func_body = (def_of_const (constr_of_monomorphic_global func)) in let func_body = EConstr.of_constr func_body in let (f_name, _, body1) = destLambda sigma func_body in let f_id = match f_name.binder_name with | Name f_id -> next_ident_away_in_goal f_id ids | Anonymous -> anomaly (Pp.str "Anonymous function.") in let n_names_types,_ = decompose_lam_n sigma nb_args body1 in let n_ids,ids = List.fold_left (fun (n_ids,ids) (n_name,_) -> match n_name.binder_name with | Name id -> let n_id = next_ident_away_in_goal id ids in n_id::n_ids,n_id::ids | _ -> anomaly (Pp.str "anonymous argument.") ) ([],(f_id::ids)) n_names_types in let rec_arg_id = List.nth n_ids (rec_arg_num - 1) in let expr = instantiate_lambda sigma func_body (mkVar f_id::(List.map mkVar n_ids)) in termination_proof_header is_mes input_type ids n_ids relation rec_arg_num rec_arg_id (fun rec_arg_id hrec acc_id acc_inv g -> (prove_terminate (fun infos -> tclIDTAC) { is_main_branch = true; (* we are on the main branche (i.e. still on a match ... with .... end *) is_final = true; (* and on leaf (more or less) *) f_terminate = delayed_force coq_O; nb_arg = nb_args; concl_tac; rec_arg_id = rec_arg_id; is_mes = is_mes; ih = hrec; f_id = f_id; f_constr = mkVar f_id; func = func; info = expr; acc_inv = acc_inv; acc_id = acc_id; values_and_bounds = []; eqs = []; forbidden_ids = []; args_assoc = [] } ) g ) (fun b ids -> Proofview.V82.of_tactic (tclUSER_if_not_mes concl_tac b ids)) g end let get_current_subgoals_types pstate = let p = Proof_global.get_proof pstate in let Proof.{ goals=sgs; sigma; _ } = Proof.data p in sigma, List.map (Goal.V82.abstract_type sigma) sgs exception EmptySubgoals let build_and_l sigma l = let and_constr = UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.and.type" in let conj_constr = Coqlib.build_coq_conj () in let mk_and p1 p2 = mkApp(EConstr.of_constr and_constr,[|p1;p2|]) in let rec is_well_founded t = match EConstr.kind sigma t with | Prod(_,_,t') -> is_well_founded t' | App(_,_) -> let (f,_) = decompose_app sigma t in EConstr.eq_constr sigma f (well_founded ()) | _ -> false in let compare t1 t2 = let b1,b2= is_well_founded t1,is_well_founded t2 in if (b1&&b2) || not (b1 || b2) then 0 else if b1 && not b2 then 1 else -1 in let l = List.sort compare l in let rec f = function | [] -> raise EmptySubgoals | [p] -> p,tclIDTAC,1 | p1::pl -> let c,tac,nb = f pl in mk_and p1 c, tclTHENS (Proofview.V82.of_tactic (apply (EConstr.of_constr (constr_of_monomorphic_global conj_constr)))) [tclIDTAC; tac ],nb+1 in f l let is_rec_res id = let rec_res_name = Id.to_string rec_res_id in let id_name = Id.to_string id in try String.equal (String.sub id_name 0 (String.length rec_res_name)) rec_res_name with Invalid_argument _ -> false let clear_goals sigma = let rec clear_goal t = match EConstr.kind sigma t with | Prod({binder_name=Name id} as na,t',b) -> let b' = clear_goal b in if noccurn sigma 1 b' && (is_rec_res id) then Vars.lift (-1) b' else if b' == b then t else mkProd(na,t',b') | _ -> EConstr.map sigma clear_goal t in List.map clear_goal let build_new_goal_type lemma = let sigma, sub_gls_types = Lemmas.pf_fold get_current_subgoals_types lemma in (* Pp.msgnl (str "sub_gls_types1 := " ++ Util.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *) let sub_gls_types = clear_goals sigma sub_gls_types in (* Pp.msgnl (str "sub_gls_types2 := " ++ Pp.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *) let res = build_and_l sigma sub_gls_types in sigma, res let is_opaque_constant c = let cb = Global.lookup_constant c in match cb.Declarations.const_body with | Declarations.OpaqueDef _ -> Proof_global.Opaque | Declarations.Undef _ -> Proof_global.Opaque | Declarations.Def _ -> Proof_global.Transparent | Declarations.Primitive _ -> Proof_global.Opaque let open_new_goal ~lemma build_proof sigma using_lemmas ref_ goal_name (gls_type,decompose_and_tac,nb_goal) = (* Pp.msgnl (str "gls_type := " ++ Printer.pr_lconstr gls_type); *) let current_proof_name = Lemmas.pf_fold Proof_global.get_proof_name lemma in let name = match goal_name with | Some s -> s | None -> try add_suffix current_proof_name "_subproof" with e when CErrors.noncritical e -> anomaly (Pp.str "open_new_goal with an unnamed theorem.") in let na = next_global_ident_away name Id.Set.empty in if Termops.occur_existential sigma gls_type then CErrors.user_err Pp.(str "\"abstract\" cannot handle existentials"); let hook _ = let opacity = let na_ref = qualid_of_ident na in let na_global = Smartlocate.global_with_alias na_ref in match na_global with GlobRef.ConstRef c -> is_opaque_constant c | _ -> anomaly ~label:"equation_lemma" (Pp.str "not a constant.") in let lemma = mkConst (Names.Constant.make1 (Lib.make_kn na)) in ref_ := Value (EConstr.Unsafe.to_constr lemma); let lid = ref [] in let h_num = ref (-1) in let env = Global.env () in let start_tac = let open Tacmach.New in let open Tacticals.New in Proofview.Goal.enter (fun gl -> let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gl) in New.observe_tclTHENLIST (fun _ _ -> mt ()) [ generalize [lemma] ; Simple.intro hid ; Proofview.Goal.enter (fun gl -> let ids = pf_ids_of_hyps gl in tclTHEN (Elim.h_decompose_and (mkVar hid)) (Proofview.Goal.enter (fun gl -> let ids' = pf_ids_of_hyps gl in lid := List.rev (List.subtract Id.equal ids' ids); if List.is_empty !lid then lid := [hid]; tclIDTAC))) ]) in let end_tac = let open Tacmach.New in let open Tacticals.New in Proofview.Goal.enter (fun gl -> let sigma = project gl in match EConstr.kind sigma (pf_concl gl) with | App(f,_) when EConstr.eq_constr sigma f (well_founded ()) -> Auto.h_auto None [] (Some []) | _ -> incr h_num; tclCOMPLETE( tclFIRST [ tclTHEN (eapply_with_bindings (mkVar (List.nth !lid !h_num), NoBindings)) e_assumption ; Eauto.eauto_with_bases (true,5) [(fun _ sigma -> (sigma, (Lazy.force refl_equal)))] [Hints.Hint_db.empty TransparentState.empty false ] ] )) in let lemma = build_proof env (Evd.from_env env) start_tac end_tac in Lemmas.save_lemma_proved ~lemma ~opaque:opacity ~idopt:None in let info = Lemmas.Info.make ~hook:(DeclareDef.Hook.make hook) () in let lemma = Lemmas.start_lemma ~name:na ~poly:false (* FIXME *) ~info sigma gls_type in let lemma = if Indfun_common.is_strict_tcc () then fst @@ Lemmas.by (Proofview.V82.tactic (tclIDTAC)) lemma else fst @@ Lemmas.by (Proofview.V82.tactic begin fun g -> tclTHEN (decompose_and_tac) (tclORELSE (tclFIRST (List.map (fun c -> Proofview.V82.of_tactic (Tacticals.New.tclTHENLIST [intros; Simple.apply (fst (interp_constr (Global.env()) Evd.empty c)) (*FIXME*); Tacticals.New.tclCOMPLETE Auto.default_auto ]) ) using_lemmas) ) tclIDTAC) g end) lemma in if Lemmas.(pf_fold Proof_global.get_open_goals) lemma = 0 then (defined lemma; None) else Some lemma let com_terminate interactive_proof tcc_lemma_name tcc_lemma_ref is_mes fonctional_ref input_type relation rec_arg_num thm_name using_lemmas nb_args ctx hook = let start_proof env ctx tac_start tac_end = let info = Lemmas.Info.make ~hook () in let lemma = Lemmas.start_lemma ~name:thm_name ~poly:false (*FIXME*) ~info ctx (EConstr.of_constr (compute_terminate_type nb_args fonctional_ref)) in let lemma = fst @@ Lemmas.by (New.observe_tac (fun _ _ -> str "starting_tac") tac_start) lemma in fst @@ Lemmas.by (Proofview.V82.tactic (observe_tac (fun _ _ -> str "whole_start") (whole_start tac_end nb_args is_mes fonctional_ref input_type relation rec_arg_num ))) lemma in let lemma = start_proof Global.(env ()) ctx Tacticals.New.tclIDTAC Tacticals.New.tclIDTAC in try let sigma, new_goal_type = build_new_goal_type lemma in let sigma = Evd.from_ctx (Evd.evar_universe_context sigma) in open_new_goal ~lemma start_proof sigma using_lemmas tcc_lemma_ref (Some tcc_lemma_name) (new_goal_type) with EmptySubgoals -> (* a non recursive function declared with measure ! *) tcc_lemma_ref := Not_needed; if interactive_proof then Some lemma else (defined lemma; None) let start_equation (f:GlobRef.t) (term_f:GlobRef.t) (cont_tactic:Id.t list -> tactic) g = let sigma = project g in let ids = pf_ids_of_hyps g in let terminate_constr = constr_of_monomorphic_global term_f in let terminate_constr = EConstr.of_constr terminate_constr in let nargs = nb_prod (project g) (EConstr.of_constr (type_of_const sigma terminate_constr)) in let x = n_x_id ids nargs in observe_tac (fun _ _ -> str "start_equation") (observe_tclTHENLIST (fun _ _ -> str "start_equation") [ h_intros x; Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference f)]); observe_tac (fun _ _ -> str "simplest_case") (Proofview.V82.of_tactic (simplest_case (mkApp (terminate_constr, Array.of_list (List.map mkVar x))))); observe_tac (fun _ _ -> str "prove_eq") (cont_tactic x)]) g;; let com_eqn uctx nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type = let open CVars in let opacity = match terminate_ref with | GlobRef.ConstRef c -> is_opaque_constant c | _ -> anomaly ~label:"terminate_lemma" (Pp.str "not a constant.") in let evd = Evd.from_ctx uctx in let f_constr = constr_of_monomorphic_global f_ref in let equation_lemma_type = subst1 f_constr equation_lemma_type in let lemma = Lemmas.start_lemma ~name:eq_name ~poly:false evd (EConstr.of_constr equation_lemma_type) in let lemma = fst @@ Lemmas.by (Proofview.V82.tactic (start_equation f_ref terminate_ref (fun x -> prove_eq (fun _ -> tclIDTAC) {nb_arg=nb_arg; f_terminate = EConstr.of_constr (constr_of_monomorphic_global terminate_ref); f_constr = EConstr.of_constr f_constr; concl_tac = Tacticals.New.tclIDTAC; func=functional_ref; info=(instantiate_lambda Evd.empty (EConstr.of_constr (def_of_const (constr_of_monomorphic_global functional_ref))) (EConstr.of_constr f_constr::List.map mkVar x) ); is_main_branch = true; is_final = true; values_and_bounds = []; eqs = []; forbidden_ids = []; acc_inv = lazy (assert false); acc_id = Id.of_string "____"; args_assoc = []; f_id = Id.of_string "______"; rec_arg_id = Id.of_string "______"; is_mes = false; ih = Id.of_string "______"; } ) )) lemma in let _ = Flags.silently (fun () -> Lemmas.save_lemma_proved ~lemma ~opaque:opacity ~idopt:None) () in () (* Pp.msgnl (fun _ _ -> str "eqn finished"); *) let recursive_definition ~interactive_proof ~is_mes function_name rec_impls type_of_f r rec_arg_num eq generate_induction_principle using_lemmas : Lemmas.t option = let open Term in let open Constr in let open CVars in let env = Global.env() in let evd = Evd.from_env env in let evd, function_type = interp_type_evars ~program_mode:false env evd type_of_f in let function_r = Sorts.Relevant in (* TODO relevance *) let env = EConstr.push_named (Context.Named.Declaration.LocalAssum (make_annot function_name function_r,function_type)) env in (* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *) let evd, ty = interp_type_evars ~program_mode:false env evd ~impls:rec_impls eq in let evd = Evd.minimize_universes evd in let equation_lemma_type = Reductionops.nf_betaiotazeta env evd (Evarutil.nf_evar evd ty) in let function_type = EConstr.to_constr ~abort_on_undefined_evars:false evd function_type in let equation_lemma_type = EConstr.Unsafe.to_constr equation_lemma_type in (* Pp.msgnl (fun _ _ -> str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *) let res_vars,eq' = decompose_prod equation_lemma_type in let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in let eq' = Reductionops.nf_zeta env_eq' evd (EConstr.of_constr eq') in let eq' = EConstr.Unsafe.to_constr eq' in let res = (* Pp.msgnl (fun _ _ -> str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *) (* Pp.msgnl (fun _ _ -> str "rec_arg_num := " ++ str (fun _ _ -> string_of_int rec_arg_num)); *) (* Pp.msgnl (fun _ _ -> str "eq' := " ++ str (fun _ _ -> string_of_int rec_arg_num)); *) match Constr.kind eq' with | App(e,[|_;_;eq_fix|]) -> mkLambda (make_annot (Name function_name) Sorts.Relevant,function_type,subst_var function_name (compose_lam res_vars eq_fix)) | _ -> failwith "Recursive Definition (res not eq)" in let pre_rec_args,function_type_before_rec_arg = decompose_prod_n (rec_arg_num - 1) function_type in let (_, rec_arg_type, _) = destProd function_type_before_rec_arg in let arg_types = List.rev_map snd (fst (decompose_prod_n (List.length res_vars) function_type)) in let equation_id = add_suffix function_name "_equation" in let functional_id = add_suffix function_name "_F" in let term_id = add_suffix function_name "_terminate" in let functional_ref = let univs = Evd.univ_entry ~poly:false evd in declare_fun functional_id Decls.(IsDefinition Definition) ~univs res in (* Refresh the global universes, now including those of _F *) let evd = Evd.from_env (Global.env ()) in let env_with_pre_rec_args = push_rel_context(List.map (function (x,t) -> LocalAssum (x,t)) pre_rec_args) env in let relation, evuctx = interp_constr env_with_pre_rec_args evd r in let evd = Evd.from_ctx evuctx in let tcc_lemma_name = add_suffix function_name "_tcc" in let tcc_lemma_constr = ref Undefined in (* let _ = Pp.msgnl (fun _ _ -> str "relation := " ++ Printer.pr_lconstr_env env_with_pre_rec_args relation) in *) let hook { DeclareDef.Hook.S.uctx ; _ } = let term_ref = Nametab.locate (qualid_of_ident term_id) in let f_ref = declare_f function_name Decls.(IsProof Lemma) arg_types term_ref in let _ = Extraction_plugin.Table.extraction_inline true [qualid_of_ident term_id] in (* message "start second proof"; *) let stop = (* XXX: What is the correct way to get sign at hook time *) try com_eqn uctx (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type); false with e when CErrors.noncritical e -> begin if do_observe () then Feedback.msg_debug (str "Cannot create equation Lemma " ++ CErrors.print e) else CErrors.user_err ~hdr:"Cannot create equation Lemma" (str "Cannot create equation lemma." ++ spc () ++ str "This may be because the function is nested-recursive.") ; true end in if not stop then let eq_ref = Nametab.locate (qualid_of_ident equation_id ) in let f_ref = destConst (constr_of_monomorphic_global f_ref) and functional_ref = destConst (constr_of_monomorphic_global functional_ref) and eq_ref = destConst (constr_of_monomorphic_global eq_ref) in generate_induction_principle f_ref tcc_lemma_constr functional_ref eq_ref rec_arg_num (EConstr.of_constr rec_arg_type) (nb_prod evd (EConstr.of_constr res)) relation in (* XXX STATE Why do we need this... why is the toplevel protection not enough *) funind_purify (fun () -> com_terminate interactive_proof tcc_lemma_name tcc_lemma_constr is_mes functional_ref (EConstr.of_constr rec_arg_type) relation rec_arg_num term_id using_lemmas (List.length res_vars) evd (DeclareDef.Hook.make hook)) () coq-8.11.0/plugins/funind/functional_principles_proofs.mli0000644000175000017500000000123013612664533023713 0ustar treinentreinenopen Names val prove_princ_for_struct : Evd.evar_map ref -> bool -> int -> Constant.t array -> EConstr.constr array -> int -> Tacmach.tactic val prove_principle_for_gen : Constant.t * Constant.t * Constant.t -> (* name of the function, the functional and the fixpoint equation *) Indfun_common.tcc_lemma_value ref -> (* a pointer to the obligation proofs lemma *) bool -> (* is that function uses measure *) int -> (* the number of recursive argument *) EConstr.types -> (* the type of the recursive argument *) EConstr.constr -> (* the wf relation used to prove the function *) Tacmach.tactic (* val is_pte : rel_declaration -> bool *) coq-8.11.0/plugins/funind/gen_principle.mli0000644000175000017500000000231413612664533020553 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Pp.t * Pp.t -> unit val warn_cannot_define_principle : ?loc:Loc.t -> Pp.t * Pp.t -> unit val do_generate_principle_interactive : Vernacexpr.fixpoint_expr list -> Lemmas.t val do_generate_principle : Vernacexpr.fixpoint_expr list -> unit val make_graph : Names.GlobRef.t -> unit (* Can be thrown by build_{,case}_scheme *) exception No_graph_found val build_scheme : (Names.Id.t * Libnames.qualid * Sorts.family) list -> unit val build_case_scheme : (Names.Id.t * Libnames.qualid * Sorts.family) -> unit coq-8.11.0/plugins/funind/functional_principles_types.mli0000644000175000017500000000142313612664533023553 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Sorts.t array -> Constr.t -> Constr.types coq-8.11.0/plugins/funind/indfun_common.mli0000644000175000017500000000726413612664533020601 0ustar treinentreinenopen Names (* The mk_?_id function build different name w.r.t. a function Each of their use is justified in the code *) val mk_rel_id : Id.t -> Id.t val mk_correct_id : Id.t -> Id.t val mk_complete_id : Id.t -> Id.t val mk_equation_id : Id.t -> Id.t val fresh_id : Id.t list -> string -> Id.t val fresh_name : Id.t list -> string -> Name.t val get_name : Id.t list -> ?default:string -> Name.t -> Name.t val array_get_start : 'a array -> 'a array val locate_ind : Libnames.qualid -> inductive val locate_constant : Libnames.qualid -> Constant.t val locate_with_msg : Pp.t -> (Libnames.qualid -> 'a) -> Libnames.qualid -> 'a val filter_map : ('a -> bool) -> ('a -> 'b) -> 'a list -> 'b list val list_union_eq : ('a -> 'a -> bool) -> 'a list -> 'a list -> 'a list val list_add_set_eq : ('a -> 'a -> bool) -> 'a -> 'a list -> 'a list val chop_rlambda_n : int -> Glob_term.glob_constr -> (Name.t*Glob_term.glob_constr*Glob_term.glob_constr option) list * Glob_term.glob_constr val chop_rprod_n : int -> Glob_term.glob_constr -> (Name.t*Glob_term.glob_constr) list * Glob_term.glob_constr val eq : EConstr.constr Lazy.t val refl_equal : EConstr.constr Lazy.t val jmeq : unit -> EConstr.constr val jmeq_refl : unit -> EConstr.constr val make_eq : unit -> EConstr.constr (* [with_full_print f a] applies [f] to [a] in full printing environment. This function preserves the print settings *) val with_full_print : ('a -> 'b) -> 'a -> 'b (*****************) type function_info = { function_constant : Constant.t; graph_ind : inductive; equation_lemma : Constant.t option; correctness_lemma : Constant.t option; completeness_lemma : Constant.t option; rect_lemma : Constant.t option; rec_lemma : Constant.t option; prop_lemma : Constant.t option; sprop_lemma : Constant.t option; is_general : bool; } val find_Function_infos : Constant.t -> function_info option val find_Function_of_graph : inductive -> function_info option (* WARNING: To be used just after the graph definition !!! *) val add_Function : bool -> Constant.t -> unit val update_Function : function_info -> unit (** debugging *) val pr_info : Environ.env -> Evd.evar_map -> function_info -> Pp.t val pr_table : Environ.env -> Evd.evar_map -> Pp.t val observe_tac : (Environ.env -> Evd.evar_map -> Pp.t) -> Tacmach.tactic -> Tacmach.tactic module New : sig val observe_tac : header:Pp.t -> (Environ.env -> Evd.evar_map -> Pp.t) -> unit Proofview.tactic -> unit Proofview.tactic end (* val function_debug : bool ref *) val observe : Pp.t -> unit val do_observe : unit -> bool val do_rewrite_dependent : unit -> bool (* To localize pb *) exception Building_graph of exn exception Defining_principle of exn exception ToShow of exn val is_strict_tcc : unit -> bool val h_intros: Names.Id.t list -> Tacmach.tactic val h_id : Names.Id.t val hrec_id : Names.Id.t val acc_inv_id : EConstr.constr Util.delayed val ltof_ref : GlobRef.t Util.delayed val well_founded_ltof : EConstr.constr Util.delayed val acc_rel : EConstr.constr Util.delayed val well_founded : EConstr.constr Util.delayed val evaluable_of_global_reference : GlobRef.t -> Names.evaluable_global_reference val list_rewrite : bool -> (EConstr.constr*bool) list -> Tacmach.tactic val decompose_lam_n : Evd.evar_map -> int -> EConstr.t -> (Names.Name.t Context.binder_annot * EConstr.t) list * EConstr.t val compose_lam : (Names.Name.t Context.binder_annot * EConstr.t) list -> EConstr.t -> EConstr.t val compose_prod : (Names.Name.t Context.binder_annot * EConstr.t) list -> EConstr.t -> EConstr.t type tcc_lemma_value = | Undefined | Value of Constr.t | Not_needed val funind_purify : ('a -> 'b) -> ('a -> 'b) coq-8.11.0/plugins/funind/invfun.ml0000644000175000017500000001567613612664533017110 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let sigma = project gl in let typ = pf_unsafe_type_of gl (mkVar hid) in match EConstr.kind sigma typ with | App(i,args) when isInd sigma i -> let ((kn',num) as ind'),u = destInd sigma i in if MutInd.equal kn kn' then (* We have generated a graph hypothesis so that we must change it if we can *) let info = match find_Function_of_graph ind' with | Some info -> info | None -> (* The graphs are mutually recursive but we cannot find one of them !*) CErrors.anomaly (Pp.str "Cannot retrieve infos about a mutual block.") in (* if we can find a completeness lemma for this function then we can come back to the functional form. If not, we do nothing *) match info.completeness_lemma with | None -> tclIDTAC | Some f_complete -> let f_args,res = Array.chop (Array.length args - 1) args in tclTHENLIST [ generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])] ; clear [hid] ; Simple.intro hid ; post_tac hid ] else tclIDTAC | _ -> tclIDTAC ) (* [functional_inversion hid fconst f_correct ] is the functional version of [inversion] [hid] is the hypothesis to invert, [fconst] is the function to invert and [f_correct] is the correctness lemma for [fconst]. The sketch is the following~: \begin{enumerate} \item Transforms the hypothesis [hid] such that its type is now $res\ =\ f\ t_1 \ldots t_n$ (fails if it is not possible) \item replace [hid] with $R\_f t_1 \ldots t_n res$ using [f_correct] \item apply [inversion] on [hid] \item finally in each branch, replace each hypothesis [R\_f ..] by [f ...] using [f_complete] (whenever such a lemma exists) \end{enumerate} *) let functional_inversion kn hid fconst f_correct = Proofview.Goal.enter (fun gl -> let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps gl) Id.Set.empty in let sigma = project gl in let type_of_h = pf_unsafe_type_of gl (mkVar hid) in match EConstr.kind sigma type_of_h with | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) -> let pre_tac,f_args,res = match EConstr.kind sigma args.(1),EConstr.kind sigma args.(2) with | App(f,f_args),_ when EConstr.eq_constr sigma f fconst -> ((fun hid -> intros_symmetry (Locusops.onHyp hid))),f_args,args.(2) |_,App(f,f_args) when EConstr.eq_constr sigma f fconst -> ((fun hid -> tclIDTAC),f_args,args.(1)) | _ -> (fun hid -> tclFAIL 1 Pp.(mt ())),[||],args.(2) in tclTHENLIST [ pre_tac hid ; generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])] ; clear [hid] ; Simple.intro hid ; Inv.inv Inv.FullInversion None (Tactypes.NamedHyp hid) ; Proofview.Goal.enter (fun gl -> let new_ids = List.filter (fun id -> not (Id.Set.mem id old_ids)) (pf_ids_of_hyps gl) in tclMAP (revert_graph kn pre_tac) (hid::new_ids) ) ] | _ -> tclFAIL 1 Pp.(mt ()) ) let invfun qhyp f = let f = match f with | GlobRef.ConstRef f -> f | _ -> CErrors.user_err Pp.(str "Not a function") in match find_Function_infos f with | None -> CErrors.user_err (Pp.str "No graph found") | Some finfos -> match finfos.correctness_lemma with | None -> CErrors.user_err (Pp.str "Cannot use equivalence with graph!") | Some f_correct -> let f_correct = mkConst f_correct and kn = fst finfos.graph_ind in Tactics.try_intros_until (fun hid -> functional_inversion kn hid (mkConst f) f_correct) qhyp let invfun qhyp f = let exception NoFunction in match f with | Some f -> invfun qhyp f | None -> let tac_action hid gl = let sigma = project gl in let hyp_typ = pf_unsafe_type_of gl (mkVar hid) in match EConstr.kind sigma hyp_typ with | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) -> begin let f1,_ = decompose_app sigma args.(1) in try if not (isConst sigma f1) then raise NoFunction; let finfos = Option.get (find_Function_infos (fst (destConst sigma f1))) in let f_correct = mkConst(Option.get finfos.correctness_lemma) and kn = fst finfos.graph_ind in functional_inversion kn hid f1 f_correct with | NoFunction | Option.IsNone -> let f2,_ = decompose_app sigma args.(2) in if isConst sigma f2 then match find_Function_infos (fst (destConst sigma f2)) with | None -> if do_observe () then CErrors.user_err (Pp.str "No graph found for any side of equality") else CErrors.user_err Pp.(str "Cannot find inversion information for hypothesis " ++ Ppconstr.pr_id hid) | Some finfos -> match finfos.correctness_lemma with | None -> if do_observe () then CErrors.user_err (Pp.str "Cannot use equivalence with graph for any side of the equality") else CErrors.user_err Pp.(str "Cannot find inversion information for hypothesis " ++ Ppconstr.pr_id hid) | Some f_correct -> let f_correct = mkConst f_correct and kn = fst finfos.graph_ind in functional_inversion kn hid f2 f_correct else (* NoFunction *) CErrors.user_err Pp.(str "Hypothesis " ++ Ppconstr.pr_id hid ++ str " must contain at least one Function") end | _ -> CErrors.user_err Pp.(Ppconstr.pr_id hid ++ str " must be an equality ") in try_intros_until (tac_action %> Proofview.Goal.enter) qhyp coq-8.11.0/plugins/funind/FunInd.v0000644000175000017500000000135213612664533016605 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt | _ -> rt,List.rev acc in decompose_rapp [] (* [glob_make_eq t1 t2] build the glob_constr corresponding to [t2 = t1] *) let glob_make_eq ?(typ= mkGHole ()) t1 t2 = mkGApp(mkGRef (Coqlib.lib_ref "core.eq.type"),[typ;t2;t1]) (* [glob_make_neq t1 t2] build the glob_constr corresponding to [t1 <> t2] *) let glob_make_neq t1 t2 = mkGApp(mkGRef (Coqlib.lib_ref "core.not.type"),[glob_make_eq t1 t2]) let remove_name_from_mapping mapping na = match na with | Anonymous -> mapping | Name id -> Id.Map.remove id mapping let change_vars = let rec change_vars mapping rt = DAst.map_with_loc (fun ?loc -> function | GRef _ as x -> x | GVar id -> let new_id = try Id.Map.find id mapping with Not_found -> id in GVar(new_id) | GEvar _ as x -> x | GPatVar _ as x -> x | GApp(rt',rtl) -> GApp(change_vars mapping rt', List.map (change_vars mapping) rtl ) | GLambda(name,k,t,b) -> GLambda(name, k, change_vars mapping t, change_vars (remove_name_from_mapping mapping name) b ) | GProd(name,k,t,b) -> GProd( name, k, change_vars mapping t, change_vars (remove_name_from_mapping mapping name) b ) | GLetIn(name,def,typ,b) -> GLetIn(name, change_vars mapping def, Option.map (change_vars mapping) typ, change_vars (remove_name_from_mapping mapping name) b ) | GLetTuple(nal,(na,rto),b,e) -> let new_mapping = List.fold_left remove_name_from_mapping mapping nal in GLetTuple(nal, (na, Option.map (change_vars mapping) rto), change_vars mapping b, change_vars new_mapping e ) | GCases(sty,infos,el,brl) -> GCases(sty, infos, List.map (fun (e,x) -> (change_vars mapping e,x)) el, List.map (change_vars_br mapping) brl ) | GIf(b,(na,e_option),lhs,rhs) -> GIf(change_vars mapping b, (na,Option.map (change_vars mapping) e_option), change_vars mapping lhs, change_vars mapping rhs ) | GRec _ -> user_err ?loc Pp.(str "Local (co)fixes are not supported") | GSort _ as x -> x | GHole _ as x -> x | GInt _ as x -> x | GFloat _ as x -> x | GCast(b,c) -> GCast(change_vars mapping b, Glob_ops.map_cast_type (change_vars mapping) c) ) rt and change_vars_br mapping ({CAst.loc;v=(idl,patl,res)} as br) = let new_mapping = List.fold_right Id.Map.remove idl mapping in if Id.Map.is_empty new_mapping then br else CAst.make ?loc (idl,patl,change_vars new_mapping res) in change_vars let rec alpha_pat excluded pat = let loc = pat.CAst.loc in match DAst.get pat with | PatVar Anonymous -> let new_id = Indfun_common.fresh_id excluded "_x" in (DAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty | PatVar(Name id) -> if Id.List.mem id excluded then let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in (DAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded), (Id.Map.add id new_id Id.Map.empty) else pat, excluded,Id.Map.empty | PatCstr(constr,patl,na) -> let new_na,new_excluded,map = match na with | Name id when Id.List.mem id excluded -> let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in Name new_id,new_id::excluded, Id.Map.add id new_id Id.Map.empty | _ -> na,excluded,Id.Map.empty in let new_patl,new_excluded,new_map = List.fold_left (fun (patl,excluded,map) pat -> let new_pat,new_excluded,new_map = alpha_pat excluded pat in (new_pat::patl,new_excluded,Id.Map.fold Id.Map.add new_map map) ) ([],new_excluded,map) patl in (DAst.make ?loc @@ PatCstr(constr,List.rev new_patl,new_na)),new_excluded,new_map let alpha_patl excluded patl = let patl,new_excluded,map = List.fold_left (fun (patl,excluded,map) pat -> let new_pat,new_excluded,new_map = alpha_pat excluded pat in new_pat::patl,new_excluded,(Id.Map.fold Id.Map.add new_map map) ) ([],excluded,Id.Map.empty) patl in (List.rev patl,new_excluded,map) let raw_get_pattern_id pat acc = let rec get_pattern_id pat = match DAst.get pat with | PatVar(Anonymous) -> assert false | PatVar(Name id) -> [id] | PatCstr(constr,patternl,_) -> List.fold_right (fun pat idl -> let idl' = get_pattern_id pat in idl'@idl ) patternl [] in (get_pattern_id pat)@acc let get_pattern_id pat = raw_get_pattern_id pat [] let rec alpha_rt excluded rt = let loc = rt.CAst.loc in let new_rt = DAst.make ?loc @@ match DAst.get rt with | GRef _ | GVar _ | GEvar _ | GPatVar _ as rt -> rt | GLambda(Anonymous,k,t,b) -> let new_id = Namegen.next_ident_away (Id.of_string "_x") (Id.Set.of_list excluded) in let new_excluded = new_id :: excluded in let new_t = alpha_rt new_excluded t in let new_b = alpha_rt new_excluded b in GLambda(Name new_id,k,new_t,new_b) | GProd(Anonymous,k,t,b) -> let new_t = alpha_rt excluded t in let new_b = alpha_rt excluded b in GProd(Anonymous,k,new_t,new_b) | GLetIn(Anonymous,b,t,c) -> let new_b = alpha_rt excluded b in let new_t = Option.map (alpha_rt excluded) t in let new_c = alpha_rt excluded c in GLetIn(Anonymous,new_b,new_t,new_c) | GLambda(Name id,k,t,b) -> let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in let t,b = if Id.equal new_id id then t, b else let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in (t,replace b) in let new_excluded = new_id::excluded in let new_t = alpha_rt new_excluded t in let new_b = alpha_rt new_excluded b in GLambda(Name new_id,k,new_t,new_b) | GProd(Name id,k,t,b) -> let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in let new_excluded = new_id::excluded in let t,b = if Id.equal new_id id then t,b else let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in (t,replace b) in let new_t = alpha_rt new_excluded t in let new_b = alpha_rt new_excluded b in GProd(Name new_id,k,new_t,new_b) | GLetIn(Name id,b,t,c) -> let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in let c = if Id.equal new_id id then c else change_vars (Id.Map.add id new_id Id.Map.empty) c in let new_excluded = new_id::excluded in let new_b = alpha_rt new_excluded b in let new_t = Option.map (alpha_rt new_excluded) t in let new_c = alpha_rt new_excluded c in GLetIn(Name new_id,new_b,new_t,new_c) | GLetTuple(nal,(na,rto),t,b) -> let rev_new_nal,new_excluded,mapping = List.fold_left (fun (nal,excluded,mapping) na -> match na with | Anonymous -> (na::nal,excluded,mapping) | Name id -> let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in if Id.equal new_id id then na::nal,id::excluded,mapping else (Name new_id)::nal,id::excluded,(Id.Map.add id new_id mapping) ) ([],excluded,Id.Map.empty) nal in let new_nal = List.rev rev_new_nal in let new_rto,new_t,new_b = if Id.Map.is_empty mapping then rto,t,b else let replace = change_vars mapping in (Option.map replace rto, t,replace b) in let new_t = alpha_rt new_excluded new_t in let new_b = alpha_rt new_excluded new_b in let new_rto = Option.map (alpha_rt new_excluded) new_rto in GLetTuple(new_nal,(na,new_rto),new_t,new_b) | GCases(sty,infos,el,brl) -> let new_el = List.map (function (rt,i) -> alpha_rt excluded rt, i) el in GCases(sty,infos,new_el,List.map (alpha_br excluded) brl) | GIf(b,(na,e_o),lhs,rhs) -> GIf(alpha_rt excluded b, (na,Option.map (alpha_rt excluded) e_o), alpha_rt excluded lhs, alpha_rt excluded rhs ) | GRec _ -> user_err Pp.(str "Not handled GRec") | GSort _ | GInt _ | GFloat _ | GHole _ as rt -> rt | GCast (b,c) -> GCast(alpha_rt excluded b, Glob_ops.map_cast_type (alpha_rt excluded) c) | GApp(f,args) -> GApp(alpha_rt excluded f, List.map (alpha_rt excluded) args ) in new_rt and alpha_br excluded {CAst.loc;v=(ids,patl,res)} = let new_patl,new_excluded,mapping = alpha_patl excluded patl in let new_ids = List.fold_right raw_get_pattern_id new_patl [] in let new_excluded = new_ids@excluded in let renamed_res = change_vars mapping res in let new_res = alpha_rt new_excluded renamed_res in CAst.make ?loc (new_ids,new_patl,new_res) (* [is_free_in id rt] checks if [id] is a free variable in [rt] *) let is_free_in id = let rec is_free_in x = DAst.with_loc_val (fun ?loc -> function | GRef _ -> false | GVar id' -> Id.compare id' id == 0 | GEvar _ -> false | GPatVar _ -> false | GApp(rt,rtl) -> List.exists is_free_in (rt::rtl) | GLambda(n,_,t,b) | GProd(n,_,t,b) -> let check_in_b = match n with | Name id' -> not (Id.equal id' id) | _ -> true in is_free_in t || (check_in_b && is_free_in b) | GLetIn(n,b,t,c) -> let check_in_c = match n with | Name id' -> not (Id.equal id' id) | _ -> true in is_free_in b || Option.cata is_free_in true t || (check_in_c && is_free_in c) | GCases(_,_,el,brl) -> (List.exists (fun (e,_) -> is_free_in e) el) || List.exists is_free_in_br brl | GLetTuple(nal,_,b,t) -> let check_in_nal = not (List.exists (function Name id' -> Id.equal id' id | _ -> false) nal) in is_free_in t || (check_in_nal && is_free_in b) | GIf(cond,_,br1,br2) -> is_free_in cond || is_free_in br1 || is_free_in br2 | GRec _ -> user_err Pp.(str "Not handled GRec") | GSort _ -> false | GHole _ -> false | GCast (b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t | GCast (b,CastCoerce) -> is_free_in b | GInt _ | GFloat _ -> false ) x and is_free_in_br {CAst.v=(ids,_,rt)} = (not (Id.List.mem id ids)) && is_free_in rt in is_free_in let rec pattern_to_term pt = DAst.with_val (function | PatVar Anonymous -> assert false | PatVar(Name id) -> mkGVar id | PatCstr(constr,patternl,_) -> let cst_narg = Inductiveops.constructor_nallargs (Global.env ()) constr in let implicit_args = Array.to_list (Array.init (cst_narg - List.length patternl) (fun _ -> mkGHole ()) ) in let patl_as_term = List.map pattern_to_term patternl in mkGApp(mkGRef(GlobRef.ConstructRef constr), implicit_args@patl_as_term ) ) pt let replace_var_by_term x_id term = let rec replace_var_by_pattern x = DAst.map (function | GVar id when Id.compare id x_id == 0 -> DAst.get term | GRef _ | GVar _ | GEvar _ | GPatVar _ as rt -> rt | GApp(rt',rtl) -> GApp(replace_var_by_pattern rt', List.map replace_var_by_pattern rtl ) | GLambda(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GLambda(name,k,t,b) -> GLambda(name, k, replace_var_by_pattern t, replace_var_by_pattern b ) | GProd(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GProd(name,k,t,b) -> GProd( name, k, replace_var_by_pattern t, replace_var_by_pattern b ) | GLetIn(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GLetIn(name,def,typ,b) -> GLetIn(name, replace_var_by_pattern def, Option.map (replace_var_by_pattern) typ, replace_var_by_pattern b ) | GLetTuple(nal,_,_,_) as rt when List.exists (function Name id -> Id.equal id x_id | _ -> false) nal -> rt | GLetTuple(nal,(na,rto),def,b) -> GLetTuple(nal, (na,Option.map replace_var_by_pattern rto), replace_var_by_pattern def, replace_var_by_pattern b ) | GCases(sty,infos,el,brl) -> GCases(sty, infos, List.map (fun (e,x) -> (replace_var_by_pattern e,x)) el, List.map replace_var_by_pattern_br brl ) | GIf(b,(na,e_option),lhs,rhs) -> GIf(replace_var_by_pattern b, (na,Option.map replace_var_by_pattern e_option), replace_var_by_pattern lhs, replace_var_by_pattern rhs ) | GRec _ -> CErrors.user_err (Pp.str "Not handled GRec") | GSort _ | GHole _ as rt -> rt | GInt _ as rt -> rt | GFloat _ as rt -> rt | GCast(b,c) -> GCast(replace_var_by_pattern b, Glob_ops.map_cast_type replace_var_by_pattern c) ) x and replace_var_by_pattern_br ({CAst.loc;v=(idl,patl,res)} as br) = if List.exists (fun id -> Id.compare id x_id == 0) idl then br else CAst.make ?loc (idl,patl,replace_var_by_pattern res) in replace_var_by_pattern (* checking unifiability of patterns *) exception NotUnifiable let rec are_unifiable_aux = function | [] -> () | (l, r) ::eqs -> match DAst.get l, DAst.get r with | PatVar _ ,_ | _, PatVar _-> are_unifiable_aux eqs | PatCstr(constructor1,cpl1,_), PatCstr(constructor2,cpl2,_) -> if not (eq_constructor constructor2 constructor1) then raise NotUnifiable else let eqs' = try (List.combine cpl1 cpl2) @ eqs with Invalid_argument _ -> anomaly (Pp.str "are_unifiable_aux.") in are_unifiable_aux eqs' let are_unifiable pat1 pat2 = try are_unifiable_aux [pat1,pat2]; true with NotUnifiable -> false let rec eq_cases_pattern_aux = function | [] -> () | (l, r) ::eqs -> match DAst.get l, DAst.get r with | PatVar _, PatVar _ -> eq_cases_pattern_aux eqs | PatCstr(constructor1,cpl1,_), PatCstr(constructor2,cpl2,_) -> if not (eq_constructor constructor2 constructor1) then raise NotUnifiable else let eqs' = try (List.combine cpl1 cpl2) @ eqs with Invalid_argument _ -> anomaly (Pp.str "eq_cases_pattern_aux.") in eq_cases_pattern_aux eqs' | _ -> raise NotUnifiable let eq_cases_pattern pat1 pat2 = try eq_cases_pattern_aux [pat1,pat2]; true with NotUnifiable -> false let ids_of_pat = let rec ids_of_pat ids = DAst.with_val (function | PatVar Anonymous -> ids | PatVar(Name id) -> Id.Set.add id ids | PatCstr(_,patl,_) -> List.fold_left ids_of_pat ids patl ) in ids_of_pat Id.Set.empty let expand_as = let rec add_as map rt = match DAst.get rt with | PatVar _ -> map | PatCstr(_,patl,Name id) -> Id.Map.add id (pattern_to_term rt) (List.fold_left add_as map patl) | PatCstr(_,patl,_) -> List.fold_left add_as map patl in let rec expand_as map = DAst.map (function | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ as rt -> rt | GVar id as rt -> begin try DAst.get (Id.Map.find id map) with Not_found -> rt end | GApp(f,args) -> GApp(expand_as map f,List.map (expand_as map) args) | GLambda(na,k,t,b) -> GLambda(na,k,expand_as map t, expand_as map b) | GProd(na,k,t,b) -> GProd(na,k,expand_as map t, expand_as map b) | GLetIn(na,v,typ,b) -> GLetIn(na, expand_as map v,Option.map (expand_as map) typ,expand_as map b) | GLetTuple(nal,(na,po),v,b) -> GLetTuple(nal,(na,Option.map (expand_as map) po), expand_as map v, expand_as map b) | GIf(e,(na,po),br1,br2) -> GIf(expand_as map e,(na,Option.map (expand_as map) po), expand_as map br1, expand_as map br2) | GRec _ -> user_err Pp.(str "Not handled GRec") | GCast(b,c) -> GCast(expand_as map b, Glob_ops.map_cast_type (expand_as map) c) | GCases(sty,po,el,brl) -> GCases(sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el, List.map (expand_as_br map) brl) ) and expand_as_br map {CAst.loc; v=(idl,cpl,rt)} = CAst.make ?loc (idl,cpl, expand_as (List.fold_left add_as map cpl) rt) in expand_as Id.Map.empty (* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution *) exception Found of Evd.evar_info let resolve_and_replace_implicits ?(flags=Pretyping.all_and_fail_flags) ?(expected_type=Pretyping.WithoutTypeConstraint) env sigma rt = let open Evd in let open Evar_kinds in (* we first (pseudo) understand [rt] and get back the computed evar_map *) (* FIXME : JF (30/03/2017) I'm not completely sure to have split understand as needed. If someone knows how to prevent solved existantial removal in understand, please do not hesitate to change the computation of [ctx] here *) let ctx,_,_ = Pretyping.ise_pretype_gen flags env sigma Glob_ops.empty_lvar expected_type rt in let ctx = Evd.minimize_universes ctx in let f c = EConstr.of_constr (Evarutil.nf_evars_universes ctx (EConstr.Unsafe.to_constr c)) in (* then we map [rt] to replace the implicit holes by their values *) let rec change rt = match DAst.get rt with | GHole(ImplicitArg(grk,pk,bk),_,_) -> (* we only want to deal with implicit arguments *) ( try (* we scan the new evar map to find the evar corresponding to this hole (by looking the source *) Evd.fold (* to simulate an iter *) (fun _ evi _ -> match evi.evar_source with | (loc_evi,ImplicitArg(gr_evi,p_evi,b_evi)) -> if GlobRef.equal grk gr_evi && pk=p_evi && bk=b_evi && rt.CAst.loc = loc_evi then raise (Found evi) | _ -> () ) ctx (); (* the hole was not solved : we do nothing *) rt with Found evi -> (* we found the evar corresponding to this hole *) match evi.evar_body with | Evar_defined c -> (* we just have to lift the solution in glob_term *) Detyping.detype Detyping.Now false Id.Set.empty env ctx (f c) | Evar_empty -> rt (* the hole was not solved : we do nothing *) ) | (GHole(BinderType na,_,_)) -> (* we only want to deal with implicit arguments *) ( let res = try (* we scan the new evar map to find the evar corresponding to this hole (by looking the source *) Evd.fold (* to simulate an iter *) (fun _ evi _ -> match evi.evar_source with | (loc_evi,BinderType na') -> if Name.equal na na' && rt.CAst.loc = loc_evi then raise (Found evi) | _ -> () ) ctx (); (* the hole was not solved : we do nothing *) rt with Found evi -> (* we found the evar corresponding to this hole *) match evi.evar_body with | Evar_defined c -> (* we just have to lift the solution in glob_term *) Detyping.detype Detyping.Now false Id.Set.empty env ctx (f c) | Evar_empty -> rt (* the hole was not solved : we d when falseo nothing *) in res ) | _ -> Glob_ops.map_glob_constr change rt in change rt coq-8.11.0/plugins/funind/Recdef.v0000644000175000017500000000316013612664533016611 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* A) -> A -> A := fun (fl : A -> A) (def : A) => match n with | O => def | S m => fl (iter m fl def) end. End Iter. Theorem le_lt_SS x y : x <= y -> x < S (S y). Proof. intros. now apply Nat.lt_succ_r, Nat.le_le_succ_r. Qed. Theorem Splus_lt x y : y < S (x + y). Proof. apply Nat.lt_succ_r. rewrite Nat.add_comm. apply Nat.le_add_r. Qed. Theorem SSplus_lt x y : x < S (S (x + y)). Proof. apply le_lt_SS, Nat.le_add_r. Qed. Inductive max_type (m n:nat) : Set := cmt : forall v, m <= v -> n <= v -> max_type m n. Definition max m n : max_type m n. Proof. destruct (Compare_dec.le_gt_dec m n) as [h|h]. - exists n; [exact h | apply le_n]. - exists m; [apply le_n | apply Nat.lt_le_incl; exact h]. Defined. Definition Acc_intro_generator_function := fun A R => @Acc_intro_generator A R 100. coq-8.11.0/plugins/funind/indfun.ml0000644000175000017500000001332413612664533017052 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* i >= min && i< max) free_rels_in_br ) in List.fold_left_i test_branche 1 false (List.rev scheme_info.Tactics.branches) let choose_dest_or_ind scheme_info args = Proofview.tclBIND Proofview.tclEVARMAP (fun sigma -> Tactics.induction_destruct (is_rec_info sigma scheme_info) false args) let functional_induction with_clean c princl pat = let open Proofview.Notations in Proofview.Goal.enter_one (fun gl -> let sigma = project gl in let f,args = decompose_app sigma c in match princl with | None -> (* No principle is given let's find the good one *) begin match EConstr.kind sigma f with | Const (c',u) -> let princ_option = let finfo = (* we first try to find out a graph on f *) match find_Function_infos c' with | Some finfo -> finfo | None -> user_err (str "Cannot find induction information on "++ Printer.pr_leconstr_env (pf_env gl) sigma (mkConst c') ) in match elimination_sort_of_goal gl with | InSProp -> finfo.sprop_lemma | InProp -> finfo.prop_lemma | InSet -> finfo.rec_lemma | InType -> finfo.rect_lemma in let princ = (* then we get the principle *) match princ_option with | Some princ -> let sigma, princ = Evd.fresh_global (pf_env gl) (project gl) (GlobRef.ConstRef princ) in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT princ | None -> (*i If there is not default lemma defined then, we cross our finger and try to find a lemma named f_ind (or f_rec, f_rect) i*) let princ_name = Indrec.make_elimination_ident (Label.to_id (Constant.label c')) (elimination_sort_of_goal gl) in let princ_ref = try Constrintern.locate_reference (Libnames.qualid_of_ident princ_name) with | Not_found -> user_err (str "Cannot find induction principle for " ++ Printer.pr_leconstr_env (pf_env gl) sigma (mkConst c') ) in let sigma, princ = Evd.fresh_global (pf_env gl) (project gl) princ_ref in Proofview.Unsafe.tclEVARS sigma >>= fun () -> Proofview.tclUNIT princ in princ >>= fun princ -> (* We need to refresh gl due to the updated evar_map in princ *) Proofview.Goal.enter_one (fun gl -> Proofview.tclUNIT (princ, Tactypes.NoBindings, pf_unsafe_type_of gl princ, args)) | _ -> CErrors.user_err (str "functional induction must be used with a function" ) end | Some ((princ,binding)) -> Proofview.tclUNIT (princ, binding, pf_unsafe_type_of gl princ, args) ) >>= fun (princ, bindings, princ_type, args) -> Proofview.Goal.enter (fun gl -> let sigma = project gl in let princ_infos = compute_elim_sig (project gl) princ_type in let args_as_induction_constr = let c_list = if princ_infos.Tactics.farg_in_concl then [c] else [] in if List.length args + List.length c_list = 0 then user_err Pp.(str "Cannot recognize a valid functional scheme" ); let encoded_pat_as_patlist = List.make (List.length args + List.length c_list - 1) None @ [pat] in List.map2 (fun c pat -> ((None, ElimOnConstr (fun env sigma -> (sigma,(c,Tactypes.NoBindings)))), (None,pat), None)) (args@c_list) encoded_pat_as_patlist in let princ' = Some (princ,bindings) in let princ_vars = List.fold_right (fun a acc -> try Id.Set.add (destVar sigma a) acc with DestKO -> acc) args Id.Set.empty in let old_idl = List.fold_right Id.Set.add (pf_ids_of_hyps gl) Id.Set.empty in let old_idl = Id.Set.diff old_idl princ_vars in let subst_and_reduce gl = if with_clean then let idl = List.filter (fun id -> not (Id.Set.mem id old_idl))(pf_ids_of_hyps gl) in let flag = Genredexpr.Cbv { Redops.all_flags with Genredexpr.rDelta = false } in tclTHEN (tclMAP (fun id -> tclTRY (Equality.subst_gen (do_rewrite_dependent ()) [id])) idl) (reduce flag Locusops.allHypsAndConcl) else tclIDTAC in tclTHEN (choose_dest_or_ind princ_infos (args_as_induction_constr,princ')) (Proofview.Goal.enter subst_and_reduce)) coq-8.11.0/plugins/funind/glob_term_to_relation.ml0000644000175000017500000017631313612664533022150 0ustar treinentreinenopen Printer open Pp open Names open Constr open Context open Vars open Glob_term open Glob_ops open Indfun_common open CErrors open Util open Glob_termops module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration let observe strm = if do_observe () then Feedback.msg_debug strm else () (*let observennl strm = if do_observe () then Pp.msg strm else ()*) type binder_type = | Lambda of Name.t | Prod of Name.t | LetIn of Name.t type glob_context = (binder_type*glob_constr) list let rec solve_trivial_holes pat_as_term e = match DAst.get pat_as_term, DAst.get e with | GHole _,_ -> e | GApp(fp,argsp),GApp(fe,argse) when glob_constr_eq fp fe -> DAst.make (GApp((solve_trivial_holes fp fe),List.map2 solve_trivial_holes argsp argse)) | _,_ -> pat_as_term (* compose_glob_context [(bt_1,n_1,t_1);......] rt returns b_1(n_1,t_1,.....,bn(n_k,t_k,rt)) where the b_i's are the binders corresponding to the bt_i's *) let compose_glob_context = let compose_binder (bt,t) acc = match bt with | Lambda n -> mkGLambda(n,t,acc) | Prod n -> mkGProd(n,t,acc) | LetIn n -> mkGLetIn(n,t,None,acc) in List.fold_right compose_binder (* The main part deals with building a list of globalized constructor expressions from the rhs of a fixpoint equation. *) type 'a build_entry_pre_return = { context : glob_context; (* the binding context of the result *) value : 'a; (* The value *) } type 'a build_entry_return = { result : 'a build_entry_pre_return list; to_avoid : Id.t list } (* [combine_results combine_fun res1 res2] combine two results [res1] and [res2] w.r.t. [combine_fun]. Informally, both [res1] and [res2] are lists of "constructors" [res1_1;...] and [res2_1,....] and we need to produce [combine_fun res1_1 res2_1;combine_fun res1_1 res2_2;........] *) let combine_results (combine_fun : 'a build_entry_pre_return -> 'b build_entry_pre_return -> 'c build_entry_pre_return ) (res1: 'a build_entry_return) (res2 : 'b build_entry_return) : 'c build_entry_return = let pre_result = List.map ( fun res1 -> (* for each result in arg_res *) List.map (* we add it in each args_res *) (fun res2 -> combine_fun res1 res2 ) res2.result ) res1.result in (* and then we flatten the map *) { result = List.concat pre_result; to_avoid = List.union Id.equal res1.to_avoid res2.to_avoid } (* The combination function for an argument with a list of argument *) let combine_args arg args = { context = arg.context@args.context; (* Note that the binding context of [arg] MUST be placed before the one of [args] in order to preserve possible type dependencies *) value = arg.value::args.value; } let ids_of_binder = function | LetIn Anonymous | Prod Anonymous | Lambda Anonymous -> Id.Set.empty | LetIn (Name id) | Prod (Name id) | Lambda (Name id) -> Id.Set.singleton id let rec change_vars_in_binder mapping = function [] -> [] | (bt,t)::l -> let new_mapping = Id.Set.fold Id.Map.remove (ids_of_binder bt) mapping in (bt,change_vars mapping t):: (if Id.Map.is_empty new_mapping then l else change_vars_in_binder new_mapping l ) let rec replace_var_by_term_in_binder x_id term = function | [] -> [] | (bt,t)::l -> (bt,replace_var_by_term x_id term t):: if Id.Set.mem x_id (ids_of_binder bt) then l else replace_var_by_term_in_binder x_id term l let add_bt_names bt = Id.Set.union (ids_of_binder bt) let apply_args ctxt body args = let need_convert_id avoid id = List.exists (is_free_in id) args || Id.Set.mem id avoid in let need_convert avoid bt = Id.Set.exists (need_convert_id avoid) (ids_of_binder bt) in let next_name_away (na:Name.t) (mapping: Id.t Id.Map.t) (avoid: Id.Set.t) = match na with | Name id when Id.Set.mem id avoid -> let new_id = Namegen.next_ident_away id avoid in Name new_id,Id.Map.add id new_id mapping,Id.Set.add new_id avoid | _ -> na,mapping,avoid in let next_bt_away bt (avoid:Id.Set.t) = match bt with | LetIn na -> let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in LetIn new_na,mapping,new_avoid | Prod na -> let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in Prod new_na,mapping,new_avoid | Lambda na -> let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in Lambda new_na,mapping,new_avoid in let rec do_apply avoid ctxt body args = match ctxt,args with | _,[] -> (* No more args *) (ctxt,body) | [],_ -> (* no more fun *) let f,args' = glob_decompose_app body in (ctxt,mkGApp(f,args'@args)) | (Lambda Anonymous,t)::ctxt',arg::args' -> do_apply avoid ctxt' body args' | (Lambda (Name id),t)::ctxt',arg::args' -> let new_avoid,new_ctxt',new_body,new_id = if need_convert_id avoid id then let new_avoid = Id.Set.add id avoid in let new_id = Namegen.next_ident_away id new_avoid in let new_avoid' = Id.Set.add new_id new_avoid in let mapping = Id.Map.add id new_id Id.Map.empty in let new_ctxt' = change_vars_in_binder mapping ctxt' in let new_body = change_vars mapping body in new_avoid',new_ctxt',new_body,new_id else Id.Set.add id avoid,ctxt',body,id in let new_body = replace_var_by_term new_id arg new_body in let new_ctxt' = replace_var_by_term_in_binder new_id arg new_ctxt' in do_apply avoid new_ctxt' new_body args' | (bt,t)::ctxt',_ -> let new_avoid,new_ctxt',new_body,new_bt = let new_avoid = add_bt_names bt avoid in if need_convert avoid bt then let new_bt,mapping,new_avoid = next_bt_away bt new_avoid in ( new_avoid, change_vars_in_binder mapping ctxt', change_vars mapping body, new_bt ) else new_avoid,ctxt',body,bt in let new_ctxt',new_body = do_apply new_avoid new_ctxt' new_body args in (new_bt,t)::new_ctxt',new_body in do_apply Id.Set.empty ctxt body args let combine_app f args = let new_ctxt,new_value = apply_args f.context f.value args.value in { (* Note that the binding context of [args] MUST be placed before the one of the applied value in order to preserve possible type dependencies *) context = args.context@new_ctxt; value = new_value; } let combine_lam n t b = { context = []; value = mkGLambda(n, compose_glob_context t.context t.value, compose_glob_context b.context b.value ) } let combine_prod2 n t b = { context = []; value = mkGProd(n, compose_glob_context t.context t.value, compose_glob_context b.context b.value ) } let combine_prod n t b = { context = t.context@((Prod n,t.value)::b.context); value = b.value} let combine_letin n t b = { context = t.context@((LetIn n,t.value)::b.context); value = b.value} let mk_result ctxt value avoid = { result = [{context = ctxt; value = value}] ; to_avoid = avoid } (************************************************* Some functions to deal with overlapping patterns **************************************************) let coq_True_ref = lazy (Coqlib.lib_ref "core.True.type") let coq_False_ref = lazy (Coqlib.lib_ref "core.False.type") (* [make_discr_match_el \[e1,...en\]] builds match e1,...,en with (the list of expressions on which we will do the matching) *) let make_discr_match_el = List.map (fun e -> (e,(Anonymous,None))) (* [make_discr_match_brl i \[pat_1,...,pat_n\]] constructs a discrimination pattern matching on the ith expression. that is. match ?????? with \\ | pat_1 => False \\ | pat_{i-1} => False \\ | pat_i => True \\ | pat_{i+1} => False \\ \vdots | pat_n => False end *) let make_discr_match_brl i = List.map_i (fun j {CAst.v=(idl,patl,_)} -> CAst.make @@ if Int.equal j i then (idl,patl, mkGRef (Lazy.force coq_True_ref)) else (idl,patl, mkGRef (Lazy.force coq_False_ref)) ) 0 (* [make_discr_match brl el i] generates an hypothesis such that it reduce to true iff brl_{i} is the first branch matched by [el] Used when we want to simulate the coq pattern matching algorithm *) let make_discr_match brl = fun el i -> mkGCases(None, make_discr_match_el el, make_discr_match_brl i brl) (**********************************************************************) (* functions used to build case expression from lettuple and if ones *) (**********************************************************************) (* [build_constructors_of_type] construct the array of pattern of its inductive argument*) let build_constructors_of_type ind' argl = let (mib,ind) = Inductive.lookup_mind_specif (Global.env()) ind' in let npar = mib.Declarations.mind_nparams in Array.mapi (fun i _ -> let construct = ind',i+1 in let constructref = GlobRef.ConstructRef(construct) in let _implicit_positions_of_cst = Impargs.implicits_of_global constructref in let cst_narg = Inductiveops.constructor_nallargs (Global.env ()) construct in let argl = if List.is_empty argl then List.make cst_narg (mkGHole ()) else List.make npar (mkGHole ()) @ argl in let pat_as_term = mkGApp(mkGRef (GlobRef.ConstructRef(ind',i+1)),argl) in cases_pattern_of_glob_constr (Global.env()) Anonymous pat_as_term ) ind.Declarations.mind_consnames (******************) (* Main functions *) (******************) let raw_push_named (na,raw_value,raw_typ) env = match na with | Anonymous -> env | Name id -> let typ,_ = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in let na = make_annot id Sorts.Relevant in (* TODO relevance *) (match raw_value with | None -> EConstr.push_named (NamedDecl.LocalAssum (na,typ)) env | Some value -> EConstr.push_named (NamedDecl.LocalDef (na, value, typ)) env) let add_pat_variables sigma pat typ env : Environ.env = let rec add_pat_variables env pat typ : Environ.env = observe (str "new rel env := " ++ Printer.pr_rel_context_of env (Evd.from_env env)); match DAst.get pat with | PatVar na -> Environ.push_rel (RelDecl.LocalAssum (make_annot na Sorts.Relevant,typ)) env | PatCstr(c,patl,na) -> let Inductiveops.IndType(indf,indargs) = try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) with Not_found -> assert false in let constructors = Inductiveops.get_constructors env indf in let constructor : Inductiveops.constructor_summary = List.find (fun cs -> eq_constructor c (fst cs.Inductiveops.cs_cstr)) (Array.to_list constructors) in let cs_args_types :types list = List.map RelDecl.get_type constructor.Inductiveops.cs_args in List.fold_left2 add_pat_variables env patl (List.rev cs_args_types) in let new_env = add_pat_variables env pat typ in let res = fst ( Context.Rel.fold_outside (fun decl (env,ctxt) -> let open Context.Rel.Declaration in match decl with | LocalAssum ({binder_name=Anonymous},_) | LocalDef ({binder_name=Anonymous},_,_) -> assert false | LocalAssum ({binder_name=Name id} as na, t) -> let na = {na with binder_name=id} in let new_t = substl ctxt t in observe (str "for variable " ++ Ppconstr.pr_id id ++ fnl () ++ str "old type := " ++ Printer.pr_lconstr_env env sigma t ++ fnl () ++ str "new type := " ++ Printer.pr_lconstr_env env sigma new_t ++ fnl () ); let open Context.Named.Declaration in (Environ.push_named (LocalAssum (na,new_t)) env,mkVar id::ctxt) | LocalDef ({binder_name=Name id} as na, v, t) -> let na = {na with binder_name=id} in let new_t = substl ctxt t in let new_v = substl ctxt v in observe (str "for variable " ++ Ppconstr.pr_id id ++ fnl () ++ str "old type := " ++ Printer.pr_lconstr_env env sigma t ++ fnl () ++ str "new type := " ++ Printer.pr_lconstr_env env sigma new_t ++ fnl () ++ str "old value := " ++ Printer.pr_lconstr_env env sigma v ++ fnl () ++ str "new value := " ++ Printer.pr_lconstr_env env sigma new_v ++ fnl () ); let open Context.Named.Declaration in (Environ.push_named (LocalDef (na,new_v,new_t)) env,mkVar id::ctxt) ) (Environ.rel_context new_env) ~init:(env,[]) ) in observe (str "new var env := " ++ Printer.pr_named_context_of res (Evd.from_env env)); res let rec pattern_to_term_and_type env typ = DAst.with_val (function | PatVar Anonymous -> assert false | PatVar (Name id) -> mkGVar id | PatCstr(constr,patternl,_) -> let cst_narg = Inductiveops.constructor_nallargs (Global.env ()) constr in let Inductiveops.IndType(indf,indargs) = try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) with Not_found -> assert false in let constructors = Inductiveops.get_constructors env indf in let constructor = List.find (fun cs -> eq_constructor (fst cs.Inductiveops.cs_cstr) constr) (Array.to_list constructors) in let cs_args_types :types list = List.map RelDecl.get_type constructor.Inductiveops.cs_args in let _,cstl = Inductiveops.dest_ind_family indf in let csta = Array.of_list cstl in let implicit_args = Array.to_list (Array.init (cst_narg - List.length patternl) (fun i -> Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) (EConstr.of_constr csta.(i))) ) in let patl_as_term = List.map2 (pattern_to_term_and_type env) (List.rev cs_args_types) patternl in mkGApp(mkGRef(GlobRef.ConstructRef constr), implicit_args@patl_as_term ) ) (* [build_entry_lc funnames avoid rt] construct the list (in fact a build_entry_return) of constructors corresponding to [rt] when replacing calls to [funnames] by calls to the corresponding graphs. The idea to transform a term [t] into a list of constructors [lc] is the following: \begin{itemize} \item if the term is a binder (bind x, body) then first compute [lc'] the list corresponding to [body] and add (bind x. _) to each elements of [lc] \item if the term has the form (g t1 ... ... tn) where g does not appears in (fnames) then compute [lc1] ... [lcn] the lists of constructors corresponding to [t1] ... [tn], then combine those lists and [g] as follows~: for each element [c1,...,cn] of [lc1\times...\times lcn], [g c1 ... cn] is an element of [lc] \item if the term has the form (f t1 .... tn) where [f] appears in [fnames] then compute [lc1] ... [lcn] the lists of constructors corresponding to [t1] ... [tn], then compute those lists and [f] as follows~: for each element [c1,...,cn] of [lc1\times...\times lcn] create a new variable [res] and [forall res, R_f c1 ... cn res] is in [lc] \item if the term is a cast just treat its body part \item if the term is a match, an if or a lettuple then compute the lists corresponding to each branch of the case and concatenate them (informally, each branch of a match produces a new constructor) \end{itemize} WARNING: The terms constructed here are only USING the glob_constr syntax but are highly bad formed. We must wait to have complete all the current calculi to set the recursive calls. At this point, each term [f t1 ... tn] (where f appears in [funnames]) is replaced by a pseudo term [forall res, res t1 ... tn, res]. A reconstruction phase is done later. We in fact not create a constructor list since then end of each constructor has not the expected form but only the value of the function *) let rec build_entry_lc env sigma funnames avoid rt : glob_constr build_entry_return = observe (str " Entering : " ++ Printer.pr_glob_constr_env env rt); let open CAst in match DAst.get rt with | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ -> (* do nothing (except changing type of course) *) mk_result [] rt avoid | GApp(_,_) -> let f,args = glob_decompose_app rt in let args_res : (glob_constr list) build_entry_return = List.fold_right (* create the arguments lists of constructors and combine them *) (fun arg ctxt_argsl -> let arg_res = build_entry_lc env sigma funnames ctxt_argsl.to_avoid arg in combine_results combine_args arg_res ctxt_argsl ) args (mk_result [] [] avoid) in begin match DAst.get f with | GLambda _ -> let rec aux t l = match l with | [] -> t | u::l -> DAst.make @@ match DAst.get t with | GLambda(na,_,nat,b) -> GLetIn(na,u,None,aux b l) | _ -> GApp(t,l) in build_entry_lc env sigma funnames avoid (aux f args) | GVar id when Id.Set.mem id funnames -> (* if we have [f t1 ... tn] with [f]$\in$[fnames] then we create a fresh variable [res], add [res] and its "value" (i.e. [res v1 ... vn]) to each pseudo constructor build for the arguments (i.e. a pseudo context [ctxt] and a pseudo value "v1 ... vn". The "value" of this branch is then simply [res] *) let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) rt_as_constr in let res_raw_type = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) rt_typ in let res = fresh_id args_res.to_avoid "_res" in let new_avoid = res::args_res.to_avoid in let res_rt = mkGVar res in let new_result = List.map (fun arg_res -> let new_hyps = [Prod (Name res),res_raw_type; Prod Anonymous,mkGApp(res_rt,(mkGVar id)::arg_res.value)] in {context = arg_res.context@new_hyps; value = res_rt } ) args_res.result in { result = new_result; to_avoid = new_avoid } | GVar _ | GEvar _ | GPatVar _ | GHole _ | GSort _ | GRef _ -> (* if have [g t1 ... tn] with [g] not appearing in [funnames] then foreach [ctxt,v1 ... vn] in [args_res] we return [ctxt, g v1 .... vn] *) { args_res with result = List.map (fun args_res -> {args_res with value = mkGApp(f,args_res.value)}) args_res.result } | GApp _ -> assert false (* we have collected all the app in [glob_decompose_app] *) | GLetIn(n,v,t,b) -> (* if we have [(let x := v in b) t1 ... tn] , we discard our work and compute the list of constructor for [let x = v in (b t1 ... tn)] up to alpha conversion *) let new_n,new_b,new_avoid = match n with | Name id when List.exists (is_free_in id) args -> (* need to alpha-convert the name *) let new_id = Namegen.next_ident_away id (Id.Set.of_list avoid) in let new_avoid = id:: avoid in let new_b = replace_var_by_term id (DAst.make @@ GVar id) b in (Name new_id,new_b,new_avoid) | _ -> n,b,avoid in build_entry_lc env sigma funnames avoid (mkGLetIn(new_n,v,t,mkGApp(new_b,args))) | GCases _ | GIf _ | GLetTuple _ -> (* we have [(match e1, ...., en with ..... end) t1 tn] we first compute the result from the case and then combine each of them with each of args one *) let f_res = build_entry_lc env sigma funnames args_res.to_avoid f in combine_results combine_app f_res args_res | GCast(b,_) -> (* for an applied cast we just trash the cast part and restart the work. WARNING: We need to restart since [b] itself should be an application term *) build_entry_lc env sigma funnames avoid (mkGApp(b,args)) | GRec _ -> user_err Pp.(str "Not handled GRec") | GProd _ -> user_err Pp.(str "Cannot apply a type") | GInt _ -> user_err Pp.(str "Cannot apply an integer") | GFloat _ -> user_err Pp.(str "Cannot apply a float") end (* end of the application treatement *) | GLambda(n,_,t,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the type and combine the two result *) let t_res = build_entry_lc env sigma funnames avoid t in let new_n = match n with | Name _ -> n | Anonymous -> Name (Indfun_common.fresh_id [] "_x") in let new_env = raw_push_named (new_n,None,t) env in let b_res = build_entry_lc new_env sigma funnames avoid b in combine_results (combine_lam new_n) t_res b_res | GProd(n,_,t,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the type and combine the two result *) let t_res = build_entry_lc env sigma funnames avoid t in let new_env = raw_push_named (n,None,t) env in let b_res = build_entry_lc new_env sigma funnames avoid b in if List.length t_res.result = 1 && List.length b_res.result = 1 then combine_results (combine_prod2 n) t_res b_res else combine_results (combine_prod n) t_res b_res | GLetIn(n,v,typ,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the value [t] and combine the two result *) let v = match typ with None -> v | Some t -> DAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in let v_res = build_entry_lc env sigma funnames avoid v in let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in let v_r = Sorts.Relevant in (* TODO relevance *) let new_env = match n with Anonymous -> env | Name id -> EConstr.push_named (NamedDecl.LocalDef (make_annot id v_r,v_as_constr,v_type)) env in let b_res = build_entry_lc new_env sigma funnames avoid b in combine_results (combine_letin n) v_res b_res | GCases(_,_,el,brl) -> (* we create the discrimination function and treat the case itself *) let make_discr = make_discr_match brl in build_entry_lc_from_case env sigma funnames make_discr el brl avoid | GIf(b,(na,e_option),lhs,rhs) -> let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in let (ind,_) = try Inductiveops.find_inductive env (Evd.from_env env) b_typ with Not_found -> user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr_env env b ++ str " in " ++ Printer.pr_glob_constr_env env rt ++ str ". try again with a cast") in let case_pats = build_constructors_of_type (fst ind) [] in assert (Int.equal (Array.length case_pats) 2); let brl = List.map_i (fun i x -> CAst.make ([],[case_pats.(i)],x)) 0 [lhs;rhs] in let match_expr = mkGCases(None,[(b,(Anonymous,None))],brl) in (* Pp.msgnl (str "new case := " ++ Printer.pr_glob_constr match_expr); *) build_entry_lc env sigma funnames avoid match_expr | GLetTuple(nal,_,b,e) -> begin let nal_as_glob_constr = List.map (function Name id -> mkGVar id | Anonymous -> mkGHole () ) nal in let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in let (ind,_) = try Inductiveops.find_inductive env (Evd.from_env env) b_typ with Not_found -> user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr_env env b ++ str " in " ++ Printer.pr_glob_constr_env env rt ++ str ". try again with a cast") in let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in assert (Int.equal (Array.length case_pats) 1); let br = CAst.make ([],[case_pats.(0)],e) in let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in build_entry_lc env sigma funnames avoid match_expr end | GRec _ -> user_err Pp.(str "Not handled GRec") | GCast(b,_) -> build_entry_lc env sigma funnames avoid b and build_entry_lc_from_case env sigma funname make_discr (el:tomatch_tuples) (brl:Glob_term.cases_clauses) avoid : glob_constr build_entry_return = match el with | [] -> assert false (* this case correspond to match with .... !*) | el -> (* this case correspond to match el with brl end we first compute the list of lists corresponding to [el] and combine them . Then for each element of the combinations, we compute the result we compute one list per branch in [brl] and finally we just concatenate those list *) let case_resl = List.fold_right (fun (case_arg,_) ctxt_argsl -> let arg_res = build_entry_lc env sigma funname ctxt_argsl.to_avoid case_arg in combine_results combine_args arg_res ctxt_argsl ) el (mk_result [] [] avoid) in let types = List.map (fun (case_arg,_) -> let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) case_arg_as_constr) ) el in (****** The next works only if the match is not dependent ****) let results = List.map (fun ca -> let res = build_entry_lc_from_case_term env sigma types funname (make_discr) [] brl case_resl.to_avoid ca in res ) case_resl.result in { result = List.concat (List.map (fun r -> r.result) results); to_avoid = List.fold_left (fun acc r -> List.union Id.equal acc r.to_avoid) [] results } and build_entry_lc_from_case_term env sigma types funname make_discr patterns_to_prevent brl avoid matched_expr = match brl with | [] -> (* computed_branches *) {result = [];to_avoid = avoid} | br::brl' -> (* alpha conversion to prevent name clashes *) let {CAst.v=(idl,patl,return)} = alpha_br avoid br in let new_avoid = idl@avoid in (* for now we can no more use idl as an identifier *) (* building a list of precondition stating that we are not in this branch (will be used in the following recursive calls) *) let new_env = List.fold_right2 (add_pat_variables sigma) patl types env in let not_those_patterns : (Id.t list -> glob_constr -> glob_constr) list = List.map2 (fun pat typ -> fun avoid pat'_as_term -> let renamed_pat,_,_ = alpha_pat avoid pat in let pat_ids = get_pattern_id renamed_pat in let env_with_pat_ids = add_pat_variables sigma pat typ new_env in List.fold_right (fun id acc -> let typ_of_id = Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (EConstr.mkVar id) in let raw_typ_of_id = Detyping.detype Detyping.Now false Id.Set.empty env_with_pat_ids (Evd.from_env env) typ_of_id in mkGProd (Name id,raw_typ_of_id,acc)) pat_ids (glob_make_neq pat'_as_term (pattern_to_term renamed_pat)) ) patl types in (* Checking if we can be in this branch (will be used in the following recursive calls) *) let unify_with_those_patterns : (cases_pattern -> bool*bool) list = List.map (fun pat pat' -> are_unifiable pat pat',eq_cases_pattern pat pat') patl in (* we first compute the other branch result (in ordrer to keep the order of the matching as much as possible) *) let brl'_res = build_entry_lc_from_case_term env sigma types funname make_discr ((unify_with_those_patterns,not_those_patterns)::patterns_to_prevent) brl' avoid matched_expr in (* We now create the precondition of this branch i.e. 1- the list of variable appearing in the different patterns of this branch and the list of equation stating than el = patl (List.flatten ...) 2- If there exists a previous branch which pattern unify with the one of this branch then a discrimination precond stating that we are not in a previous branch (if List.exists ...) *) let those_pattern_preconds = (List.flatten ( List.map3 (fun pat e typ_as_constr -> let this_pat_ids = ids_of_pat pat in let typ_as_constr = EConstr.of_constr typ_as_constr in let typ = Detyping.detype Detyping.Now false Id.Set.empty new_env (Evd.from_env env) typ_as_constr in let pat_as_term = pattern_to_term pat in (* removing trivial holes *) let pat_as_term = solve_trivial_holes pat_as_term e in (* observe (str "those_pattern_preconds" ++ spc () ++ *) (* str "pat" ++ spc () ++ pr_glob_constr pat_as_term ++ spc ()++ *) (* str "e" ++ spc () ++ pr_glob_constr e ++spc ()++ *) (* str "typ_as_constr" ++ spc () ++ pr_lconstr typ_as_constr); *) List.fold_right (fun id acc -> if Id.Set.mem id this_pat_ids then (Prod (Name id), let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (EConstr.mkVar id) in let raw_typ_of_id = Detyping.detype Detyping.Now false Id.Set.empty new_env (Evd.from_env env) typ_of_id in raw_typ_of_id )::acc else acc ) idl [(Prod Anonymous,glob_make_eq ~typ pat_as_term e)] ) patl matched_expr.value types ) ) @ (if List.exists (function (unifl,_) -> let (unif,_) = List.split (List.map2 (fun x y -> x y) unifl patl) in List.for_all (fun x -> x) unif) patterns_to_prevent then let i = List.length patterns_to_prevent in let pats_as_constr = List.map2 (pattern_to_term_and_type new_env) types patl in [(Prod Anonymous,make_discr pats_as_constr i )] else [] ) in (* We compute the result of the value returned by the branch*) let return_res = build_entry_lc new_env sigma funname new_avoid return in (* and combine it with the preconds computed for this branch *) let this_branch_res = List.map (fun res -> { context = matched_expr.context@those_pattern_preconds@res.context ; value = res.value} ) return_res.result in { brl'_res with result = this_branch_res@brl'_res.result } let is_res r = match DAst.get r with | GVar id -> begin try String.equal (String.sub (Id.to_string id) 0 4) "_res" with Invalid_argument _ -> false end | _ -> false let is_gr c gr = match DAst.get c with | GRef (r, _) -> GlobRef.equal r gr | _ -> false let is_gvar c = match DAst.get c with | GVar id -> true | _ -> false let same_raw_term rt1 rt2 = match DAst.get rt1, DAst.get rt2 with | GRef(r1,_), GRef (r2,_) -> GlobRef.equal r1 r2 | GHole _, GHole _ -> true | _ -> false let decompose_raw_eq env lhs rhs = let rec decompose_raw_eq lhs rhs acc = observe (str "decomposing eq for " ++ pr_glob_constr_env env lhs ++ str " " ++ pr_glob_constr_env env rhs); let (rhd,lrhs) = glob_decompose_app rhs in let (lhd,llhs) = glob_decompose_app lhs in observe (str "lhd := " ++ pr_glob_constr_env env lhd); observe (str "rhd := " ++ pr_glob_constr_env env rhd); observe (str "llhs := " ++ int (List.length llhs)); observe (str "lrhs := " ++ int (List.length lrhs)); let sllhs = List.length llhs in let slrhs = List.length lrhs in if same_raw_term lhd rhd && Int.equal sllhs slrhs then (* let _ = assert false in *) List.fold_right2 decompose_raw_eq llhs lrhs acc else (lhs,rhs)::acc in decompose_raw_eq lhs rhs [] exception Continue (* The second phase which reconstruct the real type of the constructor. rebuild the globalized constructors expression. eliminates some meaningless equalities, applies some rewrites...... *) let rec rebuild_cons env nb_args relname args crossed_types depth rt = observe (str "rebuilding : " ++ pr_glob_constr_env env rt); let open Context.Rel.Declaration in let open CAst in match DAst.get rt with | GProd(n,k,t,b) -> let not_free_in_t id = not (is_free_in id t) in let new_crossed_types = t::crossed_types in begin match DAst.get t with | GApp(res_rt ,args') when is_res res_rt -> begin let arg = List.hd args' in match DAst.get arg with | GVar this_relname -> (*i The next call to mk_rel_id is valid since we are constructing the graph Ensures by: obvious i*) let new_t = mkGApp(mkGVar(mk_rel_id this_relname),List.tl args'@[res_rt]) in let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname args new_crossed_types (depth + 1) b in mkGProd(n,new_t,new_b), Id.Set.filter not_free_in_t id_to_exclude | _ -> (* the first args is the name of the function! *) assert false end | GApp(eq_as_ref,[ty; id ;rt]) when is_gvar id && is_gr eq_as_ref Coqlib.(lib_ref "core.eq.type") && n == Anonymous -> let loc1 = rt.CAst.loc in let loc2 = eq_as_ref.CAst.loc in let loc3 = id.CAst.loc in let id = match DAst.get id with GVar id -> id | _ -> assert false in begin try observe (str "computing new type for eq : " ++ pr_glob_constr_env env rt); let t' = try fst (Pretyping.understand env (Evd.from_env env) t)(*FIXME*) with e when CErrors.noncritical e -> raise Continue in let is_in_b = is_free_in id b in let _keep_eq = not (List.exists (is_free_in id) args) || is_in_b || List.exists (is_free_in id) crossed_types in let new_args = List.map (replace_var_by_term id rt) args in let subst_b = if is_in_b then b else replace_var_by_term id rt b in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname new_args new_crossed_types (depth + 1) subst_b in mkGProd(n,t,new_b),id_to_exclude with Continue -> let jmeq = GlobRef.IndRef (fst (EConstr.destInd Evd.empty (jmeq ()))) in let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in let ind,args' = Inductiveops.find_inductive env Evd.(from_env env) ty' in let mib,_ = Global.lookup_inductive (fst ind) in let nparam = mib.Declarations.mind_nparams in let params,arg' = ((Util.List.chop nparam args')) in let rt_typ = DAst.make @@ GApp(DAst.make @@ GRef (GlobRef.IndRef (fst ind),None), (List.map (fun p -> Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) (EConstr.of_constr p)) params)@(Array.to_list (Array.make (List.length args' - nparam) (mkGHole ())))) in let eq' = DAst.make ?loc:loc1 @@ GApp(DAst.make ?loc:loc2 @@GRef(jmeq,None),[ty;DAst.make ?loc:loc3 @@ GVar id;rt_typ;rt]) in observe (str "computing new type for jmeq : " ++ pr_glob_constr_env env eq'); let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in observe (str " computing new type for jmeq : done") ; let sigma = Evd.(from_env env) in let new_args = match EConstr.kind sigma eq'_as_constr with | App(_,[|_;_;ty;_|]) -> let ty = Array.to_list (snd (EConstr.destApp sigma ty)) in let ty' = snd (Util.List.chop nparam ty) in List.fold_left2 (fun acc var_as_constr arg -> if isRel var_as_constr then let na = RelDecl.get_name (Environ.lookup_rel (destRel var_as_constr) env) in match na with | Anonymous -> acc | Name id' -> (id',Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) arg)::acc else if isVar var_as_constr then (destVar var_as_constr,Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) arg)::acc else acc ) [] arg' ty' | _ -> assert false in let is_in_b = is_free_in id b in let _keep_eq = not (List.exists (is_free_in id) args) || is_in_b || List.exists (is_free_in id) crossed_types in let new_args = List.fold_left (fun args (id,rt) -> List.map (replace_var_by_term id rt) args ) args ((id,rt)::new_args) in let subst_b = if is_in_b then b else replace_var_by_term id rt b in let new_env = let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in let r = Sorts.Relevant in (* TODO relevance *) EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname new_args new_crossed_types (depth + 1) subst_b in mkGProd(n,eq',new_b),id_to_exclude end (* J.F:. keep this comment it explain how to remove some meaningless equalities if keep_eq then mkGProd(n,t,new_b),id_to_exclude else new_b, Id.Set.add id id_to_exclude *) | GApp(eq_as_ref,[ty;rt1;rt2]) when is_gr eq_as_ref Coqlib.(lib_ref "core.eq.type") && n == Anonymous -> begin try let l = decompose_raw_eq env rt1 rt2 in if List.length l > 1 then let new_rt = List.fold_left (fun acc (lhs,rhs) -> mkGProd(Anonymous, mkGApp(mkGRef Coqlib.(lib_ref "core.eq.type"),[mkGHole ();lhs;rhs]),acc) ) b l in rebuild_cons env nb_args relname args crossed_types depth new_rt else raise Continue with Continue -> observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt); let t',ctx = Pretyping.understand env (Evd.from_env env) t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname args new_crossed_types (depth + 1) b in match n with | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude end | _ -> observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt); let t',ctx = Pretyping.understand env (Evd.from_env env) t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname args new_crossed_types (depth + 1) b in match n with | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude end | GLambda(n,k,t,b) -> begin let not_free_in_t id = not (is_free_in id t) in let new_crossed_types = t :: crossed_types in observe (str "computing new type for lambda : " ++ pr_glob_constr_env env rt); let t',ctx = Pretyping.understand env (Evd.from_env env) t in match n with | Name id -> let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname (args@[mkGVar id])new_crossed_types (depth + 1 ) b in if Id.Set.mem id id_to_exclude && depth >= nb_args then new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) else DAst.make @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude | _ -> anomaly (Pp.str "Should not have an anonymous function here.") (* We have renamed all the anonymous functions during alpha_renaming phase *) end | GLetIn(n,v,t,b) -> begin let t = match t with None -> v | Some t -> DAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in let not_free_in_t id = not (is_free_in id t) in let evd = (Evd.from_env env) in let t',ctx = Pretyping.understand env evd t in let evd = Evd.from_ctx ctx in let type_t' = Typing.unsafe_type_of env evd t' in let t' = EConstr.Unsafe.to_constr t' in let type_t' = EConstr.Unsafe.to_constr type_t' in let new_env = Environ.push_rel (LocalDef (make_annot n Sorts.Relevant,t',type_t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname args (t::crossed_types) (depth + 1 ) b in match n with | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) | _ -> DAst.make @@ GLetIn(n,t,None,new_b), (* HOPING IT WOULD WORK *) Id.Set.filter not_free_in_t id_to_exclude end | GLetTuple(nal,(na,rto),t,b) -> assert (Option.is_empty rto); begin let not_free_in_t id = not (is_free_in id t) in let new_t,id_to_exclude' = rebuild_cons env nb_args relname args (crossed_types) depth t in let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot na r,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname args (t::crossed_types) (depth + 1) b in (* match n with *) (* | Name id when Id.Set.mem id id_to_exclude -> *) (* new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) *) (* | _ -> *) DAst.make @@ GLetTuple(nal,(na,None),t,new_b), Id.Set.filter not_free_in_t (Id.Set.union id_to_exclude id_to_exclude') end | _ -> mkGApp(mkGVar relname,args@[rt]),Id.Set.empty (* debugging wrapper *) let rebuild_cons env nb_args relname args crossed_types rt = (* observennl (str "rebuild_cons : rt := "++ pr_glob_constr rt ++ *) (* str "nb_args := " ++ str (string_of_int nb_args)); *) let res = rebuild_cons env nb_args relname args crossed_types 0 rt in (* observe (str " leads to "++ pr_glob_constr (fst res)); *) res (* naive implementation of parameter detection. A parameter is an argument which is only preceded by parameters and whose calls are all syntactically equal. TODO: Find a valid way to deal with implicit arguments here! *) let rec compute_cst_params relnames params gt = DAst.with_val (function | GRef _ | GVar _ | GEvar _ | GPatVar _ | GInt _ | GFloat _ -> params | GApp(f,args) -> begin match DAst.get f with | GVar relname' when Id.Set.mem relname' relnames -> compute_cst_params_from_app [] (params,args) | _ -> List.fold_left (compute_cst_params relnames) params (f::args) end | GLambda(_,_,t,b) | GProd(_,_,t,b) | GLetTuple(_,_,t,b) -> let t_params = compute_cst_params relnames params t in compute_cst_params relnames t_params b | GLetIn(_,v,t,b) -> let v_params = compute_cst_params relnames params v in let t_params = Option.fold_left (compute_cst_params relnames) v_params t in compute_cst_params relnames t_params b | GCases _ -> params (* If there is still cases at this point they can only be discrimination ones *) | GSort _ -> params | GHole _ -> params | GIf _ | GRec _ | GCast _ -> CErrors.user_err ~hdr:"compute_cst_params" (str "Not handled case") ) gt and compute_cst_params_from_app acc (params,rtl) = let is_gid id c = match DAst.get c with GVar id' -> Id.equal id id' | _ -> false in match params,rtl with | _::_,[] -> assert false (* the rel has at least nargs + 1 arguments ! *) | ((Name id,_,None) as param)::params', c::rtl' when is_gid id c -> compute_cst_params_from_app (param::acc) (params',rtl') | _ -> List.rev acc let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * glob_constr option) list array) csts = let rels_params = Array.mapi (fun i args -> List.fold_left (fun params (_,cst) -> compute_cst_params relnames params cst) args csts.(i) ) args in let l = ref [] in let _ = try List.iteri (fun i ((n,nt,typ) as param) -> if Array.for_all (fun l -> let (n',nt',typ') = List.nth l i in Name.equal n n' && glob_constr_eq nt nt' && Option.equal glob_constr_eq typ typ') rels_params then l := param::!l ) rels_params.(0) with e when CErrors.noncritical e -> () in List.rev !l let rec rebuild_return_type rt = let loc = rt.CAst.loc in match rt.CAst.v with | Constrexpr.CProdN(n,t') -> CAst.make ?loc @@ Constrexpr.CProdN(n,rebuild_return_type t') | Constrexpr.CLetIn(na,v,t,t') -> CAst.make ?loc @@ Constrexpr.CLetIn(na,v,t,rebuild_return_type t') | _ -> CAst.make ?loc @@ Constrexpr.CProdN([Constrexpr.CLocalAssum ([CAst.make Anonymous], Constrexpr.Default Explicit, rt)], CAst.make @@ Constrexpr.CSort(UAnonymous {rigid=true})) let do_build_inductive evd (funconstants: pconstant list) (funsargs: (Name.t * glob_constr * glob_constr option) list list) returned_types (rtl:glob_constr list) = let _time1 = System.get_time () in let funnames = List.map (fun c -> Label.to_id (KerName.label (Constant.canonical (fst c)))) funconstants in (* Pp.msgnl (prlist_with_sep fnl Printer.pr_glob_constr rtl); *) let funnames_as_set = List.fold_right Id.Set.add funnames Id.Set.empty in let funnames = Array.of_list funnames in let funsargs = Array.of_list funsargs in let returned_types = Array.of_list returned_types in (* alpha_renaming of the body to prevent variable capture during manipulation *) let rtl_alpha = List.map (function rt -> expand_as (alpha_rt [] rt)) rtl in let rta = Array.of_list rtl_alpha in (*i The next call to mk_rel_id is valid since we are constructing the graph Ensures by: obvious i*) let relnames = Array.map mk_rel_id funnames in let relnames_as_set = Array.fold_right Id.Set.add relnames Id.Set.empty in (* Construction of the pseudo constructors *) let open Context.Named.Declaration in let evd,env = Array.fold_right2 (fun id (c, u) (evd,env) -> let u = EConstr.EInstance.make u in let evd,t = Typing.type_of env evd (EConstr.mkConstU (c, u)) in let t = EConstr.Unsafe.to_constr t in evd, Environ.push_named (LocalAssum (make_annot id Sorts.Relevant,t)) env ) funnames (Array.of_list funconstants) (evd,Global.env ()) in (* we solve and replace the implicits *) let rta = Array.mapi (fun i rt -> let _,t = Typing.type_of env evd (EConstr.of_constr (mkConstU ((Array.of_list funconstants).(i)))) in resolve_and_replace_implicits ~expected_type:(Pretyping.OfType t) env evd rt ) rta in let resa = Array.map (build_entry_lc env evd funnames_as_set []) rta in let env_with_graphs = let rel_arity i funargs = (* Rebuilding arities (with parameters) *) let rel_first_args :(Name.t * Glob_term.glob_constr * Glob_term.glob_constr option ) list = funargs in List.fold_right (fun (n,t,typ) acc -> match typ with | Some typ -> CAst.make @@ Constrexpr.CLetIn((CAst.make n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ), acc) | None -> CAst.make @@ Constrexpr.CProdN ([Constrexpr.CLocalAssum([CAst.make n],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t)], acc ) ) rel_first_args (rebuild_return_type returned_types.(i)) in (* We need to lift back our work topconstr but only with all information We mimic a Set Printing All. Then save the graphs and reset Printing options to their primitive values *) let rel_arities = Array.mapi rel_arity funsargs in Util.Array.fold_left2 (fun env rel_name rel_ar -> let rex = fst (with_full_print (Constrintern.interp_constr env evd) rel_ar) in let rex = EConstr.Unsafe.to_constr rex in let r = Sorts.Relevant in (* TODO relevance *) Environ.push_named (LocalAssum (make_annot rel_name r,rex)) env) env relnames rel_arities in (* and of the real constructors*) let constr i res = List.map (function result (* (args',concl') *) -> let rt = compose_glob_context result.context result.value in let nb_args = List.length funsargs.(i) in (* with_full_print (fun rt -> Pp.msgnl (str "glob constr " ++ pr_glob_constr rt)) rt; *) fst ( rebuild_cons env_with_graphs nb_args relnames.(i) [] [] rt ) ) res.result in (* adding names to constructors *) let next_constructor_id = ref (-1) in let mk_constructor_id i = incr next_constructor_id; (*i The next call to mk_rel_id is valid since we are constructing the graph Ensures by: obvious i*) Id.of_string ((Id.to_string (mk_rel_id funnames.(i)))^"_"^(string_of_int !next_constructor_id)) in let rel_constructors i rt : (Id.t*glob_constr) list = next_constructor_id := (-1); List.map (fun constr -> (mk_constructor_id i),constr) (constr i rt) in let rel_constructors = Array.mapi rel_constructors resa in (* Computing the set of parameters if asked *) let rels_params = compute_params_name relnames_as_set funsargs rel_constructors in let nrel_params = List.length rels_params in let rel_constructors = (* Taking into account the parameters in constructors *) Array.map (List.map (fun (id,rt) -> (id,snd (chop_rprod_n nrel_params rt)))) rel_constructors in let rel_arity i funargs = (* Reduilding arities (with parameters) *) let rel_first_args :(Name.t * Glob_term.glob_constr * Glob_term.glob_constr option ) list = (snd (List.chop nrel_params funargs)) in List.fold_right (fun (n,t,typ) acc -> match typ with | Some typ -> CAst.make @@ Constrexpr.CLetIn((CAst.make n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ), acc) | None -> CAst.make @@ Constrexpr.CProdN ([Constrexpr.CLocalAssum([CAst.make n],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t)], acc ) ) rel_first_args (rebuild_return_type returned_types.(i)) in (* We need to lift back our work topconstr but only with all information We mimic a Set Printing All. Then save the graphs and reset Printing options to their primitive values *) let rel_arities = Array.mapi rel_arity funsargs in let rel_params_ids = List.fold_left (fun acc (na,_,_) -> match na with Anonymous -> acc | Name id -> id::acc ) [] rels_params in let rel_params = List.map (fun (n,t,typ) -> match typ with | Some typ -> Constrexpr.CLocalDef((CAst.make n), Constrextern.extern_glob_constr Id.Set.empty t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ)) | None -> Constrexpr.CLocalAssum ([(CAst.make n)], Constrexpr_ops.default_binder_kind, Constrextern.extern_glob_constr Id.Set.empty t) ) rels_params in let ext_rels_constructors = Array.map (List.map (fun (id,t) -> false,((CAst.make id), with_full_print (Constrextern.extern_glob_type Id.Set.empty) ((* zeta_normalize *) (alpha_rt rel_params_ids t)) ) )) (rel_constructors) in let rel_ind i ext_rel_constructors = ((CAst.make @@ relnames.(i)), rel_params, Some rel_arities.(i), ext_rel_constructors),[] in let ext_rel_constructors = (Array.mapi rel_ind ext_rels_constructors) in let rel_inds = Array.to_list ext_rel_constructors in (* let _ = *) (* Pp.msgnl (\* observe *\) ( *) (* str "Inductive" ++ spc () ++ *) (* prlist_with_sep *) (* (fun () -> fnl ()++spc () ++ str "with" ++ spc ()) *) (* (function ((_,id),_,params,ar,constr) -> *) (* Ppconstr.pr_id id ++ spc () ++ *) (* Ppconstr.pr_binders params ++ spc () ++ *) (* str ":" ++ spc () ++ *) (* Ppconstr.pr_lconstr_expr ar ++ spc () ++ str ":=" ++ *) (* prlist_with_sep *) (* (fun _ -> fnl () ++ spc () ++ str "|" ++ spc ()) *) (* (function (_,((_,id),t)) -> *) (* Ppconstr.pr_id id ++ spc () ++ str ":" ++ spc () ++ *) (* Ppconstr.pr_lconstr_expr t) *) (* constr *) (* ) *) (* rel_inds *) (* ) *) (* in *) let _time2 = System.get_time () in try with_full_print (Flags.silently (ComInductive.do_mutual_inductive ~template:(Some false) None rel_inds ~cumulative:false ~poly:false ~private_ind:false ~uniform:ComInductive.NonUniformParameters)) Declarations.Finite with | UserError(s,msg) as e -> let _time3 = System.get_time () in (* Pp.msgnl (str "error : "++ str (string_of_float (System.time_difference time2 time3))); *) let repacked_rel_inds = List.map (fun ((a , b , c , l),ntn) -> ((false,(a,None)) , b, c , Vernacexpr.Inductive_kw, Vernacexpr.Constructors l),ntn ) rel_inds in let msg = str "while trying to define"++ spc () ++ Ppvernac.pr_vernac (CAst.make Vernacexpr.{ control = []; attrs = []; expr = VernacInductive(None,false,Declarations.Finite,repacked_rel_inds)}) ++ fnl () ++ msg in observe (msg); raise e | reraise -> let _time3 = System.get_time () in (* Pp.msgnl (str "error : "++ str (string_of_float (System.time_difference time2 time3))); *) let repacked_rel_inds = List.map (fun ((a , b , c , l),ntn) -> ((false,(a,None)) , b, c , Vernacexpr.Inductive_kw, Vernacexpr.Constructors l),ntn ) rel_inds in let msg = str "while trying to define"++ spc () ++ Ppvernac.pr_vernac (CAst.make @@ Vernacexpr.{ control = []; attrs = []; expr = VernacInductive(None,false,Declarations.Finite,repacked_rel_inds)}) ++ fnl () ++ CErrors.print reraise in observe msg; raise reraise let build_inductive evd funconstants funsargs returned_types rtl = let pu = !Detyping.print_universes in let cu = !Constrextern.print_universes in try Detyping.print_universes := true; Constrextern.print_universes := true; do_build_inductive evd funconstants funsargs returned_types rtl; Detyping.print_universes := pu; Constrextern.print_universes := cu with e when CErrors.noncritical e -> Detyping.print_universes := pu; Constrextern.print_universes := cu; raise (Building_graph e) coq-8.11.0/plugins/funind/functional_principles_types.ml0000644000175000017500000002462513612664533023413 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* [] | decl :: predicates -> (match Context.Rel.Declaration.get_name decl with | Name x -> let id = Namegen.next_ident_away x (Id.Set.of_list avoid) in Hashtbl.add tbl id x; RelDecl.set_name (Name id) decl :: change_predicates_names (id::avoid) predicates | Anonymous -> anomaly (Pp.str "Anonymous property binder.")) in let avoid = (Termops.ids_of_context env_with_params ) in let princ_type_info = { princ_type_info with predicates = change_predicates_names avoid princ_type_info.predicates } in (* observe (str "starting princ_type := " ++ pr_lconstr_env env princ_type); *) (* observe (str "princ_infos : " ++ pr_elim_scheme princ_type_info); *) let change_predicate_sort i decl = let new_sort = sorts.(i) in let args,_ = decompose_prod (EConstr.Unsafe.to_constr (RelDecl.get_type decl)) in let real_args = if princ_type_info.indarg_in_concl then List.tl args else args in Context.Named.Declaration.LocalAssum (map_annot Nameops.Name.get_id (Context.Rel.Declaration.get_annot decl), Term.compose_prod real_args (mkSort new_sort)) in let new_predicates = List.map_i change_predicate_sort 0 princ_type_info.predicates in let env_with_params_and_predicates = List.fold_right Environ.push_named new_predicates env_with_params in let rel_as_kn = fst (match princ_type_info.indref with | Some (GlobRef.IndRef ind) -> ind | _ -> user_err Pp.(str "Not a valid predicate") ) in let ptes_vars = List.map Context.Named.Declaration.get_id new_predicates in let is_pte = let set = List.fold_right Id.Set.add ptes_vars Id.Set.empty in fun t -> match Constr.kind t with | Var id -> Id.Set.mem id set | _ -> false in let pre_princ = let open EConstr in it_mkProd_or_LetIn (it_mkProd_or_LetIn (Option.fold_right mkProd_or_LetIn princ_type_info.indarg princ_type_info.concl ) princ_type_info.args ) princ_type_info.branches in let pre_princ = EConstr.Unsafe.to_constr pre_princ in let pre_princ = substl (List.map mkVar ptes_vars) pre_princ in let is_dom c = match Constr.kind c with | Ind((u,_),_) -> MutInd.equal u rel_as_kn | Construct(((u,_),_),_) -> MutInd.equal u rel_as_kn | _ -> false in let get_fun_num c = match Constr.kind c with | Ind((_,num),_) -> num | Construct(((_,num),_),_) -> num | _ -> assert false in let dummy_var = mkVar (Id.of_string "________") in let mk_replacement c i args = let res = mkApp(rel_to_fun.(i), Array.map pop (array_get_start args)) in observe (str "replacing " ++ pr_lconstr_env env Evd.empty c ++ str " by " ++ pr_lconstr_env env Evd.empty res); res in let rec compute_new_princ_type remove env pre_princ : types*(constr list) = let (new_princ_type,_) as res = match Constr.kind pre_princ with | Rel n -> begin try match Environ.lookup_rel n env with | LocalAssum (_,t) | LocalDef (_,_,t) when is_dom t -> raise Toberemoved | _ -> pre_princ,[] with Not_found -> assert false end | Prod(x,t,b) -> compute_new_princ_type_for_binder remove mkProd env x t b | Lambda(x,t,b) -> compute_new_princ_type_for_binder remove mkLambda env x t b | Ind _ | Construct _ when is_dom pre_princ -> raise Toberemoved | App(f,args) when is_dom f -> let var_to_be_removed = destRel (Array.last args) in let num = get_fun_num f in raise (Toberemoved_with_rel (var_to_be_removed,mk_replacement pre_princ num args)) | App(f,args) -> let args = if is_pte f && remove then array_get_start args else args in let new_args,binders_to_remove = Array.fold_right (compute_new_princ_type_with_acc remove env) args ([],[]) in let new_f,binders_to_remove_from_f = compute_new_princ_type remove env f in applistc new_f new_args, list_union_eq Constr.equal binders_to_remove_from_f binders_to_remove | LetIn(x,v,t,b) -> compute_new_princ_type_for_letin remove env x v t b | _ -> pre_princ,[] in (* let _ = match Constr.kind pre_princ with *) (* | Prod _ -> *) (* observe(str "compute_new_princ_type for "++ *) (* pr_lconstr_env env pre_princ ++ *) (* str" is "++ *) (* pr_lconstr_env env new_princ_type ++ fnl ()) *) (* | _ -> () in *) res and compute_new_princ_type_for_binder remove bind_fun env x t b = begin try let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in let new_x = map_annot (get_name (Termops.ids_of_context env)) x in let new_env = Environ.push_rel (LocalAssum (x,t)) env in let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in if List.exists (Constr.equal (mkRel 1)) binders_to_remove_from_b then (pop new_b), filter_map (Constr.equal (mkRel 1)) pop binders_to_remove_from_b else ( bind_fun(new_x,new_t,new_b), list_union_eq Constr.equal binders_to_remove_from_t (List.map pop binders_to_remove_from_b) ) with | Toberemoved -> (* observe (str "Decl of "++Ppconstr.Name.print x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in new_b, List.map pop binders_to_remove_from_b | Toberemoved_with_rel (n,c) -> (* observe (str "Decl of "++Ppconstr.Name.print x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in new_b, list_add_set_eq Constr.equal (mkRel n) (List.map pop binders_to_remove_from_b) end and compute_new_princ_type_for_letin remove env x v t b = begin try let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in let new_v,binders_to_remove_from_v = compute_new_princ_type remove env v in let new_x = map_annot (get_name (Termops.ids_of_context env)) x in let new_env = Environ.push_rel (LocalDef (x,v,t)) env in let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in if List.exists (Constr.equal (mkRel 1)) binders_to_remove_from_b then (pop new_b),filter_map (Constr.equal (mkRel 1)) pop binders_to_remove_from_b else ( mkLetIn(new_x,new_v,new_t,new_b), list_union_eq Constr.equal (list_union_eq Constr.equal binders_to_remove_from_t binders_to_remove_from_v) (List.map pop binders_to_remove_from_b) ) with | Toberemoved -> (* observe (str "Decl of "++Ppconstr.Name.print x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in new_b, List.map pop binders_to_remove_from_b | Toberemoved_with_rel (n,c) -> (* observe (str "Decl of "++Ppconstr.Name.print x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in new_b, list_add_set_eq Constr.equal (mkRel n) (List.map pop binders_to_remove_from_b) end and compute_new_princ_type_with_acc remove env e (c_acc,to_remove_acc) = let new_e,to_remove_from_e = compute_new_princ_type remove env e in new_e::c_acc,list_union_eq Constr.equal to_remove_from_e to_remove_acc in (* observe (str "Computing new principe from " ++ pr_lconstr_env env_with_params_and_predicates pre_princ); *) let pre_res,_ = compute_new_princ_type princ_type_info.indarg_in_concl env_with_params_and_predicates pre_princ in let pre_res = replace_vars (List.map_i (fun i id -> (id, mkRel i)) 1 ptes_vars) (lift (List.length ptes_vars) pre_res) in it_mkProd_or_LetIn (it_mkProd_or_LetIn pre_res (List.map (function | Context.Named.Declaration.LocalAssum (id,b) -> LocalAssum (map_annot (fun id -> Name.mk_name (Hashtbl.find tbl id)) id, b) | Context.Named.Declaration.LocalDef (id,t,b) -> LocalDef (map_annot (fun id -> Name.mk_name (Hashtbl.find tbl id)) id, t, b)) new_predicates) ) (List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) princ_type_info.params) coq-8.11.0/plugins/funind/gen_principle.ml0000644000175000017500000024645113612664533020416 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Pp.str s) (* Construct a fixpoint as a Glob_term and not as a constr *) let rec abstract_glob_constr c = function | [] -> c | Constrexpr.CLocalDef (x,b,t)::bl -> Constrexpr_ops.mkLetInC(x,b,t,abstract_glob_constr c bl) | Constrexpr.CLocalAssum (idl,k,t)::bl -> List.fold_right (fun x b -> Constrexpr_ops.mkLambdaC([x],k,t,b)) idl (abstract_glob_constr c bl) | Constrexpr.CLocalPattern _::bl -> assert false let interp_casted_constr_with_implicits env sigma impls c = Constrintern.intern_gen Pretyping.WithoutTypeConstraint env sigma ~impls c let build_newrecursive lnameargsardef = let env0 = Global.env() in let sigma = Evd.from_env env0 in let (rec_sign,rec_impls) = List.fold_left (fun (env,impls) { Vernacexpr.fname={CAst.v=recname}; binders; rtype } -> let arityc = Constrexpr_ops.mkCProdN binders rtype in let arity,ctx = Constrintern.interp_type env0 sigma arityc in let evd = Evd.from_env env0 in let evd, (_, (_, impls')) = Constrintern.interp_context_evars ~program_mode:false env evd binders in let impl = Constrintern.compute_internalization_data env0 evd Constrintern.Recursive arity impls' in let open Context.Named.Declaration in let r = Sorts.Relevant in (* TODO relevance *) (EConstr.push_named (LocalAssum (Context.make_annot recname r,arity)) env, Id.Map.add recname impl impls)) (env0,Constrintern.empty_internalization_env) lnameargsardef in let recdef = (* Declare local notations *) let f { Vernacexpr.binders; body_def } = match body_def with | Some body_def -> let def = abstract_glob_constr body_def binders in interp_casted_constr_with_implicits rec_sign sigma rec_impls def | None -> CErrors.user_err ~hdr:"Function" (Pp.str "Body of Function must be given") in States.with_state_protection (List.map f) lnameargsardef in recdef,rec_impls (* Checks whether or not the mutual bloc is recursive *) let is_rec names = let open Glob_term in let names = List.fold_right Id.Set.add names Id.Set.empty in let check_id id names = Id.Set.mem id names in let rec lookup names gt = match DAst.get gt with | GVar(id) -> check_id id names | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ -> false | GCast(b,_) -> lookup names b | GRec _ -> CErrors.user_err (Pp.str "GRec not handled") | GIf(b,_,lhs,rhs) -> (lookup names b) || (lookup names lhs) || (lookup names rhs) | GProd(na,_,t,b) | GLambda(na,_,t,b) -> lookup names t || lookup (Nameops.Name.fold_right Id.Set.remove na names) b | GLetIn(na,b,t,c) -> lookup names b || Option.cata (lookup names) true t || lookup (Nameops.Name.fold_right Id.Set.remove na names) c | GLetTuple(nal,_,t,b) -> lookup names t || lookup (List.fold_left (fun acc na -> Nameops.Name.fold_right Id.Set.remove na acc) names nal ) b | GApp(f,args) -> List.exists (lookup names) (f::args) | GCases(_,_,el,brl) -> List.exists (fun (e,_) -> lookup names e) el || List.exists (lookup_br names) brl and lookup_br names {CAst.v=(idl,_,rt)} = let new_names = List.fold_right Id.Set.remove idl names in lookup new_names rt in lookup names let rec rebuild_bl aux bl typ = let open Constrexpr in match bl,typ with | [], _ -> List.rev aux,typ | (CLocalAssum(nal,bk,_))::bl',typ -> rebuild_nal aux bk bl' nal typ | (CLocalDef(na,_,_))::bl',{ CAst.v = CLetIn(_,nat,ty,typ') } -> rebuild_bl (Constrexpr.CLocalDef(na,nat,ty)::aux) bl' typ' | _ -> assert false and rebuild_nal aux bk bl' nal typ = let open Constrexpr in match nal,typ with | _,{ CAst.v = CProdN([],typ) } -> rebuild_nal aux bk bl' nal typ | [], _ -> rebuild_bl aux bl' typ | na::nal,{ CAst.v = CProdN(CLocalAssum(na'::nal',bk',nal't)::rest,typ') } -> if Name.equal (na.CAst.v) (na'.CAst.v) || Name.is_anonymous (na'.CAst.v) then let assum = CLocalAssum([na],bk,nal't) in let new_rest = if nal' = [] then rest else (CLocalAssum(nal',bk',nal't)::rest) in rebuild_nal (assum::aux) bk bl' nal (CAst.make @@ CProdN(new_rest,typ')) else let assum = CLocalAssum([na'],bk,nal't) in let new_rest = if nal' = [] then rest else (CLocalAssum(nal',bk',nal't)::rest) in rebuild_nal (assum::aux) bk bl' (na::nal) (CAst.make @@ CProdN(new_rest,typ')) | _ -> assert false let rebuild_bl aux bl typ = rebuild_bl aux bl typ let recompute_binder_list fixpoint_exprl = let fixl = List.map (fun fix -> Vernacexpr.{ fix with rec_order = ComFixpoint.adjust_rec_order ~structonly:false fix.binders fix.rec_order }) fixpoint_exprl in let ((_,_,_,typel),_,ctx,_) = ComFixpoint.interp_fixpoint ~cofix:false fixl in let constr_expr_typel = with_full_print (List.map (fun c -> Constrextern.extern_constr false (Global.env ()) (Evd.from_ctx ctx) (EConstr.of_constr c))) typel in let fixpoint_exprl_with_new_bl = List.map2 (fun ({ Vernacexpr.binders } as fp) fix_typ -> let binders, rtype = rebuild_bl [] binders fix_typ in { fp with Vernacexpr.binders; rtype } ) fixpoint_exprl constr_expr_typel in fixpoint_exprl_with_new_bl let rec local_binders_length = function (* Assume that no `{ ... } contexts occur *) | [] -> 0 | Constrexpr.CLocalDef _::bl -> 1 + local_binders_length bl | Constrexpr.CLocalAssum (idl,_,_)::bl -> List.length idl + local_binders_length bl | Constrexpr.CLocalPattern _::bl -> assert false let prepare_body { Vernacexpr.binders } rt = let n = local_binders_length binders in (* Pp.msgnl (str "nb lambda to chop : " ++ str (string_of_int n) ++ fnl () ++Printer.pr_glob_constr rt); *) let fun_args,rt' = chop_rlambda_n n rt in (fun_args,rt') let build_functional_principle ?(opaque=Proof_global.Transparent) (evd:Evd.evar_map ref) interactive_proof old_princ_type sorts funs i proof_tac hook = (* First we get the type of the old graph principle *) let mutr_nparams = (Tactics.compute_elim_sig !evd (EConstr.of_constr old_princ_type)).Tactics.nparams in (* let time1 = System.get_time () in *) let new_principle_type = Functional_principles_types.compute_new_princ_type_from_rel (Array.map Constr.mkConstU funs) sorts old_princ_type in (* let time2 = System.get_time () in *) (* Pp.msgnl (str "computing principle type := " ++ System.fmt_time_difference time1 time2); *) let new_princ_name = Namegen.next_ident_away_in_goal (Id.of_string "___________princ_________") Id.Set.empty in let sigma, _ = Typing.type_of ~refresh:true (Global.env ()) !evd (EConstr.of_constr new_principle_type) in evd := sigma; let hook = DeclareDef.Hook.make (hook new_principle_type) in let lemma = Lemmas.start_lemma ~name:new_princ_name ~poly:false !evd (EConstr.of_constr new_principle_type) in (* let _tim1 = System.get_time () in *) let map (c, u) = EConstr.mkConstU (c, EConstr.EInstance.make u) in let lemma,_ = Lemmas.by (Proofview.V82.tactic (proof_tac (Array.map map funs) mutr_nparams)) lemma in (* let _tim2 = System.get_time () in *) (* begin *) (* let dur1 = System.time_difference tim1 tim2 in *) (* Pp.msgnl (str ("Time to compute proof: ") ++ str (string_of_float dur1)); *) (* end; *) let open Proof_global in let { name; entries } = Lemmas.pf_fold (close_proof ~opaque ~keep_body_ucst_separate:false (fun x -> x)) lemma in match entries with | [entry] -> entry, hook | _ -> CErrors.anomaly Pp.(str "[build_functional_principle] close_proof returned more than one proof term") let change_property_sort evd toSort princ princName = let open Context.Rel.Declaration in let princ = EConstr.of_constr princ in let princ_info = Tactics.compute_elim_sig evd princ in let change_sort_in_predicate decl = LocalAssum (get_annot decl, let args,ty = Term.decompose_prod (EConstr.Unsafe.to_constr (get_type decl)) in let s = Constr.destSort ty in Global.add_constraints (Univ.enforce_leq (Sorts.univ_of_sort toSort) (Sorts.univ_of_sort s) Univ.Constraint.empty); Term.compose_prod args (Constr.mkSort toSort) ) in let evd,princName_as_constr = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident princName)) in let init = let nargs = (princ_info.Tactics.nparams + (List.length princ_info.Tactics.predicates)) in Constr.mkApp(EConstr.Unsafe.to_constr princName_as_constr, Array.init nargs (fun i -> Constr.mkRel (nargs - i ))) in evd, Term.it_mkLambda_or_LetIn (Term.it_mkLambda_or_LetIn init (List.map change_sort_in_predicate princ_info.Tactics.predicates) ) (List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) princ_info.Tactics.params) let generate_functional_principle (evd: Evd.evar_map ref) interactive_proof old_princ_type sorts new_princ_name funs i proof_tac = try let f = funs.(i) in let sigma, type_sort = Evd.fresh_sort_in_family !evd Sorts.InType in evd := sigma; let new_sorts = match sorts with | None -> Array.make (Array.length funs) (type_sort) | Some a -> a in let base_new_princ_name,new_princ_name = match new_princ_name with | Some (id) -> id,id | None -> let id_of_f = Label.to_id (Constant.label (fst f)) in id_of_f,Indrec.make_elimination_ident id_of_f (Sorts.family type_sort) in let names = ref [new_princ_name] in let hook = fun new_principle_type _ -> if Option.is_empty sorts then (* let id_of_f = Label.to_id (con_label f) in *) let register_with_sort fam_sort = let evd' = Evd.from_env (Global.env ()) in let evd',s = Evd.fresh_sort_in_family evd' fam_sort in let name = Indrec.make_elimination_ident base_new_princ_name fam_sort in let evd',value = change_property_sort evd' s new_principle_type new_princ_name in let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr value)) in (* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *) let univs = Evd.univ_entry ~poly:false evd' in let ce = Declare.definition_entry ~univs value in ignore( Declare.declare_constant ~name ~kind:Decls.(IsDefinition Scheme) (Declare.DefinitionEntry ce) ); Declare.definition_message name; names := name :: !names in register_with_sort Sorts.InProp; register_with_sort Sorts.InSet in let entry, hook = build_functional_principle evd interactive_proof old_princ_type new_sorts funs i proof_tac hook in (* Pr 1278 : Don't forget to close the goal if an error is raised !!!! *) let uctx = Evd.evar_universe_context sigma in let hook_data = hook, uctx, [] in let _ : Names.GlobRef.t = DeclareDef.declare_definition ~name:new_princ_name ~hook_data ~scope:(DeclareDef.Global Declare.ImportDefaultBehavior) ~kind:Decls.(IsProof Theorem) UnivNames.empty_binders entry [] in () with e when CErrors.noncritical e -> raise (Defining_principle e) let generate_principle (evd:Evd.evar_map ref) pconstants on_error is_general do_built fix_rec_l recdefs interactive_proof (continue_proof : int -> Names.Constant.t array -> EConstr.constr array -> int -> Tacmach.tactic) : unit = let names = List.map (function { Vernacexpr.fname = {CAst.v=name} } -> name) fix_rec_l in let fun_bodies = List.map2 prepare_body fix_rec_l recdefs in let funs_args = List.map fst fun_bodies in let funs_types = List.map (function { Vernacexpr.rtype } -> rtype) fix_rec_l in try (* We then register the Inductive graphs of the functions *) Glob_term_to_relation.build_inductive !evd pconstants funs_args funs_types recdefs; if do_built then begin (*i The next call to mk_rel_id is valid since we have just construct the graph Ensures by : do_built i*) let f_R_mut = Libnames.qualid_of_ident @@ mk_rel_id (List.nth names 0) in let ind_kn = fst (locate_with_msg Pp.(Libnames.pr_qualid f_R_mut ++ str ": Not an inductive type!") locate_ind f_R_mut) in let fname_kn { Vernacexpr.fname } = let f_ref = Libnames.qualid_of_ident ?loc:fname.CAst.loc fname.CAst.v in locate_with_msg Pp.(Libnames.pr_qualid f_ref++str ": Not an inductive type!") locate_constant f_ref in let funs_kn = Array.of_list (List.map fname_kn fix_rec_l) in let _ = List.map_i (fun i x -> let env = Global.env () in let princ = Indrec.lookup_eliminator env (ind_kn,i) (Sorts.InProp) in let evd = ref (Evd.from_env env) in let evd',uprinc = Evd.fresh_global env !evd princ in let _ = evd := evd' in let sigma, princ_type = Typing.type_of ~refresh:true env !evd uprinc in evd := sigma; let princ_type = EConstr.Unsafe.to_constr princ_type in generate_functional_principle evd interactive_proof princ_type None None (Array.of_list pconstants) (* funs_kn *) i (continue_proof 0 [|funs_kn.(i)|]) ) 0 fix_rec_l in Array.iter (add_Function is_general) funs_kn; () end with e when CErrors.noncritical e -> on_error names e let register_struct is_rec fixpoint_exprl = let open EConstr in match fixpoint_exprl with | [{ Vernacexpr.fname; univs; binders; rtype; body_def }] when not is_rec -> let body = match body_def with | Some body -> body | None -> CErrors.user_err ~hdr:"Function" Pp.(str "Body of Function must be given") in ComDefinition.do_definition ~program_mode:false ~name:fname.CAst.v ~poly:false ~scope:(DeclareDef.Global Declare.ImportDefaultBehavior) ~kind:Decls.Definition univs binders None body (Some rtype); let evd,rev_pconstants = List.fold_left (fun (evd,l) { Vernacexpr.fname } -> let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname.CAst.v)) in let (cst, u) = destConst evd c in let u = EInstance.kind evd u in evd,((cst, u) :: l) ) (Evd.from_env (Global.env ()),[]) fixpoint_exprl in None, evd,List.rev rev_pconstants | _ -> ComFixpoint.do_fixpoint ~scope:(DeclareDef.Global Declare.ImportDefaultBehavior) ~poly:false fixpoint_exprl; let evd,rev_pconstants = List.fold_left (fun (evd,l) { Vernacexpr.fname } -> let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname.CAst.v)) in let (cst, u) = destConst evd c in let u = EInstance.kind evd u in evd,((cst, u) :: l) ) (Evd.from_env (Global.env ()),[]) fixpoint_exprl in None,evd,List.rev rev_pconstants let generate_correction_proof_wf f_ref tcc_lemma_ref is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation (_: int) (_:Names.Constant.t array) (_:EConstr.constr array) (_:int) : Tacmach.tactic = Functional_principles_proofs.prove_principle_for_gen (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation (* [generate_type g_to_f f graph i] build the completeness (resp. correctness) lemma type if [g_to_f = true] (resp. g_to_f = false) where [graph] is the graph of [f] and is the [i]th function in the block. [generate_type true f i] returns \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, graph\ x_1\ldots x_n\ res \rightarrow res = fv \] decomposed as the context and the conclusion [generate_type false f i] returns \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, res = fv \rightarrow graph\ x_1\ldots x_n\ res\] decomposed as the context and the conclusion *) let generate_type evd g_to_f f graph i = let open Context.Rel.Declaration in let open EConstr in let open EConstr.Vars in (*i we deduce the number of arguments of the function and its returned type from the graph i*) let evd',graph = Evd.fresh_global (Global.env ()) !evd (GlobRef.IndRef (fst (destInd !evd graph))) in evd:=evd'; let sigma, graph_arity = Typing.type_of (Global.env ()) !evd graph in evd := sigma; let ctxt,_ = decompose_prod_assum !evd graph_arity in let fun_ctxt,res_type = match ctxt with | [] | [_] -> CErrors.anomaly (Pp.str "Not a valid context.") | decl :: fun_ctxt -> fun_ctxt, RelDecl.get_type decl in let rec args_from_decl i accu = function | [] -> accu | LocalDef _ :: l -> args_from_decl (succ i) accu l | _ :: l -> let t = mkRel i in args_from_decl (succ i) (t :: accu) l in (*i We need to name the vars [res] and [fv] i*) let filter = fun decl -> match RelDecl.get_name decl with | Name id -> Some id | Anonymous -> None in let named_ctxt = Id.Set.of_list (List.map_filter filter fun_ctxt) in let res_id = Namegen.next_ident_away_in_goal (Id.of_string "_res") named_ctxt in let fv_id = Namegen.next_ident_away_in_goal (Id.of_string "fv") (Id.Set.add res_id named_ctxt) in (*i we can then type the argument to be applied to the function [f] i*) let args_as_rels = Array.of_list (args_from_decl 1 [] fun_ctxt) in (*i the hypothesis [res = fv] can then be computed We will need to lift it by one in order to use it as a conclusion i*) let make_eq = make_eq () in let res_eq_f_of_args = mkApp(make_eq ,[|lift 2 res_type;mkRel 1;mkRel 2|]) in (*i The hypothesis [graph\ x_1\ldots x_n\ res] can then be computed We will need to lift it by one in order to use it as a conclusion i*) let args_and_res_as_rels = Array.of_list (args_from_decl 3 [] fun_ctxt) in let args_and_res_as_rels = Array.append args_and_res_as_rels [|mkRel 1|] in let graph_applied = mkApp(graph, args_and_res_as_rels) in (*i The [pre_context] is the defined to be the context corresponding to \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \] i*) let pre_ctxt = LocalAssum (Context.make_annot (Name res_id) Sorts.Relevant, lift 1 res_type) :: LocalDef (Context.make_annot (Name fv_id) Sorts.Relevant, mkApp (f,args_as_rels), res_type) :: fun_ctxt in (*i and we can return the solution depending on which lemma type we are defining i*) if g_to_f then LocalAssum (Context.make_annot Anonymous Sorts.Relevant,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph else LocalAssum (Context.make_annot Anonymous Sorts.Relevant,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph (** [find_induction_principle f] searches and returns the [body] and the [type] of [f_rect] WARNING: while convertible, [type_of body] and [type] can be non equal *) let find_induction_principle evd f = let f_as_constant,u = match EConstr.kind !evd f with | Constr.Const c' -> c' | _ -> CErrors.user_err Pp.(str "Must be used with a function") in match find_Function_infos f_as_constant with | None -> raise Not_found | Some infos -> match infos.rect_lemma with | None -> raise Not_found | Some rect_lemma -> let evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (GlobRef.ConstRef rect_lemma) in let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' rect_lemma in evd:=evd'; rect_lemma,typ (* [prove_fun_correct funs_constr graphs_constr schemes lemmas_types_infos i ] is the tactic used to prove correctness lemma. [funs_constr], [graphs_constr] [schemes] [lemmas_types_infos] are the mutually recursive functions (resp. graphs of the functions and principles and correctness lemma types) to prove correct. [i] is the indice of the function to prove correct The lemma to prove if suppose to have been generated by [generate_type] (in $\zeta$ normal form that is it looks like~: [\forall (x_1:t_1)\ldots(x_n:t_n), forall res, res = f x_1\ldots x_n in, \rightarrow graph\ x_1\ldots x_n\ res] The sketch of the proof is the following one~: \begin{enumerate} \item intros until $x_n$ \item $functional\ induction\ (f.(i)\ x_1\ldots x_n)$ using schemes.(i) \item for each generated branch intro [res] and [hres :res = f x_1\ldots x_n], rewrite [hres] and the apply the corresponding constructor of the corresponding graph inductive. \end{enumerate} *) let rec generate_fresh_id x avoid i = if i == 0 then [] else let id = Namegen.next_ident_away_in_goal x (Id.Set.of_list avoid) in id::(generate_fresh_id x (id::avoid) (pred i)) let prove_fun_correct evd funs_constr graphs_constr schemes lemmas_types_infos i : Tacmach.tactic = let open Constr in let open EConstr in let open Context.Rel.Declaration in let open Tacmach in let open Tactics in let open Tacticals in fun g -> (* first of all we recreate the lemmas types to be used as predicates of the induction principle that is~: \[fun (x_1:t_1)\ldots(x_n:t_n)=> fun fv => fun res => res = fv \rightarrow graph\ x_1\ldots x_n\ res\] *) (* we the get the definition of the graphs block *) let graph_ind,u = destInd evd graphs_constr.(i) in let kn = fst graph_ind in let mib,_ = Global.lookup_inductive graph_ind in (* and the principle to use in this lemma in $\zeta$ normal form *) let f_principle,princ_type = schemes.(i) in let princ_type = Reductionops.nf_zeta (Global.env ()) evd princ_type in let princ_infos = Tactics.compute_elim_sig evd princ_type in (* The number of args of the function is then easily computable *) let nb_fun_args = Termops.nb_prod (project g) (pf_concl g) - 2 in let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in let ids = args_names@(pf_ids_of_hyps g) in (* Since we cannot ensure that the functional principle is defined in the environment and due to the bug #1174, we will need to pose the principle using a name *) let principle_id = Namegen.next_ident_away_in_goal (Id.of_string "princ") (Id.Set.of_list ids) in let ids = principle_id :: ids in (* We get the branches of the principle *) let branches = List.rev princ_infos.Tactics.branches in (* and built the intro pattern for each of them *) let intro_pats = List.map (fun decl -> List.map (fun id -> CAst.make @@ Tactypes.IntroNaming (Namegen.IntroIdentifier id)) (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (RelDecl.get_type decl))))) ) branches in (* before building the full intro pattern for the principle *) let eq_ind = make_eq () in let eq_construct = mkConstructUi (destInd evd eq_ind, 1) in (* The next to referencies will be used to find out which constructor to apply in each branch *) let ind_number = ref 0 and min_constr_number = ref 0 in (* The tactic to prove the ith branch of the principle *) let prove_branche i g = (* We get the identifiers of this branch *) let pre_args = List.fold_right (fun {CAst.v=pat} acc -> match pat with | Tactypes.IntroNaming (Namegen.IntroIdentifier id) -> id::acc | _ -> CErrors.anomaly (Pp.str "Not an identifier.") ) (List.nth intro_pats (pred i)) [] in (* and get the real args of the branch by unfolding the defined constant *) (* We can then recompute the arguments of the constructor. For each [hid] introduced by this branch, if [hid] has type $forall res, res=fv -> graph.(j)\ x_1\ x_n res$ the corresponding arguments of the constructor are [ fv (hid fv (refl_equal fv)) ]. If [hid] has another type the corresponding argument of the constructor is [hid] *) let constructor_args g = List.fold_right (fun hid acc -> let type_of_hid = pf_unsafe_type_of g (mkVar hid) in let sigma = project g in match EConstr.kind sigma type_of_hid with | Prod(_,_,t') -> begin match EConstr.kind sigma t' with | Prod(_,t'',t''') -> begin match EConstr.kind sigma t'',EConstr.kind sigma t''' with | App(eq,args), App(graph',_) when (EConstr.eq_constr sigma eq eq_ind) && Array.exists (EConstr.eq_constr_nounivs sigma graph') graphs_constr -> (args.(2)::(mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|])) ::acc) | _ -> mkVar hid :: acc end | _ -> mkVar hid :: acc end | _ -> mkVar hid :: acc ) pre_args [] in (* in fact we must also add the parameters to the constructor args *) let constructor_args g = let params_id = fst (List.chop princ_infos.Tactics.nparams args_names) in (List.map mkVar params_id)@((constructor_args g)) in (* We then get the constructor corresponding to this branch and modifies the references has needed i.e. if the constructor is the last one of the current inductive then add one the number of the inductive to take and add the number of constructor of the previous graph to the minimal constructor number *) let constructor = let constructor_num = i - !min_constr_number in let length = Array.length (mib.Declarations.mind_packets.(!ind_number).Declarations.mind_consnames) in if constructor_num <= length then begin (kn,!ind_number),constructor_num end else begin incr ind_number; min_constr_number := !min_constr_number + length ; (kn,!ind_number),1 end in (* we can then build the final proof term *) let app_constructor g = applist((mkConstructU(constructor,u)),constructor_args g) in (* an apply the tactic *) let res,hres = match generate_fresh_id (Id.of_string "z") (ids(* @this_branche_ids *)) 2 with | [res;hres] -> res,hres | _ -> assert false in (* observe (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor); *) ( tclTHENLIST [ observe_tac ("h_intro_patterns ") (let l = (List.nth intro_pats (pred i)) in match l with | [] -> tclIDTAC | _ -> Proofview.V82.of_tactic (intro_patterns false l)); (* unfolding of all the defined variables introduced by this branch *) (* observe_tac "unfolding" pre_tac; *) (* $zeta$ normalizing of the conclusion *) Proofview.V82.of_tactic (reduce (Genredexpr.Cbv { Redops.all_flags with Genredexpr.rDelta = false ; Genredexpr.rConst = [] } ) Locusops.onConcl); observe_tac ("toto ") tclIDTAC; (* introducing the result of the graph and the equality hypothesis *) observe_tac "introducing" (tclMAP (fun x -> Proofview.V82.of_tactic (Simple.intro x)) [res;hres]); (* replacing [res] with its value *) observe_tac "rewriting res value" (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar hres))); (* Conclusion *) observe_tac "exact" (fun g -> Proofview.V82.of_tactic (exact_check (app_constructor g)) g) ] ) g in (* end of branche proof *) let lemmas = Array.map (fun ((_,(ctxt,concl))) -> match ctxt with | [] | [_] | [_;_] -> CErrors.anomaly (Pp.str "bad context.") | hres::res::decl::ctxt -> let res = EConstr.it_mkLambda_or_LetIn (EConstr.it_mkProd_or_LetIn concl [hres;res]) (LocalAssum (RelDecl.get_annot decl, RelDecl.get_type decl) :: ctxt) in res) lemmas_types_infos in let param_names = fst (List.chop princ_infos.nparams args_names) in let params = List.map mkVar param_names in let lemmas = Array.to_list (Array.map (fun c -> applist(c,params)) lemmas) in (* The bindings of the principle that is the params of the principle and the different lemma types *) let bindings = let params_bindings,avoid = List.fold_left2 (fun (bindings,avoid) decl p -> let id = Namegen.next_ident_away (Nameops.Name.get_id (RelDecl.get_name decl)) (Id.Set.of_list avoid) in p::bindings,id::avoid ) ([],pf_ids_of_hyps g) princ_infos.params (List.rev params) in let lemmas_bindings = List.rev (fst (List.fold_left2 (fun (bindings,avoid) decl p -> let id = Namegen.next_ident_away (Nameops.Name.get_id (RelDecl.get_name decl)) (Id.Set.of_list avoid) in (Reductionops.nf_zeta (pf_env g) (project g) p)::bindings,id::avoid) ([],avoid) princ_infos.predicates (lemmas))) in (params_bindings@lemmas_bindings) in tclTHENLIST [ observe_tac "principle" (Proofview.V82.of_tactic (assert_by (Name principle_id) princ_type (exact_check f_principle))); observe_tac "intro args_names" (tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) args_names); (* observe_tac "titi" (pose_proof (Name (Id.of_string "__")) (Reductionops.nf_beta Evd.empty ((mkApp (mkVar principle_id,Array.of_list bindings))))); *) observe_tac "idtac" tclIDTAC; tclTHEN_i (observe_tac "functional_induction" ( (fun gl -> let term = mkApp (mkVar principle_id,Array.of_list bindings) in let gl', _ty = pf_eapply (Typing.type_of ~refresh:true) gl term in Proofview.V82.of_tactic (apply term) gl') )) (fun i g -> observe_tac ("proving branche "^string_of_int i) (prove_branche i) g ) ] g (* [prove_fun_complete funs graphs schemes lemmas_types_infos i] is the tactic used to prove completeness lemma. [funcs], [graphs] [schemes] [lemmas_types_infos] are the mutually recursive functions (resp. definitions of the graphs of the functions, principles and correctness lemma types) to prove correct. [i] is the indice of the function to prove complete The lemma to prove if suppose to have been generated by [generate_type] (in $\zeta$ normal form that is it looks like~: [\forall (x_1:t_1)\ldots(x_n:t_n), forall res, graph\ x_1\ldots x_n\ res, \rightarrow res = f x_1\ldots x_n in] The sketch of the proof is the following one~: \begin{enumerate} \item intros until $H:graph\ x_1\ldots x_n\ res$ \item $elim\ H$ using schemes.(i) \item for each generated branch, intro the news hyptohesis, for each such hyptohesis [h], if [h] has type [x=?] with [x] a variable, then subst [x], if [h] has type [t=?] with [t] not a variable then rewrite [t] in the subterms, else if [h] is a match then destruct it, else do just introduce it, after all intros, the conclusion should be a reflexive equality. \end{enumerate} *) let thin ids gl = Proofview.V82.of_tactic (Tactics.clear ids) gl (* [intros_with_rewrite] do the intros in each branch and treat each new hypothesis (unfolding, substituting, destructing cases \ldots) *) let tauto = let open Ltac_plugin in let dp = List.map Id.of_string ["Tauto" ; "Init"; "Coq"] in let mp = ModPath.MPfile (DirPath.make dp) in let kn = KerName.make mp (Label.make "tauto") in Proofview.tclBIND (Proofview.tclUNIT ()) begin fun () -> let body = Tacenv.interp_ltac kn in Tacinterp.eval_tactic body end (* [generalize_dependent_of x hyp g] generalize every hypothesis which depends of [x] but [hyp] *) let generalize_dependent_of x hyp g = let open Context.Named.Declaration in let open Tacmach in let open Tacticals in tclMAP (function | LocalAssum ({Context.binder_name=id},t) when not (Id.equal id hyp) && (Termops.occur_var (pf_env g) (project g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [EConstr.mkVar id])) (thin [id]) | _ -> tclIDTAC ) (pf_hyps g) g let rec intros_with_rewrite g = observe_tac "intros_with_rewrite" intros_with_rewrite_aux g and intros_with_rewrite_aux : Tacmach.tactic = let open Constr in let open EConstr in let open Tacmach in let open Tactics in let open Tacticals in fun g -> let eq_ind = make_eq () in let sigma = project g in match EConstr.kind sigma (pf_concl g) with | Prod(_,t,t') -> begin match EConstr.kind sigma t with | App(eq,args) when (EConstr.eq_constr sigma eq eq_ind) -> if Reductionops.is_conv (pf_env g) (project g) args.(1) args.(2) then let id = pf_get_new_id (Id.of_string "y") g in tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g else if isVar sigma args.(1) && (Environ.evaluable_named (destVar sigma args.(1)) (pf_env g)) then tclTHENLIST[ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))]); tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))] ((destVar sigma args.(1)),Locus.InHyp) ))) (pf_ids_of_hyps g); intros_with_rewrite ] g else if isVar sigma args.(2) && (Environ.evaluable_named (destVar sigma args.(2)) (pf_env g)) then tclTHENLIST[ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))]); tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))] ((destVar sigma args.(2)),Locus.InHyp) ))) (pf_ids_of_hyps g); intros_with_rewrite ] g else if isVar sigma args.(1) then let id = pf_get_new_id (Id.of_string "y") g in tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id); generalize_dependent_of (destVar sigma args.(1)) id; tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); intros_with_rewrite ] g else if isVar sigma args.(2) then let id = pf_get_new_id (Id.of_string "y") g in tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id); generalize_dependent_of (destVar sigma args.(2)) id; tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar id))); intros_with_rewrite ] g else begin let id = pf_get_new_id (Id.of_string "y") g in tclTHENLIST[ Proofview.V82.of_tactic (Simple.intro id); tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); intros_with_rewrite ] g end | Ind _ when EConstr.eq_constr sigma t (EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.False.type")) -> Proofview.V82.of_tactic tauto g | Case(_,_,v,_) -> tclTHENLIST[ Proofview.V82.of_tactic (simplest_case v); intros_with_rewrite ] g | LetIn _ -> tclTHENLIST[ Proofview.V82.of_tactic (reduce (Genredexpr.Cbv {Redops.all_flags with Genredexpr.rDelta = false; }) Locusops.onConcl) ; intros_with_rewrite ] g | _ -> let id = pf_get_new_id (Id.of_string "y") g in tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id);intros_with_rewrite] g end | LetIn _ -> tclTHENLIST[ Proofview.V82.of_tactic (reduce (Genredexpr.Cbv {Redops.all_flags with Genredexpr.rDelta = false; }) Locusops.onConcl) ; intros_with_rewrite ] g | _ -> tclIDTAC g let rec reflexivity_with_destruct_cases g = let open Constr in let open EConstr in let open Tacmach in let open Tactics in let open Tacticals in let destruct_case () = try match EConstr.kind (project g) (snd (destApp (project g) (pf_concl g))).(2) with | Case(_,_,v,_) -> tclTHENLIST[ Proofview.V82.of_tactic (simplest_case v); Proofview.V82.of_tactic intros; observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases ] | _ -> Proofview.V82.of_tactic reflexivity with e when CErrors.noncritical e -> Proofview.V82.of_tactic reflexivity in let eq_ind = make_eq () in let my_inj_flags = Some { Equality.keep_proof_equalities = false; injection_in_context = false; (* for compatibility, necessary *) injection_pattern_l2r_order = false; (* probably does not matter; except maybe with dependent hyps *) } in let discr_inject = Tacticals.onAllHypsAndConcl ( fun sc g -> match sc with None -> tclIDTAC g | Some id -> match EConstr.kind (project g) (pf_unsafe_type_of g (mkVar id)) with | App(eq,[|_;t1;t2|]) when EConstr.eq_constr (project g) eq eq_ind -> if Equality.discriminable (pf_env g) (project g) t1 t2 then Proofview.V82.of_tactic (Equality.discrHyp id) g else if Equality.injectable (pf_env g) (project g) ~keep_proofs:None t1 t2 then tclTHENLIST [Proofview.V82.of_tactic (Equality.injHyp my_inj_flags None id);thin [id];intros_with_rewrite] g else tclIDTAC g | _ -> tclIDTAC g ) in (tclFIRST [ observe_tac "reflexivity_with_destruct_cases : reflexivity" (Proofview.V82.of_tactic reflexivity); observe_tac "reflexivity_with_destruct_cases : destruct_case" ((destruct_case ())); (* We reach this point ONLY if the same value is matched (at least) two times along binding path. In this case, either we have a discriminable hypothesis and we are done, either at least an injectable one and we do the injection before continuing *) observe_tac "reflexivity_with_destruct_cases : others" (tclTHEN (tclPROGRESS discr_inject ) reflexivity_with_destruct_cases) ]) g let prove_fun_complete funcs graphs schemes lemmas_types_infos i : Tacmach.tactic = let open EConstr in let open Tacmach in let open Tactics in let open Tacticals in fun g -> (* We compute the types of the different mutually recursive lemmas in $\zeta$ normal form *) let lemmas = Array.map (fun (_,(ctxt,concl)) -> Reductionops.nf_zeta (pf_env g) (project g) (EConstr.it_mkLambda_or_LetIn concl ctxt)) lemmas_types_infos in (* We get the constant and the principle corresponding to this lemma *) let f = funcs.(i) in let graph_principle = Reductionops.nf_zeta (pf_env g) (project g) (EConstr.of_constr schemes.(i)) in let princ_type = pf_unsafe_type_of g graph_principle in let princ_infos = Tactics.compute_elim_sig (project g) princ_type in (* Then we get the number of argument of the function and compute a fresh name for each of them *) let nb_fun_args = Termops.nb_prod (project g) (pf_concl g) - 2 in let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in let ids = args_names@(pf_ids_of_hyps g) in (* and fresh names for res H and the principle (cf bug bug #1174) *) let res,hres,graph_principle_id = match generate_fresh_id (Id.of_string "z") ids 3 with | [res;hres;graph_principle_id] -> res,hres,graph_principle_id | _ -> assert false in let ids = res::hres::graph_principle_id::ids in (* we also compute fresh names for each hyptohesis of each branch of the principle *) let branches = List.rev princ_infos.branches in let intro_pats = List.map (fun decl -> List.map (fun id -> id) (generate_fresh_id (Id.of_string "y") ids (Termops.nb_prod (project g) (RelDecl.get_type decl))) ) branches in (* We will need to change the function by its body using [f_equation] if it is recursive (that is the graph is infinite or unfold if the graph is finite *) let rewrite_tac j ids : Tacmach.tactic = let graph_def = graphs.(j) in let infos = match find_Function_infos (fst (destConst (project g) funcs.(j))) with | None -> CErrors.user_err Pp.(str "No graph found") | Some infos -> infos in if infos.is_general || Rtree.is_infinite Declareops.eq_recarg graph_def.Declarations.mind_recargs then let eq_lemma = try Option.get (infos).equation_lemma with Option.IsNone -> CErrors.anomaly (Pp.str "Cannot find equation lemma.") in tclTHENLIST[ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) ids; Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_lemma)); (* Don't forget to $\zeta$ normlize the term since the principles have been $\zeta$-normalized *) Proofview.V82.of_tactic (reduce (Genredexpr.Cbv {Redops.all_flags with Genredexpr.rDelta = false; }) Locusops.onConcl) ; Proofview.V82.of_tactic (generalize (List.map mkVar ids)); thin ids ] else Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst (project g) f)))]) in (* The proof of each branche itself *) let ind_number = ref 0 in let min_constr_number = ref 0 in let prove_branche i g = (* we fist compute the inductive corresponding to the branch *) let this_ind_number = let constructor_num = i - !min_constr_number in let length = Array.length (graphs.(!ind_number).Declarations.mind_consnames) in if constructor_num <= length then !ind_number else begin incr ind_number; min_constr_number := !min_constr_number + length; !ind_number end in let this_branche_ids = List.nth intro_pats (pred i) in tclTHENLIST[ (* we expand the definition of the function *) observe_tac "rewrite_tac" (rewrite_tac this_ind_number this_branche_ids); (* introduce hypothesis with some rewrite *) observe_tac "intros_with_rewrite (all)" intros_with_rewrite; (* The proof is (almost) complete *) observe_tac "reflexivity" (reflexivity_with_destruct_cases) ] g in let params_names = fst (List.chop princ_infos.nparams args_names) in let open EConstr in let params = List.map mkVar params_names in tclTHENLIST [ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (args_names@[res;hres]); observe_tac "h_generalize" (Proofview.V82.of_tactic (generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)])); Proofview.V82.of_tactic (Simple.intro graph_principle_id); observe_tac "" (tclTHEN_i (observe_tac "elim" (Proofview.V82.of_tactic (elim false None (mkVar hres, Tactypes.NoBindings) (Some (mkVar graph_principle_id, Tactypes.NoBindings))))) (fun i g -> observe_tac "prove_branche" (prove_branche i) g )) ] g exception No_graph_found let get_funs_constant mp = let open Constr in let exception Not_Rec in let get_funs_constant const e : (Names.Constant.t*int) array = match Constr.kind (Term.strip_lam e) with | Fix((_,(na,_,_))) -> Array.mapi (fun i na -> match na.Context.binder_name with | Name id -> let const = Constant.make2 mp (Label.of_id id) in const,i | Anonymous -> CErrors.anomaly (Pp.str "Anonymous fix.") ) na | _ -> [|const,0|] in function const -> let find_constant_body const = match Global.body_of_constant Library.indirect_accessor const with | Some (body, _, _) -> let body = Tacred.cbv_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) (Global.env ()) (Evd.from_env (Global.env ())) (EConstr.of_constr body) in let body = EConstr.Unsafe.to_constr body in body | None -> CErrors.user_err Pp.(str ( "Cannot define a principle over an axiom ")) in let f = find_constant_body const in let l_const = get_funs_constant const f in (* We need to check that all the functions found are in the same block to prevent Reset strange thing *) let l_bodies = List.map find_constant_body (Array.to_list (Array.map fst l_const)) in let l_params,l_fixes = List.split (List.map Term.decompose_lam l_bodies) in (* all the parameters must be equal*) let _check_params = let first_params = List.hd l_params in List.iter (fun params -> if not (List.equal (fun (n1, c1) (n2, c2) -> Context.eq_annot Name.equal n1 n2 && Constr.equal c1 c2) first_params params) then CErrors.user_err Pp.(str "Not a mutal recursive block") ) l_params in (* The bodies has to be very similar *) let _check_bodies = try let extract_info is_first body = match Constr.kind body with | Fix((idxs,_),(na,ta,ca)) -> (idxs,na,ta,ca) | _ -> if is_first && Int.equal (List.length l_bodies) 1 then raise Not_Rec else CErrors.user_err Pp.(str "Not a mutal recursive block") in let first_infos = extract_info true (List.hd l_bodies) in let check body = (* Hope this is correct *) let eq_infos (ia1, na1, ta1, ca1) (ia2, na2, ta2, ca2) = Array.equal Int.equal ia1 ia2 && Array.equal (Context.eq_annot Name.equal) na1 na2 && Array.equal Constr.equal ta1 ta2 && Array.equal Constr.equal ca1 ca2 in if not (eq_infos first_infos (extract_info false body)) then CErrors.user_err Pp.(str "Not a mutal recursive block") in List.iter check l_bodies with Not_Rec -> () in l_const let make_scheme evd (fas : (Constr.pconstant * Sorts.family) list) : Evd.side_effects Declare.proof_entry list = let exception Found_type of int in let env = Global.env () in let funs = List.map fst fas in let first_fun = List.hd funs in let funs_mp = KerName.modpath (Constant.canonical (fst first_fun)) in let first_fun_kn = match find_Function_infos (fst first_fun) with | None -> raise No_graph_found | Some finfos -> fst finfos.graph_ind in let this_block_funs_indexes = get_funs_constant funs_mp (fst first_fun) in let this_block_funs = Array.map (fun (c,_) -> (c,snd first_fun)) this_block_funs_indexes in let prop_sort = Sorts.InProp in let funs_indexes = let this_block_funs_indexes = Array.to_list this_block_funs_indexes in List.map (function cst -> List.assoc_f Constant.equal (fst cst) this_block_funs_indexes) funs in let ind_list = List.map (fun (idx) -> let ind = first_fun_kn,idx in (ind,snd first_fun),true,prop_sort ) funs_indexes in let sigma, schemes = Indrec.build_mutual_induction_scheme env !evd ind_list in let _ = evd := sigma in let l_schemes = List.map (EConstr.of_constr %> Typing.unsafe_type_of env sigma %> EConstr.Unsafe.to_constr) schemes in let i = ref (-1) in let sorts = List.rev_map (fun (_,x) -> let sigma, fs = Evd.fresh_sort_in_family !evd x in evd := sigma; fs ) fas in (* We create the first principle by tactic *) let first_type,other_princ_types = match l_schemes with s::l_schemes -> s,l_schemes | _ -> CErrors.anomaly (Pp.str "") in let opaque = let finfos = match find_Function_infos (fst first_fun) with | None -> raise Not_found | Some finfos -> finfos in let open Proof_global in match finfos.equation_lemma with | None -> Transparent (* non recursive definition *) | Some equation -> if Declareops.is_opaque (Global.lookup_constant equation) then Opaque else Transparent in let entry, _hook = try build_functional_principle ~opaque evd false first_type (Array.of_list sorts) this_block_funs 0 (Functional_principles_proofs.prove_princ_for_struct evd false 0 (Array.of_list (List.map fst funs))) (fun _ _ -> ()) with e when CErrors.noncritical e -> raise (Defining_principle e) in incr i; (* The others are just deduced *) if List.is_empty other_princ_types then [entry] else let other_fun_princ_types = let funs = Array.map Constr.mkConstU this_block_funs in let sorts = Array.of_list sorts in List.map (Functional_principles_types.compute_new_princ_type_from_rel funs sorts) other_princ_types in let first_princ_body,first_princ_type = Declare.(entry.proof_entry_body, entry.proof_entry_type) in let ctxt,fix = Term.decompose_lam_assum (fst(fst(Future.force first_princ_body))) in (* the principle has for forall ...., fix .*) let (idxs,_),(_,ta,_ as decl) = Constr.destFix fix in let other_result = List.map (* we can now compute the other principles *) (fun scheme_type -> incr i; observe (Printer.pr_lconstr_env env sigma scheme_type); let type_concl = (Term.strip_prod_assum scheme_type) in let applied_f = List.hd (List.rev (snd (Constr.decompose_app type_concl))) in let f = fst (Constr.decompose_app applied_f) in try (* we search the number of the function in the fix block (name of the function) *) Array.iteri (fun j t -> let t = (Term.strip_prod_assum t) in let applied_g = List.hd (List.rev (snd (Constr.decompose_app t))) in let g = fst (Constr.decompose_app applied_g) in if Constr.equal f g then raise (Found_type j); observe Pp.(Printer.pr_lconstr_env env sigma f ++ str " <> " ++ Printer.pr_lconstr_env env sigma g) ) ta; (* If we reach this point, the two principle are not mutually recursive We fall back to the previous method *) let entry, _hook = build_functional_principle evd false (List.nth other_princ_types (!i - 1)) (Array.of_list sorts) this_block_funs !i (Functional_principles_proofs.prove_princ_for_struct evd false !i (Array.of_list (List.map fst funs))) (fun _ _ -> ()) in entry with Found_type i -> let princ_body = Termops.it_mkLambda_or_LetIn (Constr.mkFix((idxs,i),decl)) ctxt in Declare.definition_entry ~types:scheme_type princ_body ) other_fun_princ_types in entry::other_result (* [derive_correctness funs graphs] create correctness and completeness lemmas for each function in [funs] w.r.t. [graphs] *) let derive_correctness (funs: Constr.pconstant list) (graphs:inductive list) = let open EConstr in assert (funs <> []); assert (graphs <> []); let funs = Array.of_list funs and graphs = Array.of_list graphs in let map (c, u) = mkConstU (c, EInstance.make u) in let funs_constr = Array.map map funs in (* XXX STATE Why do we need this... why is the toplevel protection not enough *) funind_purify (fun () -> let env = Global.env () in let evd = ref (Evd.from_env env) in let graphs_constr = Array.map mkInd graphs in let lemmas_types_infos = Util.Array.map2_i (fun i f_constr graph -> (* let const_of_f,u = destConst f_constr in *) let (type_of_lemma_ctxt,type_of_lemma_concl,graph) = generate_type evd false f_constr graph i in let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in graphs_constr.(i) <- graph; let type_of_lemma = EConstr.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in let sigma, _ = Typing.type_of (Global.env ()) !evd type_of_lemma in evd := sigma; let type_of_lemma = Reductionops.nf_zeta (Global.env ()) !evd type_of_lemma in observe Pp.(str "type_of_lemma := " ++ Printer.pr_leconstr_env (Global.env ()) !evd type_of_lemma); type_of_lemma,type_info ) funs_constr graphs_constr in let schemes = (* The functional induction schemes are computed and not saved if there is more that one function if the block contains only one function we can safely reuse [f_rect] *) try if not (Int.equal (Array.length funs_constr) 1) then raise Not_found; [| find_induction_principle evd funs_constr.(0) |] with Not_found -> ( Array.of_list (List.map (fun entry -> (EConstr.of_constr (fst (fst (Future.force entry.Declare.proof_entry_body))), EConstr.of_constr (Option.get entry.Declare.proof_entry_type )) ) (make_scheme evd (Array.map_to_list (fun const -> const,Sorts.InType) funs)) ) ) in let proving_tac = prove_fun_correct !evd funs_constr graphs_constr schemes lemmas_types_infos in Array.iteri (fun i f_as_constant -> let f_id = Label.to_id (Constant.label (fst f_as_constant)) in (*i The next call to mk_correct_id is valid since we are constructing the lemma Ensures by: obvious i*) let lem_id = mk_correct_id f_id in let (typ,_) = lemmas_types_infos.(i) in let lemma = Lemmas.start_lemma ~name:lem_id ~poly:false !evd typ in let lemma = fst @@ Lemmas.by (Proofview.V82.tactic (proving_tac i)) lemma in let () = Lemmas.save_lemma_proved ~lemma ~opaque:Proof_global.Transparent ~idopt:None in let finfo = match find_Function_infos (fst f_as_constant) with | None -> raise Not_found | Some finfo -> finfo in (* let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in *) let _,lem_cst_constr = Evd.fresh_global (Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in let (lem_cst,_) = EConstr.destConst !evd lem_cst_constr in update_Function {finfo with correctness_lemma = Some lem_cst}; ) funs; let lemmas_types_infos = Util.Array.map2_i (fun i f_constr graph -> let (type_of_lemma_ctxt,type_of_lemma_concl,graph) = generate_type evd true f_constr graph i in let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in graphs_constr.(i) <- graph; let type_of_lemma = EConstr.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in let type_of_lemma = Reductionops.nf_zeta env !evd type_of_lemma in observe Pp.(str "type_of_lemma := " ++ Printer.pr_leconstr_env env !evd type_of_lemma); type_of_lemma,type_info ) funs_constr graphs_constr in let (kn,_) as graph_ind,u = (destInd !evd graphs_constr.(0)) in let mib,mip = Global.lookup_inductive graph_ind in let sigma, scheme = (Indrec.build_mutual_induction_scheme (Global.env ()) !evd (Array.to_list (Array.mapi (fun i _ -> ((kn,i), EInstance.kind !evd u),true, Sorts.InType) mib.Declarations.mind_packets ) ) ) in let schemes = Array.of_list scheme in let proving_tac = prove_fun_complete funs_constr mib.Declarations.mind_packets schemes lemmas_types_infos in Array.iteri (fun i f_as_constant -> let f_id = Label.to_id (Constant.label (fst f_as_constant)) in (*i The next call to mk_complete_id is valid since we are constructing the lemma Ensures by: obvious i*) let lem_id = mk_complete_id f_id in let lemma = Lemmas.start_lemma ~name:lem_id ~poly:false sigma (fst lemmas_types_infos.(i)) in let lemma = fst (Lemmas.by (Proofview.V82.tactic (observe_tac ("prove completeness ("^(Id.to_string f_id)^")") (proving_tac i))) lemma) in let () = Lemmas.save_lemma_proved ~lemma ~opaque:Proof_global.Transparent ~idopt:None in let finfo = match find_Function_infos (fst f_as_constant) with | None -> raise Not_found | Some finfo -> finfo in let _,lem_cst_constr = Evd.fresh_global (Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in let (lem_cst,_) = destConst !evd lem_cst_constr in update_Function {finfo with completeness_lemma = Some lem_cst} ) funs) () let warn_funind_cannot_build_inversion = CWarnings.create ~name:"funind-cannot-build-inversion" ~category:"funind" Pp.(fun e' -> strbrk "Cannot build inversion information" ++ if do_observe () then (fnl() ++ CErrors.print e') else mt ()) let derive_inversion fix_names = try let evd' = Evd.from_env (Global.env ()) in (* we first transform the fix_names identifier into their corresponding constant *) let evd',fix_names_as_constant = List.fold_right (fun id (evd,l) -> let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident id)) in let (cst, u) = EConstr.destConst evd c in evd, (cst, EConstr.EInstance.kind evd u) :: l ) fix_names (evd',[]) in (* Then we check that the graphs have been defined If one of the graphs haven't been defined we do nothing *) List.iter (fun c -> ignore (find_Function_infos (fst c))) fix_names_as_constant ; try let evd', lind = List.fold_right (fun id (evd,l) -> let evd,id = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident (mk_rel_id id))) in evd,(fst (EConstr.destInd evd id))::l ) fix_names (evd',[]) in derive_correctness fix_names_as_constant lind; with e when CErrors.noncritical e -> warn_funind_cannot_build_inversion e with e when CErrors.noncritical e -> warn_funind_cannot_build_inversion e let register_wf interactive_proof ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas args ret_type body pre_hook = let type_of_f = Constrexpr_ops.mkCProdN args ret_type in let rec_arg_num = let names = List.map CAst.(with_val (fun x -> x)) (Constrexpr_ops.names_of_local_assums args) in List.index Name.equal (Name wf_arg) names in let unbounded_eq = let f_app_args = CAst.make @@ Constrexpr.CAppExpl( (None, Libnames.qualid_of_ident fname,None) , (List.map (function | {CAst.v=Anonymous} -> assert false | {CAst.v=Name e} -> (Constrexpr_ops.mkIdentC e) ) (Constrexpr_ops.names_of_local_assums args) ) ) in CAst.make @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Libnames.qualid_of_string "Logic.eq")), [(f_app_args,None);(body,None)]) in let eq = Constrexpr_ops.mkCProdN args unbounded_eq in let hook ((f_ref,_) as fconst) tcc_lemma_ref (functional_ref,_) (eq_ref,_) rec_arg_num rec_arg_type nb_args relation = try pre_hook [fconst] (generate_correction_proof_wf f_ref tcc_lemma_ref is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation ); derive_inversion [fname] with e when CErrors.noncritical e -> (* No proof done *) () in Recdef.recursive_definition ~interactive_proof ~is_mes fname rec_impls type_of_f wf_rel_expr rec_arg_num eq hook using_lemmas let register_mes interactive_proof fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas args ret_type body = let wf_arg_type,wf_arg = match wf_arg with | None -> begin match args with | [Constrexpr.CLocalAssum ([{CAst.v=Name x}],k,t)] -> t,x | _ -> CErrors.user_err (Pp.str "Recursive argument must be specified") end | Some wf_args -> try match List.find (function | Constrexpr.CLocalAssum(l,k,t) -> List.exists (function {CAst.v=Name id} -> Id.equal id wf_args | _ -> false) l | _ -> false ) args with | Constrexpr.CLocalAssum(_,k,t) -> t,wf_args | _ -> assert false with Not_found -> assert false in let wf_rel_from_mes,is_mes = match wf_rel_expr_opt with | None -> let ltof = let make_dir l = DirPath.make (List.rev_map Id.of_string l) in Libnames.qualid_of_path (Libnames.make_path (make_dir ["Arith";"Wf_nat"]) (Id.of_string "ltof")) in let fun_from_mes = let applied_mes = Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC wf_arg]) in Constrexpr_ops.mkLambdaC ([CAst.make @@ Name wf_arg],Constrexpr_ops.default_binder_kind,wf_arg_type,applied_mes) in let wf_rel_from_mes = Constrexpr_ops.mkAppC(Constrexpr_ops.mkRefC ltof,[wf_arg_type;fun_from_mes]) in wf_rel_from_mes,true | Some wf_rel_expr -> let wf_rel_with_mes = let a = Names.Id.of_string "___a" in let b = Names.Id.of_string "___b" in Constrexpr_ops.mkLambdaC( [CAst.make @@ Name a; CAst.make @@ Name b], Constrexpr.Default Glob_term.Explicit, wf_arg_type, Constrexpr_ops.mkAppC(wf_rel_expr, [ Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC a]); Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC b]) ]) ) in wf_rel_with_mes,false in register_wf interactive_proof ~is_mes:is_mes fname rec_impls wf_rel_from_mes wf_arg using_lemmas args ret_type body let do_generate_principle_aux pconstants on_error register_built interactive_proof fixpoint_exprl : Lemmas.t option = List.iter (fun { Vernacexpr.notations } -> if not (List.is_empty notations) then CErrors.user_err (Pp.str "Function does not support notations for now")) fixpoint_exprl; let lemma, _is_struct = match fixpoint_exprl with | [{ Vernacexpr.rec_order = Some {CAst.v = Constrexpr.CWfRec (wf_x,wf_rel)} } as fixpoint_expr] -> let { Vernacexpr.fname; univs; binders; rtype; body_def } as fixpoint_expr = match recompute_binder_list [fixpoint_expr] with | [e] -> e | _ -> assert false in let fixpoint_exprl = [fixpoint_expr] in let body = match body_def with | Some body -> body | None -> CErrors.user_err ~hdr:"Function" (Pp.str "Body of Function must be given") in let recdefs,rec_impls = build_newrecursive fixpoint_exprl in let using_lemmas = [] in let pre_hook pconstants = generate_principle (ref (Evd.from_env (Global.env ()))) pconstants on_error true register_built fixpoint_exprl recdefs true in if register_built then register_wf interactive_proof fname.CAst.v rec_impls wf_rel wf_x.CAst.v using_lemmas binders rtype body pre_hook, false else None, false | [{ Vernacexpr.rec_order = Some {CAst.v = Constrexpr.CMeasureRec(wf_x,wf_mes,wf_rel_opt)} } as fixpoint_expr] -> let { Vernacexpr.fname; univs; binders; rtype; body_def} as fixpoint_expr = match recompute_binder_list [fixpoint_expr] with | [e] -> e | _ -> assert false in let fixpoint_exprl = [fixpoint_expr] in let recdefs,rec_impls = build_newrecursive fixpoint_exprl in let using_lemmas = [] in let body = match body_def with | Some body -> body | None -> CErrors.user_err ~hdr:"Function" Pp.(str "Body of Function must be given") in let pre_hook pconstants = generate_principle (ref (Evd.from_env (Global.env ()))) pconstants on_error true register_built fixpoint_exprl recdefs true in if register_built then register_mes interactive_proof fname.CAst.v rec_impls wf_mes wf_rel_opt (Option.map (fun x -> x.CAst.v) wf_x) using_lemmas binders rtype body pre_hook, true else None, true | _ -> List.iter (function { Vernacexpr.rec_order } -> match rec_order with | Some { CAst.v = (Constrexpr.CMeasureRec _ | Constrexpr.CWfRec _) } -> CErrors.user_err (Pp.str "Cannot use mutual definition with well-founded recursion or measure") | _ -> () ) fixpoint_exprl; let fixpoint_exprl = recompute_binder_list fixpoint_exprl in let fix_names = List.map (function { Vernacexpr.fname } -> fname.CAst.v) fixpoint_exprl in (* ok all the expressions are structural *) let recdefs,rec_impls = build_newrecursive fixpoint_exprl in let is_rec = List.exists (is_rec fix_names) recdefs in let lemma,evd,pconstants = if register_built then register_struct is_rec fixpoint_exprl else None, Evd.from_env (Global.env ()), pconstants in let evd = ref evd in generate_principle (ref !evd) pconstants on_error false register_built fixpoint_exprl recdefs interactive_proof (Functional_principles_proofs.prove_princ_for_struct evd interactive_proof); if register_built then begin derive_inversion fix_names; end; lemma, true in lemma let warn_cannot_define_graph = CWarnings.create ~name:"funind-cannot-define-graph" ~category:"funind" (fun (names,error) -> Pp.(strbrk "Cannot define graph(s) for " ++ h 1 names ++ error)) let warn_cannot_define_principle = CWarnings.create ~name:"funind-cannot-define-principle" ~category:"funind" (fun (names,error) -> Pp.(strbrk "Cannot define induction principle(s) for "++ h 1 names ++ error)) let warning_error names e = let e_explain e = match e with | ToShow e -> Pp.(spc () ++ CErrors.print e) | _ -> if do_observe () then Pp.(spc () ++ CErrors.print e) else Pp.mt () in match e with | Building_graph e -> let names = Pp.(prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) in warn_cannot_define_graph (names,e_explain e) | Defining_principle e -> let names = Pp.(prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) in warn_cannot_define_principle (names,e_explain e) | _ -> raise e let error_error names e = let e_explain e = match e with | ToShow e -> Pp.(spc () ++ CErrors.print e) | _ -> if do_observe () then Pp.(spc () ++ CErrors.print e) else Pp.mt () in match e with | Building_graph e -> CErrors.user_err Pp.(str "Cannot define graph(s) for " ++ h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++ e_explain e) | _ -> raise e (* [chop_n_arrow n t] chops the [n] first arrows in [t] Acts on Constrexpr.constr_expr *) let rec chop_n_arrow n t = let exception Stop of Constrexpr.constr_expr in let open Constrexpr in if n <= 0 then t (* If we have already removed all the arrows then return the type *) else (* If not we check the form of [t] *) match t.CAst.v with | Constrexpr.CProdN(nal_ta',t') -> (* If we have a forall, two results are possible : either we need to discard more than the number of arrows contained in this product declaration then we just recall [chop_n_arrow] on the remaining number of arrow to chop and [t'] we discard it and recall [chop_n_arrow], either this product contains more arrows than the number we need to chop and then we return the new type *) begin try let new_n = let rec aux (n:int) = function [] -> n | CLocalAssum(nal,k,t'')::nal_ta' -> let nal_l = List.length nal in if n >= nal_l then aux (n - nal_l) nal_ta' else let new_t' = CAst.make @@ Constrexpr.CProdN( CLocalAssum((snd (List.chop n nal)),k,t'')::nal_ta',t') in raise (Stop new_t') | _ -> CErrors.anomaly (Pp.str "Not enough products.") in aux n nal_ta' in chop_n_arrow new_n t' with Stop t -> t end | _ -> CErrors.anomaly (Pp.str "Not enough products.") let rec add_args id new_args = let open Libnames in let open Constrexpr in CAst.map (function | CRef (qid,_) as b -> if qualid_is_ident qid && Id.equal (qualid_basename qid) id then CAppExpl((None,qid,None),new_args) else b | CFix _ | CCoFix _ -> CErrors.anomaly ~label:"add_args " (Pp.str "todo.") | CProdN(nal,b1) -> CProdN(List.map (function CLocalAssum (nal,k,b2) -> CLocalAssum (nal,k,add_args id new_args b2) | CLocalDef (na,b1,t) -> CLocalDef (na,add_args id new_args b1,Option.map (add_args id new_args) t) | CLocalPattern _ -> CErrors.user_err (Pp.str "pattern with quote not allowed here.")) nal, add_args id new_args b1) | CLambdaN(nal,b1) -> CLambdaN(List.map (function CLocalAssum (nal,k,b2) -> CLocalAssum (nal,k,add_args id new_args b2) | CLocalDef (na,b1,t) -> CLocalDef (na,add_args id new_args b1,Option.map (add_args id new_args) t) | CLocalPattern _ -> CErrors.user_err (Pp.str "pattern with quote not allowed here.")) nal, add_args id new_args b1) | CLetIn(na,b1,t,b2) -> CLetIn(na,add_args id new_args b1,Option.map (add_args id new_args) t,add_args id new_args b2) | CAppExpl((pf,qid,us),exprl) -> if qualid_is_ident qid && Id.equal (qualid_basename qid) id then CAppExpl((pf,qid,us),new_args@(List.map (add_args id new_args) exprl)) else CAppExpl((pf,qid,us),List.map (add_args id new_args) exprl) | CApp((pf,b),bl) -> CApp((pf,add_args id new_args b), List.map (fun (e,o) -> add_args id new_args e,o) bl) | CCases(sty,b_option,cel,cal) -> CCases(sty,Option.map (add_args id new_args) b_option, List.map (fun (b,na,b_option) -> add_args id new_args b, na, b_option) cel, List.map CAst.(map (fun (cpl,e) -> (cpl,add_args id new_args e))) cal ) | CLetTuple(nal,(na,b_option),b1,b2) -> CLetTuple(nal,(na,Option.map (add_args id new_args) b_option), add_args id new_args b1, add_args id new_args b2 ) | CIf(b1,(na,b_option),b2,b3) -> CIf(add_args id new_args b1, (na,Option.map (add_args id new_args) b_option), add_args id new_args b2, add_args id new_args b3 ) | CHole _ | CPatVar _ | CEvar _ | CPrim _ | CSort _ as b -> b | CCast(b1,b2) -> CCast(add_args id new_args b1, Glob_ops.map_cast_type (add_args id new_args) b2) | CRecord pars -> CRecord (List.map (fun (e,o) -> e, add_args id new_args o) pars) | CNotation _ -> CErrors.anomaly ~label:"add_args " (Pp.str "CNotation.") | CGeneralization _ -> CErrors.anomaly ~label:"add_args " (Pp.str "CGeneralization.") | CDelimiters _ -> CErrors.anomaly ~label:"add_args " (Pp.str "CDelimiters.") ) let rec get_args b t : Constrexpr.local_binder_expr list * Constrexpr.constr_expr * Constrexpr.constr_expr = let open Constrexpr in match b.CAst.v with | Constrexpr.CLambdaN (CLocalAssum(nal,k,ta) as d::rest, b') -> begin let n = List.length nal in let nal_tas,b'',t'' = get_args (CAst.make ?loc:b.CAst.loc @@ Constrexpr.CLambdaN (rest,b')) (chop_n_arrow n t) in d :: nal_tas, b'',t'' end | Constrexpr.CLambdaN ([], b) -> [],b,t | _ -> [],b,t let make_graph (f_ref : GlobRef.t) = let open Constrexpr in let env = Global.env() in let sigma = Evd.from_env env in let c,c_body = match f_ref with | GlobRef.ConstRef c -> begin try c,Global.lookup_constant c with Not_found -> CErrors.user_err Pp.(str "Cannot find " ++ Printer.pr_leconstr_env env sigma (EConstr.mkConst c)) end | _ -> CErrors.user_err Pp.(str "Not a function reference") in (match Global.body_of_constant_body Library.indirect_accessor c_body with | None -> CErrors.user_err (Pp.str "Cannot build a graph over an axiom!") | Some (body, _, _) -> let env = Global.env () in let extern_body,extern_type = with_full_print (fun () -> (Constrextern.extern_constr false env sigma (EConstr.of_constr body), Constrextern.extern_type false env sigma (EConstr.of_constr (*FIXME*) c_body.Declarations.const_type) ) ) () in let (nal_tas,b,t) = get_args extern_body extern_type in let expr_list = match b.CAst.v with | Constrexpr.CFix(l_id,fixexprl) -> let l = List.map (fun (id,recexp,bl,t,b) -> let { CAst.loc; v=rec_id } = match Option.get recexp with | { CAst.v = CStructRec id } -> id | { CAst.v = CWfRec (id,_) } -> id | { CAst.v = CMeasureRec (oid,_,_) } -> Option.get oid in let new_args = List.flatten (List.map (function | Constrexpr.CLocalDef (na,_,_)-> [] | Constrexpr.CLocalAssum (nal,_,_) -> List.map (fun {CAst.loc;v=n} -> CAst.make ?loc @@ CRef(Libnames.qualid_of_ident ?loc @@ Nameops.Name.get_id n,None)) nal | Constrexpr.CLocalPattern _ -> assert false ) nal_tas ) in let b' = add_args id.CAst.v new_args b in { Vernacexpr.fname=id; univs=None ; rec_order = Some (CAst.make (CStructRec (CAst.make rec_id))) ; binders = nal_tas@bl; rtype=t; body_def=Some b'; notations = []} ) fixexprl in l | _ -> let fname = CAst.make (Label.to_id (Constant.label c)) in [{ Vernacexpr.fname; univs=None; rec_order = None; binders=nal_tas; rtype=t; body_def=Some b; notations=[]}] in let mp = Constant.modpath c in let pstate = do_generate_principle_aux [c,Univ.Instance.empty] error_error false false expr_list in assert (Option.is_empty pstate); (* We register the infos *) List.iter (fun { Vernacexpr.fname= {CAst.v=id} } -> add_Function false (Constant.make2 mp (Label.of_id id))) expr_list) (* *************** statically typed entrypoints ************************* *) let do_generate_principle_interactive fixl : Lemmas.t = match do_generate_principle_aux [] warning_error true true fixl with | Some lemma -> lemma | None -> CErrors.anomaly (Pp.str"indfun: leaving no open proof in interactive mode") let do_generate_principle fixl : unit = match do_generate_principle_aux [] warning_error true false fixl with | Some _lemma -> CErrors.anomaly (Pp.str"indfun: leaving a goal open in non-interactive mode") | None -> () let build_scheme fas = let evd = (ref (Evd.from_env (Global.env ()))) in let pconstants = (List.map (fun (_,f,sort) -> let f_as_constant = try Smartlocate.global_with_alias f with Not_found -> CErrors.user_err ~hdr:"FunInd.build_scheme" Pp.(str "Cannot find " ++ Libnames.pr_qualid f) in let evd',f = Evd.fresh_global (Global.env ()) !evd f_as_constant in let _ = evd := evd' in let sigma, _ = Typing.type_of ~refresh:true (Global.env ()) !evd f in evd := sigma; let c, u = try EConstr.destConst !evd f with Constr.DestKO -> CErrors.user_err Pp.(Printer.pr_econstr_env (Global.env ()) !evd f ++spc () ++ str "should be the named of a globally defined function") in (c, EConstr.EInstance.kind !evd u), sort ) fas ) in let bodies_types = make_scheme evd pconstants in List.iter2 (fun (princ_id,_,_) def_entry -> ignore (Declare.declare_constant ~name:princ_id ~kind:Decls.(IsProof Theorem) (Declare.DefinitionEntry def_entry)); Declare.definition_message princ_id ) fas bodies_types let build_case_scheme fa = let env = Global.env () and sigma = (Evd.from_env (Global.env ())) in (* let id_to_constr id = *) (* Constrintern.global_reference id *) (* in *) let funs = let (_,f,_) = fa in try (let open GlobRef in match Smartlocate.global_with_alias f with | ConstRef c -> c | IndRef _ | ConstructRef _ | VarRef _ -> assert false) with Not_found -> CErrors.user_err ~hdr:"FunInd.build_case_scheme" Pp.(str "Cannot find " ++ Libnames.pr_qualid f) in let sigma, (_,u) = Evd.fresh_constant_instance env sigma funs in let first_fun = funs in let funs_mp = Constant.modpath first_fun in let first_fun_kn = match find_Function_infos first_fun with | None -> raise No_graph_found | Some finfos -> fst finfos.graph_ind in let this_block_funs_indexes = get_funs_constant funs_mp first_fun in let this_block_funs = Array.map (fun (c,_) -> (c,u)) this_block_funs_indexes in let prop_sort = Sorts.InProp in let funs_indexes = let this_block_funs_indexes = Array.to_list this_block_funs_indexes in List.assoc_f Constant.equal funs this_block_funs_indexes in let (ind, sf) = let ind = first_fun_kn,funs_indexes in (ind,Univ.Instance.empty)(*FIXME*),prop_sort in let (sigma, scheme) = Indrec.build_case_analysis_scheme_default env sigma ind sf in let scheme_type = EConstr.Unsafe.to_constr ((Typing.unsafe_type_of env sigma) (EConstr.of_constr scheme)) in let sorts = (fun (_,_,x) -> fst @@ UnivGen.fresh_sort_in_family x ) fa in let princ_name = (fun (x,_,_) -> x) fa in let _ : unit = (* Pp.msgnl (str "Generating " ++ Ppconstr.pr_id princ_name ++str " with " ++ pr_lconstr scheme_type ++ str " and " ++ (fun a -> prlist_with_sep spc (fun c -> pr_lconstr (mkConst c)) (Array.to_list a)) this_block_funs ); *) generate_functional_principle (ref (Evd.from_env (Global.env ()))) false scheme_type (Some ([|sorts|])) (Some princ_name) this_block_funs 0 (Functional_principles_proofs.prove_princ_for_struct (ref (Evd.from_env (Global.env ()))) false 0 [|funs|]) in () coq-8.11.0/plugins/funind/glob_termops.mli0000644000175000017500000000764113612664533020441 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Id.t list (* [pattern_to_term pat] returns a glob_constr corresponding to [pat]. [pat] must not contain occurrences of anonymous pattern *) val pattern_to_term : cases_pattern -> glob_constr (* Some basic functions to rebuild glob_constr In each of them the location is Util.Loc.ghost *) val mkGRef : GlobRef.t -> glob_constr val mkGVar : Id.t -> glob_constr val mkGApp : glob_constr*(glob_constr list) -> glob_constr val mkGLambda : Name.t * glob_constr * glob_constr -> glob_constr val mkGProd : Name.t * glob_constr * glob_constr -> glob_constr val mkGLetIn : Name.t * glob_constr * glob_constr option * glob_constr -> glob_constr val mkGCases : glob_constr option * tomatch_tuples * cases_clauses -> glob_constr val mkGHole : unit -> glob_constr (* we only build Evd.BinderType Anonymous holes *) (* Some basic functions to decompose glob_constrs These are analogous to the ones constrs *) val glob_decompose_app : glob_constr -> glob_constr*(glob_constr list) (* [glob_make_eq t1 t2] build the glob_constr corresponding to [t2 = t1] *) val glob_make_eq : ?typ:glob_constr -> glob_constr -> glob_constr -> glob_constr (* [glob_make_neq t1 t2] build the glob_constr corresponding to [t1 <> t2] *) val glob_make_neq : glob_constr -> glob_constr -> glob_constr (* alpha_conversion functions *) (* Replace the var mapped in the glob_constr/context *) val change_vars : Id.t Id.Map.t -> glob_constr -> glob_constr (* [alpha_pat avoid pat] rename all the variables present in [pat] s.t. the result does not share variables with [avoid]. This function create a fresh variable for each occurrence of the anonymous pattern. Also returns a mapping from old variables to new ones and the concatenation of [avoid] with the variables appearing in the result. *) val alpha_pat : Id.Map.key list -> Glob_term.cases_pattern -> Glob_term.cases_pattern * Id.Map.key list * Id.t Id.Map.t (* [alpha_rt avoid rt] alpha convert [rt] s.t. the result respects barendregt conventions and does not share bound variables with avoid *) val alpha_rt : Id.t list -> glob_constr -> glob_constr (* same as alpha_rt but for case branches *) val alpha_br : Id.t list -> Glob_term.cases_clause -> Glob_term.cases_clause (* Reduction function *) val replace_var_by_term : Id.t -> Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr (* [is_free_in id rt] checks if [id] is a free variable in [rt] *) val is_free_in : Id.t -> glob_constr -> bool val are_unifiable : cases_pattern -> cases_pattern -> bool val eq_cases_pattern : cases_pattern -> cases_pattern -> bool (* ids_of_pat : cases_pattern -> Id.Set.t returns the set of variables appearing in a pattern *) val ids_of_pat : cases_pattern -> Id.Set.t val expand_as : glob_constr -> glob_constr (* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution *) val resolve_and_replace_implicits : ?flags:Pretyping.inference_flags -> ?expected_type:Pretyping.typing_constraint -> Environ.env -> Evd.evar_map -> glob_constr -> glob_constr coq-8.11.0/plugins/funind/glob_term_to_relation.mli0000644000175000017500000000117513612664533022312 0ustar treinentreinenopen Names (* [build_inductive parametrize funnames funargs returned_types bodies] constructs and saves the graphs of the functions [funnames] taking [funargs] as arguments and returning [returned_types] using bodies [bodies] *) val build_inductive : (* (ModPath.t * DirPath.t) option -> Id.t list -> (* The list of function name *) *) Evd.evar_map -> Constr.pconstant list -> (Name.t*Glob_term.glob_constr*Glob_term.glob_constr option) list list -> (* The list of function args *) Constrexpr.constr_expr list -> (* The list of function returned type *) Glob_term.glob_constr list -> (* the list of body *) unit coq-8.11.0/plugins/funind/indfun_common.ml0000644000175000017500000004101213612664533020415 0ustar treinentreinenopen Names open Pp open Constr open Libnames open Refiner let mk_prefix pre id = Id.of_string (pre^(Id.to_string id)) let mk_rel_id = mk_prefix "R_" let mk_correct_id id = Nameops.add_suffix (mk_rel_id id) "_correct" let mk_complete_id id = Nameops.add_suffix (mk_rel_id id) "_complete" let mk_equation_id id = Nameops.add_suffix id "_equation" let fresh_id avoid s = Namegen.next_ident_away_in_goal (Id.of_string s) (Id.Set.of_list avoid) let fresh_name avoid s = Name (fresh_id avoid s) let get_name avoid ?(default="H") = function | Anonymous -> fresh_name avoid default | Name n -> Name n let array_get_start a = Array.init (Array.length a - 1) (fun i -> a.(i)) let locate qid = Nametab.locate qid let locate_ind ref = match locate ref with | GlobRef.IndRef x -> x | _ -> raise Not_found let locate_constant ref = match locate ref with | GlobRef.ConstRef x -> x | _ -> raise Not_found let locate_with_msg msg f x = try f x with | Not_found -> CErrors.user_err msg let filter_map filter f = let rec it = function | [] -> [] | e::l -> if filter e then (f e) :: it l else it l in it let chop_rlambda_n = let rec chop_lambda_n acc n rt = if n == 0 then List.rev acc,rt else match DAst.get rt with | Glob_term.GLambda(name,k,t,b) -> chop_lambda_n ((name,t,None)::acc) (n-1) b | Glob_term.GLetIn(name,v,t,b) -> chop_lambda_n ((name,v,t)::acc) (n-1) b | _ -> CErrors.user_err ~hdr:"chop_rlambda_n" (str "chop_rlambda_n: Not enough Lambdas") in chop_lambda_n [] let chop_rprod_n = let rec chop_prod_n acc n rt = if n == 0 then List.rev acc,rt else match DAst.get rt with | Glob_term.GProd(name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b | _ -> CErrors.user_err ~hdr:"chop_rprod_n" (str "chop_rprod_n: Not enough products") in chop_prod_n [] let list_union_eq eq_fun l1 l2 = let rec urec = function | [] -> l2 | a::l -> if List.exists (eq_fun a) l2 then urec l else a::urec l in urec l1 let list_add_set_eq eq_fun x l = if List.exists (eq_fun x) l then l else x::l [@@@ocaml.warning "-3"] let coq_constant s = UnivGen.constr_of_monomorphic_global @@ Coqlib.gen_reference_in_modules "RecursiveDefinition" Coqlib.init_modules s;; let find_reference sl s = let dp = Names.DirPath.make (List.rev_map Id.of_string sl) in Nametab.locate (make_qualid dp (Id.of_string s)) let eq = lazy(EConstr.of_constr (coq_constant "eq")) let refl_equal = lazy(EConstr.of_constr (coq_constant "eq_refl")) let with_full_print f a = let old_implicit_args = Impargs.is_implicit_args () and old_strict_implicit_args = Impargs.is_strict_implicit_args () and old_contextual_implicit_args = Impargs.is_contextual_implicit_args () in let old_rawprint = !Flags.raw_print in let old_printuniverses = !Constrextern.print_universes in let old_printallowmatchdefaultclause = !Detyping.print_allow_match_default_clause in Constrextern.print_universes := true; Detyping.print_allow_match_default_clause := false; Flags.raw_print := true; Impargs.make_implicit_args false; Impargs.make_strict_implicit_args false; Impargs.make_contextual_implicit_args false; Dumpglob.pause (); try let res = f a in Impargs.make_implicit_args old_implicit_args; Impargs.make_strict_implicit_args old_strict_implicit_args; Impargs.make_contextual_implicit_args old_contextual_implicit_args; Flags.raw_print := old_rawprint; Constrextern.print_universes := old_printuniverses; Detyping.print_allow_match_default_clause := old_printallowmatchdefaultclause; Dumpglob.continue (); res with | reraise -> Impargs.make_implicit_args old_implicit_args; Impargs.make_strict_implicit_args old_strict_implicit_args; Impargs.make_contextual_implicit_args old_contextual_implicit_args; Flags.raw_print := old_rawprint; Constrextern.print_universes := old_printuniverses; Detyping.print_allow_match_default_clause := old_printallowmatchdefaultclause; Dumpglob.continue (); raise reraise (**********************) type function_info = { function_constant : Constant.t; graph_ind : inductive; equation_lemma : Constant.t option; correctness_lemma : Constant.t option; completeness_lemma : Constant.t option; rect_lemma : Constant.t option; rec_lemma : Constant.t option; prop_lemma : Constant.t option; sprop_lemma : Constant.t option; is_general : bool; (* Has this function been defined using general recursive definition *) } (* type function_db = function_info list *) (* let function_table = ref ([] : function_db) *) let from_function = Summary.ref Cmap_env.empty ~name:"functions_db_fn" let from_graph = Summary.ref Indmap.empty ~name:"functions_db_gr" (* let rec do_cache_info finfo = function | [] -> raise Not_found | (finfo'::finfos as l) -> if finfo' == finfo then l else if finfo'.function_constant = finfo.function_constant then finfo::finfos else let res = do_cache_info finfo finfos in if res == finfos then l else finfo'::l let cache_Function (_,(finfos)) = let new_tbl = try do_cache_info finfos !function_table with Not_found -> finfos::!function_table in if new_tbl != !function_table then function_table := new_tbl *) let cache_Function (_,finfos) = from_function := Cmap_env.add finfos.function_constant finfos !from_function; from_graph := Indmap.add finfos.graph_ind finfos !from_graph let subst_Function (subst,finfos) = let do_subst_con c = Mod_subst.subst_constant subst c and do_subst_ind i = Mod_subst.subst_ind subst i in let function_constant' = do_subst_con finfos.function_constant in let graph_ind' = do_subst_ind finfos.graph_ind in let equation_lemma' = Option.Smart.map do_subst_con finfos.equation_lemma in let correctness_lemma' = Option.Smart.map do_subst_con finfos.correctness_lemma in let completeness_lemma' = Option.Smart.map do_subst_con finfos.completeness_lemma in let rect_lemma' = Option.Smart.map do_subst_con finfos.rect_lemma in let rec_lemma' = Option.Smart.map do_subst_con finfos.rec_lemma in let prop_lemma' = Option.Smart.map do_subst_con finfos.prop_lemma in let sprop_lemma' = Option.Smart.map do_subst_con finfos.sprop_lemma in if function_constant' == finfos.function_constant && graph_ind' == finfos.graph_ind && equation_lemma' == finfos.equation_lemma && correctness_lemma' == finfos.correctness_lemma && completeness_lemma' == finfos.completeness_lemma && rect_lemma' == finfos.rect_lemma && rec_lemma' == finfos.rec_lemma && prop_lemma' == finfos.prop_lemma && sprop_lemma' == finfos.sprop_lemma then finfos else { function_constant = function_constant'; graph_ind = graph_ind'; equation_lemma = equation_lemma' ; correctness_lemma = correctness_lemma' ; completeness_lemma = completeness_lemma' ; rect_lemma = rect_lemma' ; rec_lemma = rec_lemma'; prop_lemma = prop_lemma'; sprop_lemma = sprop_lemma'; is_general = finfos.is_general } let discharge_Function (_,finfos) = Some finfos let pr_ocst env sigma c = Option.fold_right (fun v acc -> Printer.pr_lconstr_env env sigma (mkConst v)) c (mt ()) let pr_info env sigma f_info = str "function_constant := " ++ Printer.pr_lconstr_env env sigma (mkConst f_info.function_constant)++ fnl () ++ str "function_constant_type := " ++ (try Printer.pr_lconstr_env env sigma (fst (Typeops.type_of_global_in_context env (GlobRef.ConstRef f_info.function_constant))) with e when CErrors.noncritical e -> mt ()) ++ fnl () ++ str "equation_lemma := " ++ pr_ocst env sigma f_info.equation_lemma ++ fnl () ++ str "completeness_lemma :=" ++ pr_ocst env sigma f_info.completeness_lemma ++ fnl () ++ str "correctness_lemma := " ++ pr_ocst env sigma f_info.correctness_lemma ++ fnl () ++ str "rect_lemma := " ++ pr_ocst env sigma f_info.rect_lemma ++ fnl () ++ str "rec_lemma := " ++ pr_ocst env sigma f_info.rec_lemma ++ fnl () ++ str "prop_lemma := " ++ pr_ocst env sigma f_info.prop_lemma ++ fnl () ++ str "graph_ind := " ++ Printer.pr_lconstr_env env sigma (mkInd f_info.graph_ind) ++ fnl () let pr_table env sigma tb = let l = Cmap_env.fold (fun k v acc -> v::acc) tb [] in Pp.prlist_with_sep fnl (pr_info env sigma) l let in_Function : function_info -> Libobject.obj = let open Libobject in declare_object @@ superglobal_object "FUNCTIONS_DB" ~cache:cache_Function ~subst:(Some subst_Function) ~discharge:discharge_Function let find_or_none id = try Some (match Nametab.locate (qualid_of_ident id) with GlobRef.ConstRef c -> c | _ -> CErrors.anomaly (Pp.str "Not a constant.") ) with Not_found -> None let find_Function_infos f = Cmap_env.find_opt f !from_function let find_Function_of_graph ind = Indmap.find_opt ind !from_graph let update_Function finfo = (* Pp.msgnl (pr_info finfo); *) Lib.add_anonymous_leaf (in_Function finfo) let add_Function is_general f = let f_id = Label.to_id (Constant.label f) in let equation_lemma = find_or_none (mk_equation_id f_id) and correctness_lemma = find_or_none (mk_correct_id f_id) and completeness_lemma = find_or_none (mk_complete_id f_id) and rect_lemma = find_or_none (Nameops.add_suffix f_id "_rect") and rec_lemma = find_or_none (Nameops.add_suffix f_id "_rec") and prop_lemma = find_or_none (Nameops.add_suffix f_id "_ind") and sprop_lemma = find_or_none (Nameops.add_suffix f_id "_sind") and graph_ind = match Nametab.locate (qualid_of_ident (mk_rel_id f_id)) with | GlobRef.IndRef ind -> ind | _ -> CErrors.anomaly (Pp.str "Not an inductive.") in let finfos = { function_constant = f; equation_lemma = equation_lemma; completeness_lemma = completeness_lemma; correctness_lemma = correctness_lemma; rect_lemma = rect_lemma; rec_lemma = rec_lemma; prop_lemma = prop_lemma; sprop_lemma = sprop_lemma; graph_ind = graph_ind; is_general = is_general } in update_Function finfos let pr_table env sigma = pr_table env sigma !from_function (*********************************) (* Debugging *) let functional_induction_rewrite_dependent_proofs = ref true let function_debug = ref false open Goptions let functional_induction_rewrite_dependent_proofs_sig = { optdepr = false; optname = "Functional Induction Rewrite Dependent"; optkey = ["Functional";"Induction";"Rewrite";"Dependent"]; optread = (fun () -> !functional_induction_rewrite_dependent_proofs); optwrite = (fun b -> functional_induction_rewrite_dependent_proofs := b) } let () = declare_bool_option functional_induction_rewrite_dependent_proofs_sig let do_rewrite_dependent () = !functional_induction_rewrite_dependent_proofs = true let function_debug_sig = { optdepr = false; optname = "Function debug"; optkey = ["Function_debug"]; optread = (fun () -> !function_debug); optwrite = (fun b -> function_debug := b) } let () = declare_bool_option function_debug_sig let do_observe () = !function_debug let observe strm = if do_observe () then Feedback.msg_debug strm else () let debug_queue = Stack.create () let print_debug_queue b e = if not (Stack.is_empty debug_queue) then let lmsg,goal = Stack.pop debug_queue in (if b then Feedback.msg_debug (hov 1 (lmsg ++ (str " raised exception " ++ CErrors.print e) ++ str " on goal" ++ fnl() ++ goal)) else Feedback.msg_debug (hov 1 (str " from " ++ lmsg ++ str " on goal"++fnl() ++ goal)) (* print_debug_queue false e; *) ) let do_observe_tac s tac g = let goal = Printer.pr_goal g in let s = s (pf_env g) (project g) in let lmsg = (str "observation : ") ++ s in Stack.push (lmsg,goal) debug_queue; try let v = tac g in ignore(Stack.pop debug_queue); v with reraise -> let reraise = CErrors.push reraise in if not (Stack.is_empty debug_queue) then print_debug_queue true (fst reraise); Util.iraise reraise let observe_tac s tac g = if do_observe () then do_observe_tac s tac g else tac g module New = struct let do_observe_tac ~header s tac = let open Proofview.Notations in let open Proofview in Goal.enter begin fun gl -> let goal = Printer.pr_goal (Goal.print gl) in let env, sigma = Goal.env gl, Goal.sigma gl in let s = s env sigma in let lmsg = seq [header; str " : " ++ s] in tclLIFT (NonLogical.make (fun () -> Feedback.msg_debug (s++fnl()))) >>= fun () -> tclOR ( Stack.push (lmsg, goal) debug_queue; tac >>= fun v -> ignore(Stack.pop debug_queue); Proofview.tclUNIT v) (fun (exn, info) -> if not (Stack.is_empty debug_queue) then print_debug_queue true exn; tclZERO ~info exn) end let observe_tac ~header s tac = if do_observe () then do_observe_tac ~header s tac else tac end let strict_tcc = ref false let is_strict_tcc () = !strict_tcc let strict_tcc_sig = { optdepr = false; optname = "Raw Function Tcc"; optkey = ["Function_raw_tcc"]; optread = (fun () -> !strict_tcc); optwrite = (fun b -> strict_tcc := b) } let () = declare_bool_option strict_tcc_sig exception Building_graph of exn exception Defining_principle of exn exception ToShow of exn let jmeq () = try Coqlib.check_required_library Coqlib.jmeq_module_name; EConstr.of_constr @@ UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.JMeq.type" with e when CErrors.noncritical e -> raise (ToShow e) let jmeq_refl () = try Coqlib.check_required_library Coqlib.jmeq_module_name; EConstr.of_constr @@ UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.JMeq.refl" with e when CErrors.noncritical e -> raise (ToShow e) let h_intros l = tclMAP (fun x -> Proofview.V82.of_tactic (Tactics.Simple.intro x)) l let h_id = Id.of_string "h" let hrec_id = Id.of_string "hrec" let well_founded = function () -> EConstr.of_constr (coq_constant "well_founded") let acc_rel = function () -> EConstr.of_constr (coq_constant "Acc") let acc_inv_id = function () -> EConstr.of_constr (coq_constant "Acc_inv") [@@@ocaml.warning "-3"] let well_founded_ltof () = EConstr.of_constr @@ UnivGen.constr_of_monomorphic_global @@ Coqlib.find_reference "IndFun" ["Coq"; "Arith";"Wf_nat"] "well_founded_ltof" [@@@ocaml.warning "+3"] let ltof_ref = function () -> (find_reference ["Coq";"Arith";"Wf_nat"] "ltof") let make_eq () = try EConstr.of_constr (UnivGen.constr_of_monomorphic_global (Coqlib.lib_ref "core.eq.type")) with _ -> assert false let evaluable_of_global_reference r = (* Tacred.evaluable_of_global_reference (Global.env ()) *) match r with GlobRef.ConstRef sp -> EvalConstRef sp | GlobRef.VarRef id -> EvalVarRef id | _ -> assert false;; let list_rewrite (rev:bool) (eqs: (EConstr.constr*bool) list) = tclREPEAT (List.fold_right (fun (eq,b) i -> tclORELSE (Proofview.V82.of_tactic ((if b then Equality.rewriteLR else Equality.rewriteRL) eq)) i) (if rev then (List.rev eqs) else eqs) (tclFAIL 0 (mt())));; let decompose_lam_n sigma n = if n < 0 then CErrors.user_err Pp.(str "decompose_lam_n: integer parameter must be positive"); let rec lamdec_rec l n c = if Int.equal n 0 then l,c else match EConstr.kind sigma c with | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c | Cast (c,_,_) -> lamdec_rec l n c | _ -> CErrors.user_err Pp.(str "decompose_lam_n: not enough abstractions") in lamdec_rec [] n let lamn n env b = let open EConstr in let rec lamrec = function | (0, env, b) -> b | (n, ((v,t)::l), b) -> lamrec (n-1, l, mkLambda (v,t,b)) | _ -> assert false in lamrec (n,env,b) (* compose_lam [xn:Tn;..;x1:T1] b = [x1:T1]..[xn:Tn]b *) let compose_lam l b = lamn (List.length l) l b (* prodn n [xn:Tn;..;x1:T1;Gamma] b = (x1:T1)..(xn:Tn)b *) let prodn n env b = let open EConstr in let rec prodrec = function | (0, env, b) -> b | (n, ((v,t)::l), b) -> prodrec (n-1, l, mkProd (v,t,b)) | _ -> assert false in prodrec (n,env,b) (* compose_prod [xn:Tn;..;x1:T1] b = (x1:T1)..(xn:Tn)b *) let compose_prod l b = prodn (List.length l) l b type tcc_lemma_value = | Undefined | Value of constr | Not_needed (* We only "purify" on exceptions. XXX: What is this doing here? *) let funind_purify f x = let st = Vernacstate.freeze_interp_state ~marshallable:false in try f x with e -> let e = CErrors.push e in Vernacstate.unfreeze_interp_state st; Exninfo.iraise e coq-8.11.0/plugins/funind/g_indfun.mlg0000644000175000017500000002053013612664533017524 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* mt () | Some b -> spc () ++ hov 2 (str "using" ++ spc () ++ Miscprint.pr_with_bindings (prc env sigma) (prlc env sigma) b) (* Duplication of printing functions because "'a with_bindings" is (internally) not uniform in 'a: indeed constr_with_bindings at the "typed" level has type "open_constr with_bindings" instead of "constr with_bindings"; hence, its printer cannot be polymorphic in (prc,prlc)... *) let pr_fun_ind_using_typed prc prlc _ opt_c = match opt_c with | None -> mt () | Some b -> let env = Global.env () in let evd = Evd.from_env env in let (_, b) = b env evd in spc () ++ hov 2 (str "using" ++ spc () ++ Miscprint.pr_with_bindings (prc env evd) (prlc env evd) b) } ARGUMENT EXTEND fun_ind_using TYPED AS constr_with_bindings option PRINTED BY { pr_fun_ind_using_typed } RAW_PRINTED BY { pr_fun_ind_using env sigma } GLOB_PRINTED BY { pr_fun_ind_using env sigma } | [ "using" constr_with_bindings(c) ] -> { Some c } | [ ] -> { None } END TACTIC EXTEND newfuninv | [ "functional" "inversion" quantified_hypothesis(hyp) reference_opt(fname) ] -> { Invfun.invfun hyp fname } END { let pr_intro_as_pat _prc _ _ pat = match pat with | Some pat -> spc () ++ str "as" ++ spc () ++ (* Miscprint.pr_intro_pattern prc pat *) str"" | None -> mt () let out_disjunctive = CAst.map (function | IntroAction (IntroOrAndPattern l) -> l | _ -> CErrors.user_err Pp.(str "Disjunctive or conjunctive intro pattern expected.")) } ARGUMENT EXTEND with_names TYPED AS intro_pattern option PRINTED BY { pr_intro_as_pat } | [ "as" simple_intropattern(ipat) ] -> { Some ipat } | [] -> { None } END { let functional_induction b c x pat = functional_induction true c x (Option.map out_disjunctive pat) } TACTIC EXTEND newfunind | ["functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> { let c = match cl with | [] -> assert false | [c] -> c | c::cl -> EConstr.applist(c,cl) in Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl } END (***** debug only ***) TACTIC EXTEND snewfunind | ["soft" "functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> { let c = match cl with | [] -> assert false | [c] -> c | c::cl -> EConstr.applist(c,cl) in Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl } END { let pr_constr_comma_sequence env sigma prc _ _ = prlist_with_sep pr_comma (prc env sigma) } ARGUMENT EXTEND constr_comma_sequence' TYPED AS constr list PRINTED BY { pr_constr_comma_sequence env sigma } | [ constr(c) "," constr_comma_sequence'(l) ] -> { c::l } | [ constr(c) ] -> { [c] } END { let pr_auto_using env sigma prc _prlc _prt = Pptactic.pr_auto_using (prc env sigma) } ARGUMENT EXTEND auto_using' TYPED AS constr list PRINTED BY { pr_auto_using env sigma } | [ "using" constr_comma_sequence'(l) ] -> { l } | [ ] -> { [] } END { module Vernac = Pvernac.Vernac_ module Tactic = Pltac let (wit_function_rec_definition_loc : Vernacexpr.fixpoint_expr Loc.located Genarg.uniform_genarg_type) = Genarg.create_arg "function_rec_definition_loc" let function_rec_definition_loc = Pcoq.create_generic_entry Pcoq.utactic "function_rec_definition_loc" (Genarg.rawwit wit_function_rec_definition_loc) } GRAMMAR EXTEND Gram GLOBAL: function_rec_definition_loc ; function_rec_definition_loc: [ [ g = Vernac.rec_definition -> { Loc.tag ~loc g } ]] ; END { let () = let raw_printer env sigma _ _ _ (loc,body) = Ppvernac.pr_rec_definition body in Pptactic.declare_extra_vernac_genarg_pprule wit_function_rec_definition_loc raw_printer let is_proof_termination_interactively_checked recsl = List.exists (function | _,( Vernacexpr.{ rec_order = Some { CAst.v = CMeasureRec _ } } | Vernacexpr.{ rec_order = Some { CAst.v = CWfRec _} }) -> true | _, Vernacexpr.{ rec_order = Some { CAst.v = CStructRec _ } } | _, Vernacexpr.{ rec_order = None } -> false) recsl let classify_as_Fixpoint recsl = Vernac_classifier.classify_vernac (Vernacexpr.(CAst.make @@ { control = []; attrs = []; expr = VernacFixpoint(NoDischarge, List.map snd recsl)})) let classify_funind recsl = match classify_as_Fixpoint recsl with | Vernacextend.VtSideff (ids, _) when is_proof_termination_interactively_checked recsl -> Vernacextend.(VtStartProof (GuaranteesOpacity, ids)) | x -> x let is_interactive recsl = match classify_funind recsl with | Vernacextend.VtStartProof _ -> true | _ -> false } VERNAC COMMAND EXTEND Function STATE CUSTOM | ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")] => { classify_funind recsl } -> { if is_interactive recsl then Vernacextend.VtOpenProof (fun () -> Gen_principle.do_generate_principle_interactive (List.map snd recsl)) else Vernacextend.VtDefault (fun () -> Gen_principle.do_generate_principle (List.map snd recsl)) } END { let pr_fun_scheme_arg (princ_name,fun_name,s) = Names.Id.print princ_name ++ str " :=" ++ spc() ++ str "Induction for " ++ Libnames.pr_qualid fun_name ++ spc() ++ str "Sort " ++ Sorts.pr_sort_family s } VERNAC ARGUMENT EXTEND fun_scheme_arg PRINTED BY { pr_fun_scheme_arg } | [ ident(princ_name) ":=" "Induction" "for" reference(fun_name) "Sort" sort_family(s) ] -> { (princ_name,fun_name,s) } END { let warning_error names e = match e with | Building_graph e -> let names = pr_enum Libnames.pr_qualid names in let error = if do_observe () then (spc () ++ CErrors.print e) else mt () in Gen_principle.warn_cannot_define_graph (names,error) | Defining_principle e -> let names = pr_enum Libnames.pr_qualid names in let error = if do_observe () then CErrors.print e else mt () in Gen_principle.warn_cannot_define_principle (names,error) | _ -> raise e } VERNAC COMMAND EXTEND NewFunctionalScheme | ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ] => { Vernacextend.(VtSideff(List.map pi1 fas, VtLater)) } -> { begin try Gen_principle.build_scheme fas with | Gen_principle.No_graph_found -> begin match fas with | (_,fun_name,_)::_ -> begin Gen_principle.make_graph (Smartlocate.global_with_alias fun_name); try Gen_principle.build_scheme fas with | Gen_principle.No_graph_found -> CErrors.user_err Pp.(str "Cannot generate induction principle(s)") | e when CErrors.noncritical e -> let names = List.map (fun (_,na,_) -> na) fas in warning_error names e end | _ -> assert false (* we can only have non empty list *) end | e when CErrors.noncritical e -> let names = List.map (fun (_,na,_) -> na) fas in warning_error names e end } END (***** debug only ***) VERNAC COMMAND EXTEND NewFunctionalCase | ["Functional" "Case" fun_scheme_arg(fas) ] => { Vernacextend.(VtSideff([pi1 fas], VtLater)) } -> { Gen_principle.build_case_scheme fas } END (***** debug only ***) VERNAC COMMAND EXTEND GenerateGraph CLASSIFIED AS QUERY | ["Generate" "graph" "for" reference(c)] -> { Gen_principle.make_graph (Smartlocate.global_with_alias c) } END coq-8.11.0/plugins/funind/indfun.mli0000644000175000017500000000155313612664533017224 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* EConstr.constr -> (EConstr.constr * EConstr.constr Tactypes.bindings) option -> Ltac_plugin.Tacexpr.or_and_intro_pattern option -> unit Proofview.tactic coq-8.11.0/plugins/funind/recdef.mli0000644000175000017500000000103613612664533017165 0ustar treinentreinenopen Constr val tclUSER_if_not_mes : unit Proofview.tactic -> bool -> Names.Id.t list option -> unit Proofview.tactic val recursive_definition : interactive_proof:bool -> is_mes:bool -> Names.Id.t -> Constrintern.internalization_env -> Constrexpr.constr_expr -> Constrexpr.constr_expr -> int -> Constrexpr.constr_expr -> (pconstant -> Indfun_common.tcc_lemma_value ref -> pconstant -> pconstant -> int -> EConstr.types -> int -> EConstr.constr -> unit) -> Constrexpr.constr_expr list -> Lemmas.t option coq-8.11.0/plugins/funind/plugin_base.dune0000644000175000017500000000022613612664533020377 0ustar treinentreinen(library (name recdef_plugin) (public_name coq.plugins.funind) (synopsis "Coq's functional induction plugin") (libraries coq.plugins.extraction)) coq-8.11.0/plugins/cc/0000755000175000017500000000000013612664533014334 5ustar treinentreinencoq-8.11.0/plugins/cc/cc_plugin.mlpack0000644000175000017500000000004213612664533017464 0ustar treinentreinenCcalgo Ccproof Cctac G_congruence coq-8.11.0/plugins/cc/ccproof.ml0000644000175000017500000001235413612664533016326 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* prefl (Appli (t1,t2)) | _, _ -> {p_lhs=Appli (p1.p_lhs,p2.p_lhs); p_rhs=Appli (p1.p_rhs,p2.p_rhs); p_rule=Congr (p1,p2)} let rec ptrans p1 p3= match p1.p_rule,p3.p_rule with Refl _, _ ->p3 | _, Refl _ ->p1 | Trans(p1,p2), _ ->ptrans p1 (ptrans p2 p3) | Congr(p1,p2), Congr(p3,p4) ->pcongr (ptrans p1 p3) (ptrans p2 p4) | Congr(p1,p2), Trans({p_rule=Congr(p3,p4)},p5) -> ptrans (pcongr (ptrans p1 p3) (ptrans p2 p4)) p5 | _, _ -> if term_equal p1.p_rhs p3.p_lhs then {p_lhs=p1.p_lhs; p_rhs=p3.p_rhs; p_rule=Trans (p1,p3)} else anomaly (Pp.str "invalid cc transitivity.") let rec psym p = match p.p_rule with Refl _ -> p | SymAx s -> {p_lhs=p.p_rhs; p_rhs=p.p_lhs; p_rule=Ax s} | Ax s-> {p_lhs=p.p_rhs; p_rhs=p.p_lhs; p_rule=SymAx s} | Inject (p0,c,n,a)-> {p_lhs=p.p_rhs; p_rhs=p.p_lhs; p_rule=Inject (psym p0,c,n,a)} | Trans (p1,p2)-> ptrans (psym p2) (psym p1) | Congr (p1,p2)-> pcongr (psym p1) (psym p2) let pax axioms s = let l,r = Constrhash.find axioms s in {p_lhs=l; p_rhs=r; p_rule=Ax s} let psymax axioms s = let l,r = Constrhash.find axioms s in {p_lhs=r; p_rhs=l; p_rule=SymAx s} let rec nth_arg t n= match t with Appli (t1,t2)-> if n>0 then nth_arg t1 (n-1) else t2 | _ -> anomaly ~label:"nth_arg" (Pp.str "not enough args.") let pinject p c n a = {p_lhs=nth_arg p.p_lhs (n-a); p_rhs=nth_arg p.p_rhs (n-a); p_rule=Inject(p,c,n,a)} let rec equal_proof env sigma uf i j= debug (fun () -> str "equal_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ pr_idx_term env sigma uf j); if i=j then prefl (term uf i) else let (li,lj)=join_path uf i j in ptrans (path_proof env sigma uf i li) (psym (path_proof env sigma uf j lj)) and edge_proof env sigma uf ((i,j),eq)= debug (fun () -> str "edge_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ pr_idx_term env sigma uf j); let pi=equal_proof env sigma uf i eq.lhs in let pj=psym (equal_proof env sigma uf j eq.rhs) in let pij= match eq.rule with Axiom (s,reversed)-> if reversed then psymax (axioms uf) s else pax (axioms uf) s | Congruence ->congr_proof env sigma uf eq.lhs eq.rhs | Injection (ti,ipac,tj,jpac,k) -> (* pi_k ipac = p_k jpac *) let p=ind_proof env sigma uf ti ipac tj jpac in let cinfo= get_constructor_info uf ipac.cnode in pinject p cinfo.ci_constr cinfo.ci_nhyps k in ptrans (ptrans pi pij) pj and constr_proof env sigma uf i ipac= debug (fun () -> str "constr_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20)); let t=find_oldest_pac uf i ipac in let eq_it=equal_proof env sigma uf i t in if ipac.args=[] then eq_it else let fipac=tail_pac ipac in let (fi,arg)=subterms uf t in let targ=term uf arg in let p=constr_proof env sigma uf fi fipac in ptrans eq_it (pcongr p (prefl targ)) and path_proof env sigma uf i l= debug (fun () -> str "path_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ str "{" ++ (prlist_with_sep (fun () -> str ",") (fun ((_,j),_) -> int j) l) ++ str "}"); match l with | [] -> prefl (term uf i) | x::q->ptrans (path_proof env sigma uf (snd (fst x)) q) (edge_proof env sigma uf x) and congr_proof env sigma uf i j= debug (fun () -> str "congr_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ pr_idx_term env sigma uf j); let (i1,i2) = subterms uf i and (j1,j2) = subterms uf j in pcongr (equal_proof env sigma uf i1 j1) (equal_proof env sigma uf i2 j2) and ind_proof env sigma uf i ipac j jpac= debug (fun () -> str "ind_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ pr_idx_term env sigma uf j); let p=equal_proof env sigma uf i j and p1=constr_proof env sigma uf i ipac and p2=constr_proof env sigma uf j jpac in ptrans (psym p1) (ptrans p p2) let build_proof env sigma uf= function | `Prove (i,j) -> equal_proof env sigma uf i j | `Discr (i,ci,j,cj)-> ind_proof env sigma uf i ci j cj coq-8.11.0/plugins/cc/g_congruence.mlg0000644000175000017500000000214013612664533017470 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { congruence_tac 1000 [] } | [ "congruence" integer(n) ] -> { congruence_tac n [] } | [ "congruence" "with" ne_constr_list(l) ] -> { congruence_tac 1000 l } |[ "congruence" integer(n) "with" ne_constr_list(l) ] -> { congruence_tac n l } END TACTIC EXTEND f_equal | [ "f_equal" ] -> { f_equal } END coq-8.11.0/plugins/cc/cctac.ml0000644000175000017500000005242013612664533015746 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let tf=decompose_term env sigma f in let targs=Array.map (decompose_term env sigma) args in Array.fold_left (fun s t->Appli (s,t)) tf targs | Prod (_,a,_b) when noccurn sigma 1 _b -> let b = Termops.pop _b in let sort_b = sf_of env sigma b in let sort_a = sf_of env sigma a in Appli(Appli(Product (sort_a,sort_b) , decompose_term env sigma a), decompose_term env sigma b) | Construct c -> let (((mind,i_ind),i_con),u)= c in let u = EInstance.kind sigma u in let canon_mind = MutInd.make1 (MutInd.canonical mind) in let canon_ind = canon_mind,i_ind in let (oib,_)=Global.lookup_inductive (canon_ind) in let nargs=constructor_nallargs env (canon_ind,i_con) in Constructor {ci_constr= ((canon_ind,i_con),u); ci_arity=nargs; ci_nhyps=nargs-oib.mind_nparams} | Ind c -> let (mind,i_ind),u = c in let u = EInstance.kind sigma u in let canon_mind = MutInd.make1 (MutInd.canonical mind) in let canon_ind = canon_mind,i_ind in (Symb (Constr.mkIndU (canon_ind,u))) | Const (c,u) -> let u = EInstance.kind sigma u in let canon_const = Constant.make1 (Constant.canonical c) in (Symb (Constr.mkConstU (canon_const,u))) | Proj (p, c) -> let canon_mind kn = MutInd.make1 (MutInd.canonical kn) in let p' = Projection.map canon_mind p in let c = Retyping.expand_projection env sigma p' c [] in decompose_term env sigma c | _ -> let t = Termops.strip_outer_cast sigma t in if closed0 sigma t then Symb (EConstr.to_constr ~abort_on_undefined_evars:false sigma t) else raise Not_found (* decompose equality in members and type *) open Termops let atom_of_constr env sigma term = let wh = whd_delta env sigma term in let kot = EConstr.kind sigma wh in match kot with App (f,args)-> if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 then `Eq (args.(0), decompose_term env sigma args.(1), decompose_term env sigma args.(2)) else `Other (decompose_term env sigma term) | _ -> `Other (decompose_term env sigma term) let rec pattern_of_constr env sigma c = match EConstr.kind sigma (whd env sigma c) with App (f,args)-> let pf = decompose_term env sigma f in let pargs,lrels = List.split (Array.map_to_list (pattern_of_constr env sigma) args) in PApp (pf,List.rev pargs), List.fold_left Int.Set.union Int.Set.empty lrels | Prod (_,a,_b) when noccurn sigma 1 _b -> let b = Termops.pop _b in let pa,sa = pattern_of_constr env sigma a in let pb,sb = pattern_of_constr env sigma b in let sort_b = sf_of env sigma b in let sort_a = sf_of env sigma a in PApp(Product (sort_a,sort_b), [pa;pb]),(Int.Set.union sa sb) | Rel i -> PVar i,Int.Set.singleton i | _ -> let pf = decompose_term env sigma c in PApp (pf,[]),Int.Set.empty let non_trivial = function PVar _ -> false | _ -> true let patterns_of_constr env sigma nrels term= let f,args= try destApp sigma (whd_delta env sigma term) with DestKO -> raise Not_found in if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 then let patt1,rels1 = pattern_of_constr env sigma args.(1) and patt2,rels2 = pattern_of_constr env sigma args.(2) in let valid1 = if not (Int.equal (Int.Set.cardinal rels1) nrels) then Creates_variables else if non_trivial patt1 then Normal else Trivial (EConstr.to_constr sigma args.(0)) and valid2 = if not (Int.equal (Int.Set.cardinal rels2) nrels) then Creates_variables else if non_trivial patt2 then Normal else Trivial (EConstr.to_constr sigma args.(0)) in if valid1 != Creates_variables || valid2 != Creates_variables then nrels,valid1,patt1,valid2,patt2 else raise Not_found else raise Not_found let rec quantified_atom_of_constr env sigma nrels term = match EConstr.kind sigma (whd_delta env sigma term) with Prod (id,atom,ff) -> if is_global sigma (Lazy.force _False) ff then let patts=patterns_of_constr env sigma nrels atom in `Nrule patts else quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff | _ -> let patts=patterns_of_constr env sigma nrels term in `Rule patts let litteral_of_constr env sigma term= match EConstr.kind sigma (whd_delta env sigma term) with | Prod (id,atom,ff) -> if is_global sigma (Lazy.force _False) ff then match (atom_of_constr env sigma atom) with `Eq(t,a,b) -> `Neq(t,a,b) | `Other(p) -> `Nother(p) else begin try quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff with Not_found -> `Other (decompose_term env sigma term) end | _ -> atom_of_constr env sigma term (* store all equalities from the context *) let make_prb gls depth additionnal_terms = let open Tacmach.New in let env=pf_env gls in let sigma=project gls in let state = empty depth {it = Proofview.Goal.goal gls; sigma } in let pos_hyps = ref [] in let neg_hyps =ref [] in List.iter (fun c -> let t = decompose_term env sigma c in ignore (add_term state t)) additionnal_terms; List.iter (fun decl -> let id = NamedDecl.get_id decl in begin let cid=Constr.mkVar id in match litteral_of_constr env sigma (NamedDecl.get_type decl) with `Eq (t,a,b) -> add_equality state cid a b | `Neq (t,a,b) -> add_disequality state (Hyp cid) a b | `Other ph -> List.iter (fun (cidn,nh) -> add_disequality state (HeqnH (cid,cidn)) ph nh) !neg_hyps; pos_hyps:=(cid,ph):: !pos_hyps | `Nother nh -> List.iter (fun (cidp,ph) -> add_disequality state (HeqnH (cidp,cid)) ph nh) !pos_hyps; neg_hyps:=(cid,nh):: !neg_hyps | `Rule patts -> add_quant state id true patts | `Nrule patts -> add_quant state id false patts end) (Proofview.Goal.hyps gls); begin match atom_of_constr env sigma (pf_concl gls) with `Eq (t,a,b) -> add_disequality state Goal a b | `Other g -> List.iter (fun (idp,ph) -> add_disequality state (HeqG idp) ph g) !pos_hyps end; state (* indhyps builds the array of arrays of constructor hyps for (ind largs) *) let build_projection intype (cstr:pconstructor) special default gls= let open Tacmach.New in let ci= (snd(fst cstr)) in let sigma = project gls in let body=Equality.build_selector (pf_env gls) sigma ci (mkRel 1) intype special default in let id=pf_get_new_id (Id.of_string "t") gls in sigma, mkLambda(make_annot (Name id) Sorts.Relevant,intype,body) (* generate an adhoc tactic following the proof tree *) let app_global f args k = Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> k (mkApp (fc, args)) let rec gen_holes env sigma t n accu = if Int.equal n 0 then (sigma, List.rev accu) else match EConstr.kind sigma t with | Prod (_, u, t) -> let (sigma, ev) = Evarutil.new_evar env sigma u in let t = EConstr.Vars.subst1 ev t in gen_holes env sigma t (pred n) (ev :: accu) | _ -> assert false let app_global_with_holes f args n = Proofview.Goal.enter begin fun gl -> Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in Refine.refine ~typecheck:false begin fun sigma -> let t = Tacmach.New.pf_get_type_of gl fc in let t = Termops.prod_applist sigma t (Array.to_list args) in let ans = mkApp (fc, args) in let (sigma, holes) = gen_holes env sigma t n [] in let ans = applist (ans, holes) in let sigma = Typing.check env sigma ans concl in (sigma, ans) end end let assert_before n c = Proofview.Goal.enter begin fun gl -> let evm, _ = Tacmach.New.pf_apply type_of gl c in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evm) (assert_before n c) end let refresh_type env evm ty = Evarsolve.refresh_universes ~status:Evd.univ_flexible ~refreshset:true (Some false) env evm ty let refresh_universes ty k = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let evm = Tacmach.New.project gl in let evm, ty = refresh_type env evm ty in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evm) (k ty) end let constr_of_term c = EConstr.of_constr (constr_of_term c) let rec proof_tac p : unit Proofview.tactic = Proofview.Goal.enter begin fun gl -> let type_of t = Tacmach.New.pf_unsafe_type_of gl t in try (* type_of can raise exceptions *) match p.p_rule with Ax c -> exact_check (EConstr.of_constr c) | SymAx c -> let c = EConstr.of_constr c in let l=constr_of_term p.p_lhs and r=constr_of_term p.p_rhs in refresh_universes (type_of l) (fun typ -> app_global _sym_eq [|typ;r;l;c|] exact_check) | Refl t -> let lr = constr_of_term t in refresh_universes (type_of lr) (fun typ -> app_global _refl_equal [|typ;constr_of_term t|] exact_check) | Trans (p1,p2)-> let t1 = constr_of_term p1.p_lhs and t2 = constr_of_term p1.p_rhs and t3 = constr_of_term p2.p_rhs in refresh_universes (type_of t2) (fun typ -> let prf = app_global_with_holes _trans_eq [|typ;t1;t2;t3;|] 2 in Tacticals.New.tclTHENS prf [(proof_tac p1);(proof_tac p2)]) | Congr (p1,p2)-> let tf1=constr_of_term p1.p_lhs and tx1=constr_of_term p2.p_lhs and tf2=constr_of_term p1.p_rhs and tx2=constr_of_term p2.p_rhs in refresh_universes (type_of tf1) (fun typf -> refresh_universes (type_of tx1) (fun typx -> refresh_universes (type_of (mkApp (tf1,[|tx1|]))) (fun typfx -> let id = Tacmach.New.pf_get_new_id (Id.of_string "f") gl in let appx1 = mkLambda(make_annot (Name id) Sorts.Relevant,typf,mkApp(mkRel 1,[|tx1|])) in let lemma1 = app_global_with_holes _f_equal [|typf;typfx;appx1;tf1;tf2|] 1 in let lemma2 = app_global_with_holes _f_equal [|typx;typfx;tf2;tx1;tx2|] 1 in let prf = app_global_with_holes _trans_eq [|typfx; mkApp(tf1,[|tx1|]); mkApp(tf2,[|tx1|]); mkApp(tf2,[|tx2|])|] 2 in Tacticals.New.tclTHENS prf [Tacticals.New.tclTHEN lemma1 (proof_tac p1); Tacticals.New.tclFIRST [Tacticals.New.tclTHEN lemma2 (proof_tac p2); reflexivity; Tacticals.New.tclZEROMSG (Pp.str "I don't know how to handle dependent equality")]]))) | Inject (prf,cstr,nargs,argind) -> let ti=constr_of_term prf.p_lhs in let tj=constr_of_term prf.p_rhs in let default=constr_of_term p.p_lhs in let special=mkRel (1+nargs-argind) in refresh_universes (type_of ti) (fun intype -> refresh_universes (type_of default) (fun outtype -> let sigma, proj = build_projection intype cstr special default gl in let injt= app_global_with_holes _f_equal [|intype;outtype;proj;ti;tj|] 1 in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Tacticals.New.tclTHEN injt (proof_tac prf)))) with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e end let refute_tac c t1 t2 p = Proofview.Goal.enter begin fun gl -> let tt1=constr_of_term t1 and tt2=constr_of_term t2 in let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in let false_t=mkApp (c,[|mkVar hid|]) in let k intype = let neweq= app_global _eq [|intype;tt1;tt2|] in Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) [proof_tac p; simplest_elim false_t] in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt1) k end let refine_exact_check c = Proofview.Goal.enter begin fun gl -> let evm, _ = Tacmach.New.pf_apply type_of gl c in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evm) (exact_check c) end let convert_to_goal_tac c t1 t2 p = Proofview.Goal.enter begin fun gl -> let tt1=constr_of_term t1 and tt2=constr_of_term t2 in let k sort = let neweq= app_global _eq [|sort;tt1;tt2|] in let e = Tacmach.New.pf_get_new_id (Id.of_string "e") gl in let x = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in let identity=mkLambda (make_annot (Name x) Sorts.Relevant,sort,mkRel 1) in let endt = app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in Tacticals.New.tclTHENS (neweq (assert_before (Name e))) [proof_tac p; endt refine_exact_check] in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt2) k end let convert_to_hyp_tac c1 t1 c2 t2 p = Proofview.Goal.enter begin fun gl -> let tt2=constr_of_term t2 in let h = Tacmach.New.pf_get_new_id (Id.of_string "H") gl in let false_t=mkApp (c2,[|mkVar h|]) in Tacticals.New.tclTHENS (assert_before (Name h) tt2) [convert_to_goal_tac c1 t1 t2 p; simplest_elim false_t] end (* Essentially [assert (Heq : lhs = rhs) by proof_tac p; discriminate Heq] *) let discriminate_tac cstru p = Proofview.Goal.enter begin fun gl -> let lhs=constr_of_term p.p_lhs and rhs=constr_of_term p.p_rhs in let env = Proofview.Goal.env gl in let evm = Tacmach.New.project gl in let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl lhs) in let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in let neweq=app_global _eq [|intype;lhs;rhs|] in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS evm) (Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) [proof_tac p; Equality.discrHyp hid]) end (* wrap everything *) let build_term_to_complete uf pac = let cinfo = get_constructor_info uf pac.cnode in let real_args = List.rev_map (fun i -> constr_of_term (term uf i)) pac.args in let (kn, u) = cinfo.ci_constr in (applist (mkConstructU (kn, EInstance.make u), real_args), pac.arity) let cc_tactic depth additionnal_terms = Proofview.Goal.enter begin fun gl -> let sigma = Tacmach.New.project gl in Coqlib.(check_required_library logic_module_name); let _ = debug (fun () -> Pp.str "Reading subgoal ...") in let state = make_prb gl depth additionnal_terms in let _ = debug (fun () -> Pp.str "Problem built, solving ...") in let sol = execute true state in let _ = debug (fun () -> Pp.str "Computation completed.") in let uf=forest state in match sol with None -> Tacticals.New.tclFAIL 0 (str "congruence failed") | Some reason -> debug (fun () -> Pp.str "Goal solved, generating proof ..."); match reason with Discrimination (i,ipac,j,jpac) -> let p=build_proof (Tacmach.New.pf_env gl) sigma uf (`Discr (i,ipac,j,jpac)) in let cstr=(get_constructor_info uf ipac.cnode).ci_constr in discriminate_tac cstr p | Incomplete -> let open Glob_term in let env = Proofview.Goal.env gl in let terms_to_complete = List.map (build_term_to_complete uf) (epsilons uf) in let hole = DAst.make @@ GHole (Evar_kinds.InternalHole, Namegen.IntroAnonymous, None) in let pr_missing (c, missing) = let c = Detyping.detype Detyping.Now ~lax:true false Id.Set.empty env sigma c in let holes = List.init missing (fun _ -> hole) in Printer.pr_glob_constr_env env (DAst.make @@ GApp (c, holes)) in let msg = Pp.(str "Goal is solvable by congruence but some arguments are missing." ++ fnl () ++ str " Try " ++ hov 8 begin str "\"congruence with (" ++ prlist_with_sep (fun () -> str ")" ++ spc () ++ str "(") pr_missing terms_to_complete ++ str ")\"," end ++ str " replacing metavariables by arbitrary terms.") in Tacticals.New.tclFAIL 0 msg | Contradiction dis -> let env = Proofview.Goal.env gl in let p=build_proof env sigma uf (`Prove (dis.lhs,dis.rhs)) in let ta=term uf dis.lhs and tb=term uf dis.rhs in match dis.rule with Goal -> proof_tac p | Hyp id -> refute_tac (EConstr.of_constr id) ta tb p | HeqG id -> let id = EConstr.of_constr id in convert_to_goal_tac id ta tb p | HeqnH (ida,idb) -> let ida = EConstr.of_constr ida in let idb = EConstr.of_constr idb in convert_to_hyp_tac ida ta idb tb p end let cc_fail = Tacticals.New.tclZEROMSG (Pp.str "congruence failed.") let congruence_tac depth l = Tacticals.New.tclORELSE (Tacticals.New.tclTHEN (Tacticals.New.tclREPEAT introf) (cc_tactic depth l)) cc_fail (* Beware: reflexivity = constructor 1 = apply refl_equal might be slow now, let's rather do something equivalent to a "simple apply refl_equal" *) (* The [f_equal] tactic. It mimics the use of lemmas [f_equal], [f_equal2], etc. This isn't particularly related with congruence, apart from the fact that congruence is called internally. *) let mk_eq f c1 c2 k = Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> Proofview.Goal.enter begin fun gl -> let open Tacmach.New in let evm, ty = pf_apply type_of gl c1 in let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm ty in let term = mkApp (fc, [| ty; c1; c2 |]) in let evm, _ = type_of (pf_env gl) evm term in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evm) (k term) end let f_equal = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in let cut_eq c1 c2 = try (* type_of can raise an exception *) Tacticals.New.tclTHENS (mk_eq _eq c1 c2 Tactics.cut) [Proofview.tclUNIT ();Tacticals.New.tclTRY ((app_global _refl_equal [||]) apply)] with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e in Proofview.tclORELSE begin match EConstr.kind sigma concl with | App (r,[|_;t;t'|]) when is_global sigma (Lazy.force _eq) r -> begin match EConstr.kind sigma t, EConstr.kind sigma t' with | App (f,v), App (f',v') when Int.equal (Array.length v) (Array.length v') -> let rec cuts i = if i < 0 then Tacticals.New.tclTRY (congruence_tac 1000 []) else Tacticals.New.tclTHENFIRST (cut_eq v.(i) v'.(i)) (cuts (i-1)) in cuts (Array.length v - 1) | _ -> Proofview.tclUNIT () end | _ -> Proofview.tclUNIT () end begin function (e, info) -> match e with | Pretype_errors.PretypeError _ | Type_errors.TypeError _ -> Proofview.tclUNIT () | e -> Proofview.tclZERO ~info e end end coq-8.11.0/plugins/cc/cctac.mli0000644000175000017500000000166213612664533016121 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit Proofview.tactic val cc_tactic : int -> constr list -> unit Proofview.tactic val cc_fail : unit Proofview.tactic val congruence_tac : int -> constr list -> unit Proofview.tactic val f_equal : unit Proofview.tactic coq-8.11.0/plugins/cc/ccproof.mli0000644000175000017500000000376413612664533016504 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* proof val pcongr: proof -> proof -> proof val ptrans: proof -> proof -> proof val psym: proof -> proof val pax : (Ccalgo.term * Ccalgo.term) Ccalgo.Constrhash.t -> Ccalgo.Constrhash.key -> proof val psymax : (Ccalgo.term * Ccalgo.term) Ccalgo.Constrhash.t -> Ccalgo.Constrhash.key -> proof val pinject : proof -> pconstructor -> int -> int -> proof (** Proof building functions *) val equal_proof : Environ.env -> Evd.evar_map -> forest -> int -> int -> proof val edge_proof : Environ.env -> Evd.evar_map -> forest -> (int*int)*equality -> proof val path_proof : Environ.env -> Evd.evar_map -> forest -> int -> ((int*int)*equality) list -> proof val congr_proof : Environ.env -> Evd.evar_map -> forest -> int -> int -> proof val ind_proof : Environ.env -> Evd.evar_map -> forest -> int -> pa_constructor -> int -> pa_constructor -> proof (** Main proof building function *) val build_proof : Environ.env -> Evd.evar_map -> forest -> [ `Discr of int * pa_constructor * int * pa_constructor | `Prove of int * int ] -> proof coq-8.11.0/plugins/cc/ccalgo.mli0000644000175000017500000001372313612664533016275 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* term = app(constr,t) *) type cl = Rep of representative| Eqto of int*equality type vertex = Leaf| Node of (int*int) type node = {mutable clas:cl; mutable cpath: int; mutable constructors: int PacMap.t; vertex:vertex; term:term} type forest= {mutable max_size:int; mutable size:int; mutable map: node array; axioms: (term*term) Constrhash.t; mutable epsilons: pa_constructor list; syms: int Termhash.t} type state type explanation = Discrimination of (int*pa_constructor*int*pa_constructor) | Contradiction of disequality | Incomplete type matching_problem val term_equal : term -> term -> bool val constr_of_term : term -> constr val debug : (unit -> Pp.t) -> unit val forest : state -> forest val axioms : forest -> (term * term) Constrhash.t val epsilons : forest -> pa_constructor list val empty : int -> Goal.goal Evd.sigma -> state val add_term : state -> term -> int val add_equality : state -> constr -> term -> term -> unit val add_disequality : state -> from -> term -> term -> unit val add_quant : state -> Id.t -> bool -> int * patt_kind * ccpattern * patt_kind * ccpattern -> unit val tail_pac : pa_constructor -> pa_constructor val find : forest -> int -> int val find_oldest_pac : forest -> int -> pa_constructor -> int val term : forest -> int -> term val get_constructor_info : forest -> int -> cinfo val subterms : forest -> int -> int * int val join_path : forest -> int -> int -> ((int * int) * equality) list * ((int * int) * equality) list val make_fun_table : state -> Int.Set.t PafMap.t val do_match : state -> (quant_eq * int array) list ref -> matching_problem Stack.t -> unit val init_pb_stack : state -> matching_problem Stack.t val paf_of_patt : int Termhash.t -> ccpattern -> pa_fun val find_instances : state -> (quant_eq * int array) list val execute : bool -> state -> explanation option val pr_idx_term : Environ.env -> Evd.evar_map -> forest -> int -> Pp.t val empty_forest: unit -> forest (*type pa_constructor module PacMap:CSig.MapS with type key=pa_constructor type term = Symb of Term.constr | Eps | Appli of term * term | Constructor of Names.constructor*int*int type rule = Congruence | Axiom of Names.Id.t | Injection of int*int*int*int type equality = {lhs : int; rhs : int; rule : rule} module ST : sig type t val empty : unit -> t val enter : int -> int * int -> t -> unit val query : int * int -> t -> int val delete : int -> t -> unit val delete_list : int list -> t -> unit end module UF : sig type t exception Discriminable of int * int * int * int * t val empty : unit -> t val find : t -> int -> int val size : t -> int -> int val get_constructor : t -> int -> Names.constructor val pac_arity : t -> int -> int * int -> int val mem_node_pac : t -> int -> int * int -> int val add_pacs : t -> int -> pa_constructor PacMap.t -> int list * equality list val term : t -> int -> term val subterms : t -> int -> int * int val add : t -> term -> int val union : t -> int -> int -> equality -> int list * equality list val join_path : t -> int -> int -> ((int*int)*equality) list* ((int*int)*equality) list end val combine_rec : UF.t -> int list -> equality list val process_rec : UF.t -> equality list -> int list val cc : UF.t -> unit val make_uf : (Names.Id.t * (term * term)) list -> UF.t val add_one_diseq : UF.t -> (term * term) -> int * int val add_disaxioms : UF.t -> (Names.Id.t * (term * term)) list -> (Names.Id.t * (int * int)) list val check_equal : UF.t -> int * int -> bool val find_contradiction : UF.t -> (Names.Id.t * (int * int)) list -> (Names.Id.t * (int * int)) *) coq-8.11.0/plugins/cc/ccalgo.ml0000644000175000017500000007652413612664533016134 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* !cc_verbose); optwrite=(fun b -> cc_verbose := b)} in declare_bool_option gdopt (* Signature table *) module ST=struct (* l: sign -> term r: term -> sign *) module IntTable = Hashtbl.Make(Int) module IntPair = struct type t = int * int let equal (i1, j1) (i2, j2) = Int.equal i1 i2 && Int.equal j1 j2 let hash (i, j) = Hashset.Combine.combine (Int.hash i) (Int.hash j) end module IntPairTable = Hashtbl.Make(IntPair) type t = {toterm: int IntPairTable.t; tosign: (int * int) IntTable.t} let empty () = {toterm=IntPairTable.create init_size; tosign=IntTable.create init_size} let enter t sign st= if IntPairTable.mem st.toterm sign then anomaly ~label:"enter" (Pp.str "signature already entered.") else IntPairTable.replace st.toterm sign t; IntTable.replace st.tosign t sign let query sign st=IntPairTable.find st.toterm sign let delete st t= try let sign=IntTable.find st.tosign t in IntPairTable.remove st.toterm sign; IntTable.remove st.tosign t with Not_found -> () let delete_set st s = Int.Set.iter (delete st) s end type pa_constructor= { cnode : int; arity : int; args : int list} type pa_fun= {fsym:int; fnargs:int} type pa_mark= Fmark of pa_fun | Cmark of pa_constructor module PacOrd = struct type t = pa_constructor let compare { cnode = cnode0; arity = arity0; args = args0 } { cnode = cnode1; arity = arity1; args = args1 } = let cmp = Int.compare cnode0 cnode1 in if cmp = 0 then let cmp' = Int.compare arity0 arity1 in if cmp' = 0 then List.compare Int.compare args0 args1 else cmp' else cmp end module PafOrd = struct type t = pa_fun let compare { fsym = fsym0; fnargs = fnargs0 } { fsym = fsym1; fnargs = fnargs1 } = let cmp = Int.compare fsym0 fsym1 in if cmp = 0 then Int.compare fnargs0 fnargs1 else cmp end module PacMap=Map.Make(PacOrd) module PafMap=Map.Make(PafOrd) type cinfo= {ci_constr: pconstructor; (* inductive type *) ci_arity: int; (* # args *) ci_nhyps: int} (* # projectable args *) let family_eq f1 f2 = match f1, f2 with | Set, Set | Prop, Prop | Type _, Type _ -> true | _ -> false type term= Symb of constr | Product of Sorts.t * Sorts.t | Eps of Id.t | Appli of term*term | Constructor of cinfo (* constructor arity + nhyps *) let rec term_equal t1 t2 = match t1, t2 with | Symb c1, Symb c2 -> eq_constr_nounivs c1 c2 | Product (s1, t1), Product (s2, t2) -> family_eq s1 s2 && family_eq t1 t2 | Eps i1, Eps i2 -> Id.equal i1 i2 | Appli (t1, u1), Appli (t2, u2) -> term_equal t1 t2 && term_equal u1 u2 | Constructor {ci_constr=(c1,u1); ci_arity=i1; ci_nhyps=j1}, Constructor {ci_constr=(c2,u2); ci_arity=i2; ci_nhyps=j2} -> Int.equal i1 i2 && Int.equal j1 j2 && eq_constructor c1 c2 (* FIXME check eq? *) | _ -> false open Hashset.Combine let rec hash_term = function | Symb c -> combine 1 (Constr.hash c) | Product (s1, s2) -> combine3 2 (Sorts.hash s1) (Sorts.hash s2) | Eps i -> combine 3 (Id.hash i) | Appli (t1, t2) -> combine3 4 (hash_term t1) (hash_term t2) | Constructor {ci_constr=(c,u); ci_arity=i; ci_nhyps=j} -> combine4 5 (constructor_hash c) i j type ccpattern = PApp of term * ccpattern list (* arguments are reversed *) | PVar of int type rule= Congruence | Axiom of constr * bool | Injection of int * pa_constructor * int * pa_constructor * int type from= Goal | Hyp of constr | HeqG of constr | HeqnH of constr * constr type 'a eq = {lhs:int;rhs:int;rule:'a} type equality = rule eq type disequality = from eq type patt_kind = Normal | Trivial of types | Creates_variables type quant_eq = { qe_hyp_id: Id.t; qe_pol: bool; qe_nvars:int; qe_lhs: ccpattern; qe_lhs_valid:patt_kind; qe_rhs: ccpattern; qe_rhs_valid:patt_kind } let swap eq : equality = let swap_rule=match eq.rule with Congruence -> Congruence | Injection (i,pi,j,pj,k) -> Injection (j,pj,i,pi,k) | Axiom (id,reversed) -> Axiom (id,not reversed) in {lhs=eq.rhs;rhs=eq.lhs;rule=swap_rule} type inductive_status = Unknown | Partial of pa_constructor | Partial_applied | Total of (int * pa_constructor) type representative= {mutable weight:int; mutable lfathers:Int.Set.t; mutable fathers:Int.Set.t; mutable inductive_status: inductive_status; class_type : types; mutable functions: Int.Set.t PafMap.t} (*pac -> term = app(constr,t) *) type cl = Rep of representative| Eqto of int*equality type vertex = Leaf| Node of (int*int) type node = {mutable clas:cl; mutable cpath: int; mutable constructors: int PacMap.t; vertex:vertex; term:term} module Constrhash = Hashtbl.Make (struct type t = constr let equal = eq_constr_nounivs let hash = Constr.hash end) module Typehash = Constrhash module Termhash = Hashtbl.Make (struct type t = term let equal = term_equal let hash = hash_term end) module Identhash = Hashtbl.Make (struct type t = Id.t let equal = Id.equal let hash = Id.hash end) type forest= {mutable max_size:int; mutable size:int; mutable map: node array; axioms: (term*term) Constrhash.t; mutable epsilons: pa_constructor list; syms: int Termhash.t} type state = {uf: forest; sigtable:ST.t; mutable terms: Int.Set.t; combine: equality Queue.t; marks: (int * pa_mark) Queue.t; mutable diseq: disequality list; mutable quant: quant_eq list; mutable pa_classes: Int.Set.t; q_history: (int array) Identhash.t; mutable rew_depth:int; mutable changed:bool; by_type: Int.Set.t Typehash.t; mutable env:Environ.env; sigma:Evd.evar_map} let dummy_node = { clas=Eqto (min_int,{lhs=min_int;rhs=min_int;rule=Congruence}); cpath=min_int; constructors=PacMap.empty; vertex=Leaf; term=Symb (mkRel min_int) } let empty_forest() = { max_size=init_size; size=0; map=Array.make init_size dummy_node; epsilons=[]; axioms=Constrhash.create init_size; syms=Termhash.create init_size } let empty depth gls:state = { uf= empty_forest (); terms=Int.Set.empty; combine=Queue.create (); marks=Queue.create (); sigtable=ST.empty (); diseq=[]; quant=[]; pa_classes=Int.Set.empty; q_history=Identhash.create init_size; rew_depth=depth; by_type=Constrhash.create init_size; changed=false; env=pf_env gls; sigma=project gls } let forest state = state.uf let compress_path uf i j = uf.map.(j).cpath<-i let rec find_aux uf visited i= let j = uf.map.(i).cpath in if j<0 then let () = List.iter (compress_path uf i) visited in i else find_aux uf (i::visited) j let find uf i= find_aux uf [] i let get_representative uf i= match uf.map.(i).clas with Rep r -> r | _ -> anomaly ~label:"get_representative" (Pp.str "not a representative.") let get_constructors uf i= uf.map.(i).constructors let rec find_oldest_pac uf i pac= try PacMap.find pac (get_constructors uf i) with Not_found -> match uf.map.(i).clas with Eqto (j,_) -> find_oldest_pac uf j pac | Rep _ -> raise Not_found let get_constructor_info uf i= match uf.map.(i).term with Constructor cinfo->cinfo | _ -> anomaly ~label:"get_constructor" (Pp.str "not a constructor.") let size uf i= (get_representative uf i).weight let axioms uf = uf.axioms let epsilons uf = uf.epsilons let add_lfather uf i t= let r=get_representative uf i in r.weight<-r.weight+1; r.lfathers<-Int.Set.add t r.lfathers; r.fathers <-Int.Set.add t r.fathers let add_rfather uf i t= let r=get_representative uf i in r.weight<-r.weight+1; r.fathers <-Int.Set.add t r.fathers exception Discriminable of int * pa_constructor * int * pa_constructor let append_pac t p = {p with arity=pred p.arity;args=t::p.args} let tail_pac p= {p with arity=succ p.arity;args=List.tl p.args} let fsucc paf = {paf with fnargs=succ paf.fnargs} let add_pac node pac t = if not (PacMap.mem pac node.constructors) then node.constructors<-PacMap.add pac t node.constructors let add_paf rep paf t = let already = try PafMap.find paf rep.functions with Not_found -> Int.Set.empty in rep.functions<- PafMap.add paf (Int.Set.add t already) rep.functions let term uf i=uf.map.(i).term let subterms uf i= match uf.map.(i).vertex with Node(j,k) -> (j,k) | _ -> anomaly ~label:"subterms" (Pp.str "not a node.") let signature uf i= let j,k=subterms uf i in (find uf j,find uf k) let next uf= let size=uf.size in let nsize= succ size in if Int.equal nsize uf.max_size then let newmax=uf.max_size * 3 / 2 + 1 in let newmap=Array.make newmax dummy_node in begin uf.max_size<-newmax; Array.blit uf.map 0 newmap 0 size; uf.map<-newmap end else (); uf.size<-nsize; size let new_representative typ = {weight=0; lfathers=Int.Set.empty; fathers=Int.Set.empty; inductive_status=Unknown; class_type=typ; functions=PafMap.empty} (* rebuild a constr from an applicative term *) let _A_ = Name (Id.of_string "A") let _B_ = Name (Id.of_string "A") let _body_ = mkProd(make_annot Anonymous Sorts.Relevant,mkRel 2,mkRel 2) let cc_product s1 s2 = mkLambda(make_annot _A_ Sorts.Relevant,mkSort(s1), mkLambda(make_annot _B_ Sorts.Relevant,mkSort(s2),_body_)) let rec constr_of_term = function Symb s-> s | Product(s1,s2) -> cc_product s1 s2 | Eps id -> mkVar id | Constructor cinfo -> mkConstructU cinfo.ci_constr | Appli (s1,s2)-> make_app [(constr_of_term s2)] s1 and make_app l=function Appli (s1,s2)->make_app ((constr_of_term s2)::l) s1 | other -> Term.applist (constr_of_term other,l) let rec canonize_name sigma c = let c = EConstr.Unsafe.to_constr c in let func c = canonize_name sigma (EConstr.of_constr c) in match Constr.kind c with | Const (kn,u) -> let canon_const = Constant.make1 (Constant.canonical kn) in (mkConstU (canon_const,u)) | Ind ((kn,i),u) -> let canon_mind = MutInd.make1 (MutInd.canonical kn) in (mkIndU ((canon_mind,i),u)) | Construct (((kn,i),j),u) -> let canon_mind = MutInd.make1 (MutInd.canonical kn) in mkConstructU (((canon_mind,i),j),u) | Prod (na,t,ct) -> mkProd (na,func t, func ct) | Lambda (na,t,ct) -> mkLambda (na, func t,func ct) | LetIn (na,b,t,ct) -> mkLetIn (na, func b,func t,func ct) | App (ct,l) -> mkApp (func ct,Array.Smart.map func l) | Proj(p,c) -> let p' = Projection.map (fun kn -> MutInd.make1 (MutInd.canonical kn)) p in (mkProj (p', func c)) | _ -> c (* rebuild a term from a pattern and a substitution *) let build_subst uf subst = Array.map (fun i -> try term uf i with e when CErrors.noncritical e -> anomaly (Pp.str "incomplete matching.")) subst let rec inst_pattern subst = function PVar i -> subst.(pred i) | PApp (t, args) -> List.fold_right (fun spat f -> Appli (f,inst_pattern subst spat)) args t let pr_idx_term env sigma uf i = str "[" ++ int i ++ str ":=" ++ Printer.pr_econstr_env env sigma (EConstr.of_constr (constr_of_term (term uf i))) ++ str "]" let pr_term env sigma t = str "[" ++ Printer.pr_econstr_env env sigma (EConstr.of_constr (constr_of_term t)) ++ str "]" let rec add_term state t= let uf=state.uf in try Termhash.find uf.syms t with Not_found -> let b=next uf in let trm = constr_of_term t in let typ = Typing.unsafe_type_of state.env state.sigma (EConstr.of_constr trm) in let typ = canonize_name state.sigma typ in let new_node= match t with Symb _ | Product (_,_) -> let paf = {fsym=b; fnargs=0} in Queue.add (b,Fmark paf) state.marks; {clas= Rep (new_representative typ); cpath= -1; constructors=PacMap.empty; vertex= Leaf; term= t} | Eps id -> {clas= Rep (new_representative typ); cpath= -1; constructors=PacMap.empty; vertex= Leaf; term= t} | Appli (t1,t2) -> let i1=add_term state t1 and i2=add_term state t2 in add_lfather uf (find uf i1) b; add_rfather uf (find uf i2) b; state.terms<-Int.Set.add b state.terms; {clas= Rep (new_representative typ); cpath= -1; constructors=PacMap.empty; vertex= Node(i1,i2); term= t} | Constructor cinfo -> let paf = {fsym=b; fnargs=0} in Queue.add (b,Fmark paf) state.marks; let pac = {cnode= b; arity= cinfo.ci_arity; args=[]} in Queue.add (b,Cmark pac) state.marks; {clas=Rep (new_representative typ); cpath= -1; constructors=PacMap.empty; vertex=Leaf; term=t} in uf.map.(b)<-new_node; Termhash.add uf.syms t b; Typehash.replace state.by_type typ (Int.Set.add b (try Typehash.find state.by_type typ with Not_found -> Int.Set.empty)); b let add_equality state c s t= let i = add_term state s in let j = add_term state t in Queue.add {lhs=i;rhs=j;rule=Axiom(c,false)} state.combine; Constrhash.add state.uf.axioms c (s,t) let add_disequality state from s t = let i = add_term state s in let j = add_term state t in state.diseq<-{lhs=i;rhs=j;rule=from}::state.diseq let add_quant state id pol (nvars,valid1,patt1,valid2,patt2) = state.quant<- {qe_hyp_id= id; qe_pol= pol; qe_nvars=nvars; qe_lhs= patt1; qe_lhs_valid=valid1; qe_rhs= patt2; qe_rhs_valid=valid2}::state.quant let is_redundant state id args = try let norm_args = Array.map (find state.uf) args in let prev_args = Identhash.find_all state.q_history id in List.exists (fun old_args -> Util.Array.for_all2 (fun i j -> Int.equal i (find state.uf j)) norm_args old_args) prev_args with Not_found -> false let add_inst state (inst,int_subst) = Control.check_for_interrupt (); if state.rew_depth > 0 then if is_redundant state inst.qe_hyp_id int_subst then debug (fun () -> str "discarding redundant (dis)equality") else begin Identhash.add state.q_history inst.qe_hyp_id int_subst; let subst = build_subst (forest state) int_subst in let prfhead= mkVar inst.qe_hyp_id in let args = Array.map constr_of_term subst in let _ = Array.rev args in (* highest deBruijn index first *) let prf= mkApp(prfhead,args) in let s = inst_pattern subst inst.qe_lhs and t = inst_pattern subst inst.qe_rhs in state.changed<-true; state.rew_depth<-pred state.rew_depth; if inst.qe_pol then begin debug (fun () -> (str "Adding new equality, depth="++ int state.rew_depth) ++ fnl () ++ (str " [" ++ Printer.pr_econstr_env state.env state.sigma (EConstr.of_constr prf) ++ str " : " ++ pr_term state.env state.sigma s ++ str " == " ++ pr_term state.env state.sigma t ++ str "]")); add_equality state prf s t end else begin debug (fun () -> (str "Adding new disequality, depth="++ int state.rew_depth) ++ fnl () ++ (str " [" ++ Printer.pr_econstr_env state.env state.sigma (EConstr.of_constr prf) ++ str " : " ++ pr_term state.env state.sigma s ++ str " <> " ++ pr_term state.env state.sigma t ++ str "]")); add_disequality state (Hyp prf) s t end end let link uf i j eq = (* links i -> j *) let node=uf.map.(i) in node.clas<-Eqto (j,eq); node.cpath<-j let rec down_path uf i l= match uf.map.(i).clas with Eqto (j,eq) ->down_path uf j (((i,j),eq)::l) | Rep _ ->l let eq_pair (i1, j1) (i2, j2) = Int.equal i1 i2 && Int.equal j1 j2 let rec min_path=function ([],l2)->([],l2) | (l1,[])->(l1,[]) | (((c1,t1)::q1),((c2,t2)::q2)) when eq_pair c1 c2 -> min_path (q1,q2) | cpl -> cpl let join_path uf i j= assert (Int.equal (find uf i) (find uf j)); min_path (down_path uf i [],down_path uf j []) let union state i1 i2 eq= debug (fun () -> str "Linking " ++ pr_idx_term state.env state.sigma state.uf i1 ++ str " and " ++ pr_idx_term state.env state.sigma state.uf i2 ++ str "."); let r1= get_representative state.uf i1 and r2= get_representative state.uf i2 in link state.uf i1 i2 eq; Constrhash.replace state.by_type r1.class_type (Int.Set.remove i1 (try Constrhash.find state.by_type r1.class_type with Not_found -> Int.Set.empty)); let f= Int.Set.union r1.fathers r2.fathers in r2.weight<-Int.Set.cardinal f; r2.fathers<-f; r2.lfathers<-Int.Set.union r1.lfathers r2.lfathers; ST.delete_set state.sigtable r1.fathers; state.terms<-Int.Set.union state.terms r1.fathers; PacMap.iter (fun pac b -> Queue.add (b,Cmark pac) state.marks) state.uf.map.(i1).constructors; PafMap.iter (fun paf -> Int.Set.iter (fun b -> Queue.add (b,Fmark paf) state.marks)) r1.functions; match r1.inductive_status,r2.inductive_status with Unknown,_ -> () | Partial pac,Unknown -> r2.inductive_status<-Partial pac; state.pa_classes<-Int.Set.remove i1 state.pa_classes; state.pa_classes<-Int.Set.add i2 state.pa_classes | Partial _ ,(Partial _ |Partial_applied) -> state.pa_classes<-Int.Set.remove i1 state.pa_classes | Partial_applied,Unknown -> r2.inductive_status<-Partial_applied | Partial_applied,Partial _ -> state.pa_classes<-Int.Set.remove i2 state.pa_classes; r2.inductive_status<-Partial_applied | Total cpl,Unknown -> r2.inductive_status<-Total cpl; | Total (i,pac),Total _ -> Queue.add (i,Cmark pac) state.marks | _,_ -> () let merge eq state = (* merge and no-merge *) debug (fun () -> str "Merging " ++ pr_idx_term state.env state.sigma state.uf eq.lhs ++ str " and " ++ pr_idx_term state.env state.sigma state.uf eq.rhs ++ str "."); let uf=state.uf in let i=find uf eq.lhs and j=find uf eq.rhs in if not (Int.equal i j) then if (size uf i)<(size uf j) then union state i j eq else union state j i (swap eq) let update t state = (* update 1 and 2 *) debug (fun () -> str "Updating term " ++ pr_idx_term state.env state.sigma state.uf t ++ str "."); let (i,j) as sign = signature state.uf t in let (u,v) = subterms state.uf t in let rep = get_representative state.uf i in begin match rep.inductive_status with Partial _ -> rep.inductive_status <- Partial_applied; state.pa_classes <- Int.Set.remove i state.pa_classes | _ -> () end; PacMap.iter (fun pac _ -> Queue.add (t,Cmark (append_pac v pac)) state.marks) (get_constructors state.uf i); PafMap.iter (fun paf _ -> Queue.add (t,Fmark (fsucc paf)) state.marks) rep.functions; try let s = ST.query sign state.sigtable in Queue.add {lhs=t;rhs=s;rule=Congruence} state.combine with Not_found -> ST.enter t sign state.sigtable let process_function_mark t rep paf state = add_paf rep paf t; state.terms<-Int.Set.union rep.lfathers state.terms let process_constructor_mark t i rep pac state = add_pac state.uf.map.(i) pac t; match rep.inductive_status with Total (s,opac) -> if not (Int.equal pac.cnode opac.cnode) then (* Conflict *) raise (Discriminable (s,opac,t,pac)) else (* Match *) let cinfo = get_constructor_info state.uf pac.cnode in let rec f n oargs args= if n > 0 then match (oargs,args) with s1::q1,s2::q2-> Queue.add {lhs=s1;rhs=s2;rule=Injection(s,opac,t,pac,n)} state.combine; f (n-1) q1 q2 | _-> anomaly ~label:"add_pacs" (Pp.str "weird error in injection subterms merge.") in f cinfo.ci_nhyps opac.args pac.args | Partial_applied | Partial _ -> (* add_pac state.uf.map.(i) pac t; *) state.terms<-Int.Set.union rep.lfathers state.terms | Unknown -> if Int.equal pac.arity 0 then rep.inductive_status <- Total (t,pac) else begin (* add_pac state.uf.map.(i) pac t; *) state.terms<-Int.Set.union rep.lfathers state.terms; rep.inductive_status <- Partial pac; state.pa_classes<- Int.Set.add i state.pa_classes end let process_mark t m state = debug (fun () -> str "Processing mark for term " ++ pr_idx_term state.env state.sigma state.uf t ++ str "."); let i=find state.uf t in let rep=get_representative state.uf i in match m with Fmark paf -> process_function_mark t rep paf state | Cmark pac -> process_constructor_mark t i rep pac state type explanation = Discrimination of (int*pa_constructor*int*pa_constructor) | Contradiction of disequality | Incomplete let check_disequalities state = let uf=state.uf in let rec check_aux = function | dis::q -> let (info, ans) = if Int.equal (find uf dis.lhs) (find uf dis.rhs) then (str "Yes", Some dis) else (str "No", check_aux q) in let _ = debug (fun () -> str "Checking if " ++ pr_idx_term state.env state.sigma state.uf dis.lhs ++ str " = " ++ pr_idx_term state.env state.sigma state.uf dis.rhs ++ str " ... " ++ info) in ans | [] -> None in check_aux state.diseq let one_step state = try let eq = Queue.take state.combine in merge eq state; true with Queue.Empty -> try let (t,m) = Queue.take state.marks in process_mark t m state; true with Queue.Empty -> try let t = Int.Set.choose state.terms in state.terms<-Int.Set.remove t state.terms; update t state; true with Not_found -> false let __eps__ = Id.of_string "_eps_" let new_state_var typ state = let ids = Environ.ids_of_named_context_val (Environ.named_context_val state.env) in let id = Namegen.next_ident_away __eps__ ids in let r = Sorts.Relevant in (* TODO relevance *) state.env<- EConstr.push_named (Context.Named.Declaration.LocalAssum (make_annot id r,typ)) state.env; id let complete_one_class state i= match (get_representative state.uf i).inductive_status with Partial pac -> let rec app t typ n = if n<=0 then t else let _,etyp,rest= destProd typ in let id = new_state_var (EConstr.of_constr etyp) state in app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in let _c = Typing.unsafe_type_of state.env state.sigma (EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in let _c = EConstr.Unsafe.to_constr _c in let _args = List.map (fun i -> constr_of_term (term state.uf i)) pac.args in let typ = Term.prod_applist _c (List.rev _args) in let ct = app (term state.uf i) typ pac.arity in state.uf.epsilons <- pac :: state.uf.epsilons; ignore (add_term state ct) | _ -> anomaly (Pp.str "wrong incomplete class.") let complete state = Int.Set.iter (complete_one_class state) state.pa_classes type matching_problem = {mp_subst : int array; mp_inst : quant_eq; mp_stack : (ccpattern*int) list } let make_fun_table state = let uf= state.uf in let funtab=ref PafMap.empty in Array.iteri (fun i inode -> if i < uf.size then match inode.clas with Rep rep -> PafMap.iter (fun paf _ -> let elem = try PafMap.find paf !funtab with Not_found -> Int.Set.empty in funtab:= PafMap.add paf (Int.Set.add i elem) !funtab) rep.functions | _ -> ()) state.uf.map; !funtab let do_match state res pb_stack = let mp=Stack.pop pb_stack in match mp.mp_stack with [] -> res:= (mp.mp_inst,mp.mp_subst) :: !res | (patt,cl)::remains -> let uf=state.uf in match patt with PVar i -> if mp.mp_subst.(pred i)<0 then begin mp.mp_subst.(pred i)<- cl; (* no aliasing problem here *) Stack.push {mp with mp_stack=remains} pb_stack end else if Int.equal mp.mp_subst.(pred i) cl then Stack.push {mp with mp_stack=remains} pb_stack else (* mismatch for non-linear variable in pattern *) () | PApp (f,[]) -> begin try let j=Termhash.find uf.syms f in if Int.equal (find uf j) cl then Stack.push {mp with mp_stack=remains} pb_stack with Not_found -> () end | PApp(f, ((last_arg::rem_args) as args)) -> try let j=Termhash.find uf.syms f in let paf={fsym=j;fnargs=List.length args} in let rep=get_representative uf cl in let good_terms = PafMap.find paf rep.functions in let aux i = let (s,t) = signature state.uf i in Stack.push {mp with mp_subst=Array.copy mp.mp_subst; mp_stack= (PApp(f,rem_args),s) :: (last_arg,t) :: remains} pb_stack in Int.Set.iter aux good_terms with Not_found -> () let paf_of_patt syms = function PVar _ -> invalid_arg "paf_of_patt: pattern is trivial" | PApp (f,args) -> {fsym=Termhash.find syms f; fnargs=List.length args} let init_pb_stack state = let syms= state.uf.syms in let pb_stack = Stack.create () in let funtab = make_fun_table state in let aux inst = begin let good_classes = match inst.qe_lhs_valid with Creates_variables -> Int.Set.empty | Normal -> begin try let paf= paf_of_patt syms inst.qe_lhs in PafMap.find paf funtab with Not_found -> Int.Set.empty end | Trivial typ -> begin try Typehash.find state.by_type typ with Not_found -> Int.Set.empty end in Int.Set.iter (fun i -> Stack.push {mp_subst = Array.make inst.qe_nvars (-1); mp_inst=inst; mp_stack=[inst.qe_lhs,i]} pb_stack) good_classes end; begin let good_classes = match inst.qe_rhs_valid with Creates_variables -> Int.Set.empty | Normal -> begin try let paf= paf_of_patt syms inst.qe_rhs in PafMap.find paf funtab with Not_found -> Int.Set.empty end | Trivial typ -> begin try Typehash.find state.by_type typ with Not_found -> Int.Set.empty end in Int.Set.iter (fun i -> Stack.push {mp_subst = Array.make inst.qe_nvars (-1); mp_inst=inst; mp_stack=[inst.qe_rhs,i]} pb_stack) good_classes end in List.iter aux state.quant; pb_stack let find_instances state = let pb_stack= init_pb_stack state in let res =ref [] in let _ = debug (fun () -> str "Running E-matching algorithm ... "); try while true do Control.check_for_interrupt (); do_match state res pb_stack done; anomaly (Pp.str "get out of here!") with Stack.Empty -> () in !res let rec execute first_run state = debug (fun () -> str "Executing ... "); try while Control.check_for_interrupt (); one_step state do () done; match check_disequalities state with None -> if not(Int.Set.is_empty state.pa_classes) then begin debug (fun () -> str "First run was incomplete, completing ... "); complete state; execute false state end else if state.rew_depth>0 then let l=find_instances state in List.iter (add_inst state) l; if state.changed then begin state.changed <- false; execute true state end else begin debug (fun () -> str "Out of instances ... "); None end else begin debug (fun () -> str "Out of depth ... "); None end | Some dis -> Some begin if first_run then Contradiction dis else Incomplete end with Discriminable(s,spac,t,tpac) -> Some begin if first_run then Discrimination (s,spac,t,tpac) else Incomplete end coq-8.11.0/plugins/cc/README0000644000175000017500000000104013612664533015207 0ustar treinentreinen cctac: congruence-closure for coq author: Pierre Corbineau, Stage de DEA au LSV, ENS Cachan Thèse au LRI, Université Paris Sud XI Files : - ccalgo.ml : congruence closure algorithm - ccproof.ml : proof generation code - cctac.mlg : the tactic itself - CCSolve.v : a small Ltac tactic based on congruence Known Bugs : the congruence tactic can fail due to type dependencies. Related documents: Peter J. Downey, Ravi Sethi, and Robert E. Tarjan. Variations on the common subexpression problem. JACM, 27(4):758-771, October 1980. coq-8.11.0/plugins/cc/plugin_base.dune0000644000175000017500000000020613612664533017477 0ustar treinentreinen(library (name cc_plugin) (public_name coq.plugins.cc) (synopsis "Coq's congruence closure plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/.merlin.in0000644000175000017500000000000413612664533015635 0ustar treinentreinenREC coq-8.11.0/plugins/omega/0000755000175000017500000000000013612664533015037 5ustar treinentreinencoq-8.11.0/plugins/omega/omega_plugin.mlpack0000644000175000017500000000003013612664533020667 0ustar treinentreinenOmega Coq_omega G_omega coq-8.11.0/plugins/omega/OmegaTactic.v0000644000175000017500000000145613612664533017414 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bigint -> bool val less_than : bigint -> bigint -> bool val add : bigint -> bigint -> bigint val sub : bigint -> bigint -> bigint val mult : bigint -> bigint -> bigint val euclid : bigint -> bigint -> bigint * bigint val neg : bigint -> bigint val zero : bigint val one : bigint val to_string : bigint -> string end let debug = ref false module MakeOmegaSolver (I:INT) = struct type bigint = I.bigint let (=?) = I.equal let (?) x y = I.less_than y x let (>=?) x y = I.less_than y x || x = y let (+) = I.add let (-) = I.sub let ( * ) = I.mult let (/) x y = fst (I.euclid x y) let (mod) x y = snd (I.euclid x y) let zero = I.zero let one = I.one let two = one + one let negone = I.neg one let abs x = if I.less_than x zero then I.neg x else x let string_of_bigint = I.to_string let neg = I.neg (* To ensure that polymorphic (<) is not used mistakenly on big integers *) (* Warning: do not use (=) either on big int *) let (<) = ((<) : int -> int -> bool) let (>) = ((>) : int -> int -> bool) let (<=) = ((<=) : int -> int -> bool) let (>=) = ((>=) : int -> int -> bool) let pp i = print_int i; print_newline (); flush stdout let push v l = l := v :: !l let rec pgcd x y = if y =? zero then x else pgcd y (x mod y) let pgcd_l = function | [] -> failwith "pgcd_l" | x :: l -> List.fold_left pgcd x l let floor_div a b = match a >=? zero , b >? zero with | true,true -> a / b | false,false -> a / b | true, false -> (a-one) / b - one | false,true -> (a+one) / b - one type coeff = {c: bigint ; v: int} type linear = coeff list type eqn_kind = EQUA | INEQ | DISE type afine = { (* a number uniquely identifying the equation *) id: int ; (* a boolean true for an eq, false for an ineq (Sigma a_i x_i >= 0) *) kind: eqn_kind; (* the variables and their coefficient *) body: coeff list; (* a constant *) constant: bigint } type state_action = { st_new_eq : afine; st_def : afine; (* /!\ this represents [st_def = st_var] *) st_orig : afine; st_coef : bigint; st_var : int } type action = | DIVIDE_AND_APPROX of afine * afine * bigint * bigint | NOT_EXACT_DIVIDE of afine * bigint | FORGET_C of int | EXACT_DIVIDE of afine * bigint | SUM of int * (bigint * afine) * (bigint * afine) | STATE of state_action | HYP of afine | FORGET of int * int | FORGET_I of int * int | CONTRADICTION of afine * afine | NEGATE_CONTRADICT of afine * afine * bool | MERGE_EQ of int * afine * int | CONSTANT_NOT_NUL of int * bigint | CONSTANT_NUL of int | CONSTANT_NEG of int * bigint | SPLIT_INEQ of afine * (int * action list) * (int * action list) | WEAKEN of int * bigint exception UNSOLVABLE exception NO_CONTRADICTION let display_eq print_var (l,e) = let _ = List.fold_left (fun not_first f -> print_string (if f.c ? zero then Printf.printf "+ %s " (string_of_bigint e) else if e accu + one + trace_length l1 + trace_length l2 | _ -> accu + one in List.fold_left action_length zero l let operator_of_eq = function | EQUA -> "=" | DISE -> "!=" | INEQ -> ">=" let kind_of = function | EQUA -> "equation" | DISE -> "disequation" | INEQ -> "inequation" let display_system print_var l = List.iter (fun { kind=b; body=e; constant=c; id=id} -> Printf.printf "E%d: " id; display_eq print_var (e,c); Printf.printf "%s 0\n" (operator_of_eq b)) l; print_string "------------------------\n\n" let display_inequations print_var l = List.iter (fun e -> display_eq print_var e;print_string ">= 0\n") l; print_string "------------------------\n\n" let sbi = string_of_bigint let rec display_action print_var = function | act :: l -> begin match act with | DIVIDE_AND_APPROX (e1,e2,k,d) -> Printf.printf "Inequation E%d is divided by %s and the constant coefficient is \ rounded by subtracting %s.\n" e1.id (sbi k) (sbi d) | NOT_EXACT_DIVIDE (e,k) -> Printf.printf "Constant in equation E%d is not divisible by the pgcd \ %s of its other coefficients.\n" e.id (sbi k) | EXACT_DIVIDE (e,k) -> Printf.printf "Equation E%d is divided by the pgcd \ %s of its coefficients.\n" e.id (sbi k) | WEAKEN (e,k) -> Printf.printf "To ensure a solution in the dark shadow \ the equation E%d is weakened by %s.\n" e (sbi k) | SUM (e,(c1,e1),(c2,e2)) -> Printf.printf "We state %s E%d = %s %s E%d + %s %s E%d.\n" (kind_of e1.kind) e (sbi c1) (kind_of e1.kind) e1.id (sbi c2) (kind_of e2.kind) e2.id | STATE { st_new_eq = e } -> Printf.printf "We define a new equation E%d: " e.id; display_eq print_var (e.body,e.constant); print_string (operator_of_eq e.kind); print_string " 0" | HYP e -> Printf.printf "We define E%d: " e.id; display_eq print_var (e.body,e.constant); print_string (operator_of_eq e.kind); print_string " 0\n" | FORGET_C e -> Printf.printf "E%d is trivially satisfiable.\n" e | FORGET (e1,e2) -> Printf.printf "E%d subsumes E%d.\n" e1 e2 | FORGET_I (e1,e2) -> Printf.printf "E%d subsumes E%d.\n" e1 e2 | MERGE_EQ (e,e1,e2) -> Printf.printf "E%d and E%d can be merged into E%d.\n" e1.id e2 e | CONTRADICTION (e1,e2) -> Printf.printf "Equations E%d and E%d imply a contradiction on their \ constant factors.\n" e1.id e2.id | NEGATE_CONTRADICT(e1,e2,b) -> Printf.printf "Equations E%d and E%d state that their body is at the same time \ equal and different\n" e1.id e2.id | CONSTANT_NOT_NUL (e,k) -> Printf.printf "Equation E%d states %s = 0.\n" e (sbi k) | CONSTANT_NEG(e,k) -> Printf.printf "Equation E%d states %s >= 0.\n" e (sbi k) | CONSTANT_NUL e -> Printf.printf "Inequation E%d states 0 != 0.\n" e | SPLIT_INEQ (e,(e1,l1),(e2,l2)) -> Printf.printf "Equation E%d is split in E%d and E%d\n\n" e.id e1 e2; display_action print_var l1; print_newline (); display_action print_var l2; print_newline () end; display_action print_var l | [] -> flush stdout let default_print_var v = Printf.sprintf "X%d" v (* For debugging *) (*""*) let add_event, history, clear_history = let accu = ref [] in (fun (v:action) -> if !debug then display_action default_print_var [v]; push v accu), (fun () -> !accu), (fun () -> accu := []) let nf_linear = List.sort (fun x y -> Util.(-) y.v x.v) let nf ((b : bool),(e,(x : int))) = (b,(nf_linear e,x)) let map_eq_linear f = let rec loop = function | x :: l -> let c = f x.c in if c=?zero then loop l else {v=x.v; c=c} :: loop l | [] -> [] in loop let map_eq_afine f e = { id = e.id; kind = e.kind; body = map_eq_linear f e.body; constant = f e.constant } let negate_eq = map_eq_afine (fun x -> neg x) let rec sum p0 p1 = match (p0,p1) with | ([], l) -> l | (l, []) -> l | (((x1::l1) as l1'), ((x2::l2) as l2')) -> if x1.v = x2.v then let c = x1.c + x2.c in if c =? zero then sum l1 l2 else {v=x1.v;c=c} :: sum l1 l2 else if x1.v > x2.v then x1 :: sum l1 l2' else x2 :: sum l1' l2 let sum_afine new_eq_id eq1 eq2 = { kind = eq1.kind; id = new_eq_id (); body = sum eq1.body eq2.body; constant = eq1.constant + eq2.constant } exception FACTOR1 let rec chop_factor_1 = function | x :: l -> if abs x.c =? one then x,l else let (c',l') = chop_factor_1 l in (c',x::l') | [] -> raise FACTOR1 exception CHOPVAR let rec chop_var v = function | f :: l -> if f.v = v then f,l else let (f',l') = chop_var v l in (f',f::l') | [] -> raise CHOPVAR let normalize ({id=id; kind=eq_flag; body=e; constant =x} as eq) = if e = [] then begin match eq_flag with | EQUA -> if x =? zero then [] else begin add_event (CONSTANT_NOT_NUL(id,x)); raise UNSOLVABLE end | DISE -> if x <> zero then [] else begin add_event (CONSTANT_NUL id); raise UNSOLVABLE end | INEQ -> if x >=? zero then [] else begin add_event (CONSTANT_NEG(id,x)); raise UNSOLVABLE end end else let gcd = pgcd_l (List.map (fun f -> abs f.c) e) in if eq_flag=EQUA && x mod gcd <> zero then begin add_event (NOT_EXACT_DIVIDE (eq,gcd)); raise UNSOLVABLE end else if eq_flag=DISE && x mod gcd <> zero then begin add_event (FORGET_C eq.id); [] end else if gcd <> one then begin let c = floor_div x gcd in let d = x - c * gcd in let new_eq = {id=id; kind=eq_flag; constant=c; body=map_eq_linear (fun c -> c / gcd) e} in add_event (if eq_flag=EQUA || eq_flag = DISE then EXACT_DIVIDE(eq,gcd) else DIVIDE_AND_APPROX(eq,new_eq,gcd,d)); [new_eq] end else [eq] let eliminate_with_in new_eq_id {v=v;c=c_unite} eq2 ({body=e1; constant=c1} as eq1) = try let (f,_) = chop_var v e1 in let coeff = if c_unite=?one then neg f.c else if c_unite=? negone then f.c else failwith "eliminate_with_in" in let res = sum_afine new_eq_id eq1 (map_eq_afine (fun c -> c * coeff) eq2) in add_event (SUM (res.id,(one,eq1),(coeff,eq2))); res with CHOPVAR -> eq1 let omega_mod a b = a - b * floor_div (two * a + b) (two * b) let banerjee_step (new_eq_id,new_var_id,print_var) original l1 l2 = let e = original.body in let sigma = new_var_id () in if e == [] then begin display_system print_var [original] ; failwith "TL" end; let smallest,var = List.fold_left (fun (v,p) c -> if v >? (abs c.c) then abs c.c,c.v else (v,p)) (abs (List.hd e).c, (List.hd e).v) (List.tl e) in let m = smallest + one in let new_eq = { constant = omega_mod original.constant m; body = {c= neg m;v=sigma} :: map_eq_linear (fun a -> omega_mod a m) original.body; id = new_eq_id (); kind = EQUA } in let definition = { constant = neg (floor_div (two * original.constant + m) (two * m)); body = map_eq_linear (fun a -> neg (floor_div (two * a + m) (two * m))) original.body; id = new_eq_id (); kind = EQUA } in add_event (STATE {st_new_eq = new_eq; st_def = definition; st_orig = original; st_coef = m; st_var = sigma}); let new_eq = List.hd (normalize new_eq) in let eliminated_var, def = chop_var var new_eq.body in let other_equations = Util.List.map_append (fun e -> normalize (eliminate_with_in new_eq_id eliminated_var new_eq e)) l1 in let inequations = Util.List.map_append (fun e -> normalize (eliminate_with_in new_eq_id eliminated_var new_eq e)) l2 in let original' = eliminate_with_in new_eq_id eliminated_var new_eq original in let mod_original = map_eq_afine (fun c -> c / m) original' in add_event (EXACT_DIVIDE (original',m)); List.hd (normalize mod_original),other_equations,inequations let rec eliminate_one_equation ((new_eq_id,new_var_id,print_var) as new_ids) (e,other,ineqs) = if !debug then display_system print_var (e::other); try let v,def = chop_factor_1 e.body in (Util.List.map_append (fun e' -> normalize (eliminate_with_in new_eq_id v e e')) other, Util.List.map_append (fun e' -> normalize (eliminate_with_in new_eq_id v e e')) ineqs) with FACTOR1 -> eliminate_one_equation new_ids (banerjee_step new_ids e other ineqs) let rec banerjee ((_,_,print_var) as new_ids) (sys_eq,sys_ineq) = let rec fst_eq_1 = function (eq::l) -> if List.exists (fun x -> abs x.c =? one) eq.body then eq,l else let (eq',l') = fst_eq_1 l in (eq',eq::l') | [] -> raise Not_found in match sys_eq with [] -> if !debug then display_system print_var sys_ineq; sys_ineq | (e1::rest) -> let eq,other = try fst_eq_1 sys_eq with Not_found -> (e1,rest) in if eq.body = [] then if eq.constant =? zero then begin add_event (FORGET_C eq.id); banerjee new_ids (other,sys_ineq) end else begin add_event (CONSTANT_NOT_NUL(eq.id,eq.constant)); raise UNSOLVABLE end else banerjee new_ids (eliminate_one_equation new_ids (eq,other,sys_ineq)) type kind = INVERTED | NORMAL let redundancy_elimination new_eq_id system = let normal = function ({body=f::_} as e) when f.c negate_eq e, INVERTED | e -> e,NORMAL in let table = Hashtbl.create 7 in List.iter (fun e -> let ({body=ne} as nx) ,kind = normal e in if ne = [] then if nx.constant let kept = if v.constant Some nx,optinvert end else begin match optinvert with Some v -> let _kept = if v.constant >? nx.constant then begin add_event (FORGET_I (v.id,nx.id));v end else begin add_event (FORGET_I (nx.id,v.id));nx end in (optnormal,Some(if v.constant >? nx.constant then v else nx)) | None -> optnormal,Some nx end in begin match final with (Some high, Some low) -> if high.constant () end; Hashtbl.remove table ne; Hashtbl.add table ne final with Not_found -> Hashtbl.add table ne (if kind = NORMAL then (Some nx,None) else (None,Some nx))) system; let accu_eq = ref [] in let accu_ineq = ref [] in Hashtbl.iter (fun p0 p1 -> match (p0,p1) with | (e, (Some x, Some y)) when x.constant =? y.constant -> let id=new_eq_id () in add_event (MERGE_EQ(id,x,y.id)); push {id=id; kind=EQUA; body=x.body; constant=x.constant} accu_eq | (e, (optnorm,optinvert)) -> begin match optnorm with Some x -> push x accu_ineq | _ -> () end; begin match optinvert with Some x -> push (negate_eq x) accu_ineq | _ -> () end) table; !accu_eq,!accu_ineq exception SOLVED_SYSTEM let select_variable system = let table = Hashtbl.create 7 in let push v c= try let r = Hashtbl.find table v in r := max !r (abs c) with Not_found -> Hashtbl.add table v (ref (abs c)) in List.iter (fun {body=l} -> List.iter (fun f -> push f.v f.c) l) system; let vmin,cmin = ref (-1), ref zero in let var_cpt = ref 0 in Hashtbl.iter (fun v ({contents = c}) -> incr var_cpt; if c try let f,eq' = chop_var v eq.body in if f.c >=? zero then (not_occ,((f.c,eq) :: below),over) else (not_occ,below,((neg f.c,eq) :: over)) with CHOPVAR -> (eq::not_occ,below,over)) ([],[],[]) system let product new_eq_id dark_shadow low high = List.fold_left (fun accu (a,eq1) -> List.fold_left (fun accu (b,eq2) -> let eq = sum_afine new_eq_id (map_eq_afine (fun c -> c * b) eq1) (map_eq_afine (fun c -> c * a) eq2) in add_event(SUM(eq.id,(b,eq1),(a,eq2))); match normalize eq with | [eq] -> let final_eq = if dark_shadow then let delta = (a - one) * (b - one) in add_event(WEAKEN(eq.id,delta)); {id = eq.id; kind=INEQ; body = eq.body; constant = eq.constant - delta} else eq in final_eq :: accu | (e::_) -> failwith "Product dardk" | [] -> accu) accu high) [] low let fourier_motzkin (new_eq_id,_,print_var) dark_shadow system = let v = select_variable system in let (ineq_out, ineq_low,ineq_high) = classify v system in let expanded = ineq_out @ product new_eq_id dark_shadow ineq_low ineq_high in if !debug then display_system print_var expanded; expanded let simplify ((new_eq_id,new_var_id,print_var) as new_ids) dark_shadow system = if List.exists (fun e -> e.kind = DISE) system then failwith "disequation in simplify"; clear_history (); List.iter (fun e -> add_event (HYP e)) system; let system = Util.List.map_append normalize system in let eqs,ineqs = List.partition (fun e -> e.kind=EQUA) system in let simp_eq,simp_ineq = redundancy_elimination new_eq_id ineqs in let system = (eqs @ simp_eq,simp_ineq) in let rec loop1a system = let sys_ineq = banerjee new_ids system in loop1b sys_ineq and loop1b sys_ineq = let simp_eq,simp_ineq = redundancy_elimination new_eq_id sys_ineq in if simp_eq = [] then simp_ineq else loop1a (simp_eq,simp_ineq) in let rec loop2 system = try let expanded = fourier_motzkin new_ids dark_shadow system in loop2 (loop1b expanded) with SOLVED_SYSTEM -> if !debug then display_system print_var system; system in loop2 (loop1a system) let rec depend relie_on accu = function | act :: l -> begin match act with | DIVIDE_AND_APPROX (e,_,_,_) -> if Int.List.mem e.id relie_on then depend relie_on (act::accu) l else depend relie_on accu l | EXACT_DIVIDE (e,_) -> if Int.List.mem e.id relie_on then depend relie_on (act::accu) l else depend relie_on accu l | WEAKEN (e,_) -> if Int.List.mem e relie_on then depend relie_on (act::accu) l else depend relie_on accu l | SUM (e,(_,e1),(_,e2)) -> if Int.List.mem e relie_on then depend (e1.id::e2.id::relie_on) (act::accu) l else depend relie_on accu l | STATE {st_new_eq=e;st_orig=o} -> if Int.List.mem e.id relie_on then depend (o.id::relie_on) (act::accu) l else depend relie_on accu l | HYP e -> if Int.List.mem e.id relie_on then depend relie_on (act::accu) l else depend relie_on accu l | FORGET_C _ -> depend relie_on accu l | FORGET _ -> depend relie_on accu l | FORGET_I _ -> depend relie_on accu l | MERGE_EQ (e,e1,e2) -> if Int.List.mem e relie_on then depend (e1.id::e2::relie_on) (act::accu) l else depend relie_on accu l | NOT_EXACT_DIVIDE (e,_) -> depend (e.id::relie_on) (act::accu) l | CONTRADICTION (e1,e2) -> depend (e1.id::e2.id::relie_on) (act::accu) l | CONSTANT_NOT_NUL (e,_) -> depend (e::relie_on) (act::accu) l | CONSTANT_NEG (e,_) -> depend (e::relie_on) (act::accu) l | CONSTANT_NUL e -> depend (e::relie_on) (act::accu) l | NEGATE_CONTRADICT (e1,e2,_) -> depend (e1.id::e2.id::relie_on) (act::accu) l | SPLIT_INEQ _ -> failwith "depend" end | [] -> relie_on, accu let negation (eqs,ineqs) = let diseq,_ = List.partition (fun e -> e.kind = DISE) ineqs in let normal = function | ({body=f::_} as e) when f.c negate_eq e, INVERTED | e -> e,NORMAL in let table = Hashtbl.create 7 in List.iter (fun e -> let {body=ne;constant=c} ,kind = normal e in Hashtbl.add table (ne,c) (kind,e)) diseq; List.iter (fun e -> assert (e.kind = EQUA); let {body=ne;constant=c},kind = normal e in try let (kind',e') = Hashtbl.find table (ne,c) in add_event (NEGATE_CONTRADICT (e,e',kind=kind')); raise UNSOLVABLE with Not_found -> ()) eqs exception FULL_SOLUTION of action list * int list let simplify_strong ((new_eq_id,new_var_id,print_var) as new_ids) system = clear_history (); List.iter (fun e -> add_event (HYP e)) system; (* Initial simplification phase *) let rec loop1a system = negation system; let sys_ineq = banerjee new_ids system in loop1b sys_ineq and loop1b sys_ineq = let dise,ine = List.partition (fun e -> e.kind = DISE) sys_ineq in let simp_eq,simp_ineq = redundancy_elimination new_eq_id ine in if simp_eq = [] then dise @ simp_ineq else loop1a (simp_eq,dise @ simp_ineq) in let rec loop2 system = try let expanded = fourier_motzkin new_ids false system in loop2 (loop1b expanded) with SOLVED_SYSTEM -> if !debug then display_system print_var system; system in let rec explode_diseq = function | (de::diseq,ineqs,expl_map) -> let id1 = new_eq_id () and id2 = new_eq_id () in let e1 = {id = id1; kind=INEQ; body = de.body; constant = de.constant -one} in let e2 = {id = id2; kind=INEQ; body = map_eq_linear neg de.body; constant = neg de.constant - one} in let new_sys = List.map (fun (what,sys) -> ((de.id,id1,true)::what, e1::sys)) ineqs @ List.map (fun (what,sys) -> ((de.id,id2,false)::what,e2::sys)) ineqs in explode_diseq (diseq,new_sys,(de.id,(de,id1,id2))::expl_map) | ([],ineqs,expl_map) -> ineqs,expl_map in try let system = Util.List.map_append normalize system in let eqs,ineqs = List.partition (fun e -> e.kind=EQUA) system in let dise,ine = List.partition (fun e -> e.kind = DISE) ineqs in let simp_eq,simp_ineq = redundancy_elimination new_eq_id ine in let system = (eqs @ simp_eq,simp_ineq @ dise) in let system' = loop1a system in let diseq,ineq = List.partition (fun e -> e.kind = DISE) system' in let first_segment = history () in let sys_exploded,explode_map = explode_diseq (diseq,[[],ineq],[]) in let all_solutions = List.map (fun (decomp,sys) -> clear_history (); try let _ = loop2 sys in raise NO_CONTRADICTION with UNSOLVABLE -> let relie_on,path = depend [] [] (history ()) in let dc,_ = List.partition (fun (_,id,_) -> Int.List.mem id relie_on) decomp in let red = List.map (fun (x,_,_) -> x) dc in (red,relie_on,decomp,path)) sys_exploded in let max_count sys = let tbl = Hashtbl.create 7 in let augment x = try incr (Hashtbl.find tbl x) with Not_found -> Hashtbl.add tbl x (ref 1) in let eq = ref (-1) and c = ref 0 in List.iter (function | ([],r_on,_,path) -> raise (FULL_SOLUTION (path,r_on)) | (l,_,_,_) -> List.iter augment l) sys; Hashtbl.iter (fun x v -> if !v > !c then begin eq := x; c := !v end) tbl; !eq in let rec solve systems = try let id = max_count systems in let rec sign = function | ((id',_,b)::l) -> if id=id' then b else sign l | [] -> failwith "solve" in let s1,s2 = List.partition (fun (_,_,decomp,_) -> sign decomp) systems in let remove_int (dep,ro,dc,pa) = (Util.List.except Int.equal id dep,ro,dc,pa) in let s1' = List.map remove_int s1 in let s2' = List.map remove_int s2 in let (r1,relie1) = solve s1' and (r2,relie2) = solve s2' in let (eq,id1,id2) = Int.List.assoc id explode_map in [SPLIT_INEQ(eq,(id1,r1),(id2, r2))], eq.id :: Util.List.union Int.equal relie1 relie2 with FULL_SOLUTION (x0,x1) -> (x0,x1) in let act,relie_on = solve all_solutions in snd(depend relie_on act first_segment) with UNSOLVABLE -> snd (depend [] [] (history ())) end coq-8.11.0/plugins/omega/coq_omega.mli0000644000175000017500000000131513612664533017474 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* x mod y = 0. Proof. intro; subst; destruct x; reflexivity. Qed. Lemma div_0_r_ext x y : y = 0 -> x / y = 0. Proof. intro; subst; destruct x; reflexivity. Qed. Lemma rem_0_r_ext x y : y = 0 -> Z.rem x y = x. Proof. intro; subst; destruct x; reflexivity. Qed. Lemma quot_0_r_ext x y : y = 0 -> Z.quot x y = 0. Proof. intro; subst; destruct x; reflexivity. Qed. Lemma rem_bound_pos_pos x y : 0 < y -> 0 <= x -> 0 <= Z.rem x y < y. Proof. intros; apply Z.rem_bound_pos; assumption. Qed. Lemma rem_bound_neg_pos x y : y < 0 -> 0 <= x -> 0 <= Z.rem x y < -y. Proof. rewrite <- Z.rem_opp_r'; intros; apply Z.rem_bound_pos; rewrite ?Z.opp_pos_neg; assumption. Qed. Lemma rem_bound_pos_neg x y : 0 < y -> x <= 0 -> -y < Z.rem x y <= 0. Proof. rewrite <- (Z.opp_involutive x), Z.rem_opp_l', <- Z.opp_lt_mono, and_comm, !Z.opp_nonpos_nonneg; apply rem_bound_pos_pos. Qed. Lemma rem_bound_neg_neg x y : y < 0 -> x <= 0 -> y < Z.rem x y <= 0. Proof. rewrite <- (Z.opp_involutive x), <- (Z.opp_involutive y), Z.rem_opp_l', <- Z.opp_lt_mono, and_comm, !Z.opp_nonpos_nonneg, Z.opp_involutive; apply rem_bound_neg_pos. Qed. Ltac div_mod_to_equations_generalize x y := pose proof (Z.div_mod x y); pose proof (Z.mod_pos_bound x y); pose proof (Z.mod_neg_bound x y); pose proof (div_0_r_ext x y); pose proof (mod_0_r_ext x y); let q := fresh "q" in let r := fresh "r" in set (q := x / y) in *; set (r := x mod y) in *; clearbody q r. Ltac quot_rem_to_equations_generalize x y := pose proof (Z.quot_rem' x y); pose proof (rem_bound_pos_pos x y); pose proof (rem_bound_pos_neg x y); pose proof (rem_bound_neg_pos x y); pose proof (rem_bound_neg_neg x y); pose proof (quot_0_r_ext x y); pose proof (rem_0_r_ext x y); let q := fresh "q" in let r := fresh "r" in set (q := Z.quot x y) in *; set (r := Z.rem x y) in *; clearbody q r. Ltac div_mod_to_equations_step := match goal with | [ |- context[?x / ?y] ] => div_mod_to_equations_generalize x y | [ |- context[?x mod ?y] ] => div_mod_to_equations_generalize x y | [ H : context[?x / ?y] |- _ ] => div_mod_to_equations_generalize x y | [ H : context[?x mod ?y] |- _ ] => div_mod_to_equations_generalize x y end. Ltac quot_rem_to_equations_step := match goal with | [ |- context[Z.quot ?x ?y] ] => quot_rem_to_equations_generalize x y | [ |- context[Z.rem ?x ?y] ] => quot_rem_to_equations_generalize x y | [ H : context[Z.quot ?x ?y] |- _ ] => quot_rem_to_equations_generalize x y | [ H : context[Z.rem ?x ?y] |- _ ] => quot_rem_to_equations_generalize x y end. Ltac div_mod_to_equations' := repeat div_mod_to_equations_step. Ltac quot_rem_to_equations' := repeat quot_rem_to_equations_step. Ltac euclidean_division_equations_cleanup := repeat match goal with | [ H : ?x = ?x -> _ |- _ ] => specialize (H eq_refl) | [ H : ?x <> ?x -> _ |- _ ] => clear H | [ H : ?x < ?x -> _ |- _ ] => clear H | [ H : ?T -> _, H' : ?T |- _ ] => specialize (H H') | [ H : ?T -> _, H' : ~?T |- _ ] => clear H | [ H : ~?T -> _, H' : ?T |- _ ] => clear H | [ H : ?A -> ?x = ?x -> _ |- _ ] => specialize (fun a => H a eq_refl) | [ H : ?A -> ?x <> ?x -> _ |- _ ] => clear H | [ H : ?A -> ?x < ?x -> _ |- _ ] => clear H | [ H : ?A -> ?B -> _, H' : ?B |- _ ] => specialize (fun a => H a H') | [ H : ?A -> ?B -> _, H' : ~?B |- _ ] => clear H | [ H : ?A -> ~?B -> _, H' : ?B |- _ ] => clear H | [ H : 0 < ?x -> _, H' : ?x < 0 |- _ ] => clear H | [ H : ?x < 0 -> _, H' : 0 < ?x |- _ ] => clear H | [ H : ?A -> 0 < ?x -> _, H' : ?x < 0 |- _ ] => clear H | [ H : ?A -> ?x < 0 -> _, H' : 0 < ?x |- _ ] => clear H | [ H : 0 <= ?x -> _, H' : ?x < 0 |- _ ] => clear H | [ H : ?x <= 0 -> _, H' : 0 < ?x |- _ ] => clear H | [ H : ?A -> 0 <= ?x -> _, H' : ?x < 0 |- _ ] => clear H | [ H : ?A -> ?x <= 0 -> _, H' : 0 < ?x |- _ ] => clear H | [ H : 0 < ?x -> _, H' : ?x <= 0 |- _ ] => clear H | [ H : ?x < 0 -> _, H' : 0 <= ?x |- _ ] => clear H | [ H : ?A -> 0 < ?x -> _, H' : ?x <= 0 |- _ ] => clear H | [ H : ?A -> ?x < 0 -> _, H' : 0 <= ?x |- _ ] => clear H | [ H : 0 <= ?x -> _, H' : ?x <= 0 |- _ ] => specialize (fun pf => H (@Z.eq_le_incl 0 x (eq_sym pf))) | [ H : ?A -> 0 <= ?x -> _, H' : ?x <= 0 |- _ ] => specialize (fun a pf => H a (@Z.eq_le_incl 0 x (eq_sym pf))) | [ H : ?x <= 0 -> _, H' : 0 <= ?x |- _ ] => specialize (fun pf => H (@Z.eq_le_incl 0 x pf)) | [ H : ?A -> ?x <= 0 -> _, H' : 0 <= ?x |- _ ] => specialize (fun a pf => H a (@Z.eq_le_incl x 0 pf)) | [ H : ?x < ?y -> _, H' : ?x = ?y |- _ ] => clear H | [ H : ?x < ?y -> _, H' : ?y = ?x |- _ ] => clear H | [ H : ?A -> ?x < ?y -> _, H' : ?x = ?y |- _ ] => clear H | [ H : ?A -> ?x < ?y -> _, H' : ?y = ?x |- _ ] => clear H | [ H : ?x = ?y -> _, H' : ?x < ?y |- _ ] => clear H | [ H : ?x = ?y -> _, H' : ?y < ?x |- _ ] => clear H | [ H : ?A -> ?x = ?y -> _, H' : ?x < ?y |- _ ] => clear H | [ H : ?A -> ?x = ?y -> _, H' : ?y < ?x |- _ ] => clear H end. Ltac div_mod_to_equations := div_mod_to_equations'; euclidean_division_equations_cleanup. Ltac quot_rem_to_equations := quot_rem_to_equations'; euclidean_division_equations_cleanup. Ltac to_euclidean_division_equations := div_mod_to_equations'; quot_rem_to_equations'; euclidean_division_equations_cleanup. End Z. Set Warnings "-deprecated-tactic". (** * zify: the Z-ification tactic *) (* This tactic searches for nat and N and positive elements in the goal and translates everything into Z. It is meant as a pre-processor for (r)omega; for instance a positivity hypothesis is added whenever - a multiplication is encountered - an atom is encountered (that is a variable or an unknown construct) Recognized relations (can be handled as deeply as allowed by setoid rewrite): - { eq, le, lt, ge, gt } on { Z, positive, N, nat } Recognized operations: - on Z: Z.min, Z.max, Z.abs, Z.sgn are translated in term of <= < = - on nat: + * - S O pred min max Pos.to_nat N.to_nat Z.abs_nat - on positive: Zneg Zpos xI xO xH + * - Pos.succ Pos.pred Pos.min Pos.max Pos.of_succ_nat - on N: N0 Npos + * - N.pred N.succ N.min N.max N.of_nat Z.abs_N *) (** I) translation of Z.max, Z.min, Z.abs, Z.sgn into recognized equations *) #[deprecated( note = "Use 'zify' instead")] Ltac zify_unop_core t thm a := (* Let's introduce the specification theorem for t *) pose proof (thm a); (* Then we replace (t a) everywhere with a fresh variable *) let z := fresh "z" in set (z:=t a) in *; clearbody z. #[deprecated( note = "Use 'zify' instead")] Ltac zify_unop_var_or_term t thm a := (* If a is a variable, no need for aliasing *) let za := fresh "z" in (rename a into za; rename za into a; zify_unop_core t thm a) || (* Otherwise, a is a complex term: we alias it. *) (remember a as za; zify_unop_core t thm za). #[deprecated( note = "Use 'zify' instead")] Ltac zify_unop t thm a := (* If a is a scalar, we can simply reduce the unop. *) (* Note that simpl wasn't enough to reduce [Z.max 0 0] (#5439) *) let isz := isZcst a in match isz with | true => let u := eval compute in (t a) in change (t a) with u in * | _ => zify_unop_var_or_term t thm a end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_unop_nored t thm a := (* in this version, we don't try to reduce the unop (that can be (Z.add x)) *) let isz := isZcst a in match isz with | true => zify_unop_core t thm a | _ => zify_unop_var_or_term t thm a end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_binop t thm a b:= (* works as zify_unop, except that we should be careful when dealing with b, since it can be equal to a *) let isza := isZcst a in match isza with | true => zify_unop (t a) (thm a) b | _ => let za := fresh "z" in (rename a into za; rename za into a; zify_unop_nored (t a) (thm a) b) || (remember a as za; match goal with | H : za = b |- _ => zify_unop_nored (t za) (thm za) za | _ => zify_unop_nored (t za) (thm za) b end) end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_op_1 := match goal with | x := ?t : Z |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x | |- context [ Z.max ?a ?b ] => zify_binop Z.max Z.max_spec a b | H : context [ Z.max ?a ?b ] |- _ => zify_binop Z.max Z.max_spec a b | |- context [ Z.min ?a ?b ] => zify_binop Z.min Z.min_spec a b | H : context [ Z.min ?a ?b ] |- _ => zify_binop Z.min Z.min_spec a b | |- context [ Z.sgn ?a ] => zify_unop Z.sgn Z.sgn_spec a | H : context [ Z.sgn ?a ] |- _ => zify_unop Z.sgn Z.sgn_spec a | |- context [ Z.abs ?a ] => zify_unop Z.abs Z.abs_spec a | H : context [ Z.abs ?a ] |- _ => zify_unop Z.abs Z.abs_spec a end. Ltac zify_op := repeat zify_op_1. (** II) Conversion from nat to Z *) Definition Z_of_nat' := Z.of_nat. Ltac hide_Z_of_nat t := let z := fresh "z" in set (z:=Z.of_nat t) in *; change Z.of_nat with Z_of_nat' in z; unfold z in *; clear z. #[deprecated( note = "Use 'zify' instead")] Ltac zify_nat_rel := match goal with (* I: equalities *) | x := ?t : nat |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x | |- (@eq nat ?a ?b) => apply (Nat2Z.inj a b) (* shortcut *) | H : context [ @eq nat ?a ?b ] |- _ => rewrite <- (Nat2Z.inj_iff a b) in H | |- context [ @eq nat ?a ?b ] => rewrite <- (Nat2Z.inj_iff a b) (* II: less than *) | H : context [ lt ?a ?b ] |- _ => rewrite (Nat2Z.inj_lt a b) in H | |- context [ lt ?a ?b ] => rewrite (Nat2Z.inj_lt a b) (* III: less or equal *) | H : context [ le ?a ?b ] |- _ => rewrite (Nat2Z.inj_le a b) in H | |- context [ le ?a ?b ] => rewrite (Nat2Z.inj_le a b) (* IV: greater than *) | H : context [ gt ?a ?b ] |- _ => rewrite (Nat2Z.inj_gt a b) in H | |- context [ gt ?a ?b ] => rewrite (Nat2Z.inj_gt a b) (* V: greater or equal *) | H : context [ ge ?a ?b ] |- _ => rewrite (Nat2Z.inj_ge a b) in H | |- context [ ge ?a ?b ] => rewrite (Nat2Z.inj_ge a b) end. Ltac zify_nat_op := match goal with (* misc type conversions: positive/N/Z to nat *) | H : context [ Z.of_nat (Pos.to_nat ?a) ] |- _ => rewrite (positive_nat_Z a) in H | |- context [ Z.of_nat (Pos.to_nat ?a) ] => rewrite (positive_nat_Z a) | H : context [ Z.of_nat (N.to_nat ?a) ] |- _ => rewrite (N_nat_Z a) in H | |- context [ Z.of_nat (N.to_nat ?a) ] => rewrite (N_nat_Z a) | H : context [ Z.of_nat (Z.abs_nat ?a) ] |- _ => rewrite (Zabs2Nat.id_abs a) in H | |- context [ Z.of_nat (Z.abs_nat ?a) ] => rewrite (Zabs2Nat.id_abs a) (* plus -> Z.add *) | H : context [ Z.of_nat (plus ?a ?b) ] |- _ => rewrite (Nat2Z.inj_add a b) in H | |- context [ Z.of_nat (plus ?a ?b) ] => rewrite (Nat2Z.inj_add a b) (* min -> Z.min *) | H : context [ Z.of_nat (min ?a ?b) ] |- _ => rewrite (Nat2Z.inj_min a b) in H | |- context [ Z.of_nat (min ?a ?b) ] => rewrite (Nat2Z.inj_min a b) (* max -> Z.max *) | H : context [ Z.of_nat (max ?a ?b) ] |- _ => rewrite (Nat2Z.inj_max a b) in H | |- context [ Z.of_nat (max ?a ?b) ] => rewrite (Nat2Z.inj_max a b) (* minus -> Z.max (Z.sub ... ...) 0 *) | H : context [ Z.of_nat (minus ?a ?b) ] |- _ => rewrite (Nat2Z.inj_sub_max a b) in H | |- context [ Z.of_nat (minus ?a ?b) ] => rewrite (Nat2Z.inj_sub_max a b) (* pred -> minus ... -1 -> Z.max (Z.sub ... -1) 0 *) | H : context [ Z.of_nat (pred ?a) ] |- _ => rewrite (pred_of_minus a) in H | |- context [ Z.of_nat (pred ?a) ] => rewrite (pred_of_minus a) (* mult -> Z.mul and a positivity hypothesis *) | H : context [ Z.of_nat (mult ?a ?b) ] |- _ => pose proof (Nat2Z.is_nonneg (mult a b)); rewrite (Nat2Z.inj_mul a b) in * | |- context [ Z.of_nat (mult ?a ?b) ] => pose proof (Nat2Z.is_nonneg (mult a b)); rewrite (Nat2Z.inj_mul a b) in * (* O -> Z0 *) | H : context [ Z.of_nat O ] |- _ => change (Z.of_nat O) with Z0 in H | |- context [ Z.of_nat O ] => change (Z.of_nat O) with Z0 (* S -> number or Z.succ *) | H : context [ Z.of_nat (S ?a) ] |- _ => let isnat := isnatcst a in match isnat with | true => let t := eval compute in (Z.of_nat (S a)) in change (Z.of_nat (S a)) with t in H | _ => rewrite (Nat2Z.inj_succ a) in H | _ => (* if the [rewrite] fails (most likely a dependent occurrence of [Z.of_nat (S a)]), hide [Z.of_nat (S a)] in this one hypothesis *) change (Z.of_nat (S a)) with (Z_of_nat' (S a)) in H end | |- context [ Z.of_nat (S ?a) ] => let isnat := isnatcst a in match isnat with | true => let t := eval compute in (Z.of_nat (S a)) in change (Z.of_nat (S a)) with t | _ => rewrite (Nat2Z.inj_succ a) | _ => (* if the [rewrite] fails (most likely a dependent occurrence of [Z.of_nat (S a)]), hide [Z.of_nat (S a)] in the goal *) change (Z.of_nat (S a)) with (Z_of_nat' (S a)) end (* atoms of type nat : we add a positivity condition (if not already there) *) | _ : 0 <= Z.of_nat ?a |- _ => hide_Z_of_nat a | _ : context [ Z.of_nat ?a ] |- _ => pose proof (Nat2Z.is_nonneg a); hide_Z_of_nat a | |- context [ Z.of_nat ?a ] => pose proof (Nat2Z.is_nonneg a); hide_Z_of_nat a end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_nat := repeat zify_nat_rel; repeat zify_nat_op; unfold Z_of_nat' in *. (* III) conversion from positive to Z *) Definition Zpos' := Zpos. Definition Zneg' := Zneg. Ltac hide_Zpos t := let z := fresh "z" in set (z:=Zpos t) in *; change Zpos with Zpos' in z; unfold z in *; clear z. #[deprecated( note = "Use 'zify' instead")] Ltac zify_positive_rel := match goal with (* I: equalities *) | x := ?t : positive |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x | |- (@eq positive ?a ?b) => apply Pos2Z.inj | H : context [ @eq positive ?a ?b ] |- _ => rewrite <- (Pos2Z.inj_iff a b) in H | |- context [ @eq positive ?a ?b ] => rewrite <- (Pos2Z.inj_iff a b) (* II: less than *) | H : context [ (?a < ?b)%positive ] |- _ => change (a change (a change (a<=b)%positive with (Zpos a<=Zpos b) in H | |- context [ (?a <= ?b)%positive ] => change (a<=b)%positive with (Zpos a<=Zpos b) (* IV: greater than *) | H : context [ (?a > ?b)%positive ] |- _ => change (a>b)%positive with (Zpos a>Zpos b) in H | |- context [ (?a > ?b)%positive ] => change (a>b)%positive with (Zpos a>Zpos b) (* V: greater or equal *) | H : context [ (?a >= ?b)%positive ] |- _ => change (a>=b)%positive with (Zpos a>=Zpos b) in H | |- context [ (?a >= ?b)%positive ] => change (a>=b)%positive with (Zpos a>=Zpos b) end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_positive_op := match goal with (* Z.pow_pos -> Z.pow *) | H : context [ Z.pow_pos ?a ?b ] |- _ => change (Z.pow_pos a b) with (Z.pow a (Z.pos b)) in H | |- context [ Z.pow_pos ?a ?b ] => change (Z.pow_pos a b) with (Z.pow a (Z.pos b)) (* Zneg -> -Zpos (except for numbers) *) | H : context [ Zneg ?a ] |- _ => let isp := isPcst a in match isp with | true => change (Zneg a) with (Zneg' a) in H | _ => change (Zneg a) with (- Zpos a) in H end | |- context [ Zneg ?a ] => let isp := isPcst a in match isp with | true => change (Zneg a) with (Zneg' a) | _ => change (Zneg a) with (- Zpos a) end (* misc type conversions: nat to positive *) | H : context [ Zpos (Pos.of_succ_nat ?a) ] |- _ => rewrite (Zpos_P_of_succ_nat a) in H | |- context [ Zpos (Pos.of_succ_nat ?a) ] => rewrite (Zpos_P_of_succ_nat a) (* Z.power_pos *) | H : context [ Zpos (Pos.of_succ_nat ?a) ] |- _ => rewrite (Zpos_P_of_succ_nat a) in H | |- context [ Zpos (Pos.of_succ_nat ?a) ] => rewrite (Zpos_P_of_succ_nat a) (* Pos.add -> Z.add *) | H : context [ Zpos (?a + ?b) ] |- _ => change (Zpos (a+b)) with (Zpos a + Zpos b) in H | |- context [ Zpos (?a + ?b) ] => change (Zpos (a+b)) with (Zpos a + Zpos b) (* Pos.min -> Z.min *) | H : context [ Zpos (Pos.min ?a ?b) ] |- _ => rewrite (Pos2Z.inj_min a b) in H | |- context [ Zpos (Pos.min ?a ?b) ] => rewrite (Pos2Z.inj_min a b) (* Pos.max -> Z.max *) | H : context [ Zpos (Pos.max ?a ?b) ] |- _ => rewrite (Pos2Z.inj_max a b) in H | |- context [ Zpos (Pos.max ?a ?b) ] => rewrite (Pos2Z.inj_max a b) (* Pos.sub -> Z.max 1 (Z.sub ... ...) *) | H : context [ Zpos (Pos.sub ?a ?b) ] |- _ => rewrite (Pos2Z.inj_sub_max a b) in H | |- context [ Zpos (Pos.sub ?a ?b) ] => rewrite (Pos2Z.inj_sub_max a b) (* Pos.succ -> Z.succ *) | H : context [ Zpos (Pos.succ ?a) ] |- _ => rewrite (Pos2Z.inj_succ a) in H | |- context [ Zpos (Pos.succ ?a) ] => rewrite (Pos2Z.inj_succ a) (* Pos.pred -> Pos.sub ... -1 -> Z.max 1 (Z.sub ... - 1) *) | H : context [ Zpos (Pos.pred ?a) ] |- _ => rewrite <- (Pos.sub_1_r a) in H | |- context [ Zpos (Pos.pred ?a) ] => rewrite <- (Pos.sub_1_r a) (* Pos.mul -> Z.mul and a positivity hypothesis *) | H : context [ Zpos (?a * ?b) ] |- _ => pose proof (Pos2Z.is_pos (Pos.mul a b)); change (Zpos (a*b)) with (Zpos a * Zpos b) in * | |- context [ Zpos (?a * ?b) ] => pose proof (Pos2Z.is_pos (Pos.mul a b)); change (Zpos (a*b)) with (Zpos a * Zpos b) in * (* xO *) | H : context [ Zpos (xO ?a) ] |- _ => let isp := isPcst a in match isp with | true => change (Zpos (xO a)) with (Zpos' (xO a)) in H | _ => rewrite (Pos2Z.inj_xO a) in H end | |- context [ Zpos (xO ?a) ] => let isp := isPcst a in match isp with | true => change (Zpos (xO a)) with (Zpos' (xO a)) | _ => rewrite (Pos2Z.inj_xO a) end (* xI *) | H : context [ Zpos (xI ?a) ] |- _ => let isp := isPcst a in match isp with | true => change (Zpos (xI a)) with (Zpos' (xI a)) in H | _ => rewrite (Pos2Z.inj_xI a) in H end | |- context [ Zpos (xI ?a) ] => let isp := isPcst a in match isp with | true => change (Zpos (xI a)) with (Zpos' (xI a)) | _ => rewrite (Pos2Z.inj_xI a) end (* xI : nothing to do, just prevent adding a useless positivity condition *) | H : context [ Zpos xH ] |- _ => hide_Zpos xH | |- context [ Zpos xH ] => hide_Zpos xH (* atoms of type positive : we add a positivity condition (if not already there) *) | _ : 0 < Zpos ?a |- _ => hide_Zpos a | _ : context [ Zpos ?a ] |- _ => pose proof (Pos2Z.is_pos a); hide_Zpos a | |- context [ Zpos ?a ] => pose proof (Pos2Z.is_pos a); hide_Zpos a end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_positive := repeat zify_positive_rel; repeat zify_positive_op; unfold Zpos',Zneg' in *. (* IV) conversion from N to Z *) Definition Z_of_N' := Z.of_N. Ltac hide_Z_of_N t := let z := fresh "z" in set (z:=Z.of_N t) in *; change Z.of_N with Z_of_N' in z; unfold z in *; clear z. #[deprecated( note = "Use 'zify' instead")] Ltac zify_N_rel := match goal with (* I: equalities *) | x := ?t : N |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x | |- (@eq N ?a ?b) => apply (N2Z.inj a b) (* shortcut *) | H : context [ @eq N ?a ?b ] |- _ => rewrite <- (N2Z.inj_iff a b) in H | |- context [ @eq N ?a ?b ] => rewrite <- (N2Z.inj_iff a b) (* II: less than *) | H : context [ (?a < ?b)%N ] |- _ => rewrite (N2Z.inj_lt a b) in H | |- context [ (?a < ?b)%N ] => rewrite (N2Z.inj_lt a b) (* III: less or equal *) | H : context [ (?a <= ?b)%N ] |- _ => rewrite (N2Z.inj_le a b) in H | |- context [ (?a <= ?b)%N ] => rewrite (N2Z.inj_le a b) (* IV: greater than *) | H : context [ (?a > ?b)%N ] |- _ => rewrite (N2Z.inj_gt a b) in H | |- context [ (?a > ?b)%N ] => rewrite (N2Z.inj_gt a b) (* V: greater or equal *) | H : context [ (?a >= ?b)%N ] |- _ => rewrite (N2Z.inj_ge a b) in H | |- context [ (?a >= ?b)%N ] => rewrite (N2Z.inj_ge a b) end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_N_op := match goal with (* misc type conversions: nat to positive *) | H : context [ Z.of_N (N.of_nat ?a) ] |- _ => rewrite (nat_N_Z a) in H | |- context [ Z.of_N (N.of_nat ?a) ] => rewrite (nat_N_Z a) | H : context [ Z.of_N (Z.abs_N ?a) ] |- _ => rewrite (N2Z.inj_abs_N a) in H | |- context [ Z.of_N (Z.abs_N ?a) ] => rewrite (N2Z.inj_abs_N a) | H : context [ Z.of_N (Npos ?a) ] |- _ => rewrite (N2Z.inj_pos a) in H | |- context [ Z.of_N (Npos ?a) ] => rewrite (N2Z.inj_pos a) | H : context [ Z.of_N N0 ] |- _ => change (Z.of_N N0) with Z0 in H | |- context [ Z.of_N N0 ] => change (Z.of_N N0) with Z0 (* N.add -> Z.add *) | H : context [ Z.of_N (N.add ?a ?b) ] |- _ => rewrite (N2Z.inj_add a b) in H | |- context [ Z.of_N (N.add ?a ?b) ] => rewrite (N2Z.inj_add a b) (* N.min -> Z.min *) | H : context [ Z.of_N (N.min ?a ?b) ] |- _ => rewrite (N2Z.inj_min a b) in H | |- context [ Z.of_N (N.min ?a ?b) ] => rewrite (N2Z.inj_min a b) (* N.max -> Z.max *) | H : context [ Z.of_N (N.max ?a ?b) ] |- _ => rewrite (N2Z.inj_max a b) in H | |- context [ Z.of_N (N.max ?a ?b) ] => rewrite (N2Z.inj_max a b) (* N.sub -> Z.max 0 (Z.sub ... ...) *) | H : context [ Z.of_N (N.sub ?a ?b) ] |- _ => rewrite (N2Z.inj_sub_max a b) in H | |- context [ Z.of_N (N.sub ?a ?b) ] => rewrite (N2Z.inj_sub_max a b) (* pred -> minus ... -1 -> Z.max (Z.sub ... -1) 0 *) | H : context [ Z.of_N (N.pred ?a) ] |- _ => rewrite (N.pred_sub a) in H | |- context [ Z.of_N (N.pred ?a) ] => rewrite (N.pred_sub a) (* N.succ -> Z.succ *) | H : context [ Z.of_N (N.succ ?a) ] |- _ => rewrite (N2Z.inj_succ a) in H | |- context [ Z.of_N (N.succ ?a) ] => rewrite (N2Z.inj_succ a) (* N.mul -> Z.mul and a positivity hypothesis *) | H : context [ Z.of_N (N.mul ?a ?b) ] |- _ => pose proof (N2Z.is_nonneg (N.mul a b)); rewrite (N2Z.inj_mul a b) in * | |- context [ Z.of_N (N.mul ?a ?b) ] => pose proof (N2Z.is_nonneg (N.mul a b)); rewrite (N2Z.inj_mul a b) in * (* N.div -> Z.div and a positivity hypothesis *) | H : context [ Z.of_N (N.div ?a ?b) ] |- _ => pose proof (N2Z.is_nonneg (N.div a b)); rewrite (N2Z.inj_div a b) in * | |- context [ Z.of_N (N.div ?a ?b) ] => pose proof (N2Z.is_nonneg (N.div a b)); rewrite (N2Z.inj_div a b) in * (* N.modulo -> Z.rem / Z.modulo and a positivity hypothesis (N.modulo agrees with Z.modulo on everything except 0; so we pose both the non-zero proof for this agreement, but also replace things with [Z.rem]) *) | H : context [ Z.of_N (N.modulo ?a ?b) ] |- _ => pose proof (N2Z.is_nonneg (N.modulo a b)); pose proof (@Z.quot_div_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); pose proof (@Z.rem_mod_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); rewrite (N2Z.inj_rem a b) in * | |- context [ Z.of_N (N.div ?a ?b) ] => pose proof (N2Z.is_nonneg (N.modulo a b)); pose proof (@Z.quot_div_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); pose proof (@Z.rem_mod_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); rewrite (N2Z.inj_rem a b) in * (* atoms of type N : we add a positivity condition (if not already there) *) | _ : 0 <= Z.of_N ?a |- _ => hide_Z_of_N a | _ : context [ Z.of_N ?a ] |- _ => pose proof (N2Z.is_nonneg a); hide_Z_of_N a | |- context [ Z.of_N ?a ] => pose proof (N2Z.is_nonneg a); hide_Z_of_N a end. #[deprecated( note = "Use 'zify' instead")] Ltac zify_N := repeat zify_N_rel; repeat zify_N_op; unfold Z_of_N' in *. (** The complete Z-ification tactic *) Require Import ZifyClasses ZifyInst. Require Zify. (** [is_inj T] returns true iff the type T has an injection *) Ltac is_inj T := match T with | _ => let x := constr:(_ : InjTyp T _ ) in true | _ => false end. (* [elim_let] replaces a let binding (x := e : t) by an equation (x = e) if t is an injected type *) Ltac elim_let := repeat match goal with | x := ?t : ?ty |- _ => let b := is_inj ty in match b with | true => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x end end. Ltac zify := intros ; elim_let ; Zify.zify ; ZifyInst.saturate. coq-8.11.0/plugins/omega/OmegaLemmas.v0000644000175000017500000002400313612664533017414 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 0 <= x -> 0 <= y. Proof. now intros ->. Qed. Lemma OMEGA2 x y : 0 <= x -> 0 <= y -> 0 <= x + y. Proof. Z.order_pos. Qed. Lemma OMEGA3 x y k : k > 0 -> x = y * k -> x = 0 -> y = 0. Proof. intros LT -> EQ. apply Z.mul_eq_0 in EQ. destruct EQ; now subst. Qed. Lemma OMEGA4 x y z : x > 0 -> y > x -> z * y + x <> 0. Proof. Z.swap_greater. intros Hx Hxy. rewrite Z.add_move_0_l, <- Z.mul_opp_l. destruct (Z.lt_trichotomy (-z) 1) as [LT|[->|GT]]. - intro. revert LT. apply Z.le_ngt, (Z.le_succ_l 0). apply Z.mul_pos_cancel_r with y; Z.order. - Z.nzsimpl. Z.order. - rewrite (Z.mul_lt_mono_pos_r y), Z.mul_1_l in GT; Z.order. Qed. Lemma OMEGA5 x y z : x = 0 -> y = 0 -> x + y * z = 0. Proof. now intros -> ->. Qed. Lemma OMEGA6 x y z : 0 <= x -> y = 0 -> 0 <= x + y * z. Proof. intros H ->. now Z.nzsimpl. Qed. Lemma OMEGA7 x y z t : z > 0 -> t > 0 -> 0 <= x -> 0 <= y -> 0 <= x * z + y * t. Proof. intros. Z.swap_greater. Z.order_pos. Qed. Lemma OMEGA8 x y : 0 <= x -> 0 <= y -> x = - y -> x = 0. Proof. intros H1 H2 H3. rewrite <- Z.opp_nonpos_nonneg in H2. Z.order. Qed. Lemma OMEGA9 x y z t : y = 0 -> x = z -> y + (- x + z) * t = 0. Proof. intros. subst. now rewrite Z.add_opp_diag_l. Qed. Lemma OMEGA10 v c1 c2 l1 l2 k1 k2 : (v * c1 + l1) * k1 + (v * c2 + l2) * k2 = v * (c1 * k1 + c2 * k2) + (l1 * k1 + l2 * k2). Proof. rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc. rewrite <- !Z.add_assoc. f_equal. apply Z.add_shuffle3. Qed. Lemma OMEGA11 v1 c1 l1 l2 k1 : (v1 * c1 + l1) * k1 + l2 = v1 * (c1 * k1) + (l1 * k1 + l2). Proof. rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc. now rewrite Z.add_assoc. Qed. Lemma OMEGA12 v2 c2 l1 l2 k2 : l1 + (v2 * c2 + l2) * k2 = v2 * (c2 * k2) + (l1 + l2 * k2). Proof. rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc. apply Z.add_shuffle3. Qed. Lemma OMEGA13 (v l1 l2 : Z) (x : positive) : v * Zpos x + l1 + (v * Zneg x + l2) = l1 + l2. Proof. rewrite Z.add_shuffle1. rewrite <- Z.mul_add_distr_l, <- Pos2Z.opp_neg, Z.add_opp_diag_r. now Z.nzsimpl. Qed. Lemma OMEGA14 (v l1 l2 : Z) (x : positive) : v * Zneg x + l1 + (v * Zpos x + l2) = l1 + l2. Proof. rewrite Z.add_shuffle1. rewrite <- Z.mul_add_distr_l, <- Pos2Z.opp_neg, Z.add_opp_diag_r. now Z.nzsimpl. Qed. Lemma OMEGA15 v c1 c2 l1 l2 k2 : v * c1 + l1 + (v * c2 + l2) * k2 = v * (c1 + c2 * k2) + (l1 + l2 * k2). Proof. rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc. apply Z.add_shuffle1. Qed. Lemma OMEGA16 v c l k : (v * c + l) * k = v * (c * k) + l * k. Proof. now rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc. Qed. Lemma OMEGA17 x y z : Zne x 0 -> y = 0 -> Zne (x + y * z) 0. Proof. unfold Zne, not. intros NE EQ. subst. now Z.nzsimpl. Qed. Lemma OMEGA18 x y k : x = y * k -> Zne x 0 -> Zne y 0. Proof. unfold Zne, not. intros. subst; auto. Qed. Lemma OMEGA19 x : Zne x 0 -> 0 <= x + -1 \/ 0 <= x * -1 + -1. Proof. unfold Zne. intros Hx. apply Z.lt_gt_cases in Hx. destruct Hx as [LT|GT]. - right. change (-1) with (-(1)). rewrite Z.mul_opp_r, <- Z.opp_add_distr. Z.nzsimpl. rewrite Z.opp_nonneg_nonpos. now apply Z.le_succ_l. - left. now apply Z.lt_le_pred. Qed. Lemma OMEGA20 x y z : Zne x 0 -> y = 0 -> Zne (x + y * z) 0. Proof. unfold Zne, not. intros H1 H2 H3; apply H1; rewrite H2 in H3; simpl in H3; rewrite Z.add_0_r in H3; trivial with arith. Qed. Definition fast_Zplus_comm (x y : Z) (P : Z -> Prop) (H : P (y + x)) := eq_ind_r P H (Z.add_comm x y). Definition fast_Zplus_assoc_reverse (n m p : Z) (P : Z -> Prop) (H : P (n + (m + p))) := eq_ind_r P H (Zplus_assoc_reverse n m p). Definition fast_Zplus_assoc (n m p : Z) (P : Z -> Prop) (H : P (n + m + p)) := eq_ind_r P H (Z.add_assoc n m p). Definition fast_Zplus_permute (n m p : Z) (P : Z -> Prop) (H : P (m + (n + p))) := eq_ind_r P H (Z.add_shuffle3 n m p). Definition fast_OMEGA10 (v c1 c2 l1 l2 k1 k2 : Z) (P : Z -> Prop) (H : P (v * (c1 * k1 + c2 * k2) + (l1 * k1 + l2 * k2))) := eq_ind_r P H (OMEGA10 v c1 c2 l1 l2 k1 k2). Definition fast_OMEGA11 (v1 c1 l1 l2 k1 : Z) (P : Z -> Prop) (H : P (v1 * (c1 * k1) + (l1 * k1 + l2))) := eq_ind_r P H (OMEGA11 v1 c1 l1 l2 k1). Definition fast_OMEGA12 (v2 c2 l1 l2 k2 : Z) (P : Z -> Prop) (H : P (v2 * (c2 * k2) + (l1 + l2 * k2))) := eq_ind_r P H (OMEGA12 v2 c2 l1 l2 k2). Definition fast_OMEGA15 (v c1 c2 l1 l2 k2 : Z) (P : Z -> Prop) (H : P (v * (c1 + c2 * k2) + (l1 + l2 * k2))) := eq_ind_r P H (OMEGA15 v c1 c2 l1 l2 k2). Definition fast_OMEGA16 (v c l k : Z) (P : Z -> Prop) (H : P (v * (c * k) + l * k)) := eq_ind_r P H (OMEGA16 v c l k). Definition fast_OMEGA13 (v l1 l2 : Z) (x : positive) (P : Z -> Prop) (H : P (l1 + l2)) := eq_ind_r P H (OMEGA13 v l1 l2 x). Definition fast_OMEGA14 (v l1 l2 : Z) (x : positive) (P : Z -> Prop) (H : P (l1 + l2)) := eq_ind_r P H (OMEGA14 v l1 l2 x). Definition fast_Zred_factor0 (x : Z) (P : Z -> Prop) (H : P (x * 1)) := eq_ind_r P H (Zred_factor0 x). Definition fast_Zopp_eq_mult_neg_1 (x : Z) (P : Z -> Prop) (H : P (x * -1)) := eq_ind_r P H (Z.opp_eq_mul_m1 x). Definition fast_Zmult_comm (x y : Z) (P : Z -> Prop) (H : P (y * x)) := eq_ind_r P H (Z.mul_comm x y). Definition fast_Zopp_plus_distr (x y : Z) (P : Z -> Prop) (H : P (- x + - y)) := eq_ind_r P H (Z.opp_add_distr x y). Definition fast_Zopp_mult_distr_r (x y : Z) (P : Z -> Prop) (H : P (x * - y)) := eq_ind_r P H (Zopp_mult_distr_r x y). Definition fast_Zmult_plus_distr_l (n m p : Z) (P : Z -> Prop) (H : P (n * p + m * p)) := eq_ind_r P H (Z.mul_add_distr_r n m p). Definition fast_Zmult_assoc_reverse (n m p : Z) (P : Z -> Prop) (H : P (n * (m * p))) := eq_ind_r P H (Zmult_assoc_reverse n m p). Definition fast_Zred_factor1 (x : Z) (P : Z -> Prop) (H : P (x * 2)) := eq_ind_r P H (Zred_factor1 x). Definition fast_Zred_factor2 (x y : Z) (P : Z -> Prop) (H : P (x * (1 + y))) := eq_ind_r P H (Zred_factor2 x y). Definition fast_Zred_factor3 (x y : Z) (P : Z -> Prop) (H : P (x * (1 + y))) := eq_ind_r P H (Zred_factor3 x y). Definition fast_Zred_factor4 (x y z : Z) (P : Z -> Prop) (H : P (x * (y + z))) := eq_ind_r P H (Zred_factor4 x y z). Definition fast_Zred_factor5 (x y : Z) (P : Z -> Prop) (H : P y) := eq_ind_r P H (Zred_factor5 x y). Definition fast_Zred_factor6 (x : Z) (P : Z -> Prop) (H : P (x + 0)) := eq_ind_r P H (Zred_factor6 x). Theorem intro_Z : forall n:nat, exists y : Z, Z.of_nat n = y /\ 0 <= y * 1 + 0. Proof. intros n; exists (Z.of_nat n); split; trivial. rewrite Z.mul_1_r, Z.add_0_r. apply Nat2Z.is_nonneg. Qed. Register fast_Zplus_assoc_reverse as plugins.omega.fast_Zplus_assoc_reverse. Register fast_Zplus_assoc as plugins.omega.fast_Zplus_assoc. Register fast_Zmult_assoc_reverse as plugins.omega.fast_Zmult_assoc_reverse. Register fast_Zplus_permute as plugins.omega.fast_Zplus_permute. Register fast_Zplus_comm as plugins.omega.fast_Zplus_comm. Register fast_Zmult_comm as plugins.omega.fast_Zmult_comm. Register OMEGA1 as plugins.omega.OMEGA1. Register OMEGA2 as plugins.omega.OMEGA2. Register OMEGA3 as plugins.omega.OMEGA3. Register OMEGA4 as plugins.omega.OMEGA4. Register OMEGA5 as plugins.omega.OMEGA5. Register OMEGA6 as plugins.omega.OMEGA6. Register OMEGA7 as plugins.omega.OMEGA7. Register OMEGA8 as plugins.omega.OMEGA8. Register OMEGA9 as plugins.omega.OMEGA9. Register fast_OMEGA10 as plugins.omega.fast_OMEGA10. Register fast_OMEGA11 as plugins.omega.fast_OMEGA11. Register fast_OMEGA12 as plugins.omega.fast_OMEGA12. Register fast_OMEGA13 as plugins.omega.fast_OMEGA13. Register fast_OMEGA14 as plugins.omega.fast_OMEGA14. Register fast_OMEGA15 as plugins.omega.fast_OMEGA15. Register fast_OMEGA16 as plugins.omega.fast_OMEGA16. Register OMEGA17 as plugins.omega.OMEGA17. Register OMEGA18 as plugins.omega.OMEGA18. Register OMEGA19 as plugins.omega.OMEGA19. Register OMEGA20 as plugins.omega.OMEGA20. Register fast_Zred_factor0 as plugins.omega.fast_Zred_factor0. Register fast_Zred_factor1 as plugins.omega.fast_Zred_factor1. Register fast_Zred_factor2 as plugins.omega.fast_Zred_factor2. Register fast_Zred_factor3 as plugins.omega.fast_Zred_factor3. Register fast_Zred_factor4 as plugins.omega.fast_Zred_factor4. Register fast_Zred_factor5 as plugins.omega.fast_Zred_factor5. Register fast_Zred_factor6 as plugins.omega.fast_Zred_factor6. Register fast_Zmult_plus_distr_l as plugins.omega.fast_Zmult_plus_distr_l. Register fast_Zopp_plus_distr as plugins.omega.fast_Zopp_plus_distr. Register fast_Zopp_mult_distr_r as plugins.omega.fast_Zopp_mult_distr_r. Register fast_Zopp_eq_mult_neg_1 as plugins.omega.fast_Zopp_eq_mult_neg_1. Register new_var as plugins.omega.new_var. Register intro_Z as plugins.omega.intro_Z. coq-8.11.0/plugins/omega/Omega.v0000644000175000017500000000522213612664533016257 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* nat) => abstract omega: zarith. Hint Extern 10 (_ <= _) => abstract omega: zarith. Hint Extern 10 (_ < _) => abstract omega: zarith. Hint Extern 10 (_ >= _) => abstract omega: zarith. Hint Extern 10 (_ > _) => abstract omega: zarith. Hint Extern 10 (_ <> _ :>nat) => abstract omega: zarith. Hint Extern 10 (~ _ <= _) => abstract omega: zarith. Hint Extern 10 (~ _ < _) => abstract omega: zarith. Hint Extern 10 (~ _ >= _) => abstract omega: zarith. Hint Extern 10 (~ _ > _) => abstract omega: zarith. Hint Extern 10 (_ = _ :>Z) => abstract omega: zarith. Hint Extern 10 (_ <= _)%Z => abstract omega: zarith. Hint Extern 10 (_ < _)%Z => abstract omega: zarith. Hint Extern 10 (_ >= _)%Z => abstract omega: zarith. Hint Extern 10 (_ > _)%Z => abstract omega: zarith. Hint Extern 10 (_ <> _ :>Z) => abstract omega: zarith. Hint Extern 10 (~ (_ <= _)%Z) => abstract omega: zarith. Hint Extern 10 (~ (_ < _)%Z) => abstract omega: zarith. Hint Extern 10 (~ (_ >= _)%Z) => abstract omega: zarith. Hint Extern 10 (~ (_ > _)%Z) => abstract omega: zarith. Hint Extern 10 False => abstract omega: zarith. coq-8.11.0/plugins/omega/g_omega.mlg0000644000175000017500000000456113612664533017144 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* eval_tactic "zify_nat" | "positive" -> eval_tactic "zify_positive" | "N" -> eval_tactic "zify_N" | "Z" -> eval_tactic "zify_op" | s -> CErrors.user_err Pp.(str ("No Omega knowledge base for type "^s))) (Util.List.sort_uniquize String.compare l) in Tacticals.New.tclTHEN (Tacticals.New.tclREPEAT (Tacticals.New.tclTHENLIST tacs)) (omega_solver) let omega_with_deprecation = Deprecation.make ~since:"8.11.0" ~note:"Use lia instead." () } TACTIC EXTEND omega | [ "omega" ] -> { omega_tactic [] } END TACTIC EXTEND omega' DEPRECATED { omega_with_deprecation } | [ "omega" "with" ne_ident_list(l) ] -> { omega_tactic (List.map Names.Id.to_string l) } | [ "omega" "with" "*" ] -> { Tacticals.New.tclTHEN (eval_tactic "zify") (omega_tactic []) } END coq-8.11.0/plugins/omega/coq_omega.ml0000644000175000017500000022460013612664533017327 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let r = ref n in refs := (r,n) :: !refs; r), (fun () -> List.iter (fun (r,n) -> r:=n) !refs) let new_identifier = let cpt = intref 0 in (fun () -> let s = "Omega" ^ string_of_int !cpt in incr cpt; Id.of_string s) let new_identifier_var = let cpt = intref 0 in (fun () -> let s = "Zvar" ^ string_of_int !cpt in incr cpt; Id.of_string s) let new_id = let cpt = intref 0 in fun () -> incr cpt; !cpt let new_var_num = let cpt = intref 1000 in (fun () -> incr cpt; !cpt) let new_var = let cpt = intref 0 in fun () -> incr cpt; Nameops.make_ident "WW" (Some !cpt) let display_var i = Printf.sprintf "X%d" i let intern_id,unintern_id,reset_intern_tables = let cpt = ref 0 in let table = Hashtbl.create 7 and co_table = Hashtbl.create 7 in (fun (name : Id.t) -> try Hashtbl.find table name with Not_found -> let idx = !cpt in Hashtbl.add table name idx; Hashtbl.add co_table idx name; incr cpt; idx), (fun idx -> try Hashtbl.find co_table idx with Not_found -> let v = new_var () in Hashtbl.add table v idx; Hashtbl.add co_table idx v; v), (fun () -> cpt := 0; Hashtbl.clear table) let mk_then tacs = tclTHENLIST tacs let exists_tac c = constructor_tac false (Some 1) 1 (ImplicitBindings [c]) let generalize_tac t = generalize t let unfold s = Tactics.unfold_in_concl [Locus.AllOccurrences, Lazy.force s] let pf_nf gl c = pf_apply Tacred.simpl gl c let rev_assoc k = let rec loop = function | [] -> raise Not_found | (v,k')::_ when Int.equal k k' -> v | _ :: l -> loop l in loop let tag_hypothesis, hyp_of_tag, clear_tags = let l = ref ([]:(Id.t * int) list) in (fun h id -> l := (h,id):: !l), (fun h -> try rev_assoc h !l with Not_found -> failwith "tag_hypothesis"), (fun () -> l := []) let hide_constr,find_constr,clear_constr_tables,dump_tables = let l = ref ([]:(constr * (Id.t * Id.t * bool)) list) in (fun h id eg b -> l := (h,(id,eg,b)):: !l), (fun sigma h -> try List.assoc_f (eq_constr_nounivs sigma) h !l with Not_found -> failwith "find_contr"), (fun () -> l := []), (fun () -> !l) let reset_all () = if !reset_flag then begin reset_all_references (); reset_intern_tables (); clear_tags (); clear_constr_tables () end (* Lazy evaluation is used for Coq constants, because this code is evaluated before the compiled modules are loaded. To use the constant Zplus, one must type "Lazy.force coq_Zplus" This is the right way to access to Coq constants in tactics ML code *) let gen_constant k = lazy (k |> Coqlib.lib_ref |> UnivGen.constr_of_monomorphic_global |> EConstr.of_constr) (* Zarith *) let coq_xH = gen_constant "num.pos.xH" let coq_xO = gen_constant "num.pos.xO" let coq_xI = gen_constant "num.pos.xI" let coq_Z0 = gen_constant "num.Z.Z0" let coq_Zpos = gen_constant "num.Z.Zpos" let coq_Zneg = gen_constant "num.Z.Zneg" let coq_Z = gen_constant "num.Z.type" let coq_comparison = gen_constant "core.comparison.type" let coq_Gt = gen_constant "core.comparison.Gt" let coq_Zplus = gen_constant "num.Z.add" let coq_Zmult = gen_constant "num.Z.mul" let coq_Zopp = gen_constant "num.Z.opp" let coq_Zminus = gen_constant "num.Z.sub" let coq_Zsucc = gen_constant "num.Z.succ" let coq_Zpred = gen_constant "num.Z.pred" let coq_Z_of_nat = gen_constant "num.Z.of_nat" let coq_inj_plus = gen_constant "num.Nat2Z.inj_add" let coq_inj_mult = gen_constant "num.Nat2Z.inj_mul" let coq_inj_minus1 = gen_constant "num.Nat2Z.inj_sub" let coq_inj_minus2 = gen_constant "plugins.omega.inj_minus2" let coq_inj_S = gen_constant "num.Nat2Z.inj_succ" let coq_inj_eq = gen_constant "plugins.omega.inj_eq" let coq_inj_neq = gen_constant "plugins.omega.inj_neq" let coq_inj_le = gen_constant "plugins.omega.inj_le" let coq_inj_lt = gen_constant "plugins.omega.inj_lt" let coq_inj_ge = gen_constant "plugins.omega.inj_ge" let coq_inj_gt = gen_constant "plugins.omega.inj_gt" let coq_fast_Zplus_assoc_reverse = gen_constant "plugins.omega.fast_Zplus_assoc_reverse" let coq_fast_Zplus_assoc = gen_constant "plugins.omega.fast_Zplus_assoc" let coq_fast_Zmult_assoc_reverse = gen_constant "plugins.omega.fast_Zmult_assoc_reverse" let coq_fast_Zplus_permute = gen_constant "plugins.omega.fast_Zplus_permute" let coq_fast_Zplus_comm = gen_constant "plugins.omega.fast_Zplus_comm" let coq_fast_Zmult_comm = gen_constant "plugins.omega.fast_Zmult_comm" let coq_Zmult_le_approx = gen_constant "plugins.omega.Zmult_le_approx" let coq_OMEGA1 = gen_constant "plugins.omega.OMEGA1" let coq_OMEGA2 = gen_constant "plugins.omega.OMEGA2" let coq_OMEGA3 = gen_constant "plugins.omega.OMEGA3" let coq_OMEGA4 = gen_constant "plugins.omega.OMEGA4" let coq_OMEGA5 = gen_constant "plugins.omega.OMEGA5" let coq_OMEGA6 = gen_constant "plugins.omega.OMEGA6" let coq_OMEGA7 = gen_constant "plugins.omega.OMEGA7" let coq_OMEGA8 = gen_constant "plugins.omega.OMEGA8" let coq_OMEGA9 = gen_constant "plugins.omega.OMEGA9" let coq_fast_OMEGA10 = gen_constant "plugins.omega.fast_OMEGA10" let coq_fast_OMEGA11 = gen_constant "plugins.omega.fast_OMEGA11" let coq_fast_OMEGA12 = gen_constant "plugins.omega.fast_OMEGA12" let coq_fast_OMEGA13 = gen_constant "plugins.omega.fast_OMEGA13" let coq_fast_OMEGA14 = gen_constant "plugins.omega.fast_OMEGA14" let coq_fast_OMEGA15 = gen_constant "plugins.omega.fast_OMEGA15" let coq_fast_OMEGA16 = gen_constant "plugins.omega.fast_OMEGA16" let coq_OMEGA17 = gen_constant "plugins.omega.OMEGA17" let coq_OMEGA18 = gen_constant "plugins.omega.OMEGA18" let coq_OMEGA19 = gen_constant "plugins.omega.OMEGA19" let coq_OMEGA20 = gen_constant "plugins.omega.OMEGA20" let coq_fast_Zred_factor0 = gen_constant "plugins.omega.fast_Zred_factor0" let coq_fast_Zred_factor1 = gen_constant "plugins.omega.fast_Zred_factor1" let coq_fast_Zred_factor2 = gen_constant "plugins.omega.fast_Zred_factor2" let coq_fast_Zred_factor3 = gen_constant "plugins.omega.fast_Zred_factor3" let coq_fast_Zred_factor4 = gen_constant "plugins.omega.fast_Zred_factor4" let coq_fast_Zred_factor5 = gen_constant "plugins.omega.fast_Zred_factor5" let coq_fast_Zred_factor6 = gen_constant "plugins.omega.fast_Zred_factor6" let coq_fast_Zmult_plus_distr_l = gen_constant "plugins.omega.fast_Zmult_plus_distr_l" let coq_fast_Zopp_plus_distr = gen_constant "plugins.omega.fast_Zopp_plus_distr" let coq_fast_Zopp_mult_distr_r = gen_constant "plugins.omega.fast_Zopp_mult_distr_r" let coq_fast_Zopp_eq_mult_neg_1 = gen_constant "plugins.omega.fast_Zopp_eq_mult_neg_1" let coq_Zegal_left = gen_constant "plugins.omega.Zegal_left" let coq_Zne_left = gen_constant "plugins.omega.Zne_left" let coq_Zlt_left = gen_constant "plugins.omega.Zlt_left" let coq_Zge_left = gen_constant "plugins.omega.Zge_left" let coq_Zgt_left = gen_constant "plugins.omega.Zgt_left" let coq_Zle_left = gen_constant "plugins.omega.Zle_left" let coq_new_var = gen_constant "plugins.omega.new_var" let coq_intro_Z = gen_constant "plugins.omega.intro_Z" let coq_dec_eq = gen_constant "num.Z.eq_decidable" let coq_dec_Zne = gen_constant "plugins.omega.dec_Zne" let coq_dec_Zle = gen_constant "num.Z.le_decidable" let coq_dec_Zlt = gen_constant "num.Z.lt_decidable" let coq_dec_Zgt = gen_constant "plugins.omega.dec_Zgt" let coq_dec_Zge = gen_constant "plugins.omega.dec_Zge" let coq_not_Zeq = gen_constant "plugins.omega.not_Zeq" let coq_not_Zne = gen_constant "plugins.omega.not_Zne" let coq_Znot_le_gt = gen_constant "plugins.omega.Znot_le_gt" let coq_Znot_lt_ge = gen_constant "plugins.omega.Znot_lt_ge" let coq_Znot_ge_lt = gen_constant "plugins.omega.Znot_ge_lt" let coq_Znot_gt_le = gen_constant "plugins.omega.Znot_gt_le" let coq_neq = gen_constant "plugins.omega.neq" let coq_Zne = gen_constant "plugins.omega.Zne" let coq_Zle = gen_constant "num.Z.le" let coq_Zlt = gen_constant "num.Z.lt" let coq_Zge = gen_constant "num.Z.ge" let coq_Zgt = gen_constant "num.Z.gt" (* Peano/Datatypes *) let coq_nat = gen_constant "num.nat.type" let coq_O = gen_constant "num.nat.O" let coq_S = gen_constant "num.nat.S" let coq_le = gen_constant "num.nat.le" let coq_lt = gen_constant "num.nat.lt" let coq_ge = gen_constant "num.nat.ge" let coq_gt = gen_constant "num.nat.gt" let coq_plus = gen_constant "num.nat.add" let coq_minus = gen_constant "num.nat.sub" let coq_mult = gen_constant "num.nat.mul" let coq_pred = gen_constant "num.nat.pred" (* Compare_dec/Peano_dec/Minus *) let coq_pred_of_minus = gen_constant "num.nat.pred_of_minus" let coq_le_gt_dec = gen_constant "num.nat.le_gt_dec" let coq_dec_eq_nat = gen_constant "num.nat.eq_dec" let coq_dec_le = gen_constant "num.nat.dec_le" let coq_dec_lt = gen_constant "num.nat.dec_lt" let coq_dec_ge = gen_constant "num.nat.dec_ge" let coq_dec_gt = gen_constant "num.nat.dec_gt" let coq_not_eq = gen_constant "num.nat.not_eq" let coq_not_le = gen_constant "num.nat.not_le" let coq_not_lt = gen_constant "num.nat.not_lt" let coq_not_ge = gen_constant "num.nat.not_ge" let coq_not_gt = gen_constant "num.nat.not_gt" (* Logic/Decidable *) let coq_eq_ind_r = gen_constant "core.eq.ind_r" let coq_dec_or = gen_constant "core.dec.or" let coq_dec_and = gen_constant "core.dec.and" let coq_dec_imp = gen_constant "core.dec.imp" let coq_dec_iff = gen_constant "core.dec.iff" let coq_dec_not = gen_constant "core.dec.not" let coq_dec_False = gen_constant "core.dec.False" let coq_dec_not_not = gen_constant "core.dec.not_not" let coq_dec_True = gen_constant "core.dec.True" let coq_not_or = gen_constant "core.dec.not_or" let coq_not_and = gen_constant "core.dec.not_and" let coq_not_imp = gen_constant "core.dec.not_imp" let coq_not_iff = gen_constant "core.dec.not_iff" let coq_not_not = gen_constant "core.dec.dec_not_not" let coq_imp_simp = gen_constant "core.dec.imp_simp" let coq_iff = gen_constant "core.iff.type" let coq_not = gen_constant "core.not.type" let coq_and = gen_constant "core.and.type" let coq_or = gen_constant "core.or.type" let coq_eq = gen_constant "core.eq.type" let coq_ex = gen_constant "core.ex.type" let coq_False = gen_constant "core.False.type" let coq_True = gen_constant "core.True.type" (* uses build_coq_and, build_coq_not, build_coq_or, build_coq_ex *) (* For unfold *) let evaluable_ref_of_constr s c = let env = Global.env () in let evd = Evd.from_env env in match EConstr.kind evd (Lazy.force c) with | Const (kn,u) when Tacred.is_evaluable env (EvalConstRef kn) -> EvalConstRef kn | _ -> anomaly ~label:"Coq_omega" (Pp.str (s^" is not an evaluable constant.")) let sp_Zsucc = lazy (evaluable_ref_of_constr "Z.succ" coq_Zsucc) let sp_Zpred = lazy (evaluable_ref_of_constr "Z.pred" coq_Zpred) let sp_Zminus = lazy (evaluable_ref_of_constr "Z.sub" coq_Zminus) let sp_Zle = lazy (evaluable_ref_of_constr "Z.le" coq_Zle) let sp_Zgt = lazy (evaluable_ref_of_constr "Z.gt" coq_Zgt) let sp_not = lazy (evaluable_ref_of_constr "not" coq_not) let mk_plus t1 t2 = mkApp (Lazy.force coq_Zplus, [| t1; t2 |]) let mk_times t1 t2 = mkApp (Lazy.force coq_Zmult, [| t1; t2 |]) let mk_minus t1 t2 = mkApp (Lazy.force coq_Zminus, [| t1;t2 |]) let mk_gen_eq ty t1 t2 = mkApp (Lazy.force coq_eq, [| ty; t1; t2 |]) let mk_eq t1 t2 = mk_gen_eq (Lazy.force coq_Z) t1 t2 let mk_gt t1 t2 = mkApp (Lazy.force coq_Zgt, [| t1; t2 |]) let mk_inv t = mkApp (Lazy.force coq_Zopp, [| t |]) let mk_and t1 t2 = mkApp (Lazy.force coq_and, [| t1; t2 |]) let mk_or t1 t2 = mkApp (Lazy.force coq_or, [| t1; t2 |]) let mk_not t = mkApp (Lazy.force coq_not, [| t |]) let mk_inj t = mkApp (Lazy.force coq_Z_of_nat, [| t |]) let mk_integer n = let rec loop n = if n =? one then Lazy.force coq_xH else mkApp((if n mod two =? zero then Lazy.force coq_xO else Lazy.force coq_xI), [| loop (n/two) |]) in if n =? zero then Lazy.force coq_Z0 else mkApp ((if n >? zero then Lazy.force coq_Zpos else Lazy.force coq_Zneg), [| loop (abs n) |]) type omega_constant = | Zplus | Zmult | Zminus | Zsucc | Zopp | Zpred | Plus | Mult | Minus | Pred | S | O | Zpos | Zneg | Z0 | Z_of_nat | Eq | Neq | Zne | Zle | Zlt | Zge | Zgt | Z | Nat | And | Or | False | True | Not | Iff | Le | Lt | Ge | Gt | Other of string type result = | Kvar of Id.t | Kapp of omega_constant * constr list | Kimp of constr * constr | Kufo (* Nota: Kimp correspond to a binder (Prod), but hopefully we won't have to bother with term lifting: Kimp will correspond to anonymous product, for which (Rel 1) doesn't occur in the right term. Moreover, we'll work on fully introduced goals, hence no Rel's in the term parts that we manipulate, but rather Var's. Said otherwise: all constr manipulated here are closed *) let destructurate_prop sigma t = let eq_constr c1 c2 = eq_constr sigma c1 c2 in let c, args = decompose_app sigma t in match EConstr.kind sigma c, args with | _, [_;_;_] when eq_constr (Lazy.force coq_eq) c -> Kapp (Eq,args) | _, [_;_] when eq_constr c (Lazy.force coq_neq) -> Kapp (Neq,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zne) -> Kapp (Zne,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zle) -> Kapp (Zle,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zlt) -> Kapp (Zlt,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zge) -> Kapp (Zge,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zgt) -> Kapp (Zgt,args) | _, [_;_] when eq_constr c (Lazy.force coq_and) -> Kapp (And,args) | _, [_;_] when eq_constr c (Lazy.force coq_or) -> Kapp (Or,args) | _, [_;_] when eq_constr c (Lazy.force coq_iff) -> Kapp (Iff, args) | _, [_] when eq_constr c (Lazy.force coq_not) -> Kapp (Not,args) | _, [] when eq_constr c (Lazy.force coq_False) -> Kapp (False,args) | _, [] when eq_constr c (Lazy.force coq_True) -> Kapp (True,args) | _, [_;_] when eq_constr c (Lazy.force coq_le) -> Kapp (Le,args) | _, [_;_] when eq_constr c (Lazy.force coq_lt) -> Kapp (Lt,args) | _, [_;_] when eq_constr c (Lazy.force coq_ge) -> Kapp (Ge,args) | _, [_;_] when eq_constr c (Lazy.force coq_gt) -> Kapp (Gt,args) | Const (sp,_), args -> Kapp (Other (string_of_path (path_of_global (GlobRef.ConstRef sp))),args) | Construct (csp,_) , args -> Kapp (Other (string_of_path (path_of_global (GlobRef.ConstructRef csp))), args) | Ind (isp,_), args -> Kapp (Other (string_of_path (path_of_global (GlobRef.IndRef isp))),args) | Var id,[] -> Kvar id | Prod ({binder_name=Anonymous},typ,body), [] -> Kimp(typ,body) | Prod ({binder_name=Name _},_,_),[] -> CErrors.user_err Pp.(str "Omega: Not a quantifier-free goal") | _ -> Kufo let nf = Tacred.simpl let destructurate_type env sigma t = let is_conv = Reductionops.is_conv env sigma in let c, args = decompose_app sigma (nf env sigma t) in match EConstr.kind sigma c, args with | _, [] when is_conv c (Lazy.force coq_Z) -> Kapp (Z,args) | _, [] when is_conv c (Lazy.force coq_nat) -> Kapp (Nat,args) | _ -> Kufo let destructurate_term sigma t = let eq_constr c1 c2 = eq_constr sigma c1 c2 in let c, args = decompose_app sigma t in match EConstr.kind sigma c, args with | _, [_;_] when eq_constr c (Lazy.force coq_Zplus) -> Kapp (Zplus,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zmult) -> Kapp (Zmult,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zminus) -> Kapp (Zminus,args) | _, [_] when eq_constr c (Lazy.force coq_Zsucc) -> Kapp (Zsucc,args) | _, [_] when eq_constr c (Lazy.force coq_Zpred) -> Kapp (Zpred,args) | _, [_] when eq_constr c (Lazy.force coq_Zopp) -> Kapp (Zopp,args) | _, [_;_] when eq_constr c (Lazy.force coq_plus) -> Kapp (Plus,args) | _, [_;_] when eq_constr c (Lazy.force coq_mult) -> Kapp (Mult,args) | _, [_;_] when eq_constr c (Lazy.force coq_minus) -> Kapp (Minus,args) | _, [_] when eq_constr c (Lazy.force coq_pred) -> Kapp (Pred,args) | _, [_] when eq_constr c (Lazy.force coq_S) -> Kapp (S,args) | _, [] when eq_constr c (Lazy.force coq_O) -> Kapp (O,args) | _, [_] when eq_constr c (Lazy.force coq_Zpos) -> Kapp (Zneg,args) | _, [_] when eq_constr c (Lazy.force coq_Zneg) -> Kapp (Zpos,args) | _, [] when eq_constr c (Lazy.force coq_Z0) -> Kapp (Z0,args) | _, [_] when eq_constr c (Lazy.force coq_Z_of_nat) -> Kapp (Z_of_nat,args) | Var id,[] -> Kvar id | _ -> Kufo let recognize_number sigma t = let eq_constr c1 c2 = eq_constr sigma c1 c2 in let rec loop t = match decompose_app sigma t with | f, [t] when eq_constr f (Lazy.force coq_xI) -> one + two * loop t | f, [t] when eq_constr f (Lazy.force coq_xO) -> two * loop t | f, [] when eq_constr f (Lazy.force coq_xH) -> one | _ -> failwith "not a number" in match decompose_app sigma t with | f, [t] when eq_constr f (Lazy.force coq_Zpos) -> loop t | f, [t] when eq_constr f (Lazy.force coq_Zneg) -> neg (loop t) | f, [] when eq_constr f (Lazy.force coq_Z0) -> zero | _ -> failwith "not a number" type constr_path = | P_APP of int (* Abstraction and product *) | P_TYPE let context sigma operation path (t : constr) = let rec loop i p0 t = match (p0,EConstr.kind sigma t) with | (p, Cast (c,k,t)) -> mkCast (loop i p c,k,t) | ([], _) -> operation i t | ((P_APP n :: p), App (f,v)) -> let v' = Array.copy v in v'.(pred n) <- loop i p v'.(pred n); mkApp (f, v') | (p, Fix ((_,n as ln),(tys,lna,v))) -> let l = Array.length v in let v' = Array.copy v in v'.(n)<- loop (Util.(+) i l) p v.(n); (mkFix (ln,(tys,lna,v'))) | ((P_TYPE :: p), Prod (n,t,c)) -> (mkProd (n,loop i p t,c)) | ((P_TYPE :: p), Lambda (n,t,c)) -> (mkLambda (n,loop i p t,c)) | ((P_TYPE :: p), LetIn (n,b,t,c)) -> (mkLetIn (n,b,loop i p t,c)) | (p, _) -> failwith ("abstract_path " ^ string_of_int(List.length p)) in loop 1 path t let occurrence sigma path (t : constr) = let rec loop p0 t = match (p0,EConstr.kind sigma t) with | (p, Cast (c,_,_)) -> loop p c | ([], _) -> t | ((P_APP n :: p), App (f,v)) -> loop p v.(pred n) | (p, Fix((_,n) ,(_,_,v))) -> loop p v.(n) | ((P_TYPE :: p), Prod (n,term,c)) -> loop p term | ((P_TYPE :: p), Lambda (n,term,c)) -> loop p term | ((P_TYPE :: p), LetIn (n,b,term,c)) -> loop p term | (p, _) -> failwith ("occurrence " ^ string_of_int(List.length p)) in loop path t let abstract_path sigma typ path t = let term_occur = ref (mkRel 0) in let abstract = context sigma (fun i t -> term_occur:= t; mkRel i) path t in mkLambda (make_annot (Name (Id.of_string "x")) Sorts.Relevant, typ, abstract), !term_occur let focused_simpl path = let open Tacmach.New in Proofview.Goal.enter begin fun gl -> let newc = context (project gl) (fun i t -> pf_nf gl t) (List.rev path) (pf_concl gl) in convert_concl ~check:false newc DEFAULTcast end let focused_simpl path = focused_simpl path type oformula = | Oplus of oformula * oformula | Otimes of oformula * oformula | Oatom of Id.t | Oz of bigint | Oufo of constr let rec oprint = function | Oplus(t1,t2) -> print_string "("; oprint t1; print_string "+"; oprint t2; print_string ")" | Otimes (t1,t2) -> print_string "("; oprint t1; print_string "*"; oprint t2; print_string ")" | Oatom s -> print_string (Id.to_string s) | Oz i -> print_string (string_of_bigint i) | Oufo f -> print_string "?" let rec weight = function | Oatom c -> intern_id c | Oz _ -> -1 | Otimes(c,_) -> weight c | Oplus _ -> failwith "weight" | Oufo _ -> -1 let rec val_of = function | Oatom c -> mkVar c | Oz c -> mk_integer c | Otimes (t1,t2) -> mkApp (Lazy.force coq_Zmult, [| val_of t1; val_of t2 |]) | Oplus(t1,t2) -> mkApp (Lazy.force coq_Zplus, [| val_of t1; val_of t2 |]) | Oufo c -> c let compile name kind = let rec loop accu = function | Oplus(Otimes(Oatom v,Oz n),r) -> loop ({v=intern_id v; c=n} :: accu) r | Oz n -> let id = new_id () in tag_hypothesis name id; {kind = kind; body = List.rev accu; constant = n; id = id} | _ -> anomaly (Pp.str "compile_equation.") in loop [] let decompile af = let rec loop = function | ({v=v; c=n}::r) -> Oplus(Otimes(Oatom (unintern_id v),Oz n),loop r) | [] -> Oz af.constant in loop af.body (** Backward compat to emulate the old Refine: normalize the goal conclusion *) let new_hole env sigma c = let c = Reductionops.nf_betaiota env sigma c in Evarutil.new_evar env sigma c let clever_rewrite_base_poly typ p result theorem = let open Tacmach.New in Proofview.Goal.enter begin fun gl -> let full = pf_concl gl in let env = pf_env gl in let (abstracted,occ) = abstract_path (project gl) typ (List.rev p) full in Refine.refine ~typecheck:false begin fun sigma -> let t = applist (mkLambda (make_annot (Name (Id.of_string "P")) Sorts.Relevant, mkArrow typ Sorts.Relevant mkProp, mkLambda (make_annot (Name (Id.of_string "H")) Sorts.Relevant, applist (mkRel 1,[result]), mkApp (Lazy.force coq_eq_ind_r, [| typ; result; mkRel 2; mkRel 1; occ; theorem |]))), [abstracted]) in let argt = mkApp (abstracted, [|result|]) in let (sigma, hole) = new_hole env sigma argt in (sigma, applist (t, [hole])) end end let clever_rewrite_base p result theorem = clever_rewrite_base_poly (Lazy.force coq_Z) p result theorem let clever_rewrite_base_nat p result theorem = clever_rewrite_base_poly (Lazy.force coq_nat) p result theorem let clever_rewrite_gen p result (t,args) = let theorem = applist(t, args) in clever_rewrite_base p result theorem let clever_rewrite_gen_nat p result (t,args) = let theorem = applist(t, args) in clever_rewrite_base_nat p result theorem (** Solve using the term the term [t _] *) let refine_app gl t = let open Tacmach.New in Refine.refine ~typecheck:false begin fun sigma -> let env = pf_env gl in let ht = match EConstr.kind sigma (pf_get_type_of gl t) with | Prod (_, t, _) -> t | _ -> assert false in let (sigma, hole) = new_hole env sigma ht in (sigma, applist (t, [hole])) end let clever_rewrite p vpath t = let open Tacmach.New in Proofview.Goal.enter begin fun gl -> let full = pf_concl gl in let (abstracted,occ) = abstract_path (project gl) (Lazy.force coq_Z) (List.rev p) full in let vargs = List.map (fun p -> occurrence (project gl) p occ) vpath in let t' = applist(t, (vargs @ [abstracted])) in refine_app gl t' end (** simpl_coeffs : The subterm at location [path_init] in the current goal should look like [(v1*c1 + (v2*c2 + ... (vn*cn + k)))], and we reduce via "simpl" each [ci] and the final constant [k]. The path [path_k] gives the location of constant [k]. Earlier, the whole was a mere call to [focused_simpl], leading to reduction inside the atoms [vi], which is bad, for instance when the atom is an evaluable definition (see #4132). *) let simpl_coeffs path_init path_k = Proofview.Goal.enter begin fun gl -> let sigma = project gl in let rec loop n t = if Int.equal n 0 then pf_nf gl t else (* t should be of the form ((v * c) + ...) *) match EConstr.kind sigma t with | App(f,[|t1;t2|]) -> (match EConstr.kind sigma t1 with | App (g,[|v;c|]) -> let c' = pf_nf gl c in let t2' = loop (pred n) t2 in mkApp (f,[|mkApp (g,[|v;c'|]);t2'|]) | _ -> assert false) | _ -> assert false in let n = Util.(-) (List.length path_k) (List.length path_init) in let newc = context sigma (fun _ t -> loop n t) (List.rev path_init) (pf_concl gl) in convert_concl ~check:false newc DEFAULTcast end let rec shuffle p (t1,t2) = match t1,t2 with | Oplus(l1,r1), Oplus(l2,r2) -> if weight l1 > weight l2 then let (tac,t') = shuffle (P_APP 2 :: p) (r1,t2) in (clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: tac, Oplus(l1,t')) else let (tac,t') = shuffle (P_APP 2 :: p) (t1,r2) in (clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 1];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zplus_permute) :: tac, Oplus(l2,t')) | Oplus(l1,r1), t2 -> if weight l1 > weight t2 then let (tac,t') = shuffle (P_APP 2 :: p) (r1,t2) in clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: tac, Oplus(l1, t') else [clever_rewrite p [[P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zplus_comm)], Oplus(t2,t1) | t1,Oplus(l2,r2) -> if weight l2 > weight t1 then let (tac,t') = shuffle (P_APP 2 :: p) (t1,r2) in clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 1];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zplus_permute) :: tac, Oplus(l2,t') else [],Oplus(t1,t2) | Oz t1,Oz t2 -> [focused_simpl p], Oz(Bigint.add t1 t2) | t1,t2 -> if weight t1 < weight t2 then [clever_rewrite p [[P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zplus_comm)], Oplus(t2,t1) else [],Oplus(t1,t2) let shuffle_mult p_init k1 e1 k2 e2 = let rec loop p = function | (({c=c1;v=v1}::l1) as l1'),(({c=c2;v=v2}::l2) as l2') -> if Int.equal v1 v2 then let tac = clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 1; P_APP 2]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2]; [P_APP 2; P_APP 1; P_APP 2]; [P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA10) in if Bigint.add (Bigint.mult k1 c1) (Bigint.mult k2 c2) =? zero then let tac' = clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zred_factor5) in tac :: focused_simpl (P_APP 2::P_APP 1:: p) :: tac' :: loop p (l1,l2) else tac :: loop (P_APP 2 :: p) (l1,l2) else if v1 > v2 then clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2]; [P_APP 2]; [P_APP 1; P_APP 2]] (Lazy.force coq_fast_OMEGA11) :: loop (P_APP 2 :: p) (l1,l2') else clever_rewrite p [[P_APP 2; P_APP 1; P_APP 1; P_APP 1]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1]; [P_APP 2; P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA12) :: loop (P_APP 2 :: p) (l1',l2) | ({c=c1;v=v1}::l1), [] -> clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2]; [P_APP 2]; [P_APP 1; P_APP 2]] (Lazy.force coq_fast_OMEGA11) :: loop (P_APP 2 :: p) (l1,[]) | [],({c=c2;v=v2}::l2) -> clever_rewrite p [[P_APP 2; P_APP 1; P_APP 1; P_APP 1]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1]; [P_APP 2; P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA12) :: loop (P_APP 2 :: p) ([],l2) | [],[] -> [simpl_coeffs p_init p] in loop p_init (e1,e2) let shuffle_mult_right p_init e1 k2 e2 = let rec loop p = function | (({c=c1;v=v1}::l1) as l1'),(({c=c2;v=v2}::l2) as l2') -> if Int.equal v1 v2 then let tac = clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 2]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 2]; [P_APP 2; P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA15) in if Bigint.add c1 (Bigint.mult k2 c2) =? zero then let tac' = clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zred_factor5) in tac :: focused_simpl (P_APP 2::P_APP 1:: p) :: tac' :: loop p (l1,l2) else tac :: loop (P_APP 2 :: p) (l1,l2) else if v1 > v2 then clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: loop (P_APP 2 :: p) (l1,l2') else clever_rewrite p [[P_APP 2; P_APP 1; P_APP 1; P_APP 1]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1]; [P_APP 2; P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA12) :: loop (P_APP 2 :: p) (l1',l2) | ({c=c1;v=v1}::l1), [] -> clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: loop (P_APP 2 :: p) (l1,[]) | [],({c=c2;v=v2}::l2) -> clever_rewrite p [[P_APP 2; P_APP 1; P_APP 1; P_APP 1]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1]; [P_APP 2; P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA12) :: loop (P_APP 2 :: p) ([],l2) | [],[] -> [simpl_coeffs p_init p] in loop p_init (e1,e2) let rec shuffle_cancel p = function | [] -> [focused_simpl p] | ({c=c1}::l1) -> let tac = clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1];[P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2; P_APP 1]] (if c1 >? zero then (Lazy.force coq_fast_OMEGA13) else (Lazy.force coq_fast_OMEGA14)) in tac :: shuffle_cancel p l1 let rec scalar p n = function | Oplus(t1,t2) -> let tac1,t1' = scalar (P_APP 1 :: p) n t1 and tac2,t2' = scalar (P_APP 2 :: p) n t2 in clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 1;P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zmult_plus_distr_l) :: (tac1 @ tac2), Oplus(t1',t2') | Otimes(t1,Oz x) -> [clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 1;P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zmult_assoc_reverse); focused_simpl (P_APP 2 :: p)], Otimes(t1,Oz (n*x)) | Otimes(t1,t2) -> CErrors.user_err Pp.(str "Omega: Can't solve a goal with non-linear products") | (Oatom _ as t) -> [], Otimes(t,Oz n) | Oz i -> [focused_simpl p],Oz(n*i) | Oufo c -> [], Oufo (mkApp (Lazy.force coq_Zmult, [| mk_integer n; c |])) let scalar_norm p_init = let rec loop p = function | [] -> [simpl_coeffs p_init p] | (_::l) -> clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1];[P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_OMEGA16) :: loop (P_APP 2 :: p) l in loop p_init let norm_add p_init = let rec loop p = function | [] -> [simpl_coeffs p_init p] | _:: l -> clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: loop (P_APP 2 :: p) l in loop p_init let scalar_norm_add p_init = let rec loop p = function | [] -> [simpl_coeffs p_init p] | _ :: l -> clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2]; [P_APP 2]; [P_APP 1; P_APP 2]] (Lazy.force coq_fast_OMEGA11) :: loop (P_APP 2 :: p) l in loop p_init let rec negate p = function | Oplus(t1,t2) -> let tac1,t1' = negate (P_APP 1 :: p) t1 and tac2,t2' = negate (P_APP 2 :: p) t2 in clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 1;P_APP 2]] (Lazy.force coq_fast_Zopp_plus_distr) :: (tac1 @ tac2), Oplus(t1',t2') | Otimes(t1,Oz x) -> [clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 1;P_APP 2]] (Lazy.force coq_fast_Zopp_mult_distr_r); focused_simpl (P_APP 2 :: p)], Otimes(t1,Oz (neg x)) | Otimes(t1,t2) -> CErrors.user_err Pp.(str "Omega: Can't solve a goal with non-linear products") | (Oatom _ as t) -> let r = Otimes(t,Oz(negone)) in [clever_rewrite p [[P_APP 1]] (Lazy.force coq_fast_Zopp_eq_mult_neg_1)], r | Oz i -> [focused_simpl p],Oz(neg i) | Oufo c -> [], Oufo (mkApp (Lazy.force coq_Zopp, [| c |])) let rec transform sigma p t = let default isnat t' = try let v,th,_ = find_constr sigma t' in [clever_rewrite_base p (mkVar v) (mkVar th)], Oatom v with e when CErrors.noncritical e -> let v = new_identifier_var () and th = new_identifier () in hide_constr t' v th isnat; [clever_rewrite_base p (mkVar v) (mkVar th)], Oatom v in try match destructurate_term sigma t with | Kapp(Zplus,[t1;t2]) -> let tac1,t1' = transform sigma (P_APP 1 :: p) t1 and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in let tac,t' = shuffle p (t1',t2') in tac1 @ tac2 @ tac, t' | Kapp(Zminus,[t1;t2]) -> let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; (mkApp (Lazy.force coq_Zopp, [| t2 |])) |])) in unfold sp_Zminus :: tac,t | Kapp(Zsucc,[t1]) -> let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; mk_integer one |])) in unfold sp_Zsucc :: tac,t | Kapp(Zpred,[t1]) -> let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; mk_integer negone |])) in unfold sp_Zpred :: tac,t | Kapp(Zmult,[t1;t2]) -> let tac1,t1' = transform sigma (P_APP 1 :: p) t1 and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in begin match t1',t2' with | (_,Oz n) -> let tac,t' = scalar p n t1' in tac1 @ tac2 @ tac,t' | (Oz n,_) -> let sym = clever_rewrite p [[P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zmult_comm) in let tac,t' = scalar p n t2' in tac1 @ tac2 @ (sym :: tac),t' | _ -> default false t end | Kapp((Zpos|Zneg|Z0),_) -> (try ([],Oz(recognize_number sigma t)) with e when CErrors.noncritical e -> default false t) | Kvar s -> [],Oatom s | Kapp(Zopp,[t]) -> let tac,t' = transform sigma (P_APP 1 :: p) t in let tac',t'' = negate p t' in tac @ tac', t'' | Kapp(Z_of_nat,[t']) -> default true t' | _ -> default false t with e when catchable_exception e -> default false t let shrink_pair p f1 f2 = match f1,f2 with | Oatom v,Oatom _ -> let r = Otimes(Oatom v,Oz two) in clever_rewrite p [[P_APP 1]] (Lazy.force coq_fast_Zred_factor1), r | Oatom v, Otimes(_,c2) -> let r = Otimes(Oatom v,Oplus(c2,Oz one)) in clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zred_factor2), r | Otimes (v1,c1),Oatom v -> let r = Otimes(Oatom v,Oplus(c1,Oz one)) in clever_rewrite p [[P_APP 2];[P_APP 1;P_APP 2]] (Lazy.force coq_fast_Zred_factor3), r | Otimes (Oatom v,c1),Otimes (v2,c2) -> let r = Otimes(Oatom v,Oplus(c1,c2)) in clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 1;P_APP 2];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zred_factor4),r | t1,t2 -> begin oprint t1; print_newline (); oprint t2; print_newline (); flush stdout; CErrors.user_err Pp.(str "shrink.1") end let reduce_factor p = function | Oatom v -> let r = Otimes(Oatom v,Oz one) in [clever_rewrite p [[]] (Lazy.force coq_fast_Zred_factor0)],r | Otimes(Oatom v,Oz n) as f -> [],f | Otimes(Oatom v,c) -> let rec compute = function | Oz n -> n | Oplus(t1,t2) -> Bigint.add (compute t1) (compute t2) | _ -> CErrors.user_err Pp.(str "condense.1") in [focused_simpl (P_APP 2 :: p)], Otimes(Oatom v,Oz(compute c)) | t -> oprint t; CErrors.user_err Pp.(str "reduce_factor.1") let rec condense p = function | Oplus(f1,(Oplus(f2,r) as t)) -> if Int.equal (weight f1) (weight f2) then begin let shrink_tac,t = shrink_pair (P_APP 1 :: p) f1 f2 in let assoc_tac = clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 1];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zplus_assoc) in let tac_list,t' = condense p (Oplus(t,r)) in (assoc_tac :: shrink_tac :: tac_list), t' end else begin let tac,f = reduce_factor (P_APP 1 :: p) f1 in let tac',t' = condense (P_APP 2 :: p) t in (tac @ tac'), Oplus(f,t') end | Oplus(f1,Oz n) -> let tac,f1' = reduce_factor (P_APP 1 :: p) f1 in tac,Oplus(f1',Oz n) | Oplus(f1,f2) -> if Int.equal (weight f1) (weight f2) then begin let tac_shrink,t = shrink_pair p f1 f2 in let tac,t' = condense p t in tac_shrink :: tac,t' end else begin let tac,f = reduce_factor (P_APP 1 :: p) f1 in let tac',t' = condense (P_APP 2 :: p) f2 in (tac @ tac'),Oplus(f,t') end | Oz _ as t -> [],t | t -> let tac,t' = reduce_factor p t in let final = Oplus(t',Oz zero) in let tac' = clever_rewrite p [[]] (Lazy.force coq_fast_Zred_factor6) in tac @ [tac'], final let rec clear_zero p = function | Oplus(Otimes(Oatom v,Oz n),r) when n =? zero -> let tac = clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zred_factor5) in let tac',t = clear_zero p r in tac :: tac',t | Oplus(f,r) -> let tac,t = clear_zero (P_APP 2 :: p) r in tac,Oplus(f,t) | t -> [],t let replay_history tactic_normalisation = let aux = Id.of_string "auxiliary" in let aux1 = Id.of_string "auxiliary_1" in let aux2 = Id.of_string "auxiliary_2" in let izero = mk_integer zero in let rec loop t : unit Proofview.tactic = match t with | HYP e :: l -> begin try tclTHEN (Id.List.assoc (hyp_of_tag e.id) tactic_normalisation) (loop l) with Not_found -> loop l end | NEGATE_CONTRADICT (e2,e1,b) :: l -> let eq1 = decompile e1 and eq2 = decompile e2 in let id1 = hyp_of_tag e1.id and id2 = hyp_of_tag e2.id in let k = if b then negone else one in let p_initial = [P_APP 1;P_TYPE] in let tac= shuffle_mult_right p_initial e1.body k e2.body in tclTHENLIST [ generalize_tac [mkApp (Lazy.force coq_OMEGA17, [| val_of eq1; val_of eq2; mk_integer k; mkVar id1; mkVar id2 |])]; mk_then tac; (intros_using [aux]); resolve_id aux; reflexivity ] | CONTRADICTION (e1,e2) :: l -> let eq1 = decompile e1 and eq2 = decompile e2 in let p_initial = [P_APP 2;P_TYPE] in let tac = shuffle_cancel p_initial e1.body in let solve_le = let not_sup_sup = mkApp (Lazy.force coq_eq, [| Lazy.force coq_comparison; Lazy.force coq_Gt; Lazy.force coq_Gt |]) in tclTHENS (tclTHENLIST [ unfold sp_Zle; simpl_in_concl; intro; (absurd not_sup_sup) ]) [ assumption ; reflexivity ] in let theorem = mkApp (Lazy.force coq_OMEGA2, [| val_of eq1; val_of eq2; mkVar (hyp_of_tag e1.id); mkVar (hyp_of_tag e2.id) |]) in Proofview.tclTHEN (tclTHEN (generalize_tac [theorem]) (mk_then tac)) solve_le | DIVIDE_AND_APPROX (e1,e2,k,d) :: l -> let id = hyp_of_tag e1.id in let eq1 = val_of(decompile e1) and eq2 = val_of(decompile e2) in let kk = mk_integer k and dd = mk_integer d in let rhs = mk_plus (mk_times eq2 kk) dd in let state_eg = mk_eq eq1 rhs in let tac = scalar_norm_add [P_APP 3] e2.body in tclTHENS (cut state_eg) [ tclTHENS (tclTHENLIST [ (intros_using [aux]); (generalize_tac [mkApp (Lazy.force coq_OMEGA1, [| eq1; rhs; mkVar aux; mkVar id |])]); (clear [aux;id]); (intros_using [id]); (cut (mk_gt kk dd)) ]) [ tclTHENS (cut (mk_gt kk izero)) [ tclTHENLIST [ (intros_using [aux1; aux2]); (generalize_tac [mkApp (Lazy.force coq_Zmult_le_approx, [| kk;eq2;dd;mkVar aux1;mkVar aux2; mkVar id |])]); (clear [aux1;aux2;id]); (intros_using [id]); (loop l) ]; tclTHENLIST [ (unfold sp_Zgt); simpl_in_concl; reflexivity ] ]; tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; tclTHEN (mk_then tac) reflexivity ] | NOT_EXACT_DIVIDE (e1,k) :: l -> let c = floor_div e1.constant k in let d = Bigint.sub e1.constant (Bigint.mult c k) in let e2 = {id=e1.id; kind=EQUA;constant = c; body = map_eq_linear (fun c -> c / k) e1.body } in let eq2 = val_of(decompile e2) in let kk = mk_integer k and dd = mk_integer d in let tac = scalar_norm_add [P_APP 2] e2.body in tclTHENS (cut (mk_gt dd izero)) [ tclTHENS (cut (mk_gt kk dd)) [tclTHENLIST [ (intros_using [aux2;aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA4, [| dd;kk;eq2;mkVar aux1; mkVar aux2 |])]); (clear [aux1;aux2]); unfold sp_not; (intros_using [aux]); resolve_id aux; mk_then tac; assumption ] ; tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ] | EXACT_DIVIDE (e1,k) :: l -> let id = hyp_of_tag e1.id in let e2 = map_eq_afine (fun c -> c / k) e1 in let eq1 = val_of(decompile e1) and eq2 = val_of(decompile e2) in let kk = mk_integer k in let state_eq = mk_eq eq1 (mk_times eq2 kk) in if e1.kind == DISE then let tac = scalar_norm [P_APP 3] e2.body in tclTHENS (cut state_eq) [tclTHENLIST [ (intros_using [aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA18, [| eq1;eq2;kk;mkVar aux1; mkVar id |])]); (clear [aux1;id]); (intros_using [id]); (loop l) ]; tclTHEN (mk_then tac) reflexivity ] else let tac = scalar_norm [P_APP 3] e2.body in tclTHENS (cut state_eq) [ tclTHENS (cut (mk_gt kk izero)) [tclTHENLIST [ (intros_using [aux2;aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA3, [| eq1; eq2; kk; mkVar aux2; mkVar aux1;mkVar id|])]); (clear [aux1;aux2;id]); (intros_using [id]); (loop l) ]; tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; tclTHEN (mk_then tac) reflexivity ] | (MERGE_EQ(e3,e1,e2)) :: l -> let id = new_identifier () in tag_hypothesis id e3; let id1 = hyp_of_tag e1.id and id2 = hyp_of_tag e2 in let eq1 = val_of(decompile e1) and eq2 = val_of (decompile (negate_eq e1)) in let tac = clever_rewrite [P_APP 3] [[P_APP 1]] (Lazy.force coq_fast_Zopp_eq_mult_neg_1) :: scalar_norm [P_APP 3] e1.body in tclTHENS (cut (mk_eq eq1 (mk_inv eq2))) [tclTHENLIST [ (intros_using [aux]); (generalize_tac [mkApp (Lazy.force coq_OMEGA8, [| eq1;eq2;mkVar id1;mkVar id2; mkVar aux|])]); (clear [id1;id2;aux]); (intros_using [id]); (loop l) ]; tclTHEN (mk_then tac) reflexivity] | STATE {st_new_eq=e;st_def=def;st_orig=orig;st_coef=m;st_var=v} :: l -> let id = new_identifier () and id2 = hyp_of_tag orig.id in tag_hypothesis id e.id; let eq1 = val_of(decompile def) and eq2 = val_of(decompile orig) in let vid = unintern_id v in let theorem = mkApp (Lazy.force coq_ex, [| Lazy.force coq_Z; mkLambda (make_annot (Name vid) Sorts.Relevant, Lazy.force coq_Z, mk_eq (mkRel 1) eq1) |]) in let mm = mk_integer m in let p_initial = [P_APP 2;P_TYPE] in let tac = clever_rewrite (P_APP 1 :: P_APP 1 :: P_APP 2 :: p_initial) [[P_APP 1]] (Lazy.force coq_fast_Zopp_eq_mult_neg_1) :: shuffle_mult_right p_initial orig.body m ({c= negone;v= v}::def.body) in tclTHENS (cut theorem) [tclTHENLIST [ (intros_using [aux]); (elim_id aux); (clear [aux]); (intros_using [vid; aux]); (generalize_tac [mkApp (Lazy.force coq_OMEGA9, [| mkVar vid;eq2;eq1;mm; mkVar id2;mkVar aux |])]); mk_then tac; (clear [aux]); (intros_using [id]); (loop l) ]; tclTHEN (exists_tac eq1) reflexivity ] | SPLIT_INEQ(e,(e1,act1),(e2,act2)) :: l -> let id1 = new_identifier () and id2 = new_identifier () in tag_hypothesis id1 e1; tag_hypothesis id2 e2; let id = hyp_of_tag e.id in let tac1 = norm_add [P_APP 2;P_TYPE] e.body in let tac2 = scalar_norm_add [P_APP 2;P_TYPE] e.body in let eq = val_of(decompile e) in tclTHENS (simplest_elim (applist (Lazy.force coq_OMEGA19, [eq; mkVar id]))) [tclTHENLIST [ mk_then tac1; (intros_using [id1]); (loop act1) ]; tclTHENLIST [ mk_then tac2; (intros_using [id2]); (loop act2) ]] | SUM(e3,(k1,e1),(k2,e2)) :: l -> let id = new_identifier () in tag_hypothesis id e3; let id1 = hyp_of_tag e1.id and id2 = hyp_of_tag e2.id in let eq1 = val_of(decompile e1) and eq2 = val_of(decompile e2) in if k1 =? one && e2.kind == EQUA then let tac_thm = match e1.kind with | EQUA -> Lazy.force coq_OMEGA5 | INEQ -> Lazy.force coq_OMEGA6 | DISE -> Lazy.force coq_OMEGA20 in let kk = mk_integer k2 in let p_initial = if e1.kind == DISE then [P_APP 1; P_TYPE] else [P_APP 2; P_TYPE] in let tac = shuffle_mult_right p_initial e1.body k2 e2.body in tclTHENLIST [ (generalize_tac [mkApp (tac_thm, [| eq1; eq2; kk; mkVar id1; mkVar id2 |])]); mk_then tac; (intros_using [id]); (loop l) ] else let kk1 = mk_integer k1 and kk2 = mk_integer k2 in let p_initial = [P_APP 2;P_TYPE] in let tac= shuffle_mult p_initial k1 e1.body k2 e2.body in tclTHENS (cut (mk_gt kk1 izero)) [tclTHENS (cut (mk_gt kk2 izero)) [tclTHENLIST [ (intros_using [aux2;aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA7, [| eq1;eq2;kk1;kk2; mkVar aux1;mkVar aux2; mkVar id1;mkVar id2 |])]); (clear [aux1;aux2]); mk_then tac; (intros_using [id]); (loop l) ]; tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ] | CONSTANT_NOT_NUL(e,k) :: l -> tclTHEN ((generalize_tac [mkVar (hyp_of_tag e)])) Equality.discrConcl | CONSTANT_NUL(e) :: l -> tclTHEN (resolve_id (hyp_of_tag e)) reflexivity | CONSTANT_NEG(e,k) :: l -> tclTHENLIST [ (generalize_tac [mkVar (hyp_of_tag e)]); unfold sp_Zle; simpl_in_concl; unfold sp_not; (intros_using [aux]); resolve_id aux; reflexivity ] | _ -> Proofview.tclUNIT () in loop let normalize sigma p_initial t = let (tac,t') = transform sigma p_initial t in let (tac',t'') = condense p_initial t' in let (tac'',t''') = clear_zero p_initial t'' in tac @ tac' @ tac'' , t''' let normalize_equation sigma id flag theorem pos t t1 t2 (tactic,defs) = let p_initial = [P_APP pos ;P_TYPE] in let (tac,t') = normalize sigma p_initial t in let shift_left = tclTHEN (generalize_tac [mkApp (theorem, [| t1; t2; mkVar id |]) ]) (tclTRY (clear [id])) in if not (List.is_empty tac) then let id' = new_identifier () in ((id',(tclTHENLIST [ shift_left; mk_then tac; (intros_using [id']) ])) :: tactic, compile id' flag t' :: defs) else (tactic,defs) let destructure_omega env sigma tac_def (id,c) = if String.equal (atompart_of_id id) "State" then tac_def else try match destructurate_prop sigma c with | Kapp(Eq,[typ;t1;t2]) when begin match destructurate_type env sigma typ with Kapp(Z,[]) -> true | _ -> false end -> let t = mk_plus t1 (mk_inv t2) in normalize_equation sigma id EQUA (Lazy.force coq_Zegal_left) 2 t t1 t2 tac_def | Kapp(Zne,[t1;t2]) -> let t = mk_plus t1 (mk_inv t2) in normalize_equation sigma id DISE (Lazy.force coq_Zne_left) 1 t t1 t2 tac_def | Kapp(Zle,[t1;t2]) -> let t = mk_plus t2 (mk_inv t1) in normalize_equation sigma id INEQ (Lazy.force coq_Zle_left) 2 t t1 t2 tac_def | Kapp(Zlt,[t1;t2]) -> let t = mk_plus (mk_plus t2 (mk_integer negone)) (mk_inv t1) in normalize_equation sigma id INEQ (Lazy.force coq_Zlt_left) 2 t t1 t2 tac_def | Kapp(Zge,[t1;t2]) -> let t = mk_plus t1 (mk_inv t2) in normalize_equation sigma id INEQ (Lazy.force coq_Zge_left) 2 t t1 t2 tac_def | Kapp(Zgt,[t1;t2]) -> let t = mk_plus (mk_plus t1 (mk_integer negone)) (mk_inv t2) in normalize_equation sigma id INEQ (Lazy.force coq_Zgt_left) 2 t t1 t2 tac_def | _ -> tac_def with e when catchable_exception e -> tac_def let reintroduce id = (* [id] cannot be cleared if dependent: protect it by a try *) tclTHEN (tclTRY (clear [id])) (intro_using id) open Proofview.Notations let coq_omega = Proofview.Goal.enter begin fun gl -> clear_constr_tables (); let hyps_types = Tacmach.New.pf_hyps_types gl in let destructure_omega = Tacmach.New.pf_apply destructure_omega gl in let tactic_normalisation, system = List.fold_left destructure_omega ([],[]) hyps_types in let prelude,sys = List.fold_left (fun (tac,sys) (t,(v,th,b)) -> if b then let id = new_identifier () in let i = new_id () in tag_hypothesis id i; (tclTHENLIST [ (simplest_elim (applist (Lazy.force coq_intro_Z, [t]))); (intros_using [v; id]); (elim_id id); (clear [id]); (intros_using [th;id]); tac ]), {kind = INEQ; body = [{v=intern_id v; c=one}]; constant = zero; id = i} :: sys else (tclTHENLIST [ (simplest_elim (applist (Lazy.force coq_new_var, [t]))); (intros_using [v;th]); tac ]), sys) (Proofview.tclUNIT (),[]) (dump_tables ()) in let system = system @ sys in if !display_system_flag then display_system display_var system; if !old_style_flag then begin try let _ = simplify (new_id,new_var_num,display_var) false system in Proofview.tclUNIT () with UNSOLVABLE -> let _,path = depend [] [] (history ()) in if !display_action_flag then display_action display_var path; (tclTHEN prelude (replay_history tactic_normalisation path)) end else begin try let path = simplify_strong (new_id,new_var_num,display_var) system in if !display_action_flag then display_action display_var path; tclTHEN prelude (replay_history tactic_normalisation path) with NO_CONTRADICTION -> tclZEROMSG (Pp.str"Omega can't solve this system") end end let coq_omega = coq_omega let nat_inject = Proofview.Goal.enter begin fun gl -> let is_conv = Tacmach.New.pf_apply Reductionops.is_conv gl in let rec explore p t : unit Proofview.tactic = Proofview.tclEVARMAP >>= fun sigma -> try match destructurate_term sigma t with | Kapp(Plus,[t1;t2]) -> tclTHENLIST [ (clever_rewrite_gen p (mk_plus (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_plus),[t1;t2])); (explore (P_APP 1 :: p) t1); (explore (P_APP 2 :: p) t2) ] | Kapp(Mult,[t1;t2]) -> tclTHENLIST [ (clever_rewrite_gen p (mk_times (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_mult),[t1;t2])); (explore (P_APP 1 :: p) t1); (explore (P_APP 2 :: p) t2) ] | Kapp(Minus,[t1;t2]) -> let id = new_identifier () in tclTHENS (tclTHEN (simplest_elim (applist (Lazy.force coq_le_gt_dec, [t2;t1]))) (intros_using [id])) [ tclTHENLIST [ (clever_rewrite_gen p (mk_minus (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_minus1),[t1;t2;mkVar id])); (loop [id,mkApp (Lazy.force coq_le, [| t2;t1 |])]); (explore (P_APP 1 :: p) t1); (explore (P_APP 2 :: p) t2) ]; (tclTHEN (clever_rewrite_gen p (mk_integer zero) ((Lazy.force coq_inj_minus2),[t1;t2;mkVar id])) (loop [id,mkApp (Lazy.force coq_gt, [| t2;t1 |])])) ] | Kapp(S,[t']) -> let rec is_number t = try match destructurate_term sigma t with Kapp(S,[t]) -> is_number t | Kapp(O,[]) -> true | _ -> false with e when catchable_exception e -> false in let rec loop p t : unit Proofview.tactic = try match destructurate_term sigma t with Kapp(S,[t]) -> (tclTHEN (clever_rewrite_gen p (mkApp (Lazy.force coq_Zsucc, [| mk_inj t |])) ((Lazy.force coq_inj_S),[t])) (loop (P_APP 1 :: p) t)) | _ -> explore p t with e when catchable_exception e -> explore p t in if is_number t' then focused_simpl p else loop p t | Kapp(Pred,[t]) -> let t_minus_one = mkApp (Lazy.force coq_minus, [| t; mkApp (Lazy.force coq_S, [| Lazy.force coq_O |]) |]) in tclTHEN (clever_rewrite_gen_nat (P_APP 1 :: p) t_minus_one ((Lazy.force coq_pred_of_minus),[t])) (explore p t_minus_one) | Kapp(O,[]) -> focused_simpl p | _ -> Proofview.tclUNIT () with e when catchable_exception e -> Proofview.tclUNIT () and loop = function | [] -> Proofview.tclUNIT () | (i,t)::lit -> Proofview.tclEVARMAP >>= fun sigma -> begin try match destructurate_prop sigma t with Kapp(Le,[t1;t2]) -> tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_le, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); (explore [P_APP 2; P_TYPE] t2); (reintroduce i); (loop lit) ] | Kapp(Lt,[t1;t2]) -> tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_lt, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); (explore [P_APP 2; P_TYPE] t2); (reintroduce i); (loop lit) ] | Kapp(Ge,[t1;t2]) -> tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_ge, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); (explore [P_APP 2; P_TYPE] t2); (reintroduce i); (loop lit) ] | Kapp(Gt,[t1;t2]) -> tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_gt, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); (explore [P_APP 2; P_TYPE] t2); (reintroduce i); (loop lit) ] | Kapp(Neq,[t1;t2]) -> tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_neq, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); (explore [P_APP 2; P_TYPE] t2); (reintroduce i); (loop lit) ] | Kapp(Eq,[typ;t1;t2]) -> if is_conv typ (Lazy.force coq_nat) then tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_eq, [| t1;t2;mkVar i |]) ]); (explore [P_APP 2; P_TYPE] t1); (explore [P_APP 3; P_TYPE] t2); (reintroduce i); (loop lit) ] else loop lit | _ -> loop lit with e when catchable_exception e -> loop lit end in let hyps_types = Tacmach.New.pf_hyps_types gl in loop (List.rev hyps_types) end let dec_binop = function | Zne -> coq_dec_Zne | Zle -> coq_dec_Zle | Zlt -> coq_dec_Zlt | Zge -> coq_dec_Zge | Zgt -> coq_dec_Zgt | Le -> coq_dec_le | Lt -> coq_dec_lt | Ge -> coq_dec_ge | Gt -> coq_dec_gt | _ -> raise Not_found let not_binop = function | Zne -> coq_not_Zne | Zle -> coq_Znot_le_gt | Zlt -> coq_Znot_lt_ge | Zge -> coq_Znot_ge_lt | Zgt -> coq_Znot_gt_le | Le -> coq_not_le | Lt -> coq_not_lt | Ge -> coq_not_ge | Gt -> coq_not_gt | _ -> raise Not_found (** A decidability check : for some [t], could we build a term of type [decidable t] (i.e. [t\/~t]) ? Otherwise, we raise [Undecidable]. Note that a successful check implies that [t] has type Prop. *) exception Undecidable let rec decidability env sigma t = match destructurate_prop sigma t with | Kapp(Or,[t1;t2]) -> mkApp (Lazy.force coq_dec_or, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kapp(And,[t1;t2]) -> mkApp (Lazy.force coq_dec_and, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kapp(Iff,[t1;t2]) -> mkApp (Lazy.force coq_dec_iff, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kimp(t1,t2) -> (* This is the only situation where it's not obvious that [t] is in Prop. The recursive call on [t2] will ensure that. *) mkApp (Lazy.force coq_dec_imp, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kapp(Not,[t1]) -> mkApp (Lazy.force coq_dec_not, [| t1; decidability env sigma t1 |]) | Kapp(Eq,[typ;t1;t2]) -> begin match destructurate_type env sigma typ with | Kapp(Z,[]) -> mkApp (Lazy.force coq_dec_eq, [| t1;t2 |]) | Kapp(Nat,[]) -> mkApp (Lazy.force coq_dec_eq_nat, [| t1;t2 |]) | _ -> raise Undecidable end | Kapp(op,[t1;t2]) -> (try mkApp (Lazy.force (dec_binop op), [| t1; t2 |]) with Not_found -> raise Undecidable) | Kapp(False,[]) -> Lazy.force coq_dec_False | Kapp(True,[]) -> Lazy.force coq_dec_True | _ -> raise Undecidable let fresh_id avoid id gl = fresh_id_in_env avoid id (Proofview.Goal.env gl) let onClearedName id tac = (* We cannot ensure that hyps can be cleared (because of dependencies), *) (* so renaming may be necessary *) tclTHEN (tclTRY (clear [id])) (Proofview.Goal.enter begin fun gl -> let id = fresh_id Id.Set.empty id gl in tclTHEN (introduction id) (tac id) end) let onClearedName2 id tac = tclTHEN (tclTRY (clear [id])) (Proofview.Goal.enter begin fun gl -> let id1 = fresh_id Id.Set.empty (add_suffix id "_left") gl in let id2 = fresh_id Id.Set.empty (add_suffix id "_right") gl in tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ] end) let destructure_hyps = Proofview.Goal.enter begin fun gl -> let type_of = Tacmach.New.pf_unsafe_type_of gl in let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let decidability = decidability env sigma in let rec loop = function | [] -> (tclTHEN nat_inject coq_omega) | LocalDef (i,body,typ) :: lit when !letin_flag -> Proofview.tclEVARMAP >>= fun sigma -> begin try match destructurate_type env sigma typ with | Kapp(Nat,_) | Kapp(Z,_) -> let hid = fresh_id Id.Set.empty (add_suffix i.binder_name "_eqn") gl in let hty = mk_gen_eq typ (mkVar i.binder_name) body in tclTHEN (assert_by (Name hid) hty reflexivity) (loop (LocalAssum (make_annot hid Sorts.Relevant, hty) :: lit)) | _ -> loop lit with e when catchable_exception e -> loop lit end | decl :: lit -> (* variable without body (or !letin_flag isn't set) *) let i = NamedDecl.get_id decl in Proofview.tclEVARMAP >>= fun sigma -> begin try match destructurate_prop sigma (NamedDecl.get_type decl) with | Kapp(False,[]) -> elim_id i | Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit | Kapp(Or,[t1;t2]) -> (tclTHENS (elim_id i) [ onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t1)::lit))); onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t2)::lit))) ]) | Kapp(And,[t1;t2]) -> tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> loop (LocalAssum (make_annot i1 Sorts.Relevant,t1) :: LocalAssum (make_annot i2 Sorts.Relevant,t2) :: lit))) | Kapp(Iff,[t1;t2]) -> tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> loop (LocalAssum (make_annot i1 Sorts.Relevant,mkArrow t1 Sorts.Relevant t2) :: LocalAssum (make_annot i2 Sorts.Relevant,mkArrow t2 Sorts.Relevant t1) :: lit))) | Kimp(t1,t2) -> (* t1 and t2 might be in Type rather than Prop. For t1, the decidability check will ensure being Prop. *) if Termops.is_Prop sigma (type_of t2) then let d1 = decidability t1 in tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_imp_simp, [| t1; t2; d1; mkVar i|])]); (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_not t1) t2) :: lit)))) ] else loop lit | Kapp(Not,[t]) -> begin match destructurate_prop sigma t with Kapp(Or,[t1;t2]) -> tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]); (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_and (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(And,[t1;t2]) -> let d1 = decidability t1 in tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_and, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(Iff,[t1;t2]) -> let d1 = decidability t1 in let d2 = decidability t2 in tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_iff, [| t1; t2; d1; d2; mkVar i |])]); (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_and t1 (mk_not t2)) (mk_and (mk_not t1) t2)) :: lit)))) ] | Kimp(t1,t2) -> (* t2 must be in Prop otherwise ~(t1->t2) wouldn't be ok. For t1, being decidable implies being Prop. *) let d1 = decidability t1 in tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_imp, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_and t1 (mk_not t2)) :: lit)))) ] | Kapp(Not,[t]) -> let d = decidability t in tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]); (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t) :: lit)))) ] | Kapp(op,[t1;t2]) -> (try let thm = not_binop op in tclTHENLIST [ (generalize_tac [mkApp (Lazy.force thm, [| t1;t2;mkVar i|])]); (onClearedName i (fun _ -> loop lit)) ] with Not_found -> loop lit) | Kapp(Eq,[typ;t1;t2]) -> if !old_style_flag then begin match destructurate_type env sigma typ with | Kapp(Nat,_) -> tclTHENLIST [ (simplest_elim (mkApp (Lazy.force coq_not_eq, [|t1;t2;mkVar i|]))); (onClearedName i (fun _ -> loop lit)) ] | Kapp(Z,_) -> tclTHENLIST [ (simplest_elim (mkApp (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|]))); (onClearedName i (fun _ -> loop lit)) ] | _ -> loop lit end else begin match destructurate_type env sigma typ with | Kapp(Nat,_) -> (tclTHEN (Tactics.convert_hyp ~check:false ~reorder:false (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) decl)) (loop lit)) | Kapp(Z,_) -> (tclTHEN (Tactics.convert_hyp ~check:false ~reorder:false (NamedDecl.set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|])) decl)) (loop lit)) | _ -> loop lit end | _ -> loop lit end | _ -> loop lit with | Undecidable -> loop lit | e when catchable_exception e -> loop lit end in let hyps = Proofview.Goal.hyps gl in loop hyps end let destructure_goal = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let decidability = decidability env sigma in let rec loop t = Proofview.tclEVARMAP >>= fun sigma -> let prop () = Proofview.tclUNIT (destructurate_prop sigma t) in Proofview.V82.wrap_exceptions prop >>= fun prop -> match prop with | Kapp(Not,[t]) -> (tclTHEN (tclTHEN (unfold sp_not) intro) destructure_hyps) | Kimp(a,b) -> (tclTHEN intro (loop b)) | Kapp(False,[]) -> destructure_hyps | _ -> let goal_tac = try let dec = decidability t in tclTHEN (Proofview.Goal.enter begin fun gl -> refine_app gl (mkApp (Lazy.force coq_dec_not_not, [| t; dec |])) end) intro with Undecidable -> Tactics.elim_type (Lazy.force coq_False) | e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e in tclTHEN goal_tac destructure_hyps in (loop concl) end let destructure_goal = destructure_goal let omega_solver = Proofview.tclUNIT () >>= fun () -> (* delay for [check_required_library] *) Coqlib.check_required_library ["Coq";"omega";"Omega"]; reset_all (); destructure_goal coq-8.11.0/plugins/omega/OmegaPlugin.v0000644000175000017500000000145613612664533017443 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr -> constr -> (int*constr) list type instance= Real of (int*constr)*int (* nb trous*terme*valeur heuristique *) | Phantom of constr (* domaine de quantification *) val unif_atoms : Evd.evar_map -> metavariable -> constr -> constr -> constr -> instance option val more_general : Evd.evar_map -> (int*constr) -> (int*constr) -> bool coq-8.11.0/plugins/firstorder/rules.ml0000644000175000017500000001631413612664533017623 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* tactic) -> Sequent.t -> tactic type lseqtac= GlobRef.t -> seqtac type 'a with_backtracking = tactic -> 'a let wrap n b continue seq = Proofview.Goal.enter begin fun gls -> Control.check_for_interrupt (); let nc = Proofview.Goal.hyps gls in let env=pf_env gls in let sigma = project gls in let rec aux i nc ctx= if i<=0 then seq else match nc with []->anomaly (Pp.str "Not the expected number of hyps.") | nd::q-> let id = NamedDecl.get_id nd in if occur_var env sigma id (pf_concl gls) || List.exists (occur_var_in_decl env sigma id) ctx then (aux (i-1) q (nd::ctx)) else add_formula env sigma Hyp (GlobRef.VarRef id) (NamedDecl.get_type nd) (aux (i-1) q (nd::ctx)) in let seq1=aux n nc [] in let seq2=if b then add_formula env sigma Concl dummy_id (pf_concl gls) seq1 else seq1 in continue seq2 end let clear_global=function | GlobRef.VarRef id-> clear [id] | _->tclIDTAC (* connection rules *) let axiom_tac t seq = Proofview.Goal.enter begin fun gl -> try pf_constr_of_global (find_left (project gl) t seq) >>= fun c -> exact_no_check c with Not_found -> tclFAIL 0 (Pp.str "No axiom link") end let ll_atom_tac a backtrack id continue seq = let open EConstr in tclIFTHENELSE (tclTHENLIST [(Proofview.tclEVARMAP >>= fun sigma -> let gr = try Proofview.tclUNIT (find_left sigma a seq) with Not_found -> tclFAIL 0 (Pp.str "No link") in gr >>= fun gr -> pf_constr_of_global gr >>= fun left -> pf_constr_of_global id >>= fun id -> generalize [(mkApp(id, [|left|]))]); clear_global id; intro]) (wrap 1 false continue seq) backtrack (* right connectives rules *) let and_tac backtrack continue seq= tclIFTHENELSE simplest_split (wrap 0 true continue seq) backtrack let or_tac backtrack continue seq= tclORELSE (any_constructor false (Some (tclCOMPLETE (wrap 0 true continue seq)))) backtrack let arrow_tac backtrack continue seq= tclIFTHENELSE intro (wrap 1 true continue seq) (tclORELSE (tclTHEN introf (tclCOMPLETE (wrap 1 true continue seq))) backtrack) (* left connectives rules *) let left_and_tac ind backtrack id continue seq = Proofview.Goal.enter begin fun gl -> let n=(construct_nhyps (pf_env gl) ind).(0) in tclIFTHENELSE (tclTHENLIST [(pf_constr_of_global id >>= simplest_elim); clear_global id; tclDO n intro]) (wrap n false continue seq) backtrack end let left_or_tac ind backtrack id continue seq = Proofview.Goal.enter begin fun gl -> let v=construct_nhyps (pf_env gl) ind in let f n= tclTHENLIST [clear_global id; tclDO n intro; wrap n false continue seq] in tclIFTHENSVELSE (pf_constr_of_global id >>= simplest_elim) (Array.map f v) backtrack end let left_false_tac id= Tacticals.New.pf_constr_of_global id >>= simplest_elim (* left arrow connective rules *) (* We use this function for false, and, or, exists *) let ll_ind_tac (ind,u as indu) largs backtrack id continue seq = Proofview.Goal.enter begin fun gl -> let rcs=ind_hyps (pf_env gl) (project gl) 0 indu largs in let vargs=Array.of_list largs in (* construire le terme H->B, le generaliser etc *) let myterm idc i= let rc=rcs.(i) in let p=List.length rc in let u = EInstance.make u in let cstr=mkApp ((mkConstructU ((ind,(i+1)),u)),vargs) in let vars=Array.init p (fun j->mkRel (p-j)) in let capply=mkApp ((lift p cstr),vars) in let head=mkApp ((lift p idc),[|capply|]) in EConstr.it_mkLambda_or_LetIn head rc in let lp=Array.length rcs in let newhyps idc =List.init lp (myterm idc) in tclIFTHENELSE (tclTHENLIST [(pf_constr_of_global id >>= fun idc -> generalize (newhyps idc)); clear_global id; tclDO lp intro]) (wrap lp false continue seq) backtrack end let ll_arrow_tac a b c backtrack id continue seq= let open EConstr in let open Vars in let cc=mkProd(Context.make_annot Anonymous Sorts.Relevant,a,(lift 1 b)) in let d idc = mkLambda (Context.make_annot Anonymous Sorts.Relevant,b, mkApp (idc, [|mkLambda (Context.make_annot Anonymous Sorts.Relevant,(lift 1 a),(mkRel 2))|])) in tclORELSE (tclTHENS (cut c) [tclTHENLIST [introf; clear_global id; wrap 1 false continue seq]; tclTHENS (cut cc) [(pf_constr_of_global id >>= fun c -> exact_no_check c); tclTHENLIST [(pf_constr_of_global id >>= fun idc -> generalize [d idc]); clear_global id; introf; introf; tclCOMPLETE (wrap 2 true continue seq)]]]) backtrack (* quantifier rules (easy side) *) let forall_tac backtrack continue seq= tclORELSE (tclIFTHENELSE intro (wrap 0 true continue seq) (tclORELSE (tclTHEN introf (tclCOMPLETE (wrap 0 true continue seq))) backtrack)) (if !qflag then tclFAIL 0 (Pp.str "reversible in 1st order mode") else backtrack) let left_exists_tac ind backtrack id continue seq = Proofview.Goal.enter begin fun gl -> let n=(construct_nhyps (pf_env gl) ind).(0) in tclIFTHENELSE (Tacticals.New.pf_constr_of_global id >>= simplest_elim) (tclTHENLIST [clear_global id; tclDO n intro; (wrap (n-1) false continue seq)]) backtrack end let ll_forall_tac prod backtrack id continue seq= tclORELSE (tclTHENS (cut prod) [tclTHENLIST [intro; (pf_constr_of_global id >>= fun idc -> Proofview.Goal.enter begin fun gls-> let open EConstr in let id0 = List.nth (pf_ids_of_hyps gls) 0 in let term=mkApp(idc,[|mkVar(id0)|]) in tclTHEN (generalize [term]) (clear [id0]) end); clear_global id; intro; tclCOMPLETE (wrap 1 false continue (deepen seq))]; tclCOMPLETE (wrap 0 true continue (deepen seq))]) backtrack (* rules for instantiation with unification moved to instances.ml *) coq-8.11.0/plugins/firstorder/ground.ml0000644000175000017500000001337613612664533017774 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { accu with tr_cst = Names.Cpred.remove kn accu.tr_cst } | _ -> accu in let flags = List.fold_left f TransparentState.full (Classops.coercions ()) in red_flags:= CClosure.RedFlags.red_add_transparent CClosure.betaiotazeta flags let ground_tac solver startseq = Proofview.Goal.enter begin fun gl -> update_flags (); let rec toptac skipped seq = Proofview.Goal.enter begin fun gl -> let () = if Tacinterp.get_debug()=Tactic_debug.DebugOn 0 then let gl = { Evd.it = Proofview.Goal.goal gl; sigma = project gl } in Feedback.msg_debug (Printer.pr_goal gl) in tclORELSE (axiom_tac seq.gl seq) begin try let (hd,seq1)=take_formula (project gl) seq and re_add s=re_add_formula_list (project gl) skipped s in let continue=toptac [] and backtrack =toptac (hd::skipped) seq1 in match hd.pat with Right rpat-> begin match rpat with Rand-> and_tac backtrack continue (re_add seq1) | Rforall-> let backtrack1= if !qflag then tclFAIL 0 (Pp.str "reversible in 1st order mode") else backtrack in forall_tac backtrack1 continue (re_add seq1) | Rarrow-> arrow_tac backtrack continue (re_add seq1) | Ror-> or_tac backtrack continue (re_add seq1) | Rfalse->backtrack | Rexists(i,dom,triv)-> let (lfp,seq2)=collect_quantified (project gl) seq in let backtrack2=toptac (lfp@skipped) seq2 in if !qflag && seq.depth>0 then quantified_tac lfp backtrack2 continue (re_add seq) else backtrack2 (* need special backtracking *) end | Left lpat-> begin match lpat with Lfalse-> left_false_tac hd.id | Land ind-> left_and_tac ind backtrack hd.id continue (re_add seq1) | Lor ind-> left_or_tac ind backtrack hd.id continue (re_add seq1) | Lforall (_,_,_)-> let (lfp,seq2)=collect_quantified (project gl) seq in let backtrack2=toptac (lfp@skipped) seq2 in if !qflag && seq.depth>0 then quantified_tac lfp backtrack2 continue (re_add seq) else backtrack2 (* need special backtracking *) | Lexists ind -> if !qflag then left_exists_tac ind backtrack hd.id continue (re_add seq1) else backtrack | LA (typ,lap)-> let la_tac= begin match lap with LLatom -> backtrack | LLand (ind,largs) | LLor(ind,largs) | LLfalse (ind,largs)-> (ll_ind_tac ind largs backtrack hd.id continue (re_add seq1)) | LLforall p -> if seq.depth>0 && !qflag then (ll_forall_tac p backtrack hd.id continue (re_add seq1)) else backtrack | LLexists (ind,l) -> if !qflag then ll_ind_tac ind l backtrack hd.id continue (re_add seq1) else backtrack | LLarrow (a,b,c) -> (ll_arrow_tac a b c backtrack hd.id continue (re_add seq1)) end in ll_atom_tac typ la_tac hd.id continue (re_add seq1) end with Heap.EmptyHeap->solver end end in let n = List.length (Proofview.Goal.hyps gl) in startseq (fun seq -> wrap n true (toptac []) seq) end coq-8.11.0/plugins/firstorder/instances.mli0000644000175000017500000000167213612664533020632 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Sequent.t -> Formula.t list * Sequent.t val give_instances : Evd.evar_map -> Formula.t list -> Sequent.t -> (Unify.instance * GlobRef.t) list val quantified_tac : Formula.t list -> seqtac with_backtracking coq-8.11.0/plugins/firstorder/sequent.mli0000644000175000017500000000365713612664533020334 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr -> GlobRef.t -> GlobRef.t list CM.t -> GlobRef.t list CM.t val cm_remove : Evd.evar_map -> constr -> GlobRef.t -> GlobRef.t list CM.t -> GlobRef.t list CM.t module HP: Heap.S with type elt=Formula.t type t = {redexes:HP.t; context: GlobRef.t list CM.t; latoms:constr list; gl:types; glatom:constr option; cnt:counter; history:History.t; depth:int} val deepen: t -> t val record: h_item -> t -> t val lookup: Evd.evar_map -> h_item -> t -> bool val add_formula : Environ.env -> Evd.evar_map -> side -> GlobRef.t -> constr -> t -> t val re_add_formula_list : Evd.evar_map -> Formula.t list -> t -> t val find_left : Evd.evar_map -> constr -> t -> GlobRef.t val take_formula : Evd.evar_map -> t -> Formula.t * t val empty_seq : int -> t val extend_with_ref_list : Environ.env -> Evd.evar_map -> GlobRef.t list -> t -> t * Evd.evar_map val extend_with_auto_hints : Environ.env -> Evd.evar_map -> Hints.hint_db_name list -> t -> t * Evd.evar_map val print_cmap: GlobRef.t list CM.t -> Pp.t coq-8.11.0/plugins/firstorder/rules.mli0000644000175000017500000000316513612664533017774 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* tactic) -> Sequent.t -> tactic type lseqtac= GlobRef.t -> seqtac type 'a with_backtracking = tactic -> 'a val wrap : int -> bool -> seqtac val clear_global: GlobRef.t -> tactic val axiom_tac : constr -> Sequent.t -> tactic val ll_atom_tac : constr -> lseqtac with_backtracking val and_tac : seqtac with_backtracking val or_tac : seqtac with_backtracking val arrow_tac : seqtac with_backtracking val left_and_tac : pinductive -> lseqtac with_backtracking val left_or_tac : pinductive -> lseqtac with_backtracking val left_false_tac : GlobRef.t -> tactic val ll_ind_tac : pinductive -> constr list -> lseqtac with_backtracking val ll_arrow_tac : constr -> constr -> constr -> lseqtac with_backtracking val forall_tac : seqtac with_backtracking val left_exists_tac : pinductive -> lseqtac with_backtracking val ll_forall_tac : types -> lseqtac with_backtracking coq-8.11.0/plugins/firstorder/unify.ml0000644000175000017500000001223213612664533017616 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (m, EConstr.Unsafe.to_constr c)) subst in EConstr.of_constr (subst_meta subst (EConstr.Unsafe.to_constr t)) let unif evd t1 t2= let bige=Queue.create () and sigma=ref [] in let bind i t= sigma:=(i,t):: (List.map (function (n,tn)->(n,subst_meta [i,t] tn)) !sigma) in let rec head_reduce t= (* forbids non-sigma-normal meta in head position*) match EConstr.kind evd t with Meta i-> (try head_reduce (Int.List.assoc i !sigma) with Not_found->t) | _->t in Queue.add (t1,t2) bige; try while true do let t1,t2=Queue.take bige in let nt1=head_reduce (whd_betaiotazeta evd t1) and nt2=head_reduce (whd_betaiotazeta evd t2) in match (EConstr.kind evd nt1),(EConstr.kind evd nt2) with Meta i,Meta j-> if not (Int.equal i j) then if i let t=subst_meta !sigma nt2 in if Int.Set.is_empty (free_rels evd t) && not (occur_metavariable evd i t) then bind i t else raise (UFAIL(nt1,nt2)) | _,Meta i -> let t=subst_meta !sigma nt1 in if Int.Set.is_empty (free_rels evd t) && not (occur_metavariable evd i t) then bind i t else raise (UFAIL(nt1,nt2)) | Cast(_,_,_),_->Queue.add (strip_outer_cast evd nt1,nt2) bige | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast evd nt2) bige | (Prod(_,a,b),Prod(_,c,d))|(Lambda(_,a,b),Lambda(_,c,d))-> Queue.add (a,c) bige;Queue.add (pop b,pop d) bige | Case (_,pa,ca,va),Case (_,pb,cb,vb)-> Queue.add (pa,pb) bige; Queue.add (ca,cb) bige; let l=Array.length va in if not (Int.equal l (Array.length vb)) then raise (UFAIL (nt1,nt2)) else for i=0 to l-1 do Queue.add (va.(i),vb.(i)) bige done | App(ha,va),App(hb,vb)-> Queue.add (ha,hb) bige; let l=Array.length va in if not (Int.equal l (Array.length vb)) then raise (UFAIL (nt1,nt2)) else for i=0 to l-1 do Queue.add (va.(i),vb.(i)) bige done | _->if not (eq_constr_nounivs evd nt1 nt2) then raise (UFAIL (nt1,nt2)) done; assert false (* this place is unreachable but needed for the sake of typing *) with Queue.Empty-> !sigma let value evd i t= let add x y= if x<0 then y else if y<0 then x else x+y in let rec vaux term= if isMeta evd term && Int.equal (destMeta evd term) i then 0 else let f v t=add v (vaux t) in let vr=EConstr.fold evd f (-1) term in if vr<0 then -1 else vr+1 in vaux t type instance= Real of (int*constr)*int | Phantom of constr let mk_rel_inst evd t= let new_rel=ref 1 in let rel_env=ref [] in let rec renum_rec d t= match EConstr.kind evd t with Meta n-> (try mkRel (d+(Int.List.assoc n !rel_env)) with Not_found-> let m= !new_rel in incr new_rel; rel_env:=(n,m) :: !rel_env; mkRel (m+d)) | _ -> EConstr.map_with_binders evd succ renum_rec d t in let nt=renum_rec 0 t in (!new_rel - 1,nt) let unif_atoms evd i dom t1 t2= try let t=Int.List.assoc i (unif evd t1 t2) in if isMeta evd t then Some (Phantom dom) else Some (Real(mk_rel_inst evd t,value evd i t1)) with UFAIL(_,_) ->None | Not_found ->Some (Phantom dom) let renum_metas_from k n t= (* requires n = max (free_rels t) *) let l=List.init n (fun i->mkMeta (k+i)) in substl l t let more_general evd (m1,t1) (m2,t2)= let mt1=renum_metas_from 0 m1 t1 and mt2=renum_metas_from m1 m2 t2 in try let sigma=unif evd mt1 mt2 in let p (n,t)= nfalse coq-8.11.0/plugins/firstorder/formula.ml0000644000175000017500000002324613612664533020140 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* metavariable exception Is_atom of constr let meta_succ m = m+1 let rec nb_prod_after n c= match Constr.kind c with | Prod (_,_,b) ->if n>0 then nb_prod_after (n-1) b else 1+(nb_prod_after 0 b) | _ -> 0 let construct_nhyps env ind = let nparams = (fst (Global.lookup_inductive (fst ind))).mind_nparams in let constr_types = Inductiveops.arities_of_constructors env ind in let hyp = nb_prod_after nparams in Array.map hyp constr_types (* indhyps builds the array of arrays of constructor hyps for (ind largs)*) let ind_hyps env sigma nevar ind largs = let types= Inductiveops.arities_of_constructors env ind in let myhyps t = let t = EConstr.of_constr t in let nparam_decls = Context.Rel.length (fst (Global.lookup_inductive (fst ind))).mind_params_ctxt in let t1=Termops.prod_applist_assum sigma nparam_decls t largs in let t2=snd (decompose_prod_n_assum sigma nevar t1) in fst (decompose_prod_assum sigma t2) in Array.map myhyps types let special_nf env sigma t = Reductionops.clos_norm_flags !red_flags env sigma t let special_whd env sigma t = Reductionops.clos_whd_flags !red_flags env sigma t type kind_of_formula= Arrow of constr*constr | False of pinductive*constr list | And of pinductive*constr list*bool | Or of pinductive*constr list*bool | Exists of pinductive*constr list | Forall of constr*constr | Atom of constr let pop t = Vars.lift (-1) t let kind_of_formula env sigma term = let normalize = special_nf env sigma in let cciterm = special_whd env sigma term in match match_with_imp_term env sigma cciterm with Some (a,b)-> Arrow (a, pop b) |_-> match match_with_forall_term env sigma cciterm with Some (_,a,b)-> Forall (a, b) |_-> match match_with_nodep_ind env sigma cciterm with Some (i,l,n)-> let ind,u=EConstr.destInd sigma i in let u = EConstr.EInstance.kind sigma u in let (mib,mip) = Global.lookup_inductive ind in let nconstr=Array.length mip.mind_consnames in if Int.equal nconstr 0 then False((ind,u),l) else let has_realargs=(n>0) in let is_trivial= let is_constant n = Int.equal n 0 in Array.exists is_constant mip.mind_consnrealargs in if Inductiveops.mis_is_recursive (ind,mib,mip) || (has_realargs && not is_trivial) then Atom cciterm else if Int.equal nconstr 1 then And((ind,u),l,is_trivial) else Or((ind,u),l,is_trivial) | _ -> match match_with_sigma_type env sigma cciterm with Some (i,l)-> let (ind, u) = EConstr.destInd sigma i in let u = EConstr.EInstance.kind sigma u in Exists((ind, u), l) |_-> Atom (normalize cciterm) type atoms = {positive:constr list;negative:constr list} type side = Hyp | Concl | Hint let no_atoms = (false,{positive=[];negative=[]}) let dummy_id=GlobRef.VarRef (Id.of_string "_") (* "_" cannot be parsed *) let build_atoms env sigma metagen side cciterm = let trivial =ref false and positive=ref [] and negative=ref [] in let normalize=special_nf env sigma in let rec build_rec subst polarity cciterm= match kind_of_formula env sigma cciterm with False(_,_)->if not polarity then trivial:=true | Arrow (a,b)-> build_rec subst (not polarity) a; build_rec subst polarity b | And(i,l,b) | Or(i,l,b)-> if b then begin let unsigned=normalize (substnl subst 0 cciterm) in if polarity then positive:= unsigned :: !positive else negative:= unsigned :: !negative end; let v = ind_hyps env sigma 0 i l in let g i _ decl = build_rec subst polarity (lift i (RelDecl.get_type decl)) in let f l = List.fold_left_i g (1-(List.length l)) () l in if polarity && (* we have a constant constructor *) Array.exists (function []->true|_->false) v then trivial:=true; Array.iter f v | Exists(i,l)-> let var=mkMeta (metagen true) in let v =(ind_hyps env sigma 1 i l).(0) in let g i _ decl = build_rec (var::subst) polarity (lift i (RelDecl.get_type decl)) in List.fold_left_i g (2-(List.length l)) () v | Forall(_,b)-> let var=mkMeta (metagen true) in build_rec (var::subst) polarity b | Atom t-> let unsigned=substnl subst 0 t in if not (isMeta sigma unsigned) then (* discarding wildcard atoms *) if polarity then positive:= unsigned :: !positive else negative:= unsigned :: !negative in begin match side with Concl -> build_rec [] true cciterm | Hyp -> build_rec [] false cciterm | Hint -> let rels,head=decompose_prod sigma cciterm in let subst=List.rev_map (fun _->mkMeta (metagen true)) rels in build_rec subst false head;trivial:=false (* special for hints *) end; (!trivial, {positive= !positive; negative= !negative}) type right_pattern = Rarrow | Rand | Ror | Rfalse | Rforall | Rexists of metavariable*constr*bool type left_arrow_pattern= LLatom | LLfalse of pinductive*constr list | LLand of pinductive*constr list | LLor of pinductive*constr list | LLforall of constr | LLexists of pinductive*constr list | LLarrow of constr*constr*constr type left_pattern= Lfalse | Land of pinductive | Lor of pinductive | Lforall of metavariable*constr*bool | Lexists of pinductive | LA of constr*left_arrow_pattern type t={id:GlobRef.t; constr:constr; pat:(left_pattern,right_pattern) sum; atoms:atoms} let build_formula env sigma side nam typ metagen= let normalize = special_nf env sigma in try let m=meta_succ(metagen false) in let trivial,atoms= if !qflag then build_atoms env sigma metagen side typ else no_atoms in let pattern= match side with Concl -> let pat= match kind_of_formula env sigma typ with False(_,_) -> Rfalse | Atom a -> raise (Is_atom a) | And(_,_,_) -> Rand | Or(_,_,_) -> Ror | Exists (i,l) -> let d = RelDecl.get_type (List.last (ind_hyps env sigma 0 i l).(0)) in Rexists(m,d,trivial) | Forall (_,a) -> Rforall | Arrow (a,b) -> Rarrow in Right pat | _ -> let pat= match kind_of_formula env sigma typ with False(i,_) -> Lfalse | Atom a -> raise (Is_atom a) | And(i,_,b) -> if b then let nftyp=normalize typ in raise (Is_atom nftyp) else Land i | Or(i,_,b) -> if b then let nftyp=normalize typ in raise (Is_atom nftyp) else Lor i | Exists (ind,_) -> Lexists ind | Forall (d,_) -> Lforall(m,d,trivial) | Arrow (a,b) -> let nfa=normalize a in LA (nfa, match kind_of_formula env sigma a with False(i,l)-> LLfalse(i,l) | Atom t-> LLatom | And(i,l,_)-> LLand(i,l) | Or(i,l,_)-> LLor(i,l) | Arrow(a,c)-> LLarrow(a,c,b) | Exists(i,l)->LLexists(i,l) | Forall(_,_)->LLforall a) in Left pat in Left {id=nam; constr=normalize typ; pat=pattern; atoms=atoms} with Is_atom a-> Right a (* already in nf *) coq-8.11.0/plugins/firstorder/ground.mli0000644000175000017500000000145213612664533020135 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ((Sequent.t -> unit Proofview.tactic) -> unit Proofview.tactic) -> unit Proofview.tactic coq-8.11.0/plugins/firstorder/sequent.ml0000644000175000017500000001612413612664533020154 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* if b then incr cnt;!cnt let priority = (* pure heuristics, <=0 for non reversible *) function Right rf-> begin match rf with Rarrow -> 100 | Rand -> 40 | Ror -> -15 | Rfalse -> -50 | Rforall -> 100 | Rexists (_,_,_) -> -29 end | Left lf -> match lf with Lfalse -> 999 | Land _ -> 90 | Lor _ -> 40 | Lforall (_,_,_) -> -30 | Lexists _ -> 60 | LA(_,lap) -> match lap with LLatom -> 0 | LLfalse (_,_) -> 100 | LLand (_,_) -> 80 | LLor (_,_) -> 70 | LLforall _ -> -20 | LLexists (_,_) -> 50 | LLarrow (_,_,_) -> -10 module OrderedFormula= struct type t=Formula.t let compare e1 e2= (priority e1.pat) - (priority e2.pat) end type h_item = GlobRef.t * (int*Constr.t) option module Hitem= struct type t = h_item let compare (id1,co1) (id2,co2)= let c = GlobRef.Ordered.compare id1 id2 in if c = 0 then let cmp (i1, c1) (i2, c2) = let c = Int.compare i1 i2 in if c = 0 then Constr.compare c1 c2 else c in Option.compare cmp co1 co2 else c end module CM=Map.Make(Constr) module History=Set.Make(Hitem) let cm_add sigma typ nam cm= let typ = EConstr.to_constr ~abort_on_undefined_evars:false sigma typ in try let l=CM.find typ cm in CM.add typ (nam::l) cm with Not_found->CM.add typ [nam] cm let cm_remove sigma typ nam cm= let typ = EConstr.to_constr ~abort_on_undefined_evars:false sigma typ in try let l=CM.find typ cm in let l0=List.filter (fun id-> not (GlobRef.equal id nam)) l in match l0 with []->CM.remove typ cm | _ ->CM.add typ l0 cm with Not_found ->cm module HP=Heap.Functional(OrderedFormula) type t= {redexes:HP.t; context:(GlobRef.t list) CM.t; latoms:constr list; gl:types; glatom:constr option; cnt:counter; history:History.t; depth:int} let deepen seq={seq with depth=seq.depth-1} let record item seq={seq with history=History.add item seq.history} let lookup sigma item seq= History.mem item seq.history || match item with (_,None)->false | (id,Some (m, t))-> let p (id2,o)= match o with None -> false | Some (m2, t2)-> GlobRef.equal id id2 && m2>m && more_general sigma (m2, EConstr.of_constr t2) (m, EConstr.of_constr t) in History.exists p seq.history let add_formula env sigma side nam t seq = match build_formula env sigma side nam t seq.cnt with Left f-> begin match side with Concl -> {seq with redexes=HP.add f seq.redexes; gl=f.constr; glatom=None} | _ -> {seq with redexes=HP.add f seq.redexes; context=cm_add sigma f.constr nam seq.context} end | Right t-> match side with Concl -> {seq with gl=t;glatom=Some t} | _ -> {seq with context=cm_add sigma t nam seq.context; latoms=t::seq.latoms} let re_add_formula_list sigma lf seq= let do_one f cm= if f.id == dummy_id then cm else cm_add sigma f.constr f.id cm in {seq with redexes=List.fold_right HP.add lf seq.redexes; context=List.fold_right do_one lf seq.context} let find_left sigma t seq=List.hd (CM.find (EConstr.to_constr ~abort_on_undefined_evars:false sigma t) seq.context) (*let rev_left seq= try let lpat=(HP.maximum seq.redexes).pat in left_reversible lpat with Heap.EmptyHeap -> false *) let rec take_formula sigma seq= let hd=HP.maximum seq.redexes and hp=HP.remove seq.redexes in if hd.id == dummy_id then let nseq={seq with redexes=hp} in if seq.gl==hd.constr then hd,nseq else take_formula sigma nseq (* discarding deprecated goal *) else hd,{seq with redexes=hp; context=cm_remove sigma hd.constr hd.id seq.context} let empty_seq depth= {redexes=HP.empty; context=CM.empty; latoms=[]; gl=(mkMeta 1); glatom=None; cnt=newcnt (); history=History.empty; depth=depth} let expand_constructor_hints = List.map_append (function | GlobRef.IndRef ind -> List.init (Inductiveops.nconstructors (Global.env()) ind) (fun i -> GlobRef.ConstructRef (ind,i+1)) | gr -> [gr]) let extend_with_ref_list env sigma l seq = let l = expand_constructor_hints l in let f gr (seq, sigma) = let sigma, c = Evd.fresh_global env sigma gr in let sigma, typ= Typing.type_of env sigma c in (add_formula env sigma Hyp gr typ seq, sigma) in List.fold_right f l (seq, sigma) open Hints let extend_with_auto_hints env sigma l seq = let seqref=ref seq in let f p_a_t = match repr_hint p_a_t.code with Res_pf (c,_) | Give_exact (c,_) | Res_pf_THEN_trivial_fail (c,_) -> let (c, _, _) = c in (try let (gr, _) = Termops.global_of_constr sigma c in let typ=(Typing.unsafe_type_of env sigma c) in seqref:=add_formula env sigma Hint gr typ !seqref with Not_found->()) | _-> () in let g _ _ l = List.iter f l in let h dbname= let hdb= try searchtable_map dbname with Not_found-> user_err Pp.(str ("Firstorder: "^dbname^" : No such Hint database")) in Hint_db.iter g hdb in List.iter h l; !seqref, sigma (*FIXME: forgetting about universes*) let print_cmap map= let print_entry c l s= let env = Global.env () in let sigma = Evd.from_env env in let xc=Constrextern.extern_constr false env sigma (EConstr.of_constr c) in str "| " ++ prlist Printer.pr_global l ++ str " : " ++ Ppconstr.pr_constr_expr env sigma xc ++ cut () ++ s in (v 0 (str "-----" ++ cut () ++ CM.fold print_entry map (mt ()) ++ str "-----")) coq-8.11.0/plugins/firstorder/ground_plugin.mlpack0000644000175000017500000000006613612664533022201 0ustar treinentreinenFormula Unify Sequent Rules Instances Ground G_ground coq-8.11.0/plugins/firstorder/formula.mli0000644000175000017500000000426313612664533020307 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a -> int) -> ('b -> 'b -> int) -> 'a -> 'a -> 'b -> 'b -> int val (==?) : ('a -> 'a -> 'b ->'b -> int) -> ('c -> 'c -> int) -> 'a -> 'a -> 'b -> 'b -> 'c ->'c -> int type ('a,'b) sum = Left of 'a | Right of 'b type counter = bool -> metavariable val construct_nhyps : Environ.env -> pinductive -> int array val ind_hyps : Environ.env -> Evd.evar_map -> int -> pinductive -> constr list -> EConstr.rel_context array type atoms = {positive:constr list;negative:constr list} type side = Hyp | Concl | Hint val dummy_id: GlobRef.t val build_atoms : Environ.env -> Evd.evar_map -> counter -> side -> constr -> bool * atoms type right_pattern = Rarrow | Rand | Ror | Rfalse | Rforall | Rexists of metavariable*constr*bool type left_arrow_pattern= LLatom | LLfalse of pinductive*constr list | LLand of pinductive*constr list | LLor of pinductive*constr list | LLforall of constr | LLexists of pinductive*constr list | LLarrow of constr*constr*constr type left_pattern= Lfalse | Land of pinductive | Lor of pinductive | Lforall of metavariable*constr*bool | Lexists of pinductive | LA of constr*left_arrow_pattern type t={id: GlobRef.t; constr: constr; pat: (left_pattern,right_pattern) sum; atoms: atoms} (*exception Is_atom of constr*) val build_formula : Environ.env -> Evd.evar_map -> side -> GlobRef.t -> types -> counter -> (t,types) sum coq-8.11.0/plugins/firstorder/g_ground.mlg0000644000175000017500000001153113612664533020440 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Some !ground_depth); optwrite= (function None->ground_depth:=3 | Some i->ground_depth:=(max i 0))} in declare_int_option gdopt let default_intuition_tac = let tac _ _ = Auto.h_auto None [] None in let name = { Tacexpr.mltac_plugin = "ground_plugin"; mltac_tactic = "auto_with"; } in let entry = { Tacexpr.mltac_name = name; mltac_index = 0 } in Tacenv.register_ml_tactic name [| tac |]; Tacexpr.TacML (CAst.make (entry, [])) let (set_default_solver, default_solver, print_default_solver) = Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver" } VERNAC COMMAND EXTEND Firstorder_Set_Solver CLASSIFIED AS SIDEFF | #[ locality; ] [ "Set" "Firstorder" "Solver" tactic(t) ] -> { set_default_solver (Locality.make_section_locality locality) (Tacintern.glob_tactic t) } END VERNAC COMMAND EXTEND Firstorder_Print_Solver CLASSIFIED AS QUERY | [ "Print" "Firstorder" "Solver" ] -> { Feedback.msg_notice (Pp.(++) (Pp.str"Firstorder solver tactic is ") (print_default_solver ())) } END { let gen_ground_tac flag taco ids bases = let backup= !qflag in Proofview.tclOR begin Proofview.Goal.enter begin fun gl -> qflag:=flag; let solver= match taco with Some tac-> tac | None-> snd (default_solver ()) in let startseq k = Proofview.Goal.enter begin fun gl -> let seq=empty_seq !ground_depth in let seq, sigma = extend_with_ref_list (pf_env gl) (project gl) ids seq in let seq, sigma = extend_with_auto_hints (pf_env gl) (project gl) bases seq in tclTHEN (Proofview.Unsafe.tclEVARS sigma) (k seq) end in let result=ground_tac solver startseq in qflag := backup; result end end (fun (e, info) -> qflag := backup; Proofview.tclZERO ~info e) (* special for compatibility with Intuition let constant str = Coqlib.get_constr str let defined_connectives=lazy [[],EvalConstRef (destConst (constant "core.not.type")); [],EvalConstRef (destConst (constant "core.iff.type"))] let normalize_evaluables= onAllHypsAndConcl (function None->unfold_in_concl (Lazy.force defined_connectives) | Some id-> unfold_in_hyp (Lazy.force defined_connectives) (Tacexpr.InHypType id)) *) open Ppconstr open Printer let pr_firstorder_using_raw _ _ _ = Pptactic.pr_auto_using pr_qualid let pr_firstorder_using_glob _ _ _ = Pptactic.pr_auto_using (Pputils.pr_or_var (fun x -> pr_global (snd x))) let pr_firstorder_using_typed _ _ _ = Pptactic.pr_auto_using pr_global let warn_deprecated_syntax = CWarnings.create ~name:"firstorder-deprecated-syntax" ~category:"deprecated" (fun () -> Pp.strbrk "Deprecated syntax; use \",\" as separator") } ARGUMENT EXTEND firstorder_using TYPED AS reference list PRINTED BY { pr_firstorder_using_typed } RAW_PRINTED BY { pr_firstorder_using_raw } GLOB_PRINTED BY { pr_firstorder_using_glob } | [ "using" reference(a) ] -> { [a] } | [ "using" reference(a) "," ne_reference_list_sep(l,",") ] -> { a::l } | [ "using" reference(a) reference(b) reference_list(l) ] -> { warn_deprecated_syntax (); a::b::l } | [ ] -> { [] } END TACTIC EXTEND firstorder | [ "firstorder" tactic_opt(t) firstorder_using(l) ] -> { gen_ground_tac true (Option.map (tactic_of_value ist) t) l [] } | [ "firstorder" tactic_opt(t) "with" ne_preident_list(l) ] -> { gen_ground_tac true (Option.map (tactic_of_value ist) t) [] l } | [ "firstorder" tactic_opt(t) firstorder_using(l) "with" ne_preident_list(l') ] -> { gen_ground_tac true (Option.map (tactic_of_value ist) t) l l' } END TACTIC EXTEND gintuition | [ "gintuition" tactic_opt(t) ] -> { gen_ground_tac false (Option.map (tactic_of_value ist) t) [] [] } END coq-8.11.0/plugins/firstorder/instances.ml0000644000175000017500000001637213612664533020464 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (cmp d1 d2) | Real((m1,c1),n1),Real((m2,c2),n2)-> ((-) =? (-) ==? cmp) m2 m1 n1 n2 c1 c2 | Phantom(_),Real((m,_),_)-> if Int.equal m 0 then -1 else 1 | Real((m,_),_),Phantom(_)-> if Int.equal m 0 then 1 else -1 let compare_gr id1 id2 = if id1==id2 then 0 else if id1==dummy_id then 1 else if id2==dummy_id then -1 else GlobRef.Ordered.compare id1 id2 module OrderedInstance= struct type t=instance * GlobRef.t let compare (inst1,id1) (inst2,id2)= (compare_instance =? compare_gr) inst2 inst1 id2 id1 (* we want a __decreasing__ total order *) end module IS=Set.Make(OrderedInstance) let make_simple_atoms seq= let ratoms= match seq.glatom with Some t->[t] | None->[] in {negative=seq.latoms;positive=ratoms} let do_sequent sigma setref triv id seq i dom atoms= let flag=ref true in let phref=ref triv in let do_atoms a1 a2 = let do_pair t1 t2 = match unif_atoms sigma i dom t1 t2 with None->() | Some (Phantom _) ->phref:=true | Some c ->flag:=false;setref:=IS.add (c,id) !setref in List.iter (fun t->List.iter (do_pair t) a2.negative) a1.positive; List.iter (fun t->List.iter (do_pair t) a2.positive) a1.negative in HP.iter (fun lf->do_atoms atoms lf.atoms) seq.redexes; do_atoms atoms (make_simple_atoms seq); !flag && !phref let match_one_quantified_hyp sigma setref seq lf= match lf.pat with Left(Lforall(i,dom,triv))|Right(Rexists(i,dom,triv))-> if do_sequent sigma setref triv lf.id seq i dom lf.atoms then setref:=IS.add ((Phantom dom),lf.id) !setref | _ -> anomaly (Pp.str "can't happen.") let give_instances sigma lf seq= let setref=ref IS.empty in List.iter (match_one_quantified_hyp sigma setref seq) lf; IS.elements !setref (* collector for the engine *) let rec collect_quantified sigma seq= try let hd,seq1=take_formula sigma seq in (match hd.pat with Left(Lforall(_,_,_)) | Right(Rexists(_,_,_)) -> let (q,seq2)=collect_quantified sigma seq1 in ((hd::q),seq2) | _->[],seq) with Heap.EmptyHeap -> [],seq (* open instances processor *) let dummy_bvid=Id.of_string "x" let mk_open_instance env evmap id idc m t = let var_id= if id==dummy_id then dummy_bvid else let typ=Typing.unsafe_type_of env evmap idc in (* since we know we will get a product, reduction is not too expensive *) let (nam,_,_)=destProd evmap (whd_all env evmap typ) in match nam.Context.binder_name with Name id -> id | Anonymous -> dummy_bvid in let revt=substl (List.init m (fun i->mkRel (m-i))) t in let rec aux n avoid env evmap decls = if Int.equal n 0 then evmap, decls else let nid=(fresh_id_in_env avoid var_id env) in let (evmap, (c, _)) = Evarutil.new_type_evar env evmap Evd.univ_flexible in let decl = LocalAssum (Context.make_annot (Name nid) Sorts.Relevant, c) in aux (n-1) (Id.Set.add nid avoid) (EConstr.push_rel decl env) evmap (decl::decls) in let evmap, decls = aux m Id.Set.empty env evmap [] in (evmap, decls, revt) (* tactics *) let left_instance_tac (inst,id) continue seq= let open EConstr in Proofview.Goal.enter begin fun gl -> let sigma = project gl in match inst with Phantom dom-> if lookup sigma (id,None) seq then tclFAIL 0 (Pp.str "already done") else tclTHENS (cut dom) [tclTHENLIST [introf; (pf_constr_of_global id >>= fun idc -> Proofview.Goal.enter begin fun gl -> let id0 = List.nth (pf_ids_of_hyps gl) 0 in generalize [mkApp(idc, [|mkVar id0|])] end); introf; tclSOLVE [wrap 1 false continue (deepen (record (id,None) seq))]]; tclTRY assumption] | Real((m,t),_)-> let c = (m, EConstr.to_constr sigma t) in if lookup sigma (id,Some c) seq then tclFAIL 0 (Pp.str "already done") else let special_generalize= if m>0 then (pf_constr_of_global id >>= fun idc -> Proofview.Goal.enter begin fun gl-> let (evmap, rc, ot) = mk_open_instance (pf_env gl) (project gl) id idc m t in let gt= it_mkLambda_or_LetIn (mkApp(idc,[|ot|])) rc in let evmap, _ = try Typing.type_of (pf_env gl) evmap gt with e when CErrors.noncritical e -> user_err Pp.(str "Untypable instance, maybe higher-order non-prenex quantification") in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evmap) (generalize [gt]) end) else pf_constr_of_global id >>= fun idc -> generalize [mkApp(idc,[|t|])] in tclTHENLIST [special_generalize; introf; tclSOLVE [wrap 1 false continue (deepen (record (id,Some c) seq))]] end let right_instance_tac inst continue seq= let open EConstr in Proofview.Goal.enter begin fun gl -> match inst with Phantom dom -> tclTHENS (cut dom) [tclTHENLIST [introf; Proofview.Goal.enter begin fun gl -> let id0 = List.nth (pf_ids_of_hyps gl) 0 in split (Tactypes.ImplicitBindings [mkVar id0]) end; tclSOLVE [wrap 0 true continue (deepen seq)]]; tclTRY assumption] | Real ((0,t),_) -> (tclTHEN (split (Tactypes.ImplicitBindings [t])) (tclSOLVE [wrap 0 true continue (deepen seq)])) | Real ((m,t),_) -> tclFAIL 0 (Pp.str "not implemented ... yet") end let instance_tac inst= if (snd inst)==dummy_id then right_instance_tac (fst inst) else left_instance_tac inst let quantified_tac lf backtrack continue seq = Proofview.Goal.enter begin fun gl -> let insts=give_instances (project gl) lf seq in tclORELSE (tclFIRST (List.map (fun inst->instance_tac inst continue seq) insts)) backtrack end coq-8.11.0/plugins/firstorder/plugin_base.dune0000644000175000017500000000023013612664533021272 0ustar treinentreinen(library (name ground_plugin) (public_name coq.plugins.firstorder) (synopsis "Coq's first order logic solver plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/btauto/0000755000175000017500000000000013612664533015245 5ustar treinentreinencoq-8.11.0/plugins/btauto/refl_btauto.ml0000644000175000017500000002072013612664533020106 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* nil ty | t::q -> cons ty t (of_list ty q) end module CoqPositive = struct let _xH = bt_lib_constr "num.pos.xH" let _xO = bt_lib_constr "num.pos.xO" let _xI = bt_lib_constr "num.pos.xI" (* A coq nat from an int *) let rec of_int n = if n <= 1 then Lazy.force _xH else let ans = of_int (n / 2) in if n mod 2 = 0 then lapp _xO [|ans|] else lapp _xI [|ans|] end module Env = struct module ConstrHashtbl = Hashtbl.Make (Constr) type t = (int ConstrHashtbl.t * int ref) let add (tbl, off) (t : Constr.t) = try ConstrHashtbl.find tbl t with | Not_found -> let i = !off in let () = ConstrHashtbl.add tbl t i in let () = incr off in i let empty () = (ConstrHashtbl.create 16, ref 1) let to_list (env, _) = (* we need to get an ordered list *) let fold constr key accu = (key, constr) :: accu in let l = ConstrHashtbl.fold fold env [] in let sorted_l = List.sort (fun p1 p2 -> Int.compare (fst p1) (fst p2)) l in List.map snd sorted_l end module Bool = struct let ind = lazy (Globnames.destIndRef (Coqlib.lib_ref "core.bool.type")) let typ = bt_lib_constr "core.bool.type" let trueb = bt_lib_constr "core.bool.true" let falseb = bt_lib_constr "core.bool.false" let andb = bt_lib_constr "core.bool.andb" let orb = bt_lib_constr "core.bool.orb" let xorb = bt_lib_constr "core.bool.xorb" let negb = bt_lib_constr "core.bool.negb" type t = | Var of int | Const of bool | Andb of t * t | Orb of t * t | Xorb of t * t | Negb of t | Ifb of t * t * t let quote (env : Env.t) sigma (c : Constr.t) : t = let trueb = Lazy.force trueb in let falseb = Lazy.force falseb in let andb = Lazy.force andb in let orb = Lazy.force orb in let xorb = Lazy.force xorb in let negb = Lazy.force negb in let rec aux c = match decomp_term sigma c with | App (head, args) -> if head === andb && Array.length args = 2 then Andb (aux args.(0), aux args.(1)) else if head === orb && Array.length args = 2 then Orb (aux args.(0), aux args.(1)) else if head === xorb && Array.length args = 2 then Xorb (aux args.(0), aux args.(1)) else if head === negb && Array.length args = 1 then Negb (aux args.(0)) else Var (Env.add env c) | Case (info, r, arg, pats) -> let is_bool = let i = info.ci_ind in Names.eq_ind i (Lazy.force ind) in if is_bool then Ifb ((aux arg), (aux pats.(0)), (aux pats.(1))) else Var (Env.add env c) | _ -> if c === falseb then Const false else if c === trueb then Const true else Var (Env.add env c) in aux c end module Btauto = struct open Pp let eq = bt_lib_constr "core.eq.type" let f_var = bt_lib_constr "plugins.btauto.f_var" let f_btm = bt_lib_constr "plugins.btauto.f_btm" let f_top = bt_lib_constr "plugins.btauto.f_top" let f_cnj = bt_lib_constr "plugins.btauto.f_cnj" let f_dsj = bt_lib_constr "plugins.btauto.f_dsj" let f_neg = bt_lib_constr "plugins.btauto.f_neg" let f_xor = bt_lib_constr "plugins.btauto.f_xor" let f_ifb = bt_lib_constr "plugins.btauto.f_ifb" let eval = bt_lib_constr "plugins.btauto.eval" let witness = bt_lib_constr "plugins.btauto.witness" let soundness = bt_lib_constr "plugins.btauto.soundness" let rec convert = function | Bool.Var n -> lapp f_var [|CoqPositive.of_int n|] | Bool.Const true -> Lazy.force f_top | Bool.Const false -> Lazy.force f_btm | Bool.Andb (b1, b2) -> lapp f_cnj [|convert b1; convert b2|] | Bool.Orb (b1, b2) -> lapp f_dsj [|convert b1; convert b2|] | Bool.Negb b -> lapp f_neg [|convert b|] | Bool.Xorb (b1, b2) -> lapp f_xor [|convert b1; convert b2|] | Bool.Ifb (b1, b2, b3) -> lapp f_ifb [|convert b1; convert b2; convert b3|] let convert_env env : Constr.t = CoqList.of_list (Lazy.force Bool.typ) env let reify env t = lapp eval [|convert_env env; convert t|] let print_counterexample p penv gl = let var = lapp witness [|p|] in let var = EConstr.of_constr var in (* Compute an assignment that dissatisfies the goal *) let redfun, _ = Redexpr.reduction_of_red_expr (Refiner.pf_env gl) Genredexpr.(CbvVm None) in let _, var = redfun Refiner.(pf_env gl) Refiner.(project gl) var in let var = EConstr.Unsafe.to_constr var in let rec to_list l = match decomp_term (Tacmach.project gl) l with | App (c, _) when c === (Lazy.force CoqList._nil) -> [] | App (c, [|_; h; t|]) when c === (Lazy.force CoqList._cons) -> if h === (Lazy.force Bool.trueb) then (true :: to_list t) else if h === (Lazy.force Bool.falseb) then (false :: to_list t) else invalid_arg "to_list" | _ -> invalid_arg "to_list" in let concat sep = function | [] -> mt () | h :: t -> let rec aux = function | [] -> mt () | x :: t -> (sep ++ x ++ aux t) in h ++ aux t in let msg = try let var = to_list var in let assign = List.combine penv var in let map_msg (key, v) = let b = if v then str "true" else str "false" in let sigma, env = Tacmach.project gl, Tacmach.pf_env gl in let term = Printer.pr_constr_env env sigma key in term ++ spc () ++ str ":=" ++ spc () ++ b in let assign = List.map map_msg assign in let l = str "[" ++ (concat (str ";" ++ spc ()) assign) ++ str "]" in str "Not a tautology:" ++ spc () ++ l with e when CErrors.noncritical e -> (str "Not a tautology") in Tacticals.tclFAIL 0 msg gl let try_unification env = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let eq = Lazy.force eq in let concl = EConstr.Unsafe.to_constr concl in let t = decomp_term (Tacmach.New.project gl) concl in match t with | App (c, [|typ; p; _|]) when c === eq -> (* should be an equality [@eq poly ?p (Cst false)] *) let tac = Tacticals.New.tclORELSE0 Tactics.reflexivity (Proofview.V82.tactic (print_counterexample p env)) in tac | _ -> let msg = str "Btauto: Internal error" in Tacticals.New.tclFAIL 0 msg end let tac = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let concl = EConstr.Unsafe.to_constr concl in let sigma = Tacmach.New.project gl in let eq = Lazy.force eq in let bool = Lazy.force Bool.typ in let t = decomp_term sigma concl in match t with | App (c, [|typ; tl; tr|]) when typ === bool && c === eq -> let env = Env.empty () in let fl = Bool.quote env sigma tl in let fr = Bool.quote env sigma tr in let env = Env.to_list env in let fl = reify env fl in let fr = reify env fr in let changed_gl = Constr.mkApp (c, [|typ; fl; fr|]) in let changed_gl = EConstr.of_constr changed_gl in Tacticals.New.tclTHENLIST [ Tactics.change_concl changed_gl; Tactics.apply (EConstr.of_constr (Lazy.force soundness)); Tactics.normalise_vm_in_concl; try_unification env ] | _ -> let msg = str "Cannot recognize a boolean equality" in Tacticals.New.tclFAIL 0 msg end end coq-8.11.0/plugins/btauto/Reflect.v0000644000175000017500000003332013612664533017021 0ustar treinentreinenRequire Import Bool DecidableClass Algebra Ring PArith Omega. Section Bool. (* Boolean formulas and their evaluations *) Inductive formula := | formula_var : positive -> formula | formula_btm : formula | formula_top : formula | formula_cnj : formula -> formula -> formula | formula_dsj : formula -> formula -> formula | formula_neg : formula -> formula | formula_xor : formula -> formula -> formula | formula_ifb : formula -> formula -> formula -> formula. Fixpoint formula_eval var f := match f with | formula_var x => list_nth x var false | formula_btm => false | formula_top => true | formula_cnj fl fr => (formula_eval var fl) && (formula_eval var fr) | formula_dsj fl fr => (formula_eval var fl) || (formula_eval var fr) | formula_neg f => negb (formula_eval var f) | formula_xor fl fr => xorb (formula_eval var fl) (formula_eval var fr) | formula_ifb fc fl fr => if formula_eval var fc then formula_eval var fl else formula_eval var fr end. End Bool. (* Translation of formulas into polynomials *) Section Translation. (* This is straightforward. *) Fixpoint poly_of_formula f := match f with | formula_var x => Poly (Cst false) x (Cst true) | formula_btm => Cst false | formula_top => Cst true | formula_cnj fl fr => let pl := poly_of_formula fl in let pr := poly_of_formula fr in poly_mul pl pr | formula_dsj fl fr => let pl := poly_of_formula fl in let pr := poly_of_formula fr in poly_add (poly_add pl pr) (poly_mul pl pr) | formula_neg f => poly_add (Cst true) (poly_of_formula f) | formula_xor fl fr => poly_add (poly_of_formula fl) (poly_of_formula fr) | formula_ifb fc fl fr => let pc := poly_of_formula fc in let pl := poly_of_formula fl in let pr := poly_of_formula fr in poly_add pr (poly_add (poly_mul pc pl) (poly_mul pc pr)) end. Opaque poly_add. (* Compatibility of translation wrt evaluation *) Lemma poly_of_formula_eval_compat : forall var f, eval var (poly_of_formula f) = formula_eval var f. Proof. intros var f; induction f; simpl poly_of_formula; simpl formula_eval; auto. now simpl; match goal with [ |- ?t = ?u ] => destruct u; reflexivity end. rewrite poly_mul_compat, IHf1, IHf2; ring. repeat rewrite poly_add_compat. rewrite poly_mul_compat; try_rewrite. now match goal with [ |- ?t = ?x || ?y ] => destruct x; destruct y; reflexivity end. rewrite poly_add_compat; try_rewrite. now match goal with [ |- ?t = negb ?x ] => destruct x; reflexivity end. rewrite poly_add_compat; congruence. rewrite ?poly_add_compat, ?poly_mul_compat; try_rewrite. match goal with [ |- ?t = if ?b1 then ?b2 else ?b3 ] => destruct b1; destruct b2; destruct b3; reflexivity end. Qed. Hint Extern 5 => change 0 with (min 0 0) : core. Local Hint Resolve poly_add_valid_compat poly_mul_valid_compat : core. Local Hint Constructors valid : core. Hint Extern 5 => zify; omega : core. (* Compatibility with validity *) Lemma poly_of_formula_valid_compat : forall f, exists n, valid n (poly_of_formula f). Proof. intros f; induction f; simpl. + exists (Pos.succ p); constructor; intuition; inversion H. + exists 1%positive; auto. + exists 1%positive; auto. + destruct IHf1 as [n1 Hn1]; destruct IHf2 as [n2 Hn2]; exists (Pos.max n1 n2); auto. + destruct IHf1 as [n1 Hn1]; destruct IHf2 as [n2 Hn2]; exists (Pos.max (Pos.max n1 n2) (Pos.max n1 n2)); auto. + destruct IHf as [n Hn]; exists (Pos.max 1 n); auto. + destruct IHf1 as [n1 Hn1]; destruct IHf2 as [n2 Hn2]; exists (Pos.max n1 n2); auto. + destruct IHf1 as [n1 Hn1]; destruct IHf2 as [n2 Hn2]; destruct IHf3 as [n3 Hn3]; eexists; eauto. Qed. (* The soundness lemma ; alas not complete! *) Lemma poly_of_formula_sound : forall fl fr var, poly_of_formula fl = poly_of_formula fr -> formula_eval var fl = formula_eval var fr. Proof. intros fl fr var Heq. repeat rewrite <- poly_of_formula_eval_compat. rewrite Heq; reflexivity. Qed. End Translation. Section Completeness. (* Lemma reduce_poly_of_formula_simpl : forall fl fr var, simpl_eval (var_of_list var) (reduce (poly_of_formula fl)) = simpl_eval (var_of_list var) (reduce (poly_of_formula fr)) -> formula_eval var fl = formula_eval var fr. Proof. intros fl fr var Hrw. do 2 rewrite <- poly_of_formula_eval_compat. destruct (poly_of_formula_valid_compat fl) as [nl Hl]. destruct (poly_of_formula_valid_compat fr) as [nr Hr]. rewrite <- (reduce_eval_compat nl (poly_of_formula fl)); [|assumption]. rewrite <- (reduce_eval_compat nr (poly_of_formula fr)); [|assumption]. do 2 rewrite <- eval_simpl_eval_compat; assumption. Qed. *) (* Soundness of the method ; immediate *) Lemma reduce_poly_of_formula_sound : forall fl fr var, reduce (poly_of_formula fl) = reduce (poly_of_formula fr) -> formula_eval var fl = formula_eval var fr. Proof. intros fl fr var Heq. repeat rewrite <- poly_of_formula_eval_compat. destruct (poly_of_formula_valid_compat fl) as [nl Hl]. destruct (poly_of_formula_valid_compat fr) as [nr Hr]. rewrite <- (reduce_eval_compat nl (poly_of_formula fl)); auto. rewrite <- (reduce_eval_compat nr (poly_of_formula fr)); auto. rewrite Heq; reflexivity. Qed. Definition make_last {A} n (x def : A) := Pos.peano_rect (fun _ => list A) (cons x nil) (fun _ F => cons def F) n. (* Replace the nth element of a list *) Fixpoint list_replace l n b := match l with | nil => make_last n b false | cons a l => Pos.peano_rect _ (cons b l) (fun n _ => cons a (list_replace l n b)) n end. (** Extract a non-null witness from a polynomial *) Existing Instance Decidable_null. Fixpoint boolean_witness p := match p with | Cst c => nil | Poly p i q => if decide (null p) then let var := boolean_witness q in list_replace var i true else let var := boolean_witness p in list_replace var i false end. Lemma list_nth_base : forall A (def : A) l, list_nth 1 l def = match l with nil => def | cons x _ => x end. Proof. intros A def l; unfold list_nth. rewrite Pos.peano_rect_base; reflexivity. Qed. Lemma list_nth_succ : forall A n (def : A) l, list_nth (Pos.succ n) l def = match l with nil => def | cons _ l => list_nth n l def end. Proof. intros A def l; unfold list_nth. rewrite Pos.peano_rect_succ; reflexivity. Qed. Lemma list_nth_nil : forall A n (def : A), list_nth n nil def = def. Proof. intros A n def; induction n using Pos.peano_rect. + rewrite list_nth_base; reflexivity. + rewrite list_nth_succ; reflexivity. Qed. Lemma make_last_nth_1 : forall A n i x def, i <> n -> list_nth i (@make_last A n x def) def = def. Proof. intros A n; induction n using Pos.peano_rect; intros i x def Hd; unfold make_last; simpl. + induction i using Pos.peano_case; [elim Hd; reflexivity|]. rewrite list_nth_succ, list_nth_nil; reflexivity. + unfold make_last; rewrite Pos.peano_rect_succ; fold (make_last n x def). induction i using Pos.peano_case. - rewrite list_nth_base; reflexivity. - rewrite list_nth_succ; apply IHn; zify; omega. Qed. Lemma make_last_nth_2 : forall A n x def, list_nth n (@make_last A n x def) def = x. Proof. intros A n; induction n using Pos.peano_rect; intros x def; simpl. + reflexivity. + unfold make_last; rewrite Pos.peano_rect_succ; fold (make_last n x def). rewrite list_nth_succ; auto. Qed. Lemma list_replace_nth_1 : forall var i j x, i <> j -> list_nth i (list_replace var j x) false = list_nth i var false. Proof. intros var; induction var; intros i j x Hd; simpl. + rewrite make_last_nth_1, list_nth_nil; auto. + induction j using Pos.peano_rect. - rewrite Pos.peano_rect_base. induction i using Pos.peano_rect; [now elim Hd; auto|]. rewrite 2list_nth_succ; reflexivity. - rewrite Pos.peano_rect_succ. induction i using Pos.peano_rect. { rewrite 2list_nth_base; reflexivity. } { rewrite 2list_nth_succ; apply IHvar; zify; omega. } Qed. Lemma list_replace_nth_2 : forall var i x, list_nth i (list_replace var i x) false = x. Proof. intros var; induction var; intros i x; simpl. + now apply make_last_nth_2. + induction i using Pos.peano_rect. - rewrite Pos.peano_rect_base, list_nth_base; reflexivity. - rewrite Pos.peano_rect_succ, list_nth_succ; auto. Qed. (* The witness is correct only if the polynomial is linear *) Lemma boolean_witness_nonzero : forall k p, linear k p -> ~ null p -> eval (boolean_witness p) p = true. Proof. intros k p Hl Hp; induction Hl; simpl. destruct c; [reflexivity|elim Hp; now constructor]. case_decide. rewrite eval_null_zero; [|assumption]; rewrite list_replace_nth_2; simpl. match goal with [ |- (if ?b then true else false) = true ] => assert (Hrw : b = true); [|rewrite Hrw; reflexivity] end. erewrite eval_suffix_compat; [now eauto| |now apply linear_valid_incl; eauto]. now intros j Hd; apply list_replace_nth_1; zify; omega. rewrite list_replace_nth_2, xorb_false_r. erewrite eval_suffix_compat; [now eauto| |now apply linear_valid_incl; eauto]. now intros j Hd; apply list_replace_nth_1; zify; omega. Qed. (* This should be better when using the [vm_compute] tactic instead of plain reflexivity. *) Lemma reduce_poly_of_formula_sound_alt : forall var fl fr, reduce (poly_add (poly_of_formula fl) (poly_of_formula fr)) = Cst false -> formula_eval var fl = formula_eval var fr. Proof. intros var fl fr Heq. repeat rewrite <- poly_of_formula_eval_compat. destruct (poly_of_formula_valid_compat fl) as [nl Hl]. destruct (poly_of_formula_valid_compat fr) as [nr Hr]. rewrite <- (reduce_eval_compat nl (poly_of_formula fl)); auto. rewrite <- (reduce_eval_compat nr (poly_of_formula fr)); auto. rewrite <- xorb_false_l; change false with (eval var (Cst false)). rewrite <- poly_add_compat, <- Heq. repeat rewrite poly_add_compat. rewrite (reduce_eval_compat nl); [|assumption]. rewrite (reduce_eval_compat (Pos.max nl nr)); [|apply poly_add_valid_compat; assumption]. rewrite (reduce_eval_compat nr); [|assumption]. rewrite poly_add_compat; ring. Qed. (* The completeness lemma *) (* Lemma reduce_poly_of_formula_complete : forall fl fr, reduce (poly_of_formula fl) <> reduce (poly_of_formula fr) -> {var | formula_eval var fl <> formula_eval var fr}. Proof. intros fl fr H. pose (p := poly_add (reduce (poly_of_formula fl)) (poly_opp (reduce (poly_of_formula fr)))). pose (var := boolean_witness p). exists var. intros Hc; apply (f_equal Z_of_bool) in Hc. assert (Hfl : linear 0 (reduce (poly_of_formula fl))). now destruct (poly_of_formula_valid_compat fl) as [n Hn]; apply (linear_le_compat n); [|now auto]; apply linear_reduce; auto. assert (Hfr : linear 0 (reduce (poly_of_formula fr))). now destruct (poly_of_formula_valid_compat fr) as [n Hn]; apply (linear_le_compat n); [|now auto]; apply linear_reduce; auto. repeat rewrite <- poly_of_formula_eval_compat in Hc. define (decide (null p)) b Hb; destruct b; tac_decide. now elim H; apply (null_sub_implies_eq 0 0); fold p; auto; apply linear_valid_incl; auto. elim (boolean_witness_nonzero 0 p); auto. unfold p; rewrite <- (min_id 0); apply poly_add_linear_compat; try apply poly_opp_linear_compat; now auto. unfold p at 2; rewrite poly_add_compat, poly_opp_compat. destruct (poly_of_formula_valid_compat fl) as [nl Hnl]. destruct (poly_of_formula_valid_compat fr) as [nr Hnr]. repeat erewrite reduce_eval_compat; eauto. fold var; rewrite Hc; ring. Defined. *) End Completeness. (* Reification tactics *) (* For reflexivity purposes, that would better be transparent *) Global Transparent decide poly_add. (* Ltac append_var x l k := match l with | nil => constr: (k, cons x l) | cons x _ => constr: (k, l) | cons ?y ?l => let ans := append_var x l (S k) in match ans with (?k, ?l) => constr: (k, cons y l) end end. Ltac build_formula t l := match t with | true => constr: (formula_top, l) | false => constr: (formula_btm, l) | ?fl && ?fr => match build_formula fl l with (?tl, ?l) => match build_formula fr l with (?tr, ?l) => constr: (formula_cnj tl tr, l) end end | ?fl || ?fr => match build_formula fl l with (?tl, ?l) => match build_formula fr l with (?tr, ?l) => constr: (formula_dsj tl tr, l) end end | negb ?f => match build_formula f l with (?t, ?l) => constr: (formula_neg t, l) end | _ => let ans := append_var t l 0 in match ans with (?k, ?l) => constr: (formula_var k, l) end end. (* Extract a counterexample from a polynomial and display it *) Ltac counterexample p l := let var := constr: (boolean_witness p) in let var := eval vm_compute in var in let rec print l vl := match l with | nil => idtac | cons ?x ?l => match vl with | nil => idtac x ":=" "false"; print l (@nil bool) | cons ?v ?vl => idtac x ":=" v; print l vl end end in idtac "Counter-example:"; print l var. Ltac btauto_reify := lazymatch goal with | [ |- @eq bool ?t ?u ] => lazymatch build_formula t (@nil bool) with | (?fl, ?l) => lazymatch build_formula u l with | (?fr, ?l) => change (formula_eval l fl = formula_eval l fr) end end | _ => fail "Cannot recognize a boolean equality" end. (* The long-awaited tactic *) Ltac btauto := lazymatch goal with | [ |- @eq bool ?t ?u ] => lazymatch build_formula t (@nil bool) with | (?fl, ?l) => lazymatch build_formula u l with | (?fr, ?l) => change (formula_eval l fl = formula_eval l fr); apply reduce_poly_of_formula_sound_alt; vm_compute; (reflexivity || fail "Not a tautology") end end | _ => fail "Cannot recognize a boolean equality" end. *) Register formula_var as plugins.btauto.f_var. Register formula_btm as plugins.btauto.f_btm. Register formula_top as plugins.btauto.f_top. Register formula_cnj as plugins.btauto.f_cnj. Register formula_dsj as plugins.btauto.f_dsj. Register formula_neg as plugins.btauto.f_neg. Register formula_xor as plugins.btauto.f_xor. Register formula_ifb as plugins.btauto.f_ifb. Register formula_eval as plugins.btauto.eval. Register boolean_witness as plugins.btauto.witness. Register reduce_poly_of_formula_sound_alt as plugins.btauto.soundness. coq-8.11.0/plugins/btauto/Btauto.v0000644000175000017500000000010413612664533016665 0ustar treinentreinenRequire Import Algebra Reflect. Declare ML Module "btauto_plugin". coq-8.11.0/plugins/btauto/refl_btauto.mli0000644000175000017500000000133413612664533020257 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* apply andb_true_iff in H; destruct H | |- ?P && ?Q = true => apply <- andb_true_iff; split end. Arguments decide P /H. Hint Extern 5 => progress bool : core. Ltac define t x H := set (x := t) in *; assert (H : t = x) by reflexivity; clearbody x. Lemma Decidable_sound : forall P (H : Decidable P), decide P = true -> P. Proof. intros P H Hp; apply -> Decidable_spec; assumption. Qed. Lemma Decidable_complete : forall P (H : Decidable P), P -> decide P = true. Proof. intros P H Hp; apply <- Decidable_spec; assumption. Qed. Lemma Decidable_sound_alt : forall P (H : Decidable P), ~ P -> decide P = false. Proof. intros P [wit spec] Hd; destruct wit; simpl; tauto. Qed. Lemma Decidable_complete_alt : forall P (H : Decidable P), decide P = false -> ~ P. Proof. intros P [wit spec] Hd Hc; simpl in *; intuition congruence. Qed. Ltac try_rewrite := repeat match goal with | [ H : ?P |- _ ] => rewrite H end. (* We opacify here decide for proofs, and will make it transparent for reflexive tactics later on. *) Global Opaque decide. Ltac tac_decide := match goal with | [ H : @decide ?P ?D = true |- _ ] => apply (@Decidable_sound P D) in H | [ H : @decide ?P ?D = false |- _ ] => apply (@Decidable_complete_alt P D) in H | [ |- @decide ?P ?D = true ] => apply (@Decidable_complete P D) | [ |- @decide ?P ?D = false ] => apply (@Decidable_sound_alt P D) | [ |- negb ?b = true ] => apply negb_true_iff | [ |- negb ?b = false ] => apply negb_false_iff | [ H : negb ?b = true |- _ ] => apply negb_true_iff in H | [ H : negb ?b = false |- _ ] => apply negb_false_iff in H end. Ltac try_decide := repeat tac_decide. Ltac make_decide P := match goal with | [ |- context [@decide P ?D] ] => let b := fresh "b" in let H := fresh "H" in define (@decide P D) b H; destruct b; try_decide | [ X : context [@decide P ?D] |- _ ] => let b := fresh "b" in let H := fresh "H" in define (@decide P D) b H; destruct b; try_decide end. Ltac case_decide := match goal with | [ |- context [@decide ?P ?D] ] => let b := fresh "b" in let H := fresh "H" in define (@decide P D) b H; destruct b; try_decide | [ X : context [@decide ?P ?D] |- _ ] => let b := fresh "b" in let H := fresh "H" in define (@decide P D) b H; destruct b; try_decide | [ |- context [Pos.compare ?x ?y] ] => destruct (Pos.compare_spec x y); try lia | [ X : context [Pos.compare ?x ?y] |- _ ] => destruct (Pos.compare_spec x y); try lia end. Section Definitions. (** * Global, inductive definitions. *) (** A Horner polynomial is either a constant, or a product P × (i + Q), where i is a variable. *) Inductive poly := | Cst : bool -> poly | Poly : poly -> positive -> poly -> poly. (* TODO: We should use [positive] instead of [nat] to encode variables, for efficiency purpose. *) Inductive null : poly -> Prop := | null_intro : null (Cst false). (** Polynomials satisfy a uniqueness condition whenever they are valid. A polynomial [p] satisfies [valid n p] whenever it is well-formed and each of its variable indices is < [n]. *) Inductive valid : positive -> poly -> Prop := | valid_cst : forall k c, valid k (Cst c) | valid_poly : forall k p i q, Pos.lt i k -> ~ null q -> valid i p -> valid (Pos.succ i) q -> valid k (Poly p i q). (** Linear polynomials are valid polynomials in which every variable appears at most once. *) Inductive linear : positive -> poly -> Prop := | linear_cst : forall k c, linear k (Cst c) | linear_poly : forall k p i q, Pos.lt i k -> ~ null q -> linear i p -> linear i q -> linear k (Poly p i q). End Definitions. Section Computational. Program Instance Decidable_PosEq : forall (p q : positive), Decidable (p = q) := { Decidable_witness := Pos.eqb p q }. Next Obligation. apply Pos.eqb_eq. Qed. Program Instance Decidable_PosLt : forall p q, Decidable (Pos.lt p q) := { Decidable_witness := Pos.ltb p q }. Next Obligation. apply Pos.ltb_lt. Qed. Program Instance Decidable_PosLe : forall p q, Decidable (Pos.le p q) := { Decidable_witness := Pos.leb p q }. Next Obligation. apply Pos.leb_le. Qed. (** * The core reflexive part. *) Hint Constructors valid : core. Fixpoint beq_poly pl pr := match pl with | Cst cl => match pr with | Cst cr => decide (cl = cr) | Poly _ _ _ => false end | Poly pl il ql => match pr with | Cst _ => false | Poly pr ir qr => decide (il = ir) && beq_poly pl pr && beq_poly ql qr end end. (* We could do that with [decide equality] but dependency in proofs is heavy *) Program Instance Decidable_eq_poly : forall (p q : poly), Decidable (eq p q) := { Decidable_witness := beq_poly p q }. Next Obligation. split. revert q; induction p; intros [] ?; simpl in *; bool; try_decide; f_equal; first [intuition congruence|auto]. revert q; induction p; intros [] Heq; simpl in *; bool; try_decide; intuition; try injection Heq; first[congruence|intuition]. Qed. Program Instance Decidable_null : forall p, Decidable (null p) := { Decidable_witness := match p with Cst false => true | _ => false end }. Next Obligation. split. destruct p as [[]|]; first [discriminate|constructor]. inversion 1; trivial. Qed. Definition list_nth {A} p (l : list A) def := Pos.peano_rect (fun _ => list A -> A) (fun l => match l with nil => def | cons t l => t end) (fun _ F l => match l with nil => def | cons t l => F l end) p l. Fixpoint eval var (p : poly) := match p with | Cst c => c | Poly p i q => let vi := list_nth i var false in xorb (eval var p) (andb vi (eval var q)) end. Fixpoint valid_dec k p := match p with | Cst c => true | Poly p i q => negb (decide (null q)) && decide (i < k)%positive && valid_dec i p && valid_dec (Pos.succ i) q end. Program Instance Decidable_valid : forall n p, Decidable (valid n p) := { Decidable_witness := valid_dec n p }. Next Obligation. split. revert n; induction p; unfold valid_dec in *; intuition; bool; try_decide; auto. intros H; induction H; unfold valid_dec in *; bool; try_decide; auto. Qed. (** Basic algebra *) (* Addition of polynomials *) Fixpoint poly_add pl {struct pl} := match pl with | Cst cl => fix F pr := match pr with | Cst cr => Cst (xorb cl cr) | Poly pr ir qr => Poly (F pr) ir qr end | Poly pl il ql => fix F pr {struct pr} := match pr with | Cst cr => Poly (poly_add pl pr) il ql | Poly pr ir qr => match Pos.compare il ir with | Eq => let qs := poly_add ql qr in (* Ensure validity *) if decide (null qs) then poly_add pl pr else Poly (poly_add pl pr) il qs | Gt => Poly (poly_add pl (Poly pr ir qr)) il ql | Lt => Poly (F pr) ir qr end end end. (* Multiply a polynomial by a constant *) Fixpoint poly_mul_cst v p := match p with | Cst c => Cst (andb c v) | Poly p i q => let r := poly_mul_cst v q in (* Ensure validity *) if decide (null r) then poly_mul_cst v p else Poly (poly_mul_cst v p) i r end. (* Multiply a polynomial by a monomial *) Fixpoint poly_mul_mon k p := match p with | Cst c => if decide (null p) then p else Poly (Cst false) k p | Poly p i q => if decide (i <= k)%positive then Poly (Cst false) k (Poly p i q) else Poly (poly_mul_mon k p) i (poly_mul_mon k q) end. (* Multiplication of polynomials *) Fixpoint poly_mul pl {struct pl} := match pl with | Cst cl => poly_mul_cst cl | Poly pl il ql => fun pr => (* Multiply by a factor *) let qs := poly_mul ql pr in (* Ensure validity *) if decide (null qs) then poly_mul pl pr else poly_add (poly_mul pl pr) (poly_mul_mon il qs) end. (** Quotienting a polynomial by the relation X_i^2 ~ X_i *) (* Remove the multiple occurrences of monomials x_k *) Fixpoint reduce_aux k p := match p with | Cst c => Cst c | Poly p i q => if decide (i = k) then poly_add (reduce_aux k p) (reduce_aux k q) else let qs := reduce_aux i q in (* Ensure validity *) if decide (null qs) then (reduce_aux k p) else Poly (reduce_aux k p) i qs end. (* Rewrite any x_k ^ {n + 1} to x_k *) Fixpoint reduce p := match p with | Cst c => Cst c | Poly p i q => let qs := reduce_aux i q in (* Ensure validity *) if decide (null qs) then reduce p else Poly (reduce p) i qs end. End Computational. Section Validity. (* Decision procedure of validity *) Hint Constructors valid linear : core. Lemma valid_le_compat : forall k l p, valid k p -> (k <= l)%positive -> valid l p. Proof. intros k l p H Hl; induction H; constructor; eauto. now eapply Pos.lt_le_trans; eassumption. Qed. Lemma linear_le_compat : forall k l p, linear k p -> (k <= l)%positive -> linear l p. Proof. intros k l p H; revert l; induction H; constructor; eauto; lia. Qed. Lemma linear_valid_incl : forall k p, linear k p -> valid k p. Proof. intros k p H; induction H; constructor; auto. eapply valid_le_compat; eauto; lia. Qed. End Validity. Section Evaluation. (* Useful simple properties *) Lemma eval_null_zero : forall p var, null p -> eval var p = false. Proof. intros p var []; reflexivity. Qed. Lemma eval_extensional_eq_compat : forall p var1 var2, (forall x, list_nth x var1 false = list_nth x var2 false) -> eval var1 p = eval var2 p. Proof. intros p var1 var2 H; induction p; simpl; try_rewrite; auto. Qed. Lemma eval_suffix_compat : forall k p var1 var2, (forall i, (i < k)%positive -> list_nth i var1 false = list_nth i var2 false) -> valid k p -> eval var1 p = eval var2 p. Proof. intros k p var1 var2 Hvar Hv; revert var1 var2 Hvar. induction Hv; intros var1 var2 Hvar; simpl; [now auto|]. rewrite Hvar; [|now auto]; erewrite (IHHv1 var1 var2). + erewrite (IHHv2 var1 var2); [ring|]. intros; apply Hvar; zify; omega. + intros; apply Hvar; zify; omega. Qed. End Evaluation. Section Algebra. (* Compatibility with evaluation *) Lemma poly_add_compat : forall pl pr var, eval var (poly_add pl pr) = xorb (eval var pl) (eval var pr). Proof. intros pl; induction pl; intros pr var; simpl. + induction pr; simpl; auto; solve [try_rewrite; ring]. + induction pr; simpl; auto; try solve [try_rewrite; simpl; ring]. destruct (Pos.compare_spec p p0); repeat case_decide; simpl; first [try_rewrite; ring|idtac]. try_rewrite; ring_simplify; repeat rewrite xorb_assoc. match goal with [ |- context [xorb (andb ?b1 ?b2) (andb ?b1 ?b3)] ] => replace (xorb (andb b1 b2) (andb b1 b3)) with (andb b1 (xorb b2 b3)) by ring end. rewrite <- IHpl2. match goal with [ H : null ?p |- _ ] => rewrite (eval_null_zero _ _ H) end; ring. simpl; rewrite IHpl1; simpl; ring. Qed. Lemma poly_mul_cst_compat : forall v p var, eval var (poly_mul_cst v p) = andb v (eval var p). Proof. intros v p; induction p; intros var; simpl; [ring|]. case_decide; simpl; try_rewrite; [ring_simplify|ring]. replace (v && list_nth p2 var false && eval var p3) with (list_nth p2 var false && (v && eval var p3)) by ring. rewrite <- IHp2; inversion H; simpl; ring. Qed. Lemma poly_mul_mon_compat : forall i p var, eval var (poly_mul_mon i p) = (list_nth i var false && eval var p). Proof. intros i p var; induction p; simpl; case_decide; simpl; try_rewrite; try ring. inversion H; ring. match goal with [ |- ?u = ?t ] => set (x := t); destruct x; reflexivity end. match goal with [ |- ?u = ?t ] => set (x := t); destruct x; reflexivity end. Qed. Lemma poly_mul_compat : forall pl pr var, eval var (poly_mul pl pr) = andb (eval var pl) (eval var pr). Proof. intros pl; induction pl; intros pr var; simpl. apply poly_mul_cst_compat. case_decide; simpl. rewrite IHpl1; ring_simplify. replace (eval var pr && list_nth p var false && eval var pl2) with (list_nth p var false && (eval var pl2 && eval var pr)) by ring. now rewrite <- IHpl2; inversion H; simpl; ring. rewrite poly_add_compat, poly_mul_mon_compat, IHpl1, IHpl2; ring. Qed. Hint Extern 5 => match goal with | [ |- (Pos.max ?x ?y <= ?z)%positive ] => apply Pos.max_case_strong; intros; lia | [ |- (?z <= Pos.max ?x ?y)%positive ] => apply Pos.max_case_strong; intros; lia | [ |- (Pos.max ?x ?y < ?z)%positive ] => apply Pos.max_case_strong; intros; lia | [ |- (?z < Pos.max ?x ?y)%positive ] => apply Pos.max_case_strong; intros; lia | _ => zify; omega end : core. Hint Resolve Pos.le_max_r Pos.le_max_l : core. Hint Constructors valid linear : core. (* Compatibility of validity w.r.t algebraic operations *) Lemma poly_add_valid_compat : forall kl kr pl pr, valid kl pl -> valid kr pr -> valid (Pos.max kl kr) (poly_add pl pr). Proof. intros kl kr pl pr Hl Hr; revert kr pr Hr; induction Hl; intros kr pr Hr; simpl. { eapply valid_le_compat; [clear k|apply Pos.le_max_r]. now induction Hr; auto. } { assert (Hle : (Pos.max (Pos.succ i) kr <= Pos.max k kr)%positive) by auto. apply (valid_le_compat (Pos.max (Pos.succ i) kr)); [|assumption]. clear - IHHl1 IHHl2 Hl2 Hr H0; induction Hr. constructor; auto. now rewrite <- (Pos.max_id i); intuition. destruct (Pos.compare_spec i i0); subst; try case_decide; repeat (constructor; intuition). + apply (valid_le_compat (Pos.max i0 i0)); [now auto|]; rewrite Pos.max_id; auto. + apply (valid_le_compat (Pos.max i0 i0)); [now auto|]; rewrite Pos.max_id; lia. + apply (valid_le_compat (Pos.max (Pos.succ i0) (Pos.succ i0))); [now auto|]; rewrite Pos.max_id; lia. + apply (valid_le_compat (Pos.max (Pos.succ i) i0)); intuition. + apply (valid_le_compat (Pos.max i (Pos.succ i0))); intuition. } Qed. Lemma poly_mul_cst_valid_compat : forall k v p, valid k p -> valid k (poly_mul_cst v p). Proof. intros k v p H; induction H; simpl; [now auto|]. case_decide; [|now auto]. eapply (valid_le_compat i); [now auto|lia]. Qed. Lemma poly_mul_mon_null_compat : forall i p, null (poly_mul_mon i p) -> null p. Proof. intros i p; induction p; simpl; case_decide; simpl; inversion 1; intuition. Qed. Lemma poly_mul_mon_valid_compat : forall k i p, valid k p -> valid (Pos.max (Pos.succ i) k) (poly_mul_mon i p). Proof. intros k i p H; induction H; simpl poly_mul_mon; case_decide; intuition. + apply (valid_le_compat (Pos.succ i)); auto; constructor; intuition. - match goal with [ H : null ?p |- _ ] => solve[inversion H] end. + apply (valid_le_compat k); auto; constructor; intuition. - assert (X := poly_mul_mon_null_compat); intuition eauto. - enough (Pos.max (Pos.succ i) i0 = i0) as <-; intuition. - enough (Pos.max (Pos.succ i) (Pos.succ i0) = Pos.succ i0) as <-; intuition. Qed. Lemma poly_mul_valid_compat : forall kl kr pl pr, valid kl pl -> valid kr pr -> valid (Pos.max kl kr) (poly_mul pl pr). Proof. intros kl kr pl pr Hl Hr; revert kr pr Hr. induction Hl; intros kr pr Hr; simpl. + apply poly_mul_cst_valid_compat; auto. apply (valid_le_compat kr); now auto. + apply (valid_le_compat (Pos.max (Pos.max i kr) (Pos.max (Pos.succ i) (Pos.max (Pos.succ i) kr)))). - case_decide. { apply (valid_le_compat (Pos.max i kr)); auto. } { apply poly_add_valid_compat; auto. now apply poly_mul_mon_valid_compat; intuition. } - repeat apply Pos.max_case_strong; zify; omega. Qed. (* Compatibility of linearity wrt to linear operations *) Lemma poly_add_linear_compat : forall kl kr pl pr, linear kl pl -> linear kr pr -> linear (Pos.max kl kr) (poly_add pl pr). Proof. intros kl kr pl pr Hl; revert kr pr; induction Hl; intros kr pr Hr; simpl. + apply (linear_le_compat kr); [|apply Pos.max_case_strong; zify; omega]. now induction Hr; constructor; auto. + apply (linear_le_compat (Pos.max kr (Pos.succ i))); [|now auto]. induction Hr; simpl. - constructor; auto. replace i with (Pos.max i i) by (apply Pos.max_id); intuition. - destruct (Pos.compare_spec i i0); subst; try case_decide; repeat (constructor; intuition). { apply (linear_le_compat (Pos.max i0 i0)); [now auto|]; rewrite Pos.max_id; auto. } { apply (linear_le_compat (Pos.max i0 i0)); [now auto|]; rewrite Pos.max_id; auto. } { apply (linear_le_compat (Pos.max i0 i0)); [now auto|]; rewrite Pos.max_id; auto. } { apply (linear_le_compat (Pos.max i0 (Pos.succ i))); intuition. } { apply (linear_le_compat (Pos.max i (Pos.succ i0))); intuition. } Qed. End Algebra. Section Reduce. (* A stronger version of the next lemma *) Lemma reduce_aux_eval_compat : forall k p var, valid (Pos.succ k) p -> (list_nth k var false && eval var (reduce_aux k p) = list_nth k var false && eval var p). Proof. intros k p var; revert k; induction p; intros k Hv; simpl; auto. inversion Hv; case_decide; subst. + rewrite poly_add_compat; ring_simplify. specialize (IHp1 k); specialize (IHp2 k). destruct (list_nth k var false); ring_simplify; [|now auto]. rewrite <- (andb_true_l (eval var p1)), <- (andb_true_l (eval var p3)). rewrite <- IHp2; auto; rewrite <- IHp1; [ring|]. apply (valid_le_compat k); [now auto|zify; omega]. + remember (list_nth k var false) as b; destruct b; ring_simplify; [|now auto]. case_decide; simpl. - rewrite <- (IHp2 p2); [inversion H|now auto]; simpl. replace (eval var p1) with (list_nth k var false && eval var p1) by (rewrite <- Heqb; ring); rewrite <- (IHp1 k). { rewrite <- Heqb; ring. } { apply (valid_le_compat p2); [auto|zify; omega]. } - rewrite (IHp2 p2); [|now auto]. replace (eval var p1) with (list_nth k var false && eval var p1) by (rewrite <- Heqb; ring). rewrite <- (IHp1 k); [rewrite <- Heqb; ring|]. apply (valid_le_compat p2); [auto|zify; omega]. Qed. (* Reduction preserves evaluation by boolean assignations *) Lemma reduce_eval_compat : forall k p var, valid k p -> eval var (reduce p) = eval var p. Proof. intros k p var H; induction H; simpl; auto. case_decide; try_rewrite; simpl. + rewrite <- reduce_aux_eval_compat; auto; inversion H3; simpl; ring. + repeat rewrite reduce_aux_eval_compat; try_rewrite; now auto. Qed. Lemma reduce_aux_le_compat : forall k l p, valid k p -> (k <= l)%positive -> reduce_aux l p = reduce_aux k p. Proof. intros k l p; revert k l; induction p; intros k l H Hle; simpl; auto. inversion H; subst; repeat case_decide; subst; try (exfalso; zify; omega). + apply IHp1; [|now auto]; eapply valid_le_compat; [eauto|zify; omega]. + f_equal; apply IHp1; auto. now eapply valid_le_compat; [eauto|zify; omega]. Qed. (* Reduce projects valid polynomials into linear ones *) Lemma linear_reduce_aux : forall i p, valid (Pos.succ i) p -> linear i (reduce_aux i p). Proof. intros i p; revert i; induction p; intros i Hp; simpl. + constructor. + inversion Hp; subst; case_decide; subst. - rewrite <- (Pos.max_id i) at 1; apply poly_add_linear_compat. { apply IHp1; eapply valid_le_compat; [eassumption|zify; omega]. } { intuition. } - case_decide. { apply IHp1; eapply valid_le_compat; [eauto|zify; omega]. } { constructor; try (zify; omega); auto. erewrite (reduce_aux_le_compat p2); [|assumption|zify; omega]. apply IHp1; eapply valid_le_compat; [eauto|]; zify; omega. } Qed. Lemma linear_reduce : forall k p, valid k p -> linear k (reduce p). Proof. intros k p H; induction H; simpl. + now constructor. + case_decide. - eapply linear_le_compat; [eauto|zify; omega]. - constructor; auto. apply linear_reduce_aux; auto. Qed. End Reduce. coq-8.11.0/plugins/btauto/btauto_plugin.mlpack0000644000175000017500000000002513612664533021307 0ustar treinentreinenRefl_btauto G_btauto coq-8.11.0/plugins/btauto/g_btauto.mlg0000644000175000017500000000144313612664533017554 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Refl_btauto.Btauto.tac } END coq-8.11.0/plugins/btauto/plugin_base.dune0000644000175000017500000000020213612664533020404 0ustar treinentreinen(library (name btauto_plugin) (public_name coq.plugins.btauto) (synopsis "Coq's btauto plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/rtauto/0000755000175000017500000000000013612664533015265 5ustar treinentreinencoq-8.11.0/plugins/rtauto/refl_tauto.mli0000644000175000017500000000213213612664533020132 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Evd.evar_map -> atom_env -> EConstr.types -> Proof_search.form val make_hyps : Environ.env -> Evd.evar_map -> atom_env -> EConstr.types list -> EConstr.named_context -> (Names.Id.t Context.binder_annot * Proof_search.form) list val rtauto_tac : unit Proofview.tactic coq-8.11.0/plugins/rtauto/Bintree.v0000644000175000017500000002300313612664533017042 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (p ?= q) = Gt. Proof. intros. rewrite <- Pos.compare_succ_succ. now apply Pos.lt_gt, Pos.lt_lt_succ, Pos.gt_lt. Qed. Lemma Psucc_Gt : forall p, (Pos.succ p ?= p) = Gt. Proof. intros. apply Pos.lt_gt, Pos.lt_succ_diag_r. Qed. Fixpoint Lget (A:Set) (n:nat) (l:list A) {struct l}:option A := match l with nil => None | x::q => match n with O => Some x | S m => Lget A m q end end . Arguments Lget [A] n l. Lemma map_app : forall (A B:Set) (f:A -> B) l m, List.map f (l ++ m) = List.map f l ++ List.map f m. induction l. reflexivity. simpl. intro m ; apply f_equal;apply IHl. Qed. Lemma length_map : forall (A B:Set) (f:A -> B) l, length (List.map f l) = length l. induction l. reflexivity. simpl; apply f_equal;apply IHl. Qed. Lemma Lget_map : forall (A B:Set) (f:A -> B) i l, Lget i (List.map f l) = match Lget i l with Some a => Some (f a) | None => None end. induction i;intros [ | x l ] ;trivial. simpl;auto. Qed. Lemma Lget_app : forall (A:Set) (a:A) l i, Lget i (l ++ a :: nil) = if Arith.EqNat.beq_nat i (length l) then Some a else Lget i l. Proof. induction l;simpl Lget;simpl length. intros [ | i];simpl;reflexivity. intros [ | i];simpl. reflexivity. auto. Qed. Lemma Lget_app_Some : forall (A:Set) l delta i (a: A), Lget i l = Some a -> Lget i (l ++ delta) = Some a. induction l;destruct i;simpl;try congruence;auto. Qed. Inductive Poption {A} : Type:= PSome : A -> Poption | PNone : Poption. Arguments Poption : clear implicits. Inductive Tree {A} : Type := Tempty : Tree | Branch0 : Tree -> Tree -> Tree | Branch1 : A -> Tree -> Tree -> Tree. Arguments Tree : clear implicits. Section Store. Variable A:Type. Notation Poption := (Poption A). Notation Tree := (Tree A). Fixpoint Tget (p:positive) (T:Tree) {struct p} : Poption := match T with Tempty => PNone | Branch0 T1 T2 => match p with xI pp => Tget pp T2 | xO pp => Tget pp T1 | xH => PNone end | Branch1 a T1 T2 => match p with xI pp => Tget pp T2 | xO pp => Tget pp T1 | xH => PSome a end end. Fixpoint Tadd (p:positive) (a:A) (T:Tree) {struct p}: Tree := match T with | Tempty => match p with | xI pp => Branch0 Tempty (Tadd pp a Tempty) | xO pp => Branch0 (Tadd pp a Tempty) Tempty | xH => Branch1 a Tempty Tempty end | Branch0 T1 T2 => match p with | xI pp => Branch0 T1 (Tadd pp a T2) | xO pp => Branch0 (Tadd pp a T1) T2 | xH => Branch1 a T1 T2 end | Branch1 b T1 T2 => match p with | xI pp => Branch1 b T1 (Tadd pp a T2) | xO pp => Branch1 b (Tadd pp a T1) T2 | xH => Branch1 a T1 T2 end end. Definition mkBranch0 (T1 T2:Tree) := match T1,T2 with Tempty ,Tempty => Tempty | _,_ => Branch0 T1 T2 end. Fixpoint Tremove (p:positive) (T:Tree) {struct p}: Tree := match T with | Tempty => Tempty | Branch0 T1 T2 => match p with | xI pp => mkBranch0 T1 (Tremove pp T2) | xO pp => mkBranch0 (Tremove pp T1) T2 | xH => T end | Branch1 b T1 T2 => match p with | xI pp => Branch1 b T1 (Tremove pp T2) | xO pp => Branch1 b (Tremove pp T1) T2 | xH => mkBranch0 T1 T2 end end. Theorem Tget_Tempty: forall (p : positive), Tget p (Tempty) = PNone. destruct p;reflexivity. Qed. Theorem Tget_Tadd: forall i j a T, Tget i (Tadd j a T) = match (i ?= j) with Eq => PSome a | Lt => Tget i T | Gt => Tget i T end. Proof. intros i j. case_eq (i ?= j). intro H;rewrite (Pos.compare_eq _ _ H);intros a;clear i H. induction j;destruct T;simpl;try (apply IHj);congruence. unfold Pos.compare. generalize i;clear i;induction j;destruct T;simpl in H|-*; destruct i;simpl;try rewrite (IHj _ H);try (destruct i;simpl;congruence);reflexivity|| congruence. unfold Pos.compare. generalize i;clear i;induction j;destruct T;simpl in H|-*; destruct i;simpl;try rewrite (IHj _ H);try (destruct i;simpl;congruence);reflexivity|| congruence. Qed. Record Store : Type := mkStore {index:positive;contents:Tree}. Definition empty := mkStore xH Tempty. Definition push a S := mkStore (Pos.succ (index S)) (Tadd (index S) a (contents S)). Definition get i S := Tget i (contents S). Lemma get_empty : forall i, get i empty = PNone. intro i; case i; unfold empty,get; simpl;reflexivity. Qed. Inductive Full : Store -> Type:= F_empty : Full empty | F_push : forall a S, Full S -> Full (push a S). Theorem get_Full_Gt : forall S, Full S -> forall i, (i ?= index S) = Gt -> get i S = PNone. Proof. intros S W;induction W. unfold empty,index,get,contents;intros;apply Tget_Tempty. unfold index,get,push. simpl @contents. intros i e;rewrite Tget_Tadd. rewrite (Gt_Psucc _ _ e). unfold get in IHW. apply IHW;apply Gt_Psucc;assumption. Qed. Theorem get_Full_Eq : forall S, Full S -> get (index S) S = PNone. intros [index0 contents0] F. case F. unfold empty,index,get,contents;intros;apply Tget_Tempty. unfold push,index,get;simpl @contents. intros a S. rewrite Tget_Tadd. rewrite Psucc_Gt. intro W. change (get (Pos.succ (index S)) S =PNone). apply get_Full_Gt; auto. apply Psucc_Gt. Qed. Theorem get_push_Full : forall i a S, Full S -> get i (push a S) = match (i ?= index S) with Eq => PSome a | Lt => get i S | Gt => PNone end. Proof. intros i a S F. case_eq (i ?= index S). intro e;rewrite (Pos.compare_eq _ _ e). destruct S;unfold get,push,index;simpl @contents;rewrite Tget_Tadd. rewrite Pos.compare_refl;reflexivity. intros;destruct S;unfold get,push,index;simpl @contents;rewrite Tget_Tadd. simpl @index in H;rewrite H;reflexivity. intro H;generalize H;clear H. unfold get,push;simpl. rewrite Tget_Tadd;intro e;rewrite e. change (get i S=PNone). apply get_Full_Gt;auto. Qed. Lemma Full_push_compat : forall i a S, Full S -> forall x, get i S = PSome x -> get i (push a S) = PSome x. Proof. intros i a S F x H. case_eq (i ?= index S);intro test. rewrite (Pos.compare_eq _ _ test) in H. rewrite (get_Full_Eq _ F) in H;congruence. rewrite <- H. rewrite (get_push_Full i a). rewrite test;reflexivity. assumption. rewrite (get_Full_Gt _ F) in H;congruence. Qed. Lemma Full_index_one_empty : forall S, Full S -> index S = 1 -> S=empty. intros [ind cont] F one; inversion F. reflexivity. simpl @index in one;assert (h:=Pos.succ_not_1 (index S)). congruence. Qed. Lemma push_not_empty: forall a S, (push a S) <> empty. intros a [ind cont];unfold push,empty. intros [= H%Pos.succ_not_1]. assumption. Qed. Fixpoint In (x:A) (S:Store) (F:Full S) {struct F}: Prop := match F with F_empty => False | F_push a SS FF => x=a \/ In x SS FF end. Lemma get_In : forall (x:A) (S:Store) (F:Full S) i , get i S = PSome x -> In x S F. induction F. intro i;rewrite get_empty; congruence. intro i;rewrite get_push_Full;trivial. case_eq (i ?= index S);simpl. left;congruence. right;eauto. congruence. Qed. End Store. Arguments PNone {A}. Arguments PSome [A] _. Arguments Tempty {A}. Arguments Branch0 [A] _ _. Arguments Branch1 [A] _ _ _. Arguments Tget [A] p T. Arguments Tadd [A] p a T. Arguments Tget_Tempty [A] p. Arguments Tget_Tadd [A] i j a T. Arguments mkStore [A] index contents. Arguments index [A] s. Arguments contents [A] s. Arguments empty {A}. Arguments get [A] i S. Arguments push [A] a S. Arguments get_empty [A] i. Arguments get_push_Full [A] i a S _. Arguments Full [A] _. Arguments F_empty {A}. Arguments F_push [A] a S _. Arguments In [A] x S F. Register empty as plugins.rtauto.empty. Register push as plugins.rtauto.push. Section Map. Variables A B:Set. Variable f: A -> B. Fixpoint Tmap (T: Tree A) : Tree B := match T with Tempty => Tempty | Branch0 t1 t2 => Branch0 (Tmap t1) (Tmap t2) | Branch1 a t1 t2 => Branch1 (f a) (Tmap t1) (Tmap t2) end. Lemma Tget_Tmap: forall T i, Tget i (Tmap T)= match Tget i T with PNone => PNone | PSome a => PSome (f a) end. induction T;intro i;case i;simpl;auto. Defined. Lemma Tmap_Tadd: forall i a T, Tmap (Tadd i a T) = Tadd i (f a) (Tmap T). induction i;intros a T;case T;simpl;intros;try (rewrite IHi);simpl;reflexivity. Defined. Definition map (S:Store A) : Store B := mkStore (index S) (Tmap (contents S)). Lemma get_map: forall i S, get i (map S)= match get i S with PNone => PNone | PSome a => PSome (f a) end. destruct S;unfold get,map,contents,index;apply Tget_Tmap. Defined. Lemma map_push: forall a S, map (push a S) = push (f a) (map S). intros a S. case S. unfold push,map,contents,index. intros;rewrite Tmap_Tadd;reflexivity. Defined. Theorem Full_map : forall S, Full S -> Full (map S). intros S F. induction F. exact F_empty. rewrite map_push;constructor 2;assumption. Defined. End Map. Arguments Tmap [A B] f T. Arguments map [A B] f S. Arguments Full_map [A B f] S _. Notation "hyps \ A" := (push A hyps) (at level 72,left associativity). coq-8.11.0/plugins/rtauto/proof_search.ml0000644000175000017500000004171613612664533020302 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* !pruning); optwrite=(fun b -> pruning:=b)} let () = declare_bool_option opt_pruning type form= Atom of int | Arrow of form * form | Bot | Conjunct of form * form | Disjunct of form * form module FOrd = struct type t = form let rec compare x y = match x, y with | Bot, Bot -> 0 | Bot, _ -> -1 | Atom _, Bot -> 1 | Atom a1, Atom a2 -> Int.compare a1 a2 | Atom _, _ -> -1 | Arrow _, (Bot | Atom _) -> 1 | Arrow (f1, g1), Arrow (f2, g2) -> let cmp = compare f1 f2 in if cmp = 0 then compare g1 g2 else cmp | Arrow _, _ -> -1 | Conjunct _, (Bot | Atom _ | Arrow _) -> 1 | Conjunct (f1, g1), Conjunct (f2, g2) -> let cmp = compare f1 f2 in if cmp = 0 then compare g1 g2 else cmp | Conjunct _, _ -> -1 | Disjunct _, (Bot | Atom _ | Arrow _ | Conjunct _) -> 1 | Disjunct (f1, g1), Disjunct (f2, g2) -> let cmp = compare f1 f2 in if cmp = 0 then compare g1 g2 else cmp end module Fmap = Map.Make(FOrd) type sequent = {rev_hyps: form Int.Map.t; norev_hyps: form Int.Map.t; size:int; left:int Fmap.t; right:(int*form) list Fmap.t; cnx:(int*int*form*form) list; abs:int option; gl:form} let add_one_arrow i f1 f2 m= try Fmap.add f1 ((i,f2)::(Fmap.find f1 m)) m with Not_found -> Fmap.add f1 [i,f2] m type proof = Ax of int | I_Arrow of proof | E_Arrow of int*int*proof | D_Arrow of int*proof*proof | E_False of int | I_And of proof*proof | E_And of int*proof | D_And of int*proof | I_Or_l of proof | I_Or_r of proof | E_Or of int*proof*proof | D_Or of int*proof | Pop of int*proof type rule = SAx of int | SI_Arrow | SE_Arrow of int*int | SD_Arrow of int | SE_False of int | SI_And | SE_And of int | SD_And of int | SI_Or_l | SI_Or_r | SE_Or of int | SD_Or of int let add_step s sub = match s,sub with SAx i,[] -> Ax i | SI_Arrow,[p] -> I_Arrow p | SE_Arrow(i,j),[p] -> E_Arrow (i,j,p) | SD_Arrow i,[p1;p2] -> D_Arrow (i,p1,p2) | SE_False i,[] -> E_False i | SI_And,[p1;p2] -> I_And(p1,p2) | SE_And i,[p] -> E_And(i,p) | SD_And i,[p] -> D_And(i,p) | SI_Or_l,[p] -> I_Or_l p | SI_Or_r,[p] -> I_Or_r p | SE_Or i,[p1;p2] -> E_Or(i,p1,p2) | SD_Or i,[p] -> D_Or(i,p) | _,_ -> anomaly ~label:"add_step" (Pp.str "wrong arity.") type 'a with_deps = {dep_it:'a; dep_goal:bool; dep_hyps:Int.Set.t} type slice= {proofs_done:proof list; proofs_todo:sequent with_deps list; step:rule; needs_goal:bool; needs_hyps:Int.Set.t; changes_goal:bool; creates_hyps:Int.Set.t} type state = Complete of proof | Incomplete of sequent * slice list let project = function Complete prf -> prf | Incomplete (_,_) -> anomaly (Pp.str "not a successful state.") let pop n prf = let nprf= match prf.dep_it with Pop (i,p) -> Pop (i+n,p) | p -> Pop(n,p) in {prf with dep_it = nprf} let rec fill stack proof = match stack with [] -> Complete proof.dep_it | slice::super -> if !pruning && List.is_empty slice.proofs_done && not (slice.changes_goal && proof.dep_goal) && not (Int.Set.exists (fun i -> Int.Set.mem i proof.dep_hyps) slice.creates_hyps) then begin s_info.pruned_steps<-s_info.pruned_steps+1; s_info.pruned_branches<- s_info.pruned_branches + List.length slice.proofs_todo; let created_here=Int.Set.cardinal slice.creates_hyps in s_info.pruned_hyps<-s_info.pruned_hyps+ List.fold_left (fun sum dseq -> sum + Int.Set.cardinal dseq.dep_hyps) created_here slice.proofs_todo; fill super (pop (Int.Set.cardinal slice.creates_hyps) proof) end else let dep_hyps= Int.Set.union slice.needs_hyps (Int.Set.diff proof.dep_hyps slice.creates_hyps) in let dep_goal= slice.needs_goal || ((not slice.changes_goal) && proof.dep_goal) in let proofs_done= proof.dep_it::slice.proofs_done in match slice.proofs_todo with [] -> fill super {dep_it = add_step slice.step (List.rev proofs_done); dep_goal = dep_goal; dep_hyps = dep_hyps} | current::next -> let nslice= {proofs_done=proofs_done; proofs_todo=next; step=slice.step; needs_goal=dep_goal; needs_hyps=dep_hyps; changes_goal=current.dep_goal; creates_hyps=current.dep_hyps} in Incomplete (current.dep_it,nslice::super) let append stack (step,subgoals) = s_info.created_steps<-s_info.created_steps+1; match subgoals with [] -> s_info.branch_successes<-s_info.branch_successes+1; fill stack {dep_it=add_step step.dep_it []; dep_goal=step.dep_goal; dep_hyps=step.dep_hyps} | hd :: next -> s_info.created_branches<- s_info.created_branches+List.length next; let slice= {proofs_done=[]; proofs_todo=next; step=step.dep_it; needs_goal=step.dep_goal; needs_hyps=step.dep_hyps; changes_goal=hd.dep_goal; creates_hyps=hd.dep_hyps} in Incomplete(hd.dep_it,slice::stack) let embed seq= {dep_it=seq; dep_goal=false; dep_hyps=Int.Set.empty} let change_goal seq gl= {seq with dep_it={seq.dep_it with gl=gl}; dep_goal=true} let add_hyp seqwd f= s_info.created_hyps<-s_info.created_hyps+1; let seq=seqwd.dep_it in let num = seq.size+1 in let left = Fmap.add f num seq.left in let cnx,right= try let l=Fmap.find f seq.right in List.fold_right (fun (i,f0) l0 -> (num,i,f,f0)::l0) l seq.cnx, Fmap.remove f seq.right with Not_found -> seq.cnx,seq.right in let nseq= match f with Bot -> {seq with left=left; right=right; size=num; abs=Some num; cnx=cnx} | Atom _ -> {seq with size=num; left=left; right=right; cnx=cnx} | Conjunct (_,_) | Disjunct (_,_) -> {seq with rev_hyps=Int.Map.add num f seq.rev_hyps; size=num; left=left; right=right; cnx=cnx} | Arrow (f1,f2) -> let ncnx,nright= try let i = Fmap.find f1 seq.left in (i,num,f1,f2)::cnx,right with Not_found -> cnx,(add_one_arrow num f1 f2 right) in match f1 with Conjunct (_,_) | Disjunct (_,_) -> {seq with rev_hyps=Int.Map.add num f seq.rev_hyps; size=num; left=left; right=nright; cnx=ncnx} | Arrow(_,_) -> {seq with norev_hyps=Int.Map.add num f seq.norev_hyps; size=num; left=left; right=nright; cnx=ncnx} | _ -> {seq with size=num; left=left; right=nright; cnx=ncnx} in {seqwd with dep_it=nseq; dep_hyps=Int.Set.add num seqwd.dep_hyps} exception Here_is of (int*form) let choose m= try Int.Map.iter (fun i f -> raise (Here_is (i,f))) m; raise Not_found with Here_is (i,f) -> (i,f) let search_or seq= match seq.gl with Disjunct (f1,f2) -> [{dep_it = SI_Or_l; dep_goal = true; dep_hyps = Int.Set.empty}, [change_goal (embed seq) f1]; {dep_it = SI_Or_r; dep_goal = true; dep_hyps = Int.Set.empty}, [change_goal (embed seq) f2]] | _ -> [] let search_norev seq= let goals=ref (search_or seq) in let add_one i f= match f with Arrow (Arrow (f1,f2),f3) -> let nseq = {seq with norev_hyps=Int.Map.remove i seq.norev_hyps} in goals:= ({dep_it=SD_Arrow(i); dep_goal=false; dep_hyps=Int.Set.singleton i}, [add_hyp (add_hyp (change_goal (embed nseq) f2) (Arrow(f2,f3))) f1; add_hyp (embed nseq) f3]):: !goals | _ -> anomaly ~label:"search_no_rev" (Pp.str "can't happen.") in Int.Map.iter add_one seq.norev_hyps; List.rev !goals let search_in_rev_hyps seq= try let i,f=choose seq.rev_hyps in let make_step step= {dep_it=step; dep_goal=false; dep_hyps=Int.Set.singleton i} in let nseq={seq with rev_hyps=Int.Map.remove i seq.rev_hyps} in match f with Conjunct (f1,f2) -> [make_step (SE_And(i)), [add_hyp (add_hyp (embed nseq) f1) f2]] | Disjunct (f1,f2) -> [make_step (SE_Or(i)), [add_hyp (embed nseq) f1;add_hyp (embed nseq) f2]] | Arrow (Conjunct (f1,f2),f0) -> [make_step (SD_And(i)), [add_hyp (embed nseq) (Arrow (f1,Arrow (f2,f0)))]] | Arrow (Disjunct (f1,f2),f0) -> [make_step (SD_Or(i)), [add_hyp (add_hyp (embed nseq) (Arrow(f1,f0))) (Arrow (f2,f0))]] | _ -> anomaly ~label:"search_in_rev_hyps" (Pp.str "can't happen.") with Not_found -> search_norev seq let search_rev seq= match seq.cnx with (i,j,f1,f2)::next -> let nseq= match f1 with Conjunct (_,_) | Disjunct (_,_) -> {seq with cnx=next; rev_hyps=Int.Map.remove j seq.rev_hyps} | Arrow (_,_) -> {seq with cnx=next; norev_hyps=Int.Map.remove j seq.norev_hyps} | _ -> {seq with cnx=next} in [{dep_it=SE_Arrow(i,j); dep_goal=false; dep_hyps=Int.Set.add i (Int.Set.singleton j)}, [add_hyp (embed nseq) f2]] | [] -> match seq.gl with Arrow (f1,f2) -> [{dep_it=SI_Arrow; dep_goal=true; dep_hyps=Int.Set.empty}, [add_hyp (change_goal (embed seq) f2) f1]] | Conjunct (f1,f2) -> [{dep_it=SI_And; dep_goal=true; dep_hyps=Int.Set.empty},[change_goal (embed seq) f1; change_goal (embed seq) f2]] | _ -> search_in_rev_hyps seq let search_all seq= match seq.abs with Some i -> [{dep_it=SE_False (i); dep_goal=false; dep_hyps=Int.Set.singleton i},[]] | None -> try let ax = Fmap.find seq.gl seq.left in [{dep_it=SAx (ax); dep_goal=true; dep_hyps=Int.Set.singleton ax},[]] with Not_found -> search_rev seq let bare_sequent = embed {rev_hyps=Int.Map.empty; norev_hyps=Int.Map.empty; size=0; left=Fmap.empty; right=Fmap.empty; cnx=[]; abs=None; gl=Bot} let init_state hyps gl= let init = change_goal bare_sequent gl in let goal=List.fold_right (fun (_,f,_) seq ->add_hyp seq f) hyps init in Incomplete (goal.dep_it,[]) let success= function Complete _ -> true | Incomplete (_,_) -> false let branching = function Incomplete (seq,stack) -> Control.check_for_interrupt (); let successors = search_all seq in let _ = match successors with [] -> s_info.branch_failures<-s_info.branch_failures+1 | _::next -> s_info.nd_branching<-s_info.nd_branching+List.length next in List.map (append stack) successors | Complete prf -> anomaly (Pp.str "already succeeded.") open Pp let rec pp_form = function Arrow(f1,f2) -> (pp_or f1) ++ (str " -> ") ++ (pp_form f2) | f -> pp_or f and pp_or = function Disjunct(f1,f2) -> (pp_or f1) ++ (str " \\/ ") ++ (pp_and f2) | f -> pp_and f and pp_and = function Conjunct(f1,f2) -> (pp_and f1) ++ (str " /\\ ") ++ (pp_atom f2) | f -> pp_atom f and pp_atom= function Bot -> str "#" | Atom n -> int n | f -> str "(" ++ hv 2 (pp_form f) ++ str ")" let pr_form f = pp_form f let pp_intmap map = let pp=ref (str "") in Int.Map.iter (fun i obj -> pp:= (!pp ++ pp_form obj ++ cut ())) map; str "{ " ++ v 0 (!pp) ++ str " }" let pp_list pp_obj l= let pp=ref (str "") in List.iter (fun o -> pp := !pp ++ (pp_obj o) ++ str ", ") l; str "[ " ++ !pp ++ str "]" let pp_mapint map = let pp=ref (str "") in Fmap.iter (fun obj l -> pp:= (!pp ++ pp_form obj ++ str " => " ++ pp_list (fun (i,f) -> pp_form f) l ++ cut ()) ) map; str "{ " ++ hv 0 (!pp ++ str " }") let pp_connect (i,j,f1,f2) = pp_form f1 ++ str " => " ++ pp_form f2 let pp_gl gl= cut () ++ str "{ " ++ hv 0 ( begin match gl.abs with None -> str "" | Some i -> str "ABSURD" ++ cut () end ++ str "rev =" ++ pp_intmap gl.rev_hyps ++ cut () ++ str "norev =" ++ pp_intmap gl.norev_hyps ++ cut () ++ str "arrows=" ++ pp_mapint gl.right ++ cut () ++ str "cnx =" ++ pp_list pp_connect gl.cnx ++ cut () ++ str "goal =" ++ pp_form gl.gl ++ str " }") let pp = function Incomplete(gl,ctx) -> pp_gl gl ++ fnl () | _ -> str "" let pp_info () = let count_info = if !pruning then str "Proof steps : " ++ int s_info.created_steps ++ str " created / " ++ int s_info.pruned_steps ++ str " pruned" ++ fnl () ++ str "Proof branches : " ++ int s_info.created_branches ++ str " created / " ++ int s_info.pruned_branches ++ str " pruned" ++ fnl () ++ str "Hypotheses : " ++ int s_info.created_hyps ++ str " created / " ++ int s_info.pruned_hyps ++ str " pruned" ++ fnl () else str "Pruning is off" ++ fnl () ++ str "Proof steps : " ++ int s_info.created_steps ++ str " created" ++ fnl () ++ str "Proof branches : " ++ int s_info.created_branches ++ str " created" ++ fnl () ++ str "Hypotheses : " ++ int s_info.created_hyps ++ str " created" ++ fnl () in Feedback.msg_info ( str "Proof-search statistics :" ++ fnl () ++ count_info ++ str "Branch ends: " ++ int s_info.branch_successes ++ str " successes / " ++ int s_info.branch_failures ++ str " failures" ++ fnl () ++ str "Non-deterministic choices : " ++ int s_info.nd_branching ++ str " branches") coq-8.11.0/plugins/rtauto/rtauto_plugin.mlpack0000644000175000017500000000004113612664533021345 0ustar treinentreinenProof_search Refl_tauto G_rtauto coq-8.11.0/plugins/rtauto/refl_tauto.ml0000644000175000017500000002634213612664533017772 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constr.equal term t) atom_env.env in Atom i with Not_found -> let i=atom_env.next in atom_env.env <- (term,i)::atom_env.env; atom_env.next<- i + 1; Atom i let rec make_form env sigma atom_env term = let open EConstr in let open Vars in let normalize = special_nf env sigma in let cciterm = special_whd env sigma term in match EConstr.kind sigma cciterm with Prod(_,a,b) -> if noccurn sigma 1 b && Retyping.get_sort_family_of env sigma a == InProp then let fa = make_form env sigma atom_env a in let fb = make_form env sigma atom_env b in Arrow (fa,fb) else make_atom atom_env (normalize term) | Cast(a,_,_) -> make_form env sigma atom_env a | Ind (ind, _) -> if Names.eq_ind ind (fst (Lazy.force li_False)) then Bot else make_atom atom_env (normalize term) | App(hd,argv) when Int.equal (Array.length argv) 2 -> begin try let ind, _ = destInd sigma hd in if Names.eq_ind ind (fst (Lazy.force li_and)) then let fa = make_form env sigma atom_env argv.(0) in let fb = make_form env sigma atom_env argv.(1) in Conjunct (fa,fb) else if Names.eq_ind ind (fst (Lazy.force li_or)) then let fa = make_form env sigma atom_env argv.(0) in let fb = make_form env sigma atom_env argv.(1) in Disjunct (fa,fb) else make_atom atom_env (normalize term) with DestKO -> make_atom atom_env (normalize term) end | _ -> make_atom atom_env (normalize term) let rec make_hyps env sigma atom_env lenv = function [] -> [] | LocalDef (_,body,typ)::rest -> make_hyps env sigma atom_env (typ::body::lenv) rest | LocalAssum (id,typ)::rest -> let hrec= make_hyps env sigma atom_env (typ::lenv) rest in if List.exists (fun c -> Termops.local_occur_var Evd.empty (* FIXME *) id.binder_name c) lenv || (Retyping.get_sort_family_of env sigma typ != InProp) then hrec else (id,make_form env sigma atom_env typ)::hrec let rec build_pos n = if n<=1 then force node_count l_xH else if Int.equal (n land 1) 0 then mkApp (force node_count l_xO,[|build_pos (n asr 1)|]) else mkApp (force node_count l_xI,[|build_pos (n asr 1)|]) let rec build_form = function Atom n -> mkApp (force node_count l_Atom,[|build_pos n|]) | Arrow (f1,f2) -> mkApp (force node_count l_Arrow,[|build_form f1;build_form f2|]) | Bot -> force node_count l_Bot | Conjunct (f1,f2) -> mkApp (force node_count l_Conjunct,[|build_form f1;build_form f2|]) | Disjunct (f1,f2) -> mkApp (force node_count l_Disjunct,[|build_form f1;build_form f2|]) let rec decal k = function [] -> k | (start,delta)::rest -> if k>start then k - delta else decal k rest let add_pop size d pops= match pops with [] -> [size+d,d] | (_,sum)::_ -> (size+sum,sum+d)::pops let rec build_proof pops size = function Ax i -> mkApp (force step_count l_Ax, [|build_pos (decal i pops)|]) | I_Arrow p -> mkApp (force step_count l_I_Arrow, [|build_proof pops (size + 1) p|]) | E_Arrow(i,j,p) -> mkApp (force step_count l_E_Arrow, [|build_pos (decal i pops); build_pos (decal j pops); build_proof pops (size + 1) p|]) | D_Arrow(i,p1,p2) -> mkApp (force step_count l_D_Arrow, [|build_pos (decal i pops); build_proof pops (size + 2) p1; build_proof pops (size + 1) p2|]) | E_False i -> mkApp (force step_count l_E_False, [|build_pos (decal i pops)|]) | I_And(p1,p2) -> mkApp (force step_count l_I_And, [|build_proof pops size p1; build_proof pops size p2|]) | E_And(i,p) -> mkApp (force step_count l_E_And, [|build_pos (decal i pops); build_proof pops (size + 2) p|]) | D_And(i,p) -> mkApp (force step_count l_D_And, [|build_pos (decal i pops); build_proof pops (size + 1) p|]) | I_Or_l(p) -> mkApp (force step_count l_I_Or_l, [|build_proof pops size p|]) | I_Or_r(p) -> mkApp (force step_count l_I_Or_r, [|build_proof pops size p|]) | E_Or(i,p1,p2) -> mkApp (force step_count l_E_Or, [|build_pos (decal i pops); build_proof pops (size + 1) p1; build_proof pops (size + 1) p2|]) | D_Or(i,p) -> mkApp (force step_count l_D_Or, [|build_pos (decal i pops); build_proof pops (size + 2) p|]) | Pop(d,p) -> build_proof (add_pop size d pops) size p let build_env gamma= List.fold_right (fun (p,_) e -> mkApp(force node_count l_push,[|mkProp;p;e|])) gamma.env (mkApp (force node_count l_empty,[|mkProp|])) open Goptions let verbose = ref false let opt_verbose= {optdepr=false; optname="Rtauto Verbose"; optkey=["Rtauto";"Verbose"]; optread=(fun () -> !verbose); optwrite=(fun b -> verbose:=b)} let () = declare_bool_option opt_verbose let check = ref false let opt_check= {optdepr=false; optname="Rtauto Check"; optkey=["Rtauto";"Check"]; optread=(fun () -> !check); optwrite=(fun b -> check:=b)} let () = declare_bool_option opt_check open Pp let rtauto_tac = Proofview.Goal.enter begin fun gl -> let hyps = Proofview.Goal.hyps gl in let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in Coqlib.check_required_library ["Coq";"rtauto";"Rtauto"]; let gamma={next=1;env=[]} in let () = if Retyping.get_sort_family_of env sigma concl != InProp then user_err ~hdr:"rtauto" (Pp.str "goal should be in Prop") in let glf = make_form env sigma gamma concl in let hyps = make_hyps env sigma gamma [concl] hyps in let formula= List.fold_left (fun gl (_,f)-> Arrow (f,gl)) glf hyps in let search_fun = match Tacinterp.get_debug() with | Tactic_debug.DebugOn 0 -> Search.debug_depth_first | _ -> Search.depth_first in let () = begin reset_info (); if !verbose then Feedback.msg_info (str "Starting proof-search ..."); end in let search_start_time = System.get_time () in let prf = try project (search_fun (init_state [] formula)) with Not_found -> user_err ~hdr:"rtauto" (Pp.str "rtauto couldn't find any proof") in let search_end_time = System.get_time () in let () = if !verbose then begin Feedback.msg_info (str "Proof tree found in " ++ System.fmt_time_difference search_start_time search_end_time); pp_info (); Feedback.msg_info (str "Building proof term ... ") end in let build_start_time=System.get_time () in let () = step_count := 0; node_count := 0 in let main = mkApp (force node_count l_Reflect, [|build_env gamma; build_form formula; build_proof [] 0 prf|]) in let term= applistc main (List.rev_map (fun (id,_) -> mkVar id.binder_name) hyps) in let build_end_time=System.get_time () in let () = if !verbose then begin Feedback.msg_info (str "Proof term built in " ++ System.fmt_time_difference build_start_time build_end_time ++ fnl () ++ str "Proof size : " ++ int !step_count ++ str " steps" ++ fnl () ++ str "Proof term size : " ++ int (!step_count+ !node_count) ++ str " nodes (constants)" ++ fnl () ++ str "Giving proof term to Coq ... ") end in let tac_start_time = System.get_time () in let term = EConstr.of_constr term in let result= if !check then Tactics.exact_check term else Tactics.exact_no_check term in let tac_end_time = System.get_time () in let () = if !check then Feedback.msg_info (str "Proof term type-checking is on"); if !verbose then Feedback.msg_info (str "Internal tactic executed in " ++ System.fmt_time_difference tac_start_time tac_end_time) in result end coq-8.11.0/plugins/rtauto/g_rtauto.mlg0000644000175000017500000000144213612664533017613 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Refl_tauto.rtauto_tac } END coq-8.11.0/plugins/rtauto/Rtauto.v0000644000175000017500000002650713612664533016744 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* form | Arrow : form -> form -> form | Bot | Conjunct : form -> form -> form | Disjunct : form -> form -> form. Notation "[ n ]":=(Atom n). Notation "A =>> B":= (Arrow A B) (at level 59, right associativity). Notation "#" := Bot. Notation "A //\\ B" := (Conjunct A B) (at level 57, left associativity). Notation "A \\// B" := (Disjunct A B) (at level 58, left associativity). Definition ctx := Store form. Fixpoint pos_eq (m n:positive) {struct m} :bool := match m with xI mm => match n with xI nn => pos_eq mm nn | _ => false end | xO mm => match n with xO nn => pos_eq mm nn | _ => false end | xH => match n with xH => true | _ => false end end. Theorem pos_eq_refl : forall m n, pos_eq m n = true -> m = n. induction m;simpl;destruct n;congruence || (intro e;apply f_equal;auto). Qed. Fixpoint form_eq (p q:form) {struct p} :bool := match p with Atom m => match q with Atom n => pos_eq m n | _ => false end | Arrow p1 p2 => match q with Arrow q1 q2 => form_eq p1 q1 && form_eq p2 q2 | _ => false end | Bot => match q with Bot => true | _ => false end | Conjunct p1 p2 => match q with Conjunct q1 q2 => form_eq p1 q1 && form_eq p2 q2 | _ => false end | Disjunct p1 p2 => match q with Disjunct q1 q2 => form_eq p1 q1 && form_eq p2 q2 | _ => false end end. Theorem form_eq_refl: forall p q, form_eq p q = true -> p = q. induction p;destruct q;simpl;clean. intro h;generalize (pos_eq_refl _ _ h);congruence. case_eq (form_eq p1 q1);clean. intros e1 e2;generalize (IHp1 _ e1) (IHp2 _ e2);congruence. case_eq (form_eq p1 q1);clean. intros e1 e2;generalize (IHp1 _ e1) (IHp2 _ e2);congruence. case_eq (form_eq p1 q1);clean. intros e1 e2;generalize (IHp1 _ e1) (IHp2 _ e2);congruence. Qed. Arguments form_eq_refl [p q] _. Section with_env. Variable env:Store Prop. Fixpoint interp_form (f:form): Prop := match f with [n]=> match get n env with PNone => True | PSome P => P end | A =>> B => (interp_form A) -> (interp_form B) | # => False | A //\\ B => (interp_form A) /\ (interp_form B) | A \\// B => (interp_form A) \/ (interp_form B) end. Notation "[[ A ]]" := (interp_form A). Fixpoint interp_ctx (hyps:ctx) (F:Full hyps) (G:Prop) {struct F} : Prop := match F with F_empty => G | F_push H hyps0 F0 => interp_ctx hyps0 F0 ([[H]] -> G) end. Ltac wipe := intros;simpl;constructor. Lemma compose0 : forall hyps F (A:Prop), A -> (interp_ctx hyps F A). induction F;intros A H;simpl;auto. Qed. Lemma compose1 : forall hyps F (A B:Prop), (A -> B) -> (interp_ctx hyps F A) -> (interp_ctx hyps F B). induction F;intros A B H;simpl;auto. apply IHF;auto. Qed. Theorem compose2 : forall hyps F (A B C:Prop), (A -> B -> C) -> (interp_ctx hyps F A) -> (interp_ctx hyps F B) -> (interp_ctx hyps F C). induction F;intros A B C H;simpl;auto. apply IHF;auto. Qed. Theorem compose3 : forall hyps F (A B C D:Prop), (A -> B -> C -> D) -> (interp_ctx hyps F A) -> (interp_ctx hyps F B) -> (interp_ctx hyps F C) -> (interp_ctx hyps F D). induction F;intros A B C D H;simpl;auto. apply IHF;auto. Qed. Lemma weaken : forall hyps F f G, (interp_ctx hyps F G) -> (interp_ctx (hyps\f) (F_push f hyps F) G). induction F;simpl;intros;auto. apply compose1 with ([[a]]-> G);auto. Qed. Theorem project_In : forall hyps F g, In g hyps F -> interp_ctx hyps F [[g]]. induction F;simpl. contradiction. intros g H;destruct H. subst;apply compose0;simpl;trivial. apply compose1 with [[g]];auto. Qed. Theorem project : forall hyps F p g, get p hyps = PSome g-> interp_ctx hyps F [[g]]. intros hyps F p g e; apply project_In. apply get_In with p;assumption. Qed. Arguments project [hyps] F [p g] _. Inductive proof:Set := Ax : positive -> proof | I_Arrow : proof -> proof | E_Arrow : positive -> positive -> proof -> proof | D_Arrow : positive -> proof -> proof -> proof | E_False : positive -> proof | I_And: proof -> proof -> proof | E_And: positive -> proof -> proof | D_And: positive -> proof -> proof | I_Or_l: proof -> proof | I_Or_r: proof -> proof | E_Or: positive -> proof -> proof -> proof | D_Or: positive -> proof -> proof | Cut: form -> proof -> proof -> proof. Notation "hyps \ A" := (push A hyps) (at level 72,left associativity). Fixpoint check_proof (hyps:ctx) (gl:form) (P:proof) {struct P}: bool := match P with Ax i => match get i hyps with PSome F => form_eq F gl | _ => false end | I_Arrow p => match gl with A =>> B => check_proof (hyps \ A) B p | _ => false end | E_Arrow i j p => match get i hyps,get j hyps with PSome A,PSome (B =>>C) => form_eq A B && check_proof (hyps \ C) (gl) p | _,_ => false end | D_Arrow i p1 p2 => match get i hyps with PSome ((A =>>B)=>>C) => (check_proof ( hyps \ B =>> C \ A) B p1) && (check_proof (hyps \ C) gl p2) | _ => false end | E_False i => match get i hyps with PSome # => true | _ => false end | I_And p1 p2 => match gl with A //\\ B => check_proof hyps A p1 && check_proof hyps B p2 | _ => false end | E_And i p => match get i hyps with PSome (A //\\ B) => check_proof (hyps \ A \ B) gl p | _=> false end | D_And i p => match get i hyps with PSome (A //\\ B =>> C) => check_proof (hyps \ A=>>B=>>C) gl p | _=> false end | I_Or_l p => match gl with (A \\// B) => check_proof hyps A p | _ => false end | I_Or_r p => match gl with (A \\// B) => check_proof hyps B p | _ => false end | E_Or i p1 p2 => match get i hyps with PSome (A \\// B) => check_proof (hyps \ A) gl p1 && check_proof (hyps \ B) gl p2 | _=> false end | D_Or i p => match get i hyps with PSome (A \\// B =>> C) => (check_proof (hyps \ A=>>C \ B=>>C) gl p) | _=> false end | Cut A p1 p2 => check_proof hyps A p1 && check_proof (hyps \ A) gl p2 end. Theorem interp_proof: forall p hyps F gl, check_proof hyps gl p = true -> interp_ctx hyps F [[gl]]. induction p; intros hyps F gl. - (* Axiom *) simpl;case_eq (get p hyps);clean. intros f nth_f e;rewrite <- (form_eq_refl e). apply project with p;trivial. - (* Arrow_Intro *) destruct gl; clean. simpl; intros. change (interp_ctx (hyps\gl1) (F_push gl1 hyps F) [[gl2]]). apply IHp; try constructor; trivial. - (* Arrow_Elim *) simpl check_proof; case_eq (get p hyps); clean. intros f ef; case_eq (get p0 hyps); clean. intros f0 ef0; destruct f0; clean. case_eq (form_eq f f0_1); clean. simpl; intros e check_p1. generalize (project F ef) (project F ef0) (IHp (hyps \ f0_2) (F_push f0_2 hyps F) gl check_p1); clear check_p1 IHp p p0 p1 ef ef0. simpl. apply compose3. rewrite (form_eq_refl e). auto. - (* Arrow_Destruct *) simpl; case_eq (get p1 hyps); clean. intros f ef; destruct f; clean. destruct f1; clean. case_eq (check_proof (hyps \ f1_2 =>> f2 \ f1_1) f1_2 p2); clean. intros check_p1 check_p2. generalize (project F ef) (IHp1 (hyps \ f1_2 =>> f2 \ f1_1) (F_push f1_1 (hyps \ f1_2 =>> f2) (F_push (f1_2 =>> f2) hyps F)) f1_2 check_p1) (IHp2 (hyps \ f2) (F_push f2 hyps F) gl check_p2). simpl; apply compose3; auto. - (* False_Elim *) simpl; case_eq (get p hyps); clean. intros f ef; destruct f; clean. intros _; generalize (project F ef). apply compose1; apply False_ind. - (* And_Intro *) simpl; destruct gl; clean. case_eq (check_proof hyps gl1 p1); clean. intros Hp1 Hp2;generalize (IHp1 hyps F gl1 Hp1) (IHp2 hyps F gl2 Hp2). apply compose2 ; simpl; auto. - (* And_Elim *) simpl; case_eq (get p hyps); clean. intros f ef; destruct f; clean. intro check_p; generalize (project F ef) (IHp (hyps \ f1 \ f2) (F_push f2 (hyps \ f1) (F_push f1 hyps F)) gl check_p). simpl; apply compose2; intros [h1 h2]; auto. - (* And_Destruct*) simpl; case_eq (get p hyps); clean. intros f ef; destruct f; clean. destruct f1; clean. intro H; generalize (project F ef) (IHp (hyps \ f1_1 =>> f1_2 =>> f2) (F_push (f1_1 =>> f1_2 =>> f2) hyps F) gl H); clear H; simpl. apply compose2; auto. - (* Or_Intro_left *) destruct gl; clean. intro Hp; generalize (IHp hyps F gl1 Hp). apply compose1; simpl; auto. - (* Or_Intro_right *) destruct gl; clean. intro Hp; generalize (IHp hyps F gl2 Hp). apply compose1; simpl; auto. - (* Or_elim *) simpl; case_eq (get p1 hyps); clean. intros f ef; destruct f; clean. case_eq (check_proof (hyps \ f1) gl p2); clean. intros check_p1 check_p2; generalize (project F ef) (IHp1 (hyps \ f1) (F_push f1 hyps F) gl check_p1) (IHp2 (hyps \ f2) (F_push f2 hyps F) gl check_p2); simpl; apply compose3; simpl; intro h; destruct h; auto. - (* Or_Destruct *) simpl; case_eq (get p hyps); clean. intros f ef; destruct f; clean. destruct f1; clean. intro check_p0; generalize (project F ef) (IHp (hyps \ f1_1 =>> f2 \ f1_2 =>> f2) (F_push (f1_2 =>> f2) (hyps \ f1_1 =>> f2) (F_push (f1_1 =>> f2) hyps F)) gl check_p0); simpl. apply compose2; auto. - (* Cut *) simpl; case_eq (check_proof hyps f p1); clean. intros check_p1 check_p2; generalize (IHp1 hyps F f check_p1) (IHp2 (hyps\f) (F_push f hyps F) gl check_p2); simpl; apply compose2; auto. Qed. Theorem Reflect: forall gl prf, if check_proof empty gl prf then [[gl]] else True. intros gl prf;case_eq (check_proof empty gl prf);intro check_prf. change (interp_ctx empty F_empty [[gl]]) ; apply interp_proof with prf;assumption. trivial. Qed. End with_env. (* (* A small example *) Parameters A B C D:Prop. Theorem toto:A /\ (B \/ C) -> (A /\ B) \/ (A /\ C). exact (Reflect (empty \ A \ B \ C) ([1] //\\ ([2] \\// [3]) =>> [1] //\\ [2] \\// [1] //\\ [3]) (I_Arrow (E_And 1 (E_Or 3 (I_Or_l (I_And (Ax 2) (Ax 4))) (I_Or_r (I_And (Ax 2) (Ax 4))))))). Qed. Print toto. *) Register Reflect as plugins.rtauto.Reflect. Register Atom as plugins.rtauto.Atom. Register Arrow as plugins.rtauto.Arrow. Register Bot as plugins.rtauto.Bot. Register Conjunct as plugins.rtauto.Conjunct. Register Disjunct as plugins.rtauto.Disjunct. Register Ax as plugins.rtauto.Ax. Register I_Arrow as plugins.rtauto.I_Arrow. Register E_Arrow as plugins.rtauto.E_Arrow. Register D_Arrow as plugins.rtauto.D_Arrow. Register E_False as plugins.rtauto.E_False. Register I_And as plugins.rtauto.I_And. Register E_And as plugins.rtauto.E_And. Register D_And as plugins.rtauto.D_And. Register I_Or_l as plugins.rtauto.I_Or_l. Register I_Or_r as plugins.rtauto.I_Or_r. Register E_Or as plugins.rtauto.E_Or. Register D_Or as plugins.rtauto.D_Or. coq-8.11.0/plugins/rtauto/proof_search.mli0000644000175000017500000000253713612664533020451 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* proof val init_state : ('a * form * 'b) list -> form -> state val branching: state -> state list val success: state -> bool val pp: state -> Pp.t val pr_form : form -> Pp.t val reset_info : unit -> unit val pp_info : unit -> unit coq-8.11.0/plugins/rtauto/plugin_base.dune0000644000175000017500000000020213612664533020424 0ustar treinentreinen(library (name rtauto_plugin) (public_name coq.plugins.rtauto) (synopsis "Coq's rtauto plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/ssr/0000755000175000017500000000000013612664533014556 5ustar treinentreinencoq-8.11.0/plugins/ssr/ssripats.mli0000644000175000017500000000623613612664533017140 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ] tactical - the [:] pseudo-tactical for move, case, elim and abstract Putting these two features in the same file lets one hide much of the interaction between [tac E:] and [=>] ([E] has to be processed by [=>], not by [:] *) open Ssrast (* Atomic operations for the IPat machine. Use this if you are "patching" an * ipat written by the user, since patching it at he AST level and then * compiling it may have tricky effects, eg adding a clear in front of a view * also has the effect of disposing the view (the compilation phase takes care * of this, by using the compiled ipats you can be more precise *) type ssriop = | IOpId of Names.Id.t | IOpDrop | IOpTemporay | IOpInaccessible of string option | IOpInaccessibleAll | IOpAbstractVars of Names.Id.t list | IOpFastNondep | IOpInj of ssriops list | IOpDispatchBlock of id_block | IOpDispatchBranches of ssriops list | IOpCaseBlock of id_block | IOpCaseBranches of ssriops list | IOpRewrite of ssrocc * ssrdir | IOpView of ssrclear option * ssrview (* extra clears to be performed *) | IOpClear of ssrclear * ssrhyp option | IOpSimpl of ssrsimpl | IOpEqGen of unit Proofview.tactic (* generation of eqn *) | IOpNoop and ssriops = ssriop list val tclCompileIPats : ssripats -> ssriops (* The => tactical *) val tclIPAT : ssriops -> unit Proofview.tactic (* As above but with the SSR exception: first case is dispatch *) val tclIPATssr : ssripats -> unit Proofview.tactic (* Wrappers to deal with : and eqn generation/naming: [tac E: gens => ipats] where [E] is injected into [ipats] (at the right place) and [gens] are generalized before calling [tac] *) val ssrmovetac : ssrdgens ssrmovearg -> unit Proofview.tactic val ssrsmovetac : unit Proofview.tactic val ssrelimtac : ssrdgens ssrmovearg -> unit Proofview.tactic val ssrselimtoptac : unit Proofview.tactic val ssrcasetac : ssrdgens ssrmovearg -> unit Proofview.tactic val ssrscasetoptac : unit Proofview.tactic (* The implementation of abstract: is half here, for the [[: var ]] * ipat, and in ssrfwd for the integration with [have] *) val ssrabstract : ssrdgens -> unit Proofview.tactic (* Handling of [[:var]], needed in ssrfwd. Since ssrfwd is still outside the * tactic monad we export code with the V82 interface *) module Internal : sig val examine_abstract : EConstr.t -> Goal.goal Evd.sigma -> EConstr.types * EConstr.t array val pf_find_abstract_proof : bool -> Goal.goal Evd.sigma -> Constr.constr -> Evar.t end coq-8.11.0/plugins/ssr/ssrclasses.v0000644000175000017500000000265513612664533017142 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* .doc { font-family: monospace; white-space: pre; } # **) (** Compatibility layer for [under] and [setoid_rewrite]. Note: this file does not require [ssreflect]; it is both required by [ssrsetoid] and required by [ssrunder]. Redefine [Coq.Classes.RelationClasses.Reflexive] here, so that doing [Require Import ssreflect] does not [Require Import RelationClasses], and conversely. **) Section Defs. Context {A : Type}. Class Reflexive (R : A -> A -> Prop) := reflexivity : forall x : A, R x x. End Defs. Register Reflexive as plugins.ssreflect.reflexive_type. Register reflexivity as plugins.ssreflect.reflexive_proof. Instance eq_Reflexive {A : Type} : Reflexive (@eq A) := @eq_refl A. Instance iff_Reflexive : Reflexive iff := iff_refl. coq-8.11.0/plugins/ssr/ssrequality.mli0000644000175000017500000000467313612664533017660 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int val notimes : int val nomult : ssrmult val mkocc : ssrocc -> ssrdocc val mkclr : ssrclear -> ssrdocc val nodocc : ssrdocc val noclr : ssrdocc val simpltac : Ssrast.ssrsimpl -> Tacmach.tactic val newssrcongrtac : int * Ssrast.ssrterm -> Ltac_plugin.Tacinterp.interp_sign -> Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val mk_rwarg : Ssrast.ssrdir * (int * Ssrast.ssrmmod) -> (Ssrast.ssrclear option * Ssrast.ssrocc) * Ssrmatching.rpattern option -> ssrwkind * Ssrast.ssrterm -> ssrrwarg val norwmult : ssrdir * ssrmult val norwocc : (Ssrast.ssrclear option * Ssrast.ssrocc) * Ssrmatching.rpattern option val ssr_is_setoid : Environ.env -> Evd.evar_map -> EConstr.t -> EConstr.t array -> bool val ssrinstancesofrule : Ltac_plugin.Tacinterp.interp_sign -> Ssrast.ssrdir -> Ssrast.ssrterm -> Goal.goal Evd.sigma -> Goal.goal list Evd.sigma (* map_redex (by default the identity on after) is called on the * redex (before) and its replacement (after). It is used to * "rename" binders by the under tactic *) val ssrrewritetac : ?under:bool -> ?map_redex:(Environ.env -> Evd.evar_map -> before:EConstr.t -> after:EConstr.t -> Evd.evar_map * EConstr.t) -> Ltac_plugin.Tacinterp.interp_sign -> ssrrwarg list -> Tacmach.tactic val ipat_rewrite : ssrocc -> ssrdir -> EConstr.t -> Tacmach.tactic val unlocktac : Ltac_plugin.Tacinterp.interp_sign -> (Ssrmatching.occ * Ssrast.ssrterm) list -> Goal.goal Evd.sigma -> Goal.goal list Evd.sigma coq-8.11.0/plugins/ssr/ssrbool.v0000644000175000017500000026276313612664533016450 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* .doc { font-family: monospace; white-space: pre; } # **) Require Bool. Require Import ssreflect ssrfun. (** A theory of boolean predicates and operators. A large part of this file is concerned with boolean reflection. Definitions and notations: is_true b == the coercion of b : bool to Prop (:= b = true). This is just input and displayed as `b''. reflect P b == the reflection inductive predicate, asserting that the logical proposition P : prop with the formula b : bool. Lemmas asserting reflect P b are often referred to as "views". iffP, appP, sameP, rwP :: lemmas for direct manipulation of reflection views: iffP is used to prove reflection from logical equivalence, appP to compose views, and sameP and rwP to perform boolean and setoid rewriting. elimT :: coercion reflect >-> Funclass, which allows the direct application of `reflect' views to boolean assertions. decidable P <-> P is effectively decidable (:= {P} + {~ P}. contra, contraL, ... :: contraposition lemmas. altP my_viewP :: natural alternative for reflection; given lemma myviewP: reflect my_Prop my_formula, have #[#myP | not_myP#]# := altP my_viewP. generates two subgoals, in which my_formula has been replaced by true and false, resp., with new assumptions myP : my_Prop and not_myP: ~~ my_formula. Caveat: my_formula must be an APPLICATION, not a variable, constant, let-in, etc. (due to the poor behaviour of dependent index matching). boolP my_formula :: boolean disjunction, equivalent to altP (idP my_formula) but circumventing the dependent index capture issue; destructing boolP my_formula generates two subgoals with assumptions my_formula and ~~ myformula. As with altP, my_formula must be an application. \unless C, P <-> we can assume property P when a something that holds under condition C (such as C itself). := forall G : Prop, (C -> G) -> (P -> G) -> G. This is just C \/ P or rather its impredicative encoding, whose usage better fits the above description: given a lemma UCP whose conclusion is \unless C, P we can assume P by writing: wlog hP: / P by apply/UCP; (prove C -> goal). or even apply: UCP id _ => hP if the goal is C. classically P <-> we can assume P when proving is_true b. := forall b : bool, (P -> b) -> b. This is equivalent to ~ (~ P) when P : Prop. implies P Q == wrapper variant type that coerces to P -> Q and can be used as a P -> Q view unambiguously. Useful to avoid spurious insertion of <-> views when Q is a conjunction of foralls, as in Lemma all_and2 below; conversely, avoids confusion in apply views for impredicative properties, such as \unless C, P. Also supports contrapositives. a && b == the boolean conjunction of a and b. a || b == the boolean disjunction of a and b. a ==> b == the boolean implication of b by a. ~~ a == the boolean negation of a. a (+) b == the boolean exclusive or (or sum) of a and b. #[# /\ P1 , P2 & P3 #]# == multiway logical conjunction, up to 5 terms. #[# \/ P1 , P2 | P3 #]# == multiway logical disjunction, up to 4 terms. #[#&& a, b, c & d#]# == iterated, right associative boolean conjunction with arbitrary arity. #[#|| a, b, c | d#]# == iterated, right associative boolean disjunction with arbitrary arity. #[#==> a, b, c => d#]# == iterated, right associative boolean implication with arbitrary arity. and3P, ... == specific reflection lemmas for iterated connectives. andTb, orbAC, ... == systematic names for boolean connective properties (see suffix conventions below). prop_congr == a tactic to move a boolean equality from its coerced form in Prop to the equality in bool. bool_congr == resolution tactic for blindly weeding out like terms from boolean equalities (can fail). This file provides a theory of boolean predicates and relations: pred T == the type of bool predicates (:= T -> bool). simpl_pred T == the type of simplifying bool predicates, based on the simpl_fun type from ssrfun.v. mem_pred T == a specialized form of simpl_pred for "collective" predicates (see below). rel T == the type of bool relations. := T -> pred T or T -> T -> bool. simpl_rel T == type of simplifying relations. := T -> simpl_pred T predType == the generic predicate interface, supported for for lists and sets. pred_sort == the predType >-> Type projection; pred_sort is itself a Coercion target class. Declaring a coercion to pred_sort is an alternative way of equiping a type with a predType structure, which interoperates better with coercion subtyping. This is used, e.g., for finite sets, so that finite groups inherit the membership operation by coercing to sets. {pred T} == a type convertible to pred T, but whose head constant is pred_sort. This type should be used for parameters that can be used as collective predicates (see below), as this will allow passing in directly collections that implement predType by coercion as described above, e.g., finite sets. := pred_sort (predPredType T) If P is a predicate the proposition "x satisfies P" can be written applicatively as (P x), or using an explicit connective as (x \in P); in the latter case we say that P is a "collective" predicate. We use A, B rather than P, Q for collective predicates: x \in A == x satisfies the (collective) predicate A. x \notin A == x doesn't satisfy the (collective) predicate A. The pred T type can be used as a generic predicate type for either kind, but the two kinds of predicates should not be confused. When a "generic" pred T value of one type needs to be passed as the other the following conversions should be used explicitly: SimplPred P == a (simplifying) applicative equivalent of P. mem A == an applicative equivalent of collective predicate A: mem A x simplifies to x \in A, as mem A has in fact type mem_pred T. --> In user notation collective predicates _only_ occur as arguments to mem: A only appears as (mem A). This is hidden by notation, e.g., x \in A := in_mem x (mem A) here, enum A := enum_mem (mem A) in fintype. This makes it possible to unify the various ways in which A can be interpreted as a predicate, for both pattern matching and display. Alternatively one can use the syntax for explicit simplifying predicates and relations (in the following x is bound in E): #[#pred x | E#]# == simplifying (see ssrfun) predicate x => E. #[#pred x : T | E#]# == predicate x => E, with a cast on the argument. #[#pred : T | P#]# == constant predicate P on type T. #[#pred x | E1 & E2#]# == #[#pred x | E1 && E2#]#; an x : T cast is allowed. #[#pred x in A#]# == #[#pred x | x in A#]#. #[#pred x in A | E#]# == #[#pred x | x in A & E#]#. #[#pred x in A | E1 & E2#]# == #[#pred x in A | E1 && E2#]#. #[#predU A & B#]# == union of two collective predicates A and B. #[#predI A & B#]# == intersection of collective predicates A and B. #[#predD A & B#]# == difference of collective predicates A and B. #[#predC A#]# == complement of the collective predicate A. #[#preim f of A#]# == preimage under f of the collective predicate A. predU P Q, ..., preim f P == union, etc of applicative predicates. pred0 == the empty predicate. predT == the total (always true) predicate. if T : predArgType, then T coerces to predT. {: T} == T cast to predArgType (e.g., {: bool * nat}). In the following, x and y are bound in E: #[#rel x y | E#]# == simplifying relation x, y => E. #[#rel x y : T | E#]# == simplifying relation with arguments cast. #[#rel x y in A & B | E#]# == #[#rel x y | #[#&& x \in A, y \in B & E#]# #]#. #[#rel x y in A & B#]# == #[#rel x y | (x \in A) && (y \in B) #]#. #[#rel x y in A | E#]# == #[#rel x y in A & A | E#]#. #[#rel x y in A#]# == #[#rel x y in A & A#]#. relU R S == union of relations R and S. relpre f R == preimage of relation R under f. xpredU, ..., xrelpre == lambda terms implementing predU, ..., etc. Explicit values of type pred T (i.e., lamdba terms) should always be used applicatively, while values of collection types implementing the predType interface, such as sequences or sets should always be used as collective predicates. Defined constants and functions of type pred T or simpl_pred T as well as the explicit simpl_pred T values described below, can generally be used either way. Note however that x \in A will not auto-simplify when A is an explicit simpl_pred T value; the generic simplification rule inE must be used (when A : pred T, the unfold_in rule can be used). Constants of type pred T with an explicit simpl_pred value do not auto-simplify when used applicatively, but can still be expanded with inE. This behavior can be controlled as follows: Let A : collective_pred T := #[#pred x | ... #]#. The collective_pred T type is just an alias for pred T, but this cast stops rewrite inE from expanding the definition of A, thus treating A into an abstract collection (unfold_in or in_collective can be used to expand manually). Let A : applicative_pred T := #[#pred x | ... #]#. This cast causes inE to turn x \in A into the applicative A x form; A will then have to unfolded explicitly with the /A rule. This will also apply to any definition that reduces to A (e.g., Let B := A). Canonical A_app_pred := ApplicativePred A. This declaration, given after definition of A, similarly causes inE to turn x \in A into A x, but in addition allows the app_predE rule to turn A x back into x \in A; it can be used for any definition of type pred T, which makes it especially useful for ambivalent predicates as the relational transitive closure connect, that are used in both applicative and collective styles. Purely for aesthetics, we provide a subtype of collective predicates: qualifier q T == a pred T pretty-printing wrapper. An A : qualifier q T coerces to pred_sort and thus behaves as a collective predicate, but x \in A and x \notin A are displayed as: x \is A and x \isn't A when q = 0, x \is a A and x \isn't a A when q = 1, x \is an A and x \isn't an A when q = 2, respectively. #[#qualify x | P#]# := Qualifier 0 (fun x => P), constructor for the above. #[#qualify x : T | P#]#, #[#qualify a x | P#]#, #[#qualify an X | P#]#, etc. variants of the above with type constraints and different values of q. We provide an internal interface to support attaching properties (such as being multiplicative) to predicates: pred_key p == phantom type that will serve as a support for properties to be attached to p : {pred _}; instances should be created with Fact/Qed so as to be opaque. KeyedPred k_p == an instance of the interface structure that attaches (k_p : pred_key P) to P; the structure projection is a coercion to pred_sort. KeyedQualifier k_q == an instance of the interface structure that attaches (k_q : pred_key q) to (q : qualifier n T). DefaultPredKey p == a default value for pred_key p; the vernacular command Import DefaultKeying attaches this key to all predicates that are not explicitly keyed. Keys can be used to attach properties to predicates, qualifiers and generic nouns in a way that allows them to be used transparently. The key projection of a predicate property structure such as unsignedPred should be a pred_key, not a pred, and corresponding lemmas will have the form Lemma rpredN R S (oppS : @opprPred R S) (kS : keyed_pred oppS) : {mono -%%R: x / x \in kS}. Because x \in kS will be displayed as x \in S (or x \is S, etc), the canonical instance of opprPred will not normally be exposed (it will also be erased by /= simplification). In addition each predicate structure should have a DefaultPredKey Canonical instance that simply issues the property as a proof obligation (which can be caught by the Prop-irrelevant feature of the ssreflect plugin). Some properties of predicates and relations: A =i B <-> A and B are extensionally equivalent. {subset A <= B} <-> A is a (collective) subpredicate of B. subpred P Q <-> P is an (applicative) subpredicate or Q. subrel R S <-> R is a subrelation of S. In the following R is in rel T: reflexive R <-> R is reflexive. irreflexive R <-> R is irreflexive. symmetric R <-> R (in rel T) is symmetric (equation). pre_symmetric R <-> R is symmetric (implication). antisymmetric R <-> R is antisymmetric. total R <-> R is total. transitive R <-> R is transitive. left_transitive R <-> R is a congruence on its left hand side. right_transitive R <-> R is a congruence on its right hand side. equivalence_rel R <-> R is an equivalence relation. Localization of (Prop) predicates; if P1 is convertible to forall x, Qx, P2 to forall x y, Qxy and P3 to forall x y z, Qxyz : {for y, P1} <-> Qx{y / x}. {in A, P1} <-> forall x, x \in A -> Qx. {in A1 & A2, P2} <-> forall x y, x \in A1 -> y \in A2 -> Qxy. {in A &, P2} <-> forall x y, x \in A -> y \in A -> Qxy. {in A1 & A2 & A3, Q3} <-> forall x y z, x \in A1 -> y \in A2 -> z \in A3 -> Qxyz. {in A1 & A2 &, Q3} := {in A1 & A2 & A2, Q3}. {in A1 && A3, Q3} := {in A1 & A1 & A3, Q3}. {in A &&, Q3} := {in A & A & A, Q3}. {in A, bijective f} <-> f has a right inverse in A. {on C, P1} <-> forall x, (f x) \in C -> Qx when P1 is also convertible to Pf f, e.g., {on C, involutive f}. {on C &, P2} == forall x y, f x \in C -> f y \in C -> Qxy when P2 is also convertible to Pf f, e.g., {on C &, injective f}. {on C, P1' & g} == forall x, (f x) \in cd -> Qx when P1' is convertible to Pf f and P1' g is convertible to forall x, Qx, e.g., {on C, cancel f & g}. {on C, bijective f} == f has a right inverse on C. This file extends the lemma name suffix conventions of ssrfun as follows: A -- associativity, as in andbA : associative andb. AC -- right commutativity. ACA -- self-interchange (inner commutativity), e.g., orbACA : (a || b) || (c || d) = (a || c) || (b || d). b -- a boolean argument, as in andbb : idempotent andb. C -- commutativity, as in andbC : commutative andb, or predicate complement, as in predC. CA -- left commutativity. D -- predicate difference, as in predD. E -- elimination, as in negbFE : ~~ b = false -> b. F or f -- boolean false, as in andbF : b && false = false. I -- left/right injectivity, as in addbI : right_injective addb, or predicate intersection, as in predI. l -- a left-hand operation, as andb_orl : left_distributive andb orb. N or n -- boolean negation, as in andbN : a && (~~ a) = false. P -- a characteristic property, often a reflection lemma, as in andP : reflect (a /\ b) (a && b). r -- a right-hand operation, as orb_andr : rightt_distributive orb andb. T or t -- boolean truth, as in andbT: right_id true andb. U -- predicate union, as in predU. W -- weakening, as in in1W : {in D, forall x, P} -> forall x, P. **) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Set Warnings "-projection-no-head-constant". Notation reflect := Bool.reflect. Notation ReflectT := Bool.ReflectT. Notation ReflectF := Bool.ReflectF. Reserved Notation "~~ b" (at level 35, right associativity). Reserved Notation "b ==> c" (at level 55, right associativity). Reserved Notation "b1 (+) b2" (at level 50, left associativity). Reserved Notation "x \in A" (at level 70, no associativity, format "'[hv' x '/ ' \in A ']'"). Reserved Notation "x \notin A" (at level 70, no associativity, format "'[hv' x '/ ' \notin A ']'"). Reserved Notation "x \is A" (at level 70, no associativity, format "'[hv' x '/ ' \is A ']'"). Reserved Notation "x \isn't A" (at level 70, no associativity, format "'[hv' x '/ ' \isn't A ']'"). Reserved Notation "x \is 'a' A" (at level 70, no associativity, format "'[hv' x '/ ' \is 'a' A ']'"). Reserved Notation "x \isn't 'a' A" (at level 70, no associativity, format "'[hv' x '/ ' \isn't 'a' A ']'"). Reserved Notation "x \is 'an' A" (at level 70, no associativity, format "'[hv' x '/ ' \is 'an' A ']'"). Reserved Notation "x \isn't 'an' A" (at level 70, no associativity, format "'[hv' x '/ ' \isn't 'an' A ']'"). Reserved Notation "p1 =i p2" (at level 70, no associativity, format "'[hv' p1 '/ ' =i p2 ']'"). Reserved Notation "{ 'subset' A <= B }" (at level 0, A, B at level 69, format "'[hv' { 'subset' A '/ ' <= B } ']'"). Reserved Notation "{ : T }" (at level 0, format "{ : T }"). Reserved Notation "{ 'pred' T }" (at level 0, format "{ 'pred' T }"). Reserved Notation "[ 'predType' 'of' T ]" (at level 0, format "[ 'predType' 'of' T ]"). Reserved Notation "[ 'pred' : T | E ]" (at level 0, format "'[hv' [ 'pred' : T | '/ ' E ] ']'"). Reserved Notation "[ 'pred' x | E ]" (at level 0, x ident, format "'[hv' [ 'pred' x | '/ ' E ] ']'"). Reserved Notation "[ 'pred' x : T | E ]" (at level 0, x ident, format "'[hv' [ 'pred' x : T | '/ ' E ] ']'"). Reserved Notation "[ 'pred' x | E1 & E2 ]" (at level 0, x ident, format "'[hv' [ 'pred' x | '/ ' E1 & '/ ' E2 ] ']'"). Reserved Notation "[ 'pred' x : T | E1 & E2 ]" (at level 0, x ident, format "'[hv' [ 'pred' x : T | '/ ' E1 & E2 ] ']'"). Reserved Notation "[ 'pred' x 'in' A ]" (at level 0, x ident, format "'[hv' [ 'pred' x 'in' A ] ']'"). Reserved Notation "[ 'pred' x 'in' A | E ]" (at level 0, x ident, format "'[hv' [ 'pred' x 'in' A | '/ ' E ] ']'"). Reserved Notation "[ 'pred' x 'in' A | E1 & E2 ]" (at level 0, x ident, format "'[hv' [ 'pred' x 'in' A | '/ ' E1 & '/ ' E2 ] ']'"). Reserved Notation "[ 'qualify' x | P ]" (at level 0, x at level 99, format "'[hv' [ 'qualify' x | '/ ' P ] ']'"). Reserved Notation "[ 'qualify' x : T | P ]" (at level 0, x at level 99, format "'[hv' [ 'qualify' x : T | '/ ' P ] ']'"). Reserved Notation "[ 'qualify' 'a' x | P ]" (at level 0, x at level 99, format "'[hv' [ 'qualify' 'a' x | '/ ' P ] ']'"). Reserved Notation "[ 'qualify' 'a' x : T | P ]" (at level 0, x at level 99, format "'[hv' [ 'qualify' 'a' x : T | '/ ' P ] ']'"). Reserved Notation "[ 'qualify' 'an' x | P ]" (at level 0, x at level 99, format "'[hv' [ 'qualify' 'an' x | '/ ' P ] ']'"). Reserved Notation "[ 'qualify' 'an' x : T | P ]" (at level 0, x at level 99, format "'[hv' [ 'qualify' 'an' x : T | '/ ' P ] ']'"). Reserved Notation "[ 'rel' x y | E ]" (at level 0, x ident, y ident, format "'[hv' [ 'rel' x y | '/ ' E ] ']'"). Reserved Notation "[ 'rel' x y : T | E ]" (at level 0, x ident, y ident, format "'[hv' [ 'rel' x y : T | '/ ' E ] ']'"). Reserved Notation "[ 'rel' x y 'in' A & B | E ]" (at level 0, x ident, y ident, format "'[hv' [ 'rel' x y 'in' A & B | '/ ' E ] ']'"). Reserved Notation "[ 'rel' x y 'in' A & B ]" (at level 0, x ident, y ident, format "'[hv' [ 'rel' x y 'in' A & B ] ']'"). Reserved Notation "[ 'rel' x y 'in' A | E ]" (at level 0, x ident, y ident, format "'[hv' [ 'rel' x y 'in' A | '/ ' E ] ']'"). Reserved Notation "[ 'rel' x y 'in' A ]" (at level 0, x ident, y ident, format "'[hv' [ 'rel' x y 'in' A ] ']'"). Reserved Notation "[ 'mem' A ]" (at level 0, format "[ 'mem' A ]"). Reserved Notation "[ 'predI' A & B ]" (at level 0, format "[ 'predI' A & B ]"). Reserved Notation "[ 'predU' A & B ]" (at level 0, format "[ 'predU' A & B ]"). Reserved Notation "[ 'predD' A & B ]" (at level 0, format "[ 'predD' A & B ]"). Reserved Notation "[ 'predC' A ]" (at level 0, format "[ 'predC' A ]"). Reserved Notation "[ 'preim' f 'of' A ]" (at level 0, format "[ 'preim' f 'of' A ]"). Reserved Notation "\unless C , P" (at level 200, C at level 100, format "'[hv' \unless C , '/ ' P ']'"). Reserved Notation "{ 'for' x , P }" (at level 0, format "'[hv' { 'for' x , '/ ' P } ']'"). Reserved Notation "{ 'in' d , P }" (at level 0, format "'[hv' { 'in' d , '/ ' P } ']'"). Reserved Notation "{ 'in' d1 & d2 , P }" (at level 0, format "'[hv' { 'in' d1 & d2 , '/ ' P } ']'"). Reserved Notation "{ 'in' d & , P }" (at level 0, format "'[hv' { 'in' d & , '/ ' P } ']'"). Reserved Notation "{ 'in' d1 & d2 & d3 , P }" (at level 0, format "'[hv' { 'in' d1 & d2 & d3 , '/ ' P } ']'"). Reserved Notation "{ 'in' d1 & & d3 , P }" (at level 0, format "'[hv' { 'in' d1 & & d3 , '/ ' P } ']'"). Reserved Notation "{ 'in' d1 & d2 & , P }" (at level 0, format "'[hv' { 'in' d1 & d2 & , '/ ' P } ']'"). Reserved Notation "{ 'in' d & & , P }" (at level 0, format "'[hv' { 'in' d & & , '/ ' P } ']'"). Reserved Notation "{ 'on' cd , P }" (at level 0, format "'[hv' { 'on' cd , '/ ' P } ']'"). Reserved Notation "{ 'on' cd & , P }" (at level 0, format "'[hv' { 'on' cd & , '/ ' P } ']'"). Reserved Notation "{ 'on' cd , P & g }" (at level 0, g at level 8, format "'[hv' { 'on' cd , '/ ' P & g } ']'"). Reserved Notation "{ 'in' d , 'bijective' f }" (at level 0, f at level 8, format "'[hv' { 'in' d , '/ ' 'bijective' f } ']'"). Reserved Notation "{ 'on' cd , 'bijective' f }" (at level 0, f at level 8, format "'[hv' { 'on' cd , '/ ' 'bijective' f } ']'"). (** We introduce a number of n-ary "list-style" notations that share a common format, namely #[#op arg1, arg2, ... last_separator last_arg#]# This usually denotes a right-associative applications of op, e.g., #[#&& a, b, c & d#]# denotes a && (b && (c && d)) The last_separator must be a non-operator token. Here we use &, | or =>; our default is &, but we try to match the intended meaning of op. The separator is a workaround for limitations of the parsing engine; the same limitations mean the separator cannot be omitted even when last_arg can. The Notation declarations are complicated by the separate treatment for some fixed arities (binary for bool operators, and all arities for Prop operators). We also use the square brackets in comprehension-style notations #[#type var separator expr#]# where "type" is the type of the comprehension (e.g., pred) and "separator" is | or => . It is important that in other notations a leading square bracket #[# is always followed by an operator symbol or a fixed identifier. **) Reserved Notation "[ /\ P1 & P2 ]" (at level 0, only parsing). Reserved Notation "[ /\ P1 , P2 & P3 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 ']' '/ ' & P3 ] ']'"). Reserved Notation "[ /\ P1 , P2 , P3 & P4 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 ']' '/ ' & P4 ] ']'"). Reserved Notation "[ /\ P1 , P2 , P3 , P4 & P5 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 ']' '/ ' & P5 ] ']'"). Reserved Notation "[ \/ P1 | P2 ]" (at level 0, only parsing). Reserved Notation "[ \/ P1 , P2 | P3 ]" (at level 0, format "'[hv' [ \/ '[' P1 , '/' P2 ']' '/ ' | P3 ] ']'"). Reserved Notation "[ \/ P1 , P2 , P3 | P4 ]" (at level 0, format "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 ']' '/ ' | P4 ] ']'"). Reserved Notation "[ && b1 & c ]" (at level 0, only parsing). Reserved Notation "[ && b1 , b2 , .. , bn & c ]" (at level 0, format "'[hv' [ && '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/ ' & c ] ']'"). Reserved Notation "[ || b1 | c ]" (at level 0, only parsing). Reserved Notation "[ || b1 , b2 , .. , bn | c ]" (at level 0, format "'[hv' [ || '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/ ' | c ] ']'"). Reserved Notation "[ ==> b1 => c ]" (at level 0, only parsing). Reserved Notation "[ ==> b1 , b2 , .. , bn => c ]" (at level 0, format "'[hv' [ ==> '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/' => c ] ']'"). (** Shorter delimiter **) Delimit Scope bool_scope with B. Open Scope bool_scope. (** An alternative to xorb that behaves somewhat better wrt simplification. **) Definition addb b := if b then negb else id. (** Notation for && and || is declared in Init.Datatypes. **) Notation "~~ b" := (negb b) : bool_scope. Notation "b ==> c" := (implb b c) : bool_scope. Notation "b1 (+) b2" := (addb b1 b2) : bool_scope. (** Constant is_true b := b = true is defined in Init.Datatypes. **) Coercion is_true : bool >-> Sortclass. (* Prop *) Lemma prop_congr : forall b b' : bool, b = b' -> b = b' :> Prop. Proof. by move=> b b' ->. Qed. Ltac prop_congr := apply: prop_congr. (** Lemmas for trivial. **) Lemma is_true_true : true. Proof. by []. Qed. Lemma not_false_is_true : ~ false. Proof. by []. Qed. Lemma is_true_locked_true : locked true. Proof. by unlock. Qed. Hint Resolve is_true_true not_false_is_true is_true_locked_true : core. (** Shorter names. **) Definition isT := is_true_true. Definition notF := not_false_is_true. (** Negation lemmas. **) (** We generally take NEGATION as the standard form of a false condition: negative boolean hypotheses should be of the form ~~ b, rather than ~ b or b = false, as much as possible. **) Lemma negbT b : b = false -> ~~ b. Proof. by case: b. Qed. Lemma negbTE b : ~~ b -> b = false. Proof. by case: b. Qed. Lemma negbF b : (b : bool) -> ~~ b = false. Proof. by case: b. Qed. Lemma negbFE b : ~~ b = false -> b. Proof. by case: b. Qed. Lemma negbK : involutive negb. Proof. by case. Qed. Lemma negbNE b : ~~ ~~ b -> b. Proof. by case: b. Qed. Lemma negb_inj : injective negb. Proof. exact: can_inj negbK. Qed. Lemma negbLR b c : b = ~~ c -> ~~ b = c. Proof. exact: canLR negbK. Qed. Lemma negbRL b c : ~~ b = c -> b = ~~ c. Proof. exact: canRL negbK. Qed. Lemma contra (c b : bool) : (c -> b) -> ~~ b -> ~~ c. Proof. by case: b => //; case: c. Qed. Definition contraNN := contra. Lemma contraL (c b : bool) : (c -> ~~ b) -> b -> ~~ c. Proof. by case: b => //; case: c. Qed. Definition contraTN := contraL. Lemma contraR (c b : bool) : (~~ c -> b) -> ~~ b -> c. Proof. by case: b => //; case: c. Qed. Definition contraNT := contraR. Lemma contraLR (c b : bool) : (~~ c -> ~~ b) -> b -> c. Proof. by case: b => //; case: c. Qed. Definition contraTT := contraLR. Lemma contraT b : (~~ b -> false) -> b. Proof. by case: b => // ->. Qed. Lemma wlog_neg b : (~~ b -> b) -> b. Proof. by case: b => // ->. Qed. Lemma contraFT (c b : bool) : (~~ c -> b) -> b = false -> c. Proof. by move/contraR=> notb_c /negbT. Qed. Lemma contraFN (c b : bool) : (c -> b) -> b = false -> ~~ c. Proof. by move/contra=> notb_notc /negbT. Qed. Lemma contraTF (c b : bool) : (c -> ~~ b) -> b -> c = false. Proof. by move/contraL=> b_notc /b_notc/negbTE. Qed. Lemma contraNF (c b : bool) : (c -> b) -> ~~ b -> c = false. Proof. by move/contra=> notb_notc /notb_notc/negbTE. Qed. Lemma contraFF (c b : bool) : (c -> b) -> b = false -> c = false. Proof. by move/contraFN=> bF_notc /bF_notc/negbTE. Qed. (** Coercion of sum-style datatypes into bool, which makes it possible to use ssr's boolean if rather than Coq's "generic" if. **) Coercion isSome T (u : option T) := if u is Some _ then true else false. Coercion is_inl A B (u : A + B) := if u is inl _ then true else false. Coercion is_left A B (u : {A} + {B}) := if u is left _ then true else false. Coercion is_inleft A B (u : A + {B}) := if u is inleft _ then true else false. Prenex Implicits isSome is_inl is_left is_inleft. Definition decidable P := {P} + {~ P}. (** Lemmas for ifs with large conditions, which allow reasoning about the condition without repeating it inside the proof (the latter IS preferable when the condition is short). Usage : if the goal contains (if cond then ...) = ... case: ifP => Hcond. generates two subgoal, with the assumption Hcond : cond = true/false Rewrite if_same eliminates redundant ifs Rewrite (fun_if f) moves a function f inside an if Rewrite if_arg moves an argument inside a function-valued if **) Section BoolIf. Variables (A B : Type) (x : A) (f : A -> B) (b : bool) (vT vF : A). Variant if_spec (not_b : Prop) : bool -> A -> Set := | IfSpecTrue of b : if_spec not_b true vT | IfSpecFalse of not_b : if_spec not_b false vF. Lemma ifP : if_spec (b = false) b (if b then vT else vF). Proof. by case def_b: b; constructor. Qed. Lemma ifPn : if_spec (~~ b) b (if b then vT else vF). Proof. by case def_b: b; constructor; rewrite ?def_b. Qed. Lemma ifT : b -> (if b then vT else vF) = vT. Proof. by move->. Qed. Lemma ifF : b = false -> (if b then vT else vF) = vF. Proof. by move->. Qed. Lemma ifN : ~~ b -> (if b then vT else vF) = vF. Proof. by move/negbTE->. Qed. Lemma if_same : (if b then vT else vT) = vT. Proof. by case b. Qed. Lemma if_neg : (if ~~ b then vT else vF) = if b then vF else vT. Proof. by case b. Qed. Lemma fun_if : f (if b then vT else vF) = if b then f vT else f vF. Proof. by case b. Qed. Lemma if_arg (fT fF : A -> B) : (if b then fT else fF) x = if b then fT x else fF x. Proof. by case b. Qed. (** Turning a boolean "if" form into an application. **) Definition if_expr := if b then vT else vF. Lemma ifE : (if b then vT else vF) = if_expr. Proof. by []. Qed. End BoolIf. (** Core (internal) reflection lemmas, used for the three kinds of views. **) Section ReflectCore. Variables (P Q : Prop) (b c : bool). Hypothesis Hb : reflect P b. Lemma introNTF : (if c then ~ P else P) -> ~~ b = c. Proof. by case c; case Hb. Qed. Lemma introTF : (if c then P else ~ P) -> b = c. Proof. by case c; case Hb. Qed. Lemma elimNTF : ~~ b = c -> if c then ~ P else P. Proof. by move <-; case Hb. Qed. Lemma elimTF : b = c -> if c then P else ~ P. Proof. by move <-; case Hb. Qed. Lemma equivPif : (Q -> P) -> (P -> Q) -> if b then Q else ~ Q. Proof. by case Hb; auto. Qed. Lemma xorPif : Q \/ P -> ~ (Q /\ P) -> if b then ~ Q else Q. Proof. by case Hb => [? _ H ? | ? H _]; case: H. Qed. End ReflectCore. (** Internal negated reflection lemmas **) Section ReflectNegCore. Variables (P Q : Prop) (b c : bool). Hypothesis Hb : reflect P (~~ b). Lemma introTFn : (if c then ~ P else P) -> b = c. Proof. by move/(introNTF Hb) <-; case b. Qed. Lemma elimTFn : b = c -> if c then ~ P else P. Proof. by move <-; apply: (elimNTF Hb); case b. Qed. Lemma equivPifn : (Q -> P) -> (P -> Q) -> if b then ~ Q else Q. Proof. by rewrite -if_neg; apply: equivPif. Qed. Lemma xorPifn : Q \/ P -> ~ (Q /\ P) -> if b then Q else ~ Q. Proof. by rewrite -if_neg; apply: xorPif. Qed. End ReflectNegCore. (** User-oriented reflection lemmas **) Section Reflect. Variables (P Q : Prop) (b b' c : bool). Hypotheses (Pb : reflect P b) (Pb' : reflect P (~~ b')). Lemma introT : P -> b. Proof. exact: introTF true _. Qed. Lemma introF : ~ P -> b = false. Proof. exact: introTF false _. Qed. Lemma introN : ~ P -> ~~ b. Proof. exact: introNTF true _. Qed. Lemma introNf : P -> ~~ b = false. Proof. exact: introNTF false _. Qed. Lemma introTn : ~ P -> b'. Proof. exact: introTFn true _. Qed. Lemma introFn : P -> b' = false. Proof. exact: introTFn false _. Qed. Lemma elimT : b -> P. Proof. exact: elimTF true _. Qed. Lemma elimF : b = false -> ~ P. Proof. exact: elimTF false _. Qed. Lemma elimN : ~~ b -> ~P. Proof. exact: elimNTF true _. Qed. Lemma elimNf : ~~ b = false -> P. Proof. exact: elimNTF false _. Qed. Lemma elimTn : b' -> ~ P. Proof. exact: elimTFn true _. Qed. Lemma elimFn : b' = false -> P. Proof. exact: elimTFn false _. Qed. Lemma introP : (b -> Q) -> (~~ b -> ~ Q) -> reflect Q b. Proof. by case b; constructor; auto. Qed. Lemma iffP : (P -> Q) -> (Q -> P) -> reflect Q b. Proof. by case: Pb; constructor; auto. Qed. Lemma equivP : (P <-> Q) -> reflect Q b. Proof. by case; apply: iffP. Qed. Lemma sumboolP (decQ : decidable Q) : reflect Q decQ. Proof. by case: decQ; constructor. Qed. Lemma appP : reflect Q b -> P -> Q. Proof. by move=> Qb; move/introT; case: Qb. Qed. Lemma sameP : reflect P c -> b = c. Proof. by case; [apply: introT | apply: introF]. Qed. Lemma decPcases : if b then P else ~ P. Proof. by case Pb. Qed. Definition decP : decidable P. by case: b decPcases; [left | right]. Defined. Lemma rwP : P <-> b. Proof. by split; [apply: introT | apply: elimT]. Qed. Lemma rwP2 : reflect Q b -> (P <-> Q). Proof. by move=> Qb; split=> ?; [apply: appP | apply: elimT; case: Qb]. Qed. (** Predicate family to reflect excluded middle in bool. **) Variant alt_spec : bool -> Type := | AltTrue of P : alt_spec true | AltFalse of ~~ b : alt_spec false. Lemma altP : alt_spec b. Proof. by case def_b: b / Pb; constructor; rewrite ?def_b. Qed. End Reflect. Hint View for move/ elimTF|3 elimNTF|3 elimTFn|3 introT|2 introTn|2 introN|2. Hint View for apply/ introTF|3 introNTF|3 introTFn|3 elimT|2 elimTn|2 elimN|2. Hint View for apply// equivPif|3 xorPif|3 equivPifn|3 xorPifn|3. (** Allow the direct application of a reflection lemma to a boolean assertion. **) Coercion elimT : reflect >-> Funclass. #[universes(template)] Variant implies P Q := Implies of P -> Q. Lemma impliesP P Q : implies P Q -> P -> Q. Proof. by case. Qed. Lemma impliesPn (P Q : Prop) : implies P Q -> ~ Q -> ~ P. Proof. by case=> iP ? /iP. Qed. Coercion impliesP : implies >-> Funclass. Hint View for move/ impliesPn|2 impliesP|2. Hint View for apply/ impliesPn|2 impliesP|2. (** Impredicative or, which can emulate a classical not-implies. **) Definition unless condition property : Prop := forall goal : Prop, (condition -> goal) -> (property -> goal) -> goal. Notation "\unless C , P" := (unless C P) : type_scope. Lemma unlessL C P : implies C (\unless C, P). Proof. by split=> hC G /(_ hC). Qed. Lemma unlessR C P : implies P (\unless C, P). Proof. by split=> hP G _ /(_ hP). Qed. Lemma unless_sym C P : implies (\unless C, P) (\unless P, C). Proof. by split; apply; [apply/unlessR | apply/unlessL]. Qed. Lemma unlessP (C P : Prop) : (\unless C, P) <-> C \/ P. Proof. by split=> [|[/unlessL | /unlessR]]; apply; [left | right]. Qed. Lemma bind_unless C P {Q} : implies (\unless C, P) (\unless (\unless C, Q), P). Proof. by split; apply=> [hC|hP]; [apply/unlessL/unlessL | apply/unlessR]. Qed. Lemma unless_contra b C : implies (~~ b -> C) (\unless C, b). Proof. by split; case: b => [_ | hC]; [apply/unlessR | apply/unlessL/hC]. Qed. (** Classical reasoning becomes directly accessible for any bool subgoal. Note that we cannot use "unless" here for lack of universe polymorphism. **) Definition classically P : Prop := forall b : bool, (P -> b) -> b. Lemma classicP (P : Prop) : classically P <-> ~ ~ P. Proof. split=> [cP nP | nnP [] // nP]; last by case nnP; move/nP. by have: P -> false; [move/nP | move/cP]. Qed. Lemma classicW P : P -> classically P. Proof. by move=> hP _ ->. Qed. Lemma classic_bind P Q : (P -> classically Q) -> classically P -> classically Q. Proof. by move=> iPQ cP b /iPQ-/cP. Qed. Lemma classic_EM P : classically (decidable P). Proof. by case=> // undecP; apply/undecP; right=> notP; apply/notF/undecP; left. Qed. Lemma classic_pick T P : classically ({x : T | P x} + (forall x, ~ P x)). Proof. case=> // undecP; apply/undecP; right=> x Px. by apply/notF/undecP; left; exists x. Qed. Lemma classic_imply P Q : (P -> classically Q) -> classically (P -> Q). Proof. move=> iPQ []// notPQ; apply/notPQ=> /iPQ-cQ. by case: notF; apply: cQ => hQ; apply: notPQ. Qed. (** List notations for wider connectives; the Prop connectives have a fixed width so as to avoid iterated destruction (we go up to width 5 for /\, and width 4 for or). The bool connectives have arbitrary widths, but denote expressions that associate to the RIGHT. This is consistent with the right associativity of list expressions and thus more convenient in most proofs. **) Inductive and3 (P1 P2 P3 : Prop) : Prop := And3 of P1 & P2 & P3. Inductive and4 (P1 P2 P3 P4 : Prop) : Prop := And4 of P1 & P2 & P3 & P4. Inductive and5 (P1 P2 P3 P4 P5 : Prop) : Prop := And5 of P1 & P2 & P3 & P4 & P5. Inductive or3 (P1 P2 P3 : Prop) : Prop := Or31 of P1 | Or32 of P2 | Or33 of P3. Inductive or4 (P1 P2 P3 P4 : Prop) : Prop := Or41 of P1 | Or42 of P2 | Or43 of P3 | Or44 of P4. Notation "[ /\ P1 & P2 ]" := (and P1 P2) (only parsing) : type_scope. Notation "[ /\ P1 , P2 & P3 ]" := (and3 P1 P2 P3) : type_scope. Notation "[ /\ P1 , P2 , P3 & P4 ]" := (and4 P1 P2 P3 P4) : type_scope. Notation "[ /\ P1 , P2 , P3 , P4 & P5 ]" := (and5 P1 P2 P3 P4 P5) : type_scope. Notation "[ \/ P1 | P2 ]" := (or P1 P2) (only parsing) : type_scope. Notation "[ \/ P1 , P2 | P3 ]" := (or3 P1 P2 P3) : type_scope. Notation "[ \/ P1 , P2 , P3 | P4 ]" := (or4 P1 P2 P3 P4) : type_scope. Notation "[ && b1 & c ]" := (b1 && c) (only parsing) : bool_scope. Notation "[ && b1 , b2 , .. , bn & c ]" := (b1 && (b2 && .. (bn && c) .. )) : bool_scope. Notation "[ || b1 | c ]" := (b1 || c) (only parsing) : bool_scope. Notation "[ || b1 , b2 , .. , bn | c ]" := (b1 || (b2 || .. (bn || c) .. )) : bool_scope. Notation "[ ==> b1 , b2 , .. , bn => c ]" := (b1 ==> (b2 ==> .. (bn ==> c) .. )) : bool_scope. Notation "[ ==> b1 => c ]" := (b1 ==> c) (only parsing) : bool_scope. Section AllAnd. Variables (T : Type) (P1 P2 P3 P4 P5 : T -> Prop). Local Notation a P := (forall x, P x). Lemma all_and2 : implies (forall x, [/\ P1 x & P2 x]) [/\ a P1 & a P2]. Proof. by split=> haveP; split=> x; case: (haveP x). Qed. Lemma all_and3 : implies (forall x, [/\ P1 x, P2 x & P3 x]) [/\ a P1, a P2 & a P3]. Proof. by split=> haveP; split=> x; case: (haveP x). Qed. Lemma all_and4 : implies (forall x, [/\ P1 x, P2 x, P3 x & P4 x]) [/\ a P1, a P2, a P3 & a P4]. Proof. by split=> haveP; split=> x; case: (haveP x). Qed. Lemma all_and5 : implies (forall x, [/\ P1 x, P2 x, P3 x, P4 x & P5 x]) [/\ a P1, a P2, a P3, a P4 & a P5]. Proof. by split=> haveP; split=> x; case: (haveP x). Qed. End AllAnd. Arguments all_and2 {T P1 P2}. Arguments all_and3 {T P1 P2 P3}. Arguments all_and4 {T P1 P2 P3 P4}. Arguments all_and5 {T P1 P2 P3 P4 P5}. Lemma pair_andP P Q : P /\ Q <-> P * Q. Proof. by split; case. Qed. Section ReflectConnectives. Variable b1 b2 b3 b4 b5 : bool. Lemma idP : reflect b1 b1. Proof. by case b1; constructor. Qed. Lemma boolP : alt_spec b1 b1 b1. Proof. exact: (altP idP). Qed. Lemma idPn : reflect (~~ b1) (~~ b1). Proof. by case b1; constructor. Qed. Lemma negP : reflect (~ b1) (~~ b1). Proof. by case b1; constructor; auto. Qed. Lemma negPn : reflect b1 (~~ ~~ b1). Proof. by case b1; constructor. Qed. Lemma negPf : reflect (b1 = false) (~~ b1). Proof. by case b1; constructor. Qed. Lemma andP : reflect (b1 /\ b2) (b1 && b2). Proof. by case b1; case b2; constructor=> //; case. Qed. Lemma and3P : reflect [/\ b1, b2 & b3] [&& b1, b2 & b3]. Proof. by case b1; case b2; case b3; constructor; try by case. Qed. Lemma and4P : reflect [/\ b1, b2, b3 & b4] [&& b1, b2, b3 & b4]. Proof. by case b1; case b2; case b3; case b4; constructor; try by case. Qed. Lemma and5P : reflect [/\ b1, b2, b3, b4 & b5] [&& b1, b2, b3, b4 & b5]. Proof. by case b1; case b2; case b3; case b4; case b5; constructor; try by case. Qed. Lemma orP : reflect (b1 \/ b2) (b1 || b2). Proof. by case b1; case b2; constructor; auto; case. Qed. Lemma or3P : reflect [\/ b1, b2 | b3] [|| b1, b2 | b3]. Proof. case b1; first by constructor; constructor 1. case b2; first by constructor; constructor 2. case b3; first by constructor; constructor 3. by constructor; case. Qed. Lemma or4P : reflect [\/ b1, b2, b3 | b4] [|| b1, b2, b3 | b4]. Proof. case b1; first by constructor; constructor 1. case b2; first by constructor; constructor 2. case b3; first by constructor; constructor 3. case b4; first by constructor; constructor 4. by constructor; case. Qed. Lemma nandP : reflect (~~ b1 \/ ~~ b2) (~~ (b1 && b2)). Proof. by case b1; case b2; constructor; auto; case; auto. Qed. Lemma norP : reflect (~~ b1 /\ ~~ b2) (~~ (b1 || b2)). Proof. by case b1; case b2; constructor; auto; case; auto. Qed. Lemma implyP : reflect (b1 -> b2) (b1 ==> b2). Proof. by case b1; case b2; constructor; auto. Qed. End ReflectConnectives. Arguments idP {b1}. Arguments idPn {b1}. Arguments negP {b1}. Arguments negPn {b1}. Arguments negPf {b1}. Arguments andP {b1 b2}. Arguments and3P {b1 b2 b3}. Arguments and4P {b1 b2 b3 b4}. Arguments and5P {b1 b2 b3 b4 b5}. Arguments orP {b1 b2}. Arguments or3P {b1 b2 b3}. Arguments or4P {b1 b2 b3 b4}. Arguments nandP {b1 b2}. Arguments norP {b1 b2}. Arguments implyP {b1 b2}. Prenex Implicits idP idPn negP negPn negPf. Prenex Implicits andP and3P and4P and5P orP or3P or4P nandP norP implyP. (** Shorter, more systematic names for the boolean connectives laws. **) Lemma andTb : left_id true andb. Proof. by []. Qed. Lemma andFb : left_zero false andb. Proof. by []. Qed. Lemma andbT : right_id true andb. Proof. by case. Qed. Lemma andbF : right_zero false andb. Proof. by case. Qed. Lemma andbb : idempotent andb. Proof. by case. Qed. Lemma andbC : commutative andb. Proof. by do 2!case. Qed. Lemma andbA : associative andb. Proof. by do 3!case. Qed. Lemma andbCA : left_commutative andb. Proof. by do 3!case. Qed. Lemma andbAC : right_commutative andb. Proof. by do 3!case. Qed. Lemma andbACA : interchange andb andb. Proof. by do 4!case. Qed. Lemma orTb : forall b, true || b. Proof. by []. Qed. Lemma orFb : left_id false orb. Proof. by []. Qed. Lemma orbT : forall b, b || true. Proof. by case. Qed. Lemma orbF : right_id false orb. Proof. by case. Qed. Lemma orbb : idempotent orb. Proof. by case. Qed. Lemma orbC : commutative orb. Proof. by do 2!case. Qed. Lemma orbA : associative orb. Proof. by do 3!case. Qed. Lemma orbCA : left_commutative orb. Proof. by do 3!case. Qed. Lemma orbAC : right_commutative orb. Proof. by do 3!case. Qed. Lemma orbACA : interchange orb orb. Proof. by do 4!case. Qed. Lemma andbN b : b && ~~ b = false. Proof. by case: b. Qed. Lemma andNb b : ~~ b && b = false. Proof. by case: b. Qed. Lemma orbN b : b || ~~ b = true. Proof. by case: b. Qed. Lemma orNb b : ~~ b || b = true. Proof. by case: b. Qed. Lemma andb_orl : left_distributive andb orb. Proof. by do 3!case. Qed. Lemma andb_orr : right_distributive andb orb. Proof. by do 3!case. Qed. Lemma orb_andl : left_distributive orb andb. Proof. by do 3!case. Qed. Lemma orb_andr : right_distributive orb andb. Proof. by do 3!case. Qed. Lemma andb_idl (a b : bool) : (b -> a) -> a && b = b. Proof. by case: a; case: b => // ->. Qed. Lemma andb_idr (a b : bool) : (a -> b) -> a && b = a. Proof. by case: a; case: b => // ->. Qed. Lemma andb_id2l (a b c : bool) : (a -> b = c) -> a && b = a && c. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma andb_id2r (a b c : bool) : (b -> a = c) -> a && b = c && b. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma orb_idl (a b : bool) : (a -> b) -> a || b = b. Proof. by case: a; case: b => // ->. Qed. Lemma orb_idr (a b : bool) : (b -> a) -> a || b = a. Proof. by case: a; case: b => // ->. Qed. Lemma orb_id2l (a b c : bool) : (~~ a -> b = c) -> a || b = a || c. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma orb_id2r (a b c : bool) : (~~ b -> a = c) -> a || b = c || b. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma negb_and (a b : bool) : ~~ (a && b) = ~~ a || ~~ b. Proof. by case: a; case: b. Qed. Lemma negb_or (a b : bool) : ~~ (a || b) = ~~ a && ~~ b. Proof. by case: a; case: b. Qed. (** Pseudo-cancellation -- i.e, absorption **) Lemma andbK a b : a && b || a = a. Proof. by case: a; case: b. Qed. Lemma andKb a b : a || b && a = a. Proof. by case: a; case: b. Qed. Lemma orbK a b : (a || b) && a = a. Proof. by case: a; case: b. Qed. Lemma orKb a b : a && (b || a) = a. Proof. by case: a; case: b. Qed. (** Imply **) Lemma implybT b : b ==> true. Proof. by case: b. Qed. Lemma implybF b : (b ==> false) = ~~ b. Proof. by case: b. Qed. Lemma implyFb b : false ==> b. Proof. by []. Qed. Lemma implyTb b : (true ==> b) = b. Proof. by []. Qed. Lemma implybb b : b ==> b. Proof. by case: b. Qed. Lemma negb_imply a b : ~~ (a ==> b) = a && ~~ b. Proof. by case: a; case: b. Qed. Lemma implybE a b : (a ==> b) = ~~ a || b. Proof. by case: a; case: b. Qed. Lemma implyNb a b : (~~ a ==> b) = a || b. Proof. by case: a; case: b. Qed. Lemma implybN a b : (a ==> ~~ b) = (b ==> ~~ a). Proof. by case: a; case: b. Qed. Lemma implybNN a b : (~~ a ==> ~~ b) = b ==> a. Proof. by case: a; case: b. Qed. Lemma implyb_idl (a b : bool) : (~~ a -> b) -> (a ==> b) = b. Proof. by case: a; case: b => // ->. Qed. Lemma implyb_idr (a b : bool) : (b -> ~~ a) -> (a ==> b) = ~~ a. Proof. by case: a; case: b => // ->. Qed. Lemma implyb_id2l (a b c : bool) : (a -> b = c) -> (a ==> b) = (a ==> c). Proof. by case: a; case: b; case: c => // ->. Qed. (** Addition (xor) **) Lemma addFb : left_id false addb. Proof. by []. Qed. Lemma addbF : right_id false addb. Proof. by case. Qed. Lemma addbb : self_inverse false addb. Proof. by case. Qed. Lemma addbC : commutative addb. Proof. by do 2!case. Qed. Lemma addbA : associative addb. Proof. by do 3!case. Qed. Lemma addbCA : left_commutative addb. Proof. by do 3!case. Qed. Lemma addbAC : right_commutative addb. Proof. by do 3!case. Qed. Lemma addbACA : interchange addb addb. Proof. by do 4!case. Qed. Lemma andb_addl : left_distributive andb addb. Proof. by do 3!case. Qed. Lemma andb_addr : right_distributive andb addb. Proof. by do 3!case. Qed. Lemma addKb : left_loop id addb. Proof. by do 2!case. Qed. Lemma addbK : right_loop id addb. Proof. by do 2!case. Qed. Lemma addIb : left_injective addb. Proof. by do 3!case. Qed. Lemma addbI : right_injective addb. Proof. by do 3!case. Qed. Lemma addTb b : true (+) b = ~~ b. Proof. by []. Qed. Lemma addbT b : b (+) true = ~~ b. Proof. by case: b. Qed. Lemma addbN a b : a (+) ~~ b = ~~ (a (+) b). Proof. by case: a; case: b. Qed. Lemma addNb a b : ~~ a (+) b = ~~ (a (+) b). Proof. by case: a; case: b. Qed. Lemma addbP a b : reflect (~~ a = b) (a (+) b). Proof. by case: a; case: b; constructor. Qed. Arguments addbP {a b}. (** Resolution tactic for blindly weeding out common terms from boolean equalities. When faced with a goal of the form (andb/orb/addb b1 b2) = b3 they will try to locate b1 in b3 and remove it. This can fail! **) Ltac bool_congr := match goal with | |- (?X1 && ?X2 = ?X3) => first [ symmetry; rewrite -1?(andbC X1) -?(andbCA X1); congr 1 (andb X1); symmetry | case: (X1); [ rewrite ?andTb ?andbT // | by rewrite ?andbF /= ] ] | |- (?X1 || ?X2 = ?X3) => first [ symmetry; rewrite -1?(orbC X1) -?(orbCA X1); congr 1 (orb X1); symmetry | case: (X1); [ by rewrite ?orbT //= | rewrite ?orFb ?orbF ] ] | |- (?X1 (+) ?X2 = ?X3) => symmetry; rewrite -1?(addbC X1) -?(addbCA X1); congr 1 (addb X1); symmetry | |- (~~ ?X1 = ?X2) => congr 1 negb end. (** Predicates, i.e., packaged functions to bool. - pred T, the basic type for predicates over a type T, is simply an alias for T -> bool. We actually distinguish two kinds of predicates, which we call applicative and collective, based on the syntax used to test them at some x in T: - For an applicative predicate P, one uses prefix syntax: P x Also, most operations on applicative predicates use prefix syntax as well (e.g., predI P Q). - For a collective predicate A, one uses infix syntax: x \in A and all operations on collective predicates use infix syntax as well (e.g., #[#predI A & B#]#). There are only two kinds of applicative predicates: - pred T, the alias for T -> bool mentioned above - simpl_pred T, an alias for simpl_fun T bool with a coercion to pred T that auto-simplifies on application (see ssrfun). On the other hand, the set of collective predicate types is open-ended via - predType T, a Structure that can be used to put Canonical collective predicate interpretation on other types, such as lists, tuples, finite sets, etc. Indeed, we define such interpretations for applicative predicate types, which can therefore also be used with the infix syntax, e.g., x \in predI P Q Moreover these infix forms are convertible to their prefix counterpart (e.g., predI P Q x which in turn simplifies to P x && Q x). The converse is not true, however; collective predicate types cannot, in general, be used applicatively, because of restrictions on implicit coercions. However, we do define an explicit generic coercion - mem : forall (pT : predType), pT -> mem_pred T where mem_pred T is a variant of simpl_pred T that preserves the infix syntax, i.e., mem A x auto-simplifies to x \in A. Indeed, the infix "collective" operators are notation for a prefix operator with arguments of type mem_pred T or pred T, applied to coerced collective predicates, e.g., Notation "x \in A" := (in_mem x (mem A)). This prevents the variability in the predicate type from interfering with the application of generic lemmas. Moreover this also makes it much easier to define generic lemmas, because the simplest type -- pred T -- can be used as the type of generic collective predicates, provided one takes care not to use it applicatively; this avoids the burden of having to declare a different predicate type for each predicate parameter of each section or lemma. In detail, we ensure that the head normal form of mem A is always of the eta-long MemPred (fun x => pA x) form, where pA is the pred interpretation of A following its predType pT, i.e., the _expansion_ of topred A. For a pred T evar ?P, (mem ?P) converts MemPred (fun x => ?P x), whose argument is a Miller pattern and therefore always unify: unifying (mem A) with (mem ?P) always yields ?P = pA, because the rigid constant MemPred aligns the unification. Furthermore, we ensure pA is always either A or toP .... A where toP ... is the expansion of @topred T pT, and toP is declared as a Coercion, so pA will _display_ as A in either case, and the instances of @mem T (predPredType T) pA appearing in the premises or right-hand side of a generic lemma parametrized by ?P will be indistinguishable from @mem T pT A. Users should take care not to inadvertently "strip" (mem A) down to the coerced A, since this will expose the internal toP coercion: Coq could then display terms A x that cannot be typed as such. The topredE lemma can be used to restore the x \in A syntax in this case. While -topredE can conversely be used to change x \in P into P x for an applicative P, it is safer to use the inE, unfold_in or and memE lemmas instead, as they do not run the risk of exposing internal coercions. As a consequence it is better to explicitly cast a generic applicative predicate to simpl_pred using the SimplPred constructor when it is used as a collective predicate (see, e.g., Lemma eq_big in bigop). We also sometimes "instantiate" the predType structure by defining a coercion to the sort of the predPredType structure, conveniently denoted {pred T}. This works better for types such as {set T} that have subtypes that coerce to them, since the same coercion will be inserted by the application of mem, or of any lemma that expects a generic collective predicates with type {pred T} := pred_sort (predPredType T) = pred T; thus {pred T} should be the preferred type for generic collective predicate parameters. This device also lets us turn any Type aT : predArgType into the total predicate over that type, i.e., fun _: aT => true. This allows us to write, e.g., ##|'I_n| for the cardinal of the (finite) type of integers less than n. **) (** Boolean predicates. *) Definition pred T := T -> bool. Identity Coercion fun_of_pred : pred >-> Funclass. Definition subpred T (p1 p2 : pred T) := forall x : T, p1 x -> p2 x. (* Notation for some manifest predicates. *) Notation xpred0 := (fun=> false). Notation xpredT := (fun=> true). Notation xpredI := (fun (p1 p2 : pred _) x => p1 x && p2 x). Notation xpredU := (fun (p1 p2 : pred _) x => p1 x || p2 x). Notation xpredC := (fun (p : pred _) x => ~~ p x). Notation xpredD := (fun (p1 p2 : pred _) x => ~~ p2 x && p1 x). Notation xpreim := (fun f (p : pred _) x => p (f x)). (** The packed class interface for pred-like types. **) Structure predType T := PredType {pred_sort :> Type; topred : pred_sort -> pred T}. Definition clone_pred T U := fun pT & @pred_sort T pT -> U => fun toP (pT' := @PredType T U toP) & phant_id pT' pT => pT'. Notation "[ 'predType' 'of' T ]" := (@clone_pred _ T _ id _ id) : form_scope. Canonical predPredType T := PredType (@id (pred T)). Canonical boolfunPredType T := PredType (@id (T -> bool)). (** The type of abstract collective predicates. While {pred T} is contertible to pred T, it presents the pred_sort coercion class, which crucially does _not_ coerce to Funclass. Term whose type P coerces to {pred T} cannot be applied to arguments, but they _can_ be used as if P had a canonical predType instance, as the coercion will be inserted if the unification P =~= pred_sort ?pT fails, changing the problem into the trivial {pred T} =~= pred_sort ?pT (solution ?pT := predPredType P). Additional benefits of this approach are that any type coercing to P will also inherit this behaviour, and that the coercion will be apparent in the elaborated expression. The latter may be important if the coercion is also a canonical structure projector - see mathcomp/fingroup/fingroup.v. The main drawback of implementing predType by coercion in this way is that the type of the value must be known when the unification constraint is imposed: if we only register the constraint and then later discover later that the expression had type P it will be too late of insert a coercion, whereas a canonical instance of predType fo P would have solved the deferred constraint. Finally, definitions, lemmas and sections should use type {pred T} for their generic collective type parameters, as this will make it possible to apply such definitions and lemmas directly to values of types that implement predType by coercion to {pred T} (values of types that implement predType without coercing to {pred T} will have to be coerced explicitly using topred). **) Notation "{ 'pred' T }" := (pred_sort (predPredType T)) : type_scope. (** The type of self-simplifying collective predicates. **) Definition simpl_pred T := simpl_fun T bool. Definition SimplPred {T} (p : pred T) : simpl_pred T := SimplFun p. (** Some simpl_pred constructors. **) Definition pred0 {T} := @SimplPred T xpred0. Definition predT {T} := @SimplPred T xpredT. Definition predI {T} (p1 p2 : pred T) := SimplPred (xpredI p1 p2). Definition predU {T} (p1 p2 : pred T) := SimplPred (xpredU p1 p2). Definition predC {T} (p : pred T) := SimplPred (xpredC p). Definition predD {T} (p1 p2 : pred T) := SimplPred (xpredD p1 p2). Definition preim {aT rT} (f : aT -> rT) (d : pred rT) := SimplPred (xpreim f d). Notation "[ 'pred' : T | E ]" := (SimplPred (fun _ : T => E%B)) : fun_scope. Notation "[ 'pred' x | E ]" := (SimplPred (fun x => E%B)) : fun_scope. Notation "[ 'pred' x | E1 & E2 ]" := [pred x | E1 && E2 ] : fun_scope. Notation "[ 'pred' x : T | E ]" := (SimplPred (fun x : T => E%B)) (only parsing) : fun_scope. Notation "[ 'pred' x : T | E1 & E2 ]" := [pred x : T | E1 && E2 ] (only parsing) : fun_scope. (** Coercions for simpl_pred. As simpl_pred T values are used both applicatively and collectively we need simpl_pred to coerce to both pred T _and_ {pred T}. However it is undesirable to have two distinct constants for what are essentially identical coercion functions, as this confuses the SSReflect keyed matching algorithm. While the Coq Coercion declarations appear to disallow such Coercion aliasing, it is possible to work around this limitation with a combination of modules and functors, which we do below. In addition we also give a predType instance for simpl_pred, which will be preferred to the {pred T} coercion to solve simpl_pred T =~= pred_sort ?pT constraints; not however that the pred_of_simpl coercion _will_ be used when a simpl_pred T is passed as a {pred T}, since the simplPredType T structure for simpl_pred T is _not_ convertible to predPredType T. **) Module PredOfSimpl. Definition coerce T (sp : simpl_pred T) : pred T := fun_of_simpl sp. End PredOfSimpl. Notation pred_of_simpl := PredOfSimpl.coerce. Coercion pred_of_simpl : simpl_pred >-> pred. Canonical simplPredType T := PredType (@pred_of_simpl T). Module Type PredSortOfSimplSignature. Parameter coerce : forall T, simpl_pred T -> {pred T}. End PredSortOfSimplSignature. Module DeclarePredSortOfSimpl (PredSortOfSimpl : PredSortOfSimplSignature). Coercion PredSortOfSimpl.coerce : simpl_pred >-> pred_sort. End DeclarePredSortOfSimpl. Module Export PredSortOfSimplCoercion := DeclarePredSortOfSimpl PredOfSimpl. (** Type to pred coercion. This lets us use types of sort predArgType as a synonym for their universal predicate. We define this predicate as a simpl_pred T rather than a pred T or a {pred T} so that /= and inE reduce (T x) and x \in T to true, respectively. Unfortunately, this can't be used for existing types like bool whose sort is already fixed (at least, not without redefining bool, true, false and all bool operations and lemmas); we provide syntax to recast a given type in predArgType as a workaround. **) Definition predArgType := Type. Bind Scope type_scope with predArgType. Identity Coercion sort_of_predArgType : predArgType >-> Sortclass. Coercion pred_of_argType (T : predArgType) : simpl_pred T := predT. Notation "{ : T }" := (T%type : predArgType) : type_scope. (** Boolean relations. Simplifying relations follow the coding pattern of 2-argument simplifying functions: the simplifying type constructor is applied to the _last_ argument. This design choice will let the in_simpl componenent of inE expand membership in simpl_rel as well. We provide an explicit coercion to rel T to avoid eta-expansion during coercion; this coercion self-simplifies so it should be invisible. **) Definition rel T := T -> pred T. Identity Coercion fun_of_rel : rel >-> Funclass. Definition subrel T (r1 r2 : rel T) := forall x y : T, r1 x y -> r2 x y. Definition simpl_rel T := T -> simpl_pred T. Coercion rel_of_simpl T (sr : simpl_rel T) : rel T := fun x : T => sr x. Arguments rel_of_simpl {T} sr x /. Notation xrelU := (fun (r1 r2 : rel _) x y => r1 x y || r2 x y). Notation xrelpre := (fun f (r : rel _) x y => r (f x) (f y)). Definition SimplRel {T} (r : rel T) : simpl_rel T := fun x => SimplPred (r x). Definition relU {T} (r1 r2 : rel T) := SimplRel (xrelU r1 r2). Definition relpre {aT rT} (f : aT -> rT) (r : rel rT) := SimplRel (xrelpre f r). Notation "[ 'rel' x y | E ]" := (SimplRel (fun x y => E%B)) : fun_scope. Notation "[ 'rel' x y : T | E ]" := (SimplRel (fun x y : T => E%B)) (only parsing) : fun_scope. Lemma subrelUl T (r1 r2 : rel T) : subrel r1 (relU r1 r2). Proof. by move=> x y r1xy; apply/orP; left. Qed. Lemma subrelUr T (r1 r2 : rel T) : subrel r2 (relU r1 r2). Proof. by move=> x y r2xy; apply/orP; right. Qed. (** Variant of simpl_pred specialised to the membership operator. **) Variant mem_pred T := Mem of pred T. (** We mainly declare pred_of_mem as a coercion so that it is not displayed. Similarly to pred_of_simpl, it will usually not be inserted by type inference, as all mem_pred mp =~= pred_sort ?pT unification problems will be solve by the memPredType instance below; pred_of_mem will however be used if a mem_pred T is used as a {pred T}, which is desirable as it will avoid a redundant mem in a collective, e.g., passing (mem A) to a lemma exception a generic collective predicate p : {pred T} and premise x \in P will display a subgoal x \in A rathere than x \in mem A. Conversely, pred_of_mem will _not_ if it is used id (mem A) is used applicatively or as a pred T; there the simpl_of_mem coercion defined below will be used, resulting in a subgoal that displays as mem A x by simplifies to x \in A. **) Coercion pred_of_mem {T} mp : {pred T} := let: Mem p := mp in [eta p]. Canonical memPredType T := PredType (@pred_of_mem T). Definition in_mem {T} (x : T) mp := pred_of_mem mp x. Definition eq_mem {T} mp1 mp2 := forall x : T, in_mem x mp1 = in_mem x mp2. Definition sub_mem {T} mp1 mp2 := forall x : T, in_mem x mp1 -> in_mem x mp2. Arguments in_mem {T} x mp : simpl never. Typeclasses Opaque eq_mem. Typeclasses Opaque sub_mem. (** The [simpl_of_mem; pred_of_simpl] path provides a new mem_pred >-> pred coercion, but does _not_ override the pred_of_mem : mem_pred >-> pred_sort explicit coercion declaration above. **) Coercion simpl_of_mem {T} mp := SimplPred (fun x : T => in_mem x mp). Lemma sub_refl T (mp : mem_pred T) : sub_mem mp mp. Proof. by []. Qed. Arguments sub_refl {T mp} [x] mp_x. (** It is essential to interlock the production of the Mem constructor inside the branch of the predType match, to ensure that unifying mem A with Mem [eta ?p] sets ?p := toP A (or ?p := P if toP = id and A = [eta P]), rather than topred pT A, had we put mem A := Mem (topred A). **) Definition mem T (pT : predType T) : pT -> mem_pred T := let: PredType toP := pT in fun A => Mem [eta toP A]. Arguments mem {T pT} A : rename, simpl never. Notation "x \in A" := (in_mem x (mem A)) : bool_scope. Notation "x \in A" := (in_mem x (mem A)) : bool_scope. Notation "x \notin A" := (~~ (x \in A)) : bool_scope. Notation "A =i B" := (eq_mem (mem A) (mem B)) : type_scope. Notation "{ 'subset' A <= B }" := (sub_mem (mem A) (mem B)) : type_scope. Notation "[ 'mem' A ]" := (pred_of_simpl (simpl_of_mem (mem A))) (only parsing) : fun_scope. Notation "[ 'predI' A & B ]" := (predI [mem A] [mem B]) : fun_scope. Notation "[ 'predU' A & B ]" := (predU [mem A] [mem B]) : fun_scope. Notation "[ 'predD' A & B ]" := (predD [mem A] [mem B]) : fun_scope. Notation "[ 'predC' A ]" := (predC [mem A]) : fun_scope. Notation "[ 'preim' f 'of' A ]" := (preim f [mem A]) : fun_scope. Notation "[ 'pred' x 'in' A ]" := [pred x | x \in A] : fun_scope. Notation "[ 'pred' x 'in' A | E ]" := [pred x | x \in A & E] : fun_scope. Notation "[ 'pred' x 'in' A | E1 & E2 ]" := [pred x | x \in A & E1 && E2 ] : fun_scope. Notation "[ 'rel' x y 'in' A & B | E ]" := [rel x y | (x \in A) && (y \in B) && E] : fun_scope. Notation "[ 'rel' x y 'in' A & B ]" := [rel x y | (x \in A) && (y \in B)] : fun_scope. Notation "[ 'rel' x y 'in' A | E ]" := [rel x y in A & A | E] : fun_scope. Notation "[ 'rel' x y 'in' A ]" := [rel x y in A & A] : fun_scope. (** Aliases of pred T that let us tag instances of simpl_pred as applicative or collective, via bespoke coercions. This tagging will give control over the simplification behaviour of inE and othe rewriting lemmas below. For this control to work it is crucial that collective_of_simpl _not_ be convertible to either applicative_of_simpl or pred_of_simpl. Indeed they differ here by a commutattive conversion (of the match and lambda). **) Definition applicative_pred T := pred T. Definition collective_pred T := pred T. Coercion applicative_pred_of_simpl T (sp : simpl_pred T) : applicative_pred T := fun_of_simpl sp. Coercion collective_pred_of_simpl T (sp : simpl_pred T) : collective_pred T := let: SimplFun p := sp in p. (** Explicit simplification rules for predicate application and membership. **) Section PredicateSimplification. Variables T : Type. Implicit Types (p : pred T) (pT : predType T) (sp : simpl_pred T). Implicit Types (mp : mem_pred T). (** The following four bespoke structures provide fine-grained control over matching the various predicate forms. While all four follow a common pattern of using a canonical projection to match a particular form of predicate (in pred T, simpl_pred, mem_pred and mem_pred, respectively), and display the matched predicate in the structure type, each is in fact used for a different, specific purpose: - registered_applicative_pred: this user-facing structure is used to declare values of type pred T meant to be used applicatively. The structure parameter merely displays this same value, and is used to avoid undesirable, visible occurrence of the structure in the right hand side of rewrite rules such as app_predE. There is a canonical instance of registered_applicative_pred for values of the applicative_of_simpl coercion, which handles the Definition Apred : applicative_pred T := [pred x | ...] idiom. This instance is mainly intended for the in_applicative component of inE, in conjunction with manifest_mem_pred and applicative_mem_pred. - manifest_simpl_pred: the only instance of this structure matches manifest simpl_pred values of the form SimplPred p, displaying p in the structure type. This structure is used in in_simpl to detect and selectively expand collective predicates of this form. An explicit SimplPred p pattern would _NOT_ work for this purpose, as then the left-hand side of in_simpl would reduce to in_mem ?x (Mem [eta ?p]) and would thus match _any_ instance of \in, not just those arising from a manifest simpl_pred. - manifest_mem_pred: similar to manifest_simpl_pred, the one instance of this structure matches manifest mem_pred values of the form Mem [eta ?p]. The purpose is different however: to match and display in ?p the actual predicate appearing in an ... \in ... expression matched by the left hand side of the in_applicative component of inE; then - applicative_mem_pred is a telescope refinement of manifest_mem_pred p with a default constructor that checks that the predicate p is the value of a registered_applicative_pred; any unfolding occurring during this check does _not_ affect the value of p passed to in_applicative, since that has been fixed earlier by the manifest_mem_pred match. In particular the definition of a predicate using the applicative_pred_of_simpl idiom above will not be expanded - this very case is the reason in_applicative uses a mem_pred telescope in its left hand side. The more straightforward ?x \in applicative_pred_value ?ap (equivalent to in_mem ?x (Mem ?ap)) with ?ap : registered_applicative_pred ?p would set ?p := [pred x | ...] rather than ?p := Apred in the example above. Also note that the in_applicative component of inE must be come before the in_simpl one, as the latter also matches terms of the form x \in Apred. Finally, no component of inE matches x \in Acoll, when Definition Acoll : collective_pred T := [pred x | ...]. as the collective_pred_of_simpl is _not_ convertible to pred_of_simpl. **) Structure registered_applicative_pred p := RegisteredApplicativePred { applicative_pred_value :> pred T; _ : applicative_pred_value = p }. Definition ApplicativePred p := RegisteredApplicativePred (erefl p). Canonical applicative_pred_applicative sp := ApplicativePred (applicative_pred_of_simpl sp). Structure manifest_simpl_pred p := ManifestSimplPred { simpl_pred_value :> simpl_pred T; _ : simpl_pred_value = SimplPred p }. Canonical expose_simpl_pred p := ManifestSimplPred (erefl (SimplPred p)). Structure manifest_mem_pred p := ManifestMemPred { mem_pred_value :> mem_pred T; _ : mem_pred_value = Mem [eta p] }. Canonical expose_mem_pred p := ManifestMemPred (erefl (Mem [eta p])). Structure applicative_mem_pred p := ApplicativeMemPred {applicative_mem_pred_value :> manifest_mem_pred p}. Canonical check_applicative_mem_pred p (ap : registered_applicative_pred p) := [eta @ApplicativeMemPred ap]. Lemma mem_topred pT (pp : pT) : mem (topred pp) = mem pp. Proof. by case: pT pp. Qed. Lemma topredE pT x (pp : pT) : topred pp x = (x \in pp). Proof. by rewrite -mem_topred. Qed. Lemma app_predE x p (ap : registered_applicative_pred p) : ap x = (x \in p). Proof. by case: ap => _ /= ->. Qed. Lemma in_applicative x p (amp : applicative_mem_pred p) : in_mem x amp = p x. Proof. by case: amp => -[_ /= ->]. Qed. Lemma in_collective x p (msp : manifest_simpl_pred p) : (x \in collective_pred_of_simpl msp) = p x. Proof. by case: msp => _ /= ->. Qed. Lemma in_simpl x p (msp : manifest_simpl_pred p) : in_mem x (Mem [eta pred_of_simpl msp]) = p x. Proof. by case: msp => _ /= ->. Qed. (** Because of the explicit eta expansion in the left-hand side, this lemma should only be used in the left-to-right direction. **) Lemma unfold_in x p : (x \in ([eta p] : pred T)) = p x. Proof. by []. Qed. Lemma simpl_predE p : SimplPred p =1 p. Proof. by []. Qed. Definition inE := (in_applicative, in_simpl, simpl_predE). (* to be extended *) Lemma mem_simpl sp : mem sp = sp :> pred T. Proof. by []. Qed. Definition memE := mem_simpl. (* could be extended *) Lemma mem_mem mp : (mem mp = mp) * (mem (mp : simpl_pred T) = mp) * (mem (mp : pred T) = mp). Proof. by case: mp. Qed. End PredicateSimplification. (** Qualifiers and keyed predicates. **) Variant qualifier (q : nat) T := Qualifier of {pred T}. Coercion has_quality n T (q : qualifier n T) : {pred T} := fun x => let: Qualifier _ p := q in p x. Arguments has_quality n {T}. Lemma qualifE n T p x : (x \in @Qualifier n T p) = p x. Proof. by []. Qed. Notation "x \is A" := (x \in has_quality 0 A) : bool_scope. Notation "x \is 'a' A" := (x \in has_quality 1 A) : bool_scope. Notation "x \is 'an' A" := (x \in has_quality 2 A) : bool_scope. Notation "x \isn't A" := (x \notin has_quality 0 A) : bool_scope. Notation "x \isn't 'a' A" := (x \notin has_quality 1 A) : bool_scope. Notation "x \isn't 'an' A" := (x \notin has_quality 2 A) : bool_scope. Notation "[ 'qualify' x | P ]" := (Qualifier 0 (fun x => P%B)) : form_scope. Notation "[ 'qualify' x : T | P ]" := (Qualifier 0 (fun x : T => P%B)) (only parsing) : form_scope. Notation "[ 'qualify' 'a' x | P ]" := (Qualifier 1 (fun x => P%B)) : form_scope. Notation "[ 'qualify' 'a' x : T | P ]" := (Qualifier 1 (fun x : T => P%B)) (only parsing) : form_scope. Notation "[ 'qualify' 'an' x | P ]" := (Qualifier 2 (fun x => P%B)) : form_scope. Notation "[ 'qualify' 'an' x : T | P ]" := (Qualifier 2 (fun x : T => P%B)) (only parsing) : form_scope. (** Keyed predicates: support for property-bearing predicate interfaces. **) Section KeyPred. Variable T : Type. #[universes(template)] Variant pred_key (p : {pred T}) := DefaultPredKey. Variable p : {pred T}. Structure keyed_pred (k : pred_key p) := PackKeyedPred {unkey_pred :> {pred T}; _ : unkey_pred =i p}. Variable k : pred_key p. Definition KeyedPred := @PackKeyedPred k p (frefl _). Variable k_p : keyed_pred k. Lemma keyed_predE : k_p =i p. Proof. by case: k_p. Qed. (** Instances that strip the mem cast; the first one has "pred_of_mem" as its projection head value, while the second has "pred_of_simpl". The latter has the side benefit of preempting accidental misdeclarations. Note: pred_of_mem is the registered mem >-> pred_sort coercion, while [simpl_of_mem; pred_of_simpl] is the mem >-> pred >=> Funclass coercion. We must write down the coercions explicitly as the Canonical head constant computation does not strip casts. **) Canonical keyed_mem := @PackKeyedPred k (pred_of_mem (mem k_p)) keyed_predE. Canonical keyed_mem_simpl := @PackKeyedPred k (pred_of_simpl (mem k_p)) keyed_predE. End KeyPred. Local Notation in_unkey x S := (x \in @unkey_pred _ S _ _) (only parsing). Notation "x \in S" := (in_unkey x S) (only printing) : bool_scope. Section KeyedQualifier. Variables (T : Type) (n : nat) (q : qualifier n T). Structure keyed_qualifier (k : pred_key q) := PackKeyedQualifier {unkey_qualifier; _ : unkey_qualifier = q}. Definition KeyedQualifier k := PackKeyedQualifier k (erefl q). Variables (k : pred_key q) (k_q : keyed_qualifier k). Fact keyed_qualifier_suproof : unkey_qualifier k_q =i q. Proof. by case: k_q => /= _ ->. Qed. Canonical keyed_qualifier_keyed := PackKeyedPred k keyed_qualifier_suproof. End KeyedQualifier. Notation "x \is A" := (in_unkey x (has_quality 0 A)) (only printing) : bool_scope. Notation "x \is 'a' A" := (in_unkey x (has_quality 1 A)) (only printing) : bool_scope. Notation "x \is 'an' A" := (in_unkey x (has_quality 2 A)) (only printing) : bool_scope. Module DefaultKeying. Canonical default_keyed_pred T p := KeyedPred (@DefaultPredKey T p). Canonical default_keyed_qualifier T n (q : qualifier n T) := KeyedQualifier (DefaultPredKey q). End DefaultKeying. (** Skolemizing with conditions. **) Lemma all_tag_cond_dep I T (C : pred I) U : (forall x, T x) -> (forall x, C x -> {y : T x & U x y}) -> {f : forall x, T x & forall x, C x -> U x (f x)}. Proof. move=> f0 fP; apply: all_tag (fun x y => C x -> U x y) _ => x. by case Cx: (C x); [case/fP: Cx => y; exists y | exists (f0 x)]. Qed. Lemma all_tag_cond I T (C : pred I) U : T -> (forall x, C x -> {y : T & U x y}) -> {f : I -> T & forall x, C x -> U x (f x)}. Proof. by move=> y0; apply: all_tag_cond_dep. Qed. Lemma all_sig_cond_dep I T (C : pred I) P : (forall x, T x) -> (forall x, C x -> {y : T x | P x y}) -> {f : forall x, T x | forall x, C x -> P x (f x)}. Proof. by move=> f0 /(all_tag_cond_dep f0)[f]; exists f. Qed. Lemma all_sig_cond I T (C : pred I) P : T -> (forall x, C x -> {y : T | P x y}) -> {f : I -> T | forall x, C x -> P x (f x)}. Proof. by move=> y0; apply: all_sig_cond_dep. Qed. Section RelationProperties. (** Caveat: reflexive should not be used to state lemmas, as auto and trivial will not expand the constant. **) Variable T : Type. Variable R : rel T. Definition total := forall x y, R x y || R y x. Definition transitive := forall y x z, R x y -> R y z -> R x z. Definition symmetric := forall x y, R x y = R y x. Definition antisymmetric := forall x y, R x y && R y x -> x = y. Definition pre_symmetric := forall x y, R x y -> R y x. Lemma symmetric_from_pre : pre_symmetric -> symmetric. Proof. by move=> symR x y; apply/idP/idP; apply: symR. Qed. Definition reflexive := forall x, R x x. Definition irreflexive := forall x, R x x = false. Definition left_transitive := forall x y, R x y -> R x =1 R y. Definition right_transitive := forall x y, R x y -> R^~ x =1 R^~ y. Section PER. Hypotheses (symR : symmetric) (trR : transitive). Lemma sym_left_transitive : left_transitive. Proof. by move=> x y Rxy z; apply/idP/idP; apply: trR; rewrite // symR. Qed. Lemma sym_right_transitive : right_transitive. Proof. by move=> x y /sym_left_transitive Rxy z; rewrite !(symR z) Rxy. Qed. End PER. (** We define the equivalence property with prenex quantification so that it can be localized using the {in ..., ..} form defined below. **) Definition equivalence_rel := forall x y z, R z z * (R x y -> R x z = R y z). Lemma equivalence_relP : equivalence_rel <-> reflexive /\ left_transitive. Proof. split=> [eqiR | [Rxx trR] x y z]; last by split=> [|/trR->]. by split=> [x | x y Rxy z]; [rewrite (eqiR x x x) | rewrite (eqiR x y z)]. Qed. End RelationProperties. Lemma rev_trans T (R : rel T) : transitive R -> transitive (fun x y => R y x). Proof. by move=> trR x y z Ryx Rzy; apply: trR Rzy Ryx. Qed. (** Property localization **) Local Notation "{ 'all1' P }" := (forall x, P x : Prop) (at level 0). Local Notation "{ 'all2' P }" := (forall x y, P x y : Prop) (at level 0). Local Notation "{ 'all3' P }" := (forall x y z, P x y z: Prop) (at level 0). Local Notation ph := (phantom _). Section LocalProperties. Variables T1 T2 T3 : Type. Variables (d1 : mem_pred T1) (d2 : mem_pred T2) (d3 : mem_pred T3). Local Notation ph := (phantom Prop). Definition prop_for (x : T1) P & ph {all1 P} := P x. Lemma forE x P phP : @prop_for x P phP = P x. Proof. by []. Qed. Definition prop_in1 P & ph {all1 P} := forall x, in_mem x d1 -> P x. Definition prop_in11 P & ph {all2 P} := forall x y, in_mem x d1 -> in_mem y d2 -> P x y. Definition prop_in2 P & ph {all2 P} := forall x y, in_mem x d1 -> in_mem y d1 -> P x y. Definition prop_in111 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d2 -> in_mem z d3 -> P x y z. Definition prop_in12 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d2 -> in_mem z d2 -> P x y z. Definition prop_in21 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d1 -> in_mem z d2 -> P x y z. Definition prop_in3 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d1 -> in_mem z d1 -> P x y z. Variable f : T1 -> T2. Definition prop_on1 Pf P & phantom T3 (Pf f) & ph {all1 P} := forall x, in_mem (f x) d2 -> P x. Definition prop_on2 Pf P & phantom T3 (Pf f) & ph {all2 P} := forall x y, in_mem (f x) d2 -> in_mem (f y) d2 -> P x y. End LocalProperties. Definition inPhantom := Phantom Prop. Definition onPhantom {T} P (x : T) := Phantom Prop (P x). Definition bijective_in aT rT (d : mem_pred aT) (f : aT -> rT) := exists2 g, prop_in1 d (inPhantom (cancel f g)) & prop_on1 d (Phantom _ (cancel g)) (onPhantom (cancel g) f). Definition bijective_on aT rT (cd : mem_pred rT) (f : aT -> rT) := exists2 g, prop_on1 cd (Phantom _ (cancel f)) (onPhantom (cancel f) g) & prop_in1 cd (inPhantom (cancel g f)). Notation "{ 'for' x , P }" := (prop_for x (inPhantom P)) : type_scope. Notation "{ 'in' d , P }" := (prop_in1 (mem d) (inPhantom P)) : type_scope. Notation "{ 'in' d1 & d2 , P }" := (prop_in11 (mem d1) (mem d2) (inPhantom P)) : type_scope. Notation "{ 'in' d & , P }" := (prop_in2 (mem d) (inPhantom P)) : type_scope. Notation "{ 'in' d1 & d2 & d3 , P }" := (prop_in111 (mem d1) (mem d2) (mem d3) (inPhantom P)) : type_scope. Notation "{ 'in' d1 & & d3 , P }" := (prop_in21 (mem d1) (mem d3) (inPhantom P)) : type_scope. Notation "{ 'in' d1 & d2 & , P }" := (prop_in12 (mem d1) (mem d2) (inPhantom P)) : type_scope. Notation "{ 'in' d & & , P }" := (prop_in3 (mem d) (inPhantom P)) : type_scope. Notation "{ 'on' cd , P }" := (prop_on1 (mem cd) (inPhantom P) (inPhantom P)) : type_scope. Notation "{ 'on' cd & , P }" := (prop_on2 (mem cd) (inPhantom P) (inPhantom P)) : type_scope. Local Arguments onPhantom : clear scopes. Notation "{ 'on' cd , P & g }" := (prop_on1 (mem cd) (Phantom (_ -> Prop) P) (onPhantom P g)) : type_scope. Notation "{ 'in' d , 'bijective' f }" := (bijective_in (mem d) f) : type_scope. Notation "{ 'on' cd , 'bijective' f }" := (bijective_on (mem cd) f) : type_scope. (** Weakening and monotonicity lemmas for localized predicates. Note that using these lemmas in backward reasoning will force expansion of the predicate definition, as Coq needs to expose the quantifier to apply these lemmas. We define a few specialized variants to avoid this for some of the ssrfun predicates. **) Section LocalGlobal. Variables T1 T2 T3 : predArgType. Variables (D1 : {pred T1}) (D2 : {pred T2}) (D3 : {pred T3}). Variables (d1 d1' : mem_pred T1) (d2 d2' : mem_pred T2) (d3 d3' : mem_pred T3). Variables (f f' : T1 -> T2) (g : T2 -> T1) (h : T3). Variables (P1 : T1 -> Prop) (P2 : T1 -> T2 -> Prop). Variable P3 : T1 -> T2 -> T3 -> Prop. Variable Q1 : (T1 -> T2) -> T1 -> Prop. Variable Q1l : (T1 -> T2) -> T3 -> T1 -> Prop. Variable Q2 : (T1 -> T2) -> T1 -> T1 -> Prop. Hypothesis sub1 : sub_mem d1 d1'. Hypothesis sub2 : sub_mem d2 d2'. Hypothesis sub3 : sub_mem d3 d3'. Lemma in1W : {all1 P1} -> {in D1, {all1 P1}}. Proof. by move=> ? ?. Qed. Lemma in2W : {all2 P2} -> {in D1 & D2, {all2 P2}}. Proof. by move=> ? ?. Qed. Lemma in3W : {all3 P3} -> {in D1 & D2 & D3, {all3 P3}}. Proof. by move=> ? ?. Qed. Lemma in1T : {in T1, {all1 P1}} -> {all1 P1}. Proof. by move=> ? ?; auto. Qed. Lemma in2T : {in T1 & T2, {all2 P2}} -> {all2 P2}. Proof. by move=> ? ?; auto. Qed. Lemma in3T : {in T1 & T2 & T3, {all3 P3}} -> {all3 P3}. Proof. by move=> ? ?; auto. Qed. Lemma sub_in1 (Ph : ph {all1 P1}) : prop_in1 d1' Ph -> prop_in1 d1 Ph. Proof. by move=> allP x /sub1; apply: allP. Qed. Lemma sub_in11 (Ph : ph {all2 P2}) : prop_in11 d1' d2' Ph -> prop_in11 d1 d2 Ph. Proof. by move=> allP x1 x2 /sub1 d1x1 /sub2; apply: allP. Qed. Lemma sub_in111 (Ph : ph {all3 P3}) : prop_in111 d1' d2' d3' Ph -> prop_in111 d1 d2 d3 Ph. Proof. by move=> allP x1 x2 x3 /sub1 d1x1 /sub2 d2x2 /sub3; apply: allP. Qed. Let allQ1 f'' := {all1 Q1 f''}. Let allQ1l f'' h' := {all1 Q1l f'' h'}. Let allQ2 f'' := {all2 Q2 f''}. Lemma on1W : allQ1 f -> {on D2, allQ1 f}. Proof. by move=> ? ?. Qed. Lemma on1lW : allQ1l f h -> {on D2, allQ1l f & h}. Proof. by move=> ? ?. Qed. Lemma on2W : allQ2 f -> {on D2 &, allQ2 f}. Proof. by move=> ? ?. Qed. Lemma on1T : {on T2, allQ1 f} -> allQ1 f. Proof. by move=> ? ?; auto. Qed. Lemma on1lT : {on T2, allQ1l f & h} -> allQ1l f h. Proof. by move=> ? ?; auto. Qed. Lemma on2T : {on T2 &, allQ2 f} -> allQ2 f. Proof. by move=> ? ?; auto. Qed. Lemma subon1 (Phf : ph (allQ1 f)) (Ph : ph (allQ1 f)) : prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph. Proof. by move=> allQ x /sub2; apply: allQ. Qed. Lemma subon1l (Phf : ph (allQ1l f)) (Ph : ph (allQ1l f h)) : prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph. Proof. by move=> allQ x /sub2; apply: allQ. Qed. Lemma subon2 (Phf : ph (allQ2 f)) (Ph : ph (allQ2 f)) : prop_on2 d2' Phf Ph -> prop_on2 d2 Phf Ph. Proof. by move=> allQ x y /sub2=> d2fx /sub2; apply: allQ. Qed. Lemma can_in_inj : {in D1, cancel f g} -> {in D1 &, injective f}. Proof. by move=> fK x y /fK{2}<- /fK{2}<- ->. Qed. Lemma canLR_in x y : {in D1, cancel f g} -> y \in D1 -> x = f y -> g x = y. Proof. by move=> fK D1y ->; rewrite fK. Qed. Lemma canRL_in x y : {in D1, cancel f g} -> x \in D1 -> f x = y -> x = g y. Proof. by move=> fK D1x <-; rewrite fK. Qed. Lemma on_can_inj : {on D2, cancel f & g} -> {on D2 &, injective f}. Proof. by move=> fK x y /fK{2}<- /fK{2}<- ->. Qed. Lemma canLR_on x y : {on D2, cancel f & g} -> f y \in D2 -> x = f y -> g x = y. Proof. by move=> fK D2fy ->; rewrite fK. Qed. Lemma canRL_on x y : {on D2, cancel f & g} -> f x \in D2 -> f x = y -> x = g y. Proof. by move=> fK D2fx <-; rewrite fK. Qed. Lemma inW_bij : bijective f -> {in D1, bijective f}. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma onW_bij : bijective f -> {on D2, bijective f}. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma inT_bij : {in T1, bijective f} -> bijective f. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma onT_bij : {on T2, bijective f} -> bijective f. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma sub_in_bij (D1' : pred T1) : {subset D1 <= D1'} -> {in D1', bijective f} -> {in D1, bijective f}. Proof. by move=> subD [g' fK g'K]; exists g' => x; move/subD; [apply: fK | apply: g'K]. Qed. Lemma subon_bij (D2' : pred T2) : {subset D2 <= D2'} -> {on D2', bijective f} -> {on D2, bijective f}. Proof. by move=> subD [g' fK g'K]; exists g' => x; move/subD; [apply: fK | apply: g'K]. Qed. End LocalGlobal. Lemma sub_in2 T d d' (P : T -> T -> Prop) : sub_mem d d' -> forall Ph : ph {all2 P}, prop_in2 d' Ph -> prop_in2 d Ph. Proof. by move=> /= sub_dd'; apply: sub_in11. Qed. Lemma sub_in3 T d d' (P : T -> T -> T -> Prop) : sub_mem d d' -> forall Ph : ph {all3 P}, prop_in3 d' Ph -> prop_in3 d Ph. Proof. by move=> /= sub_dd'; apply: sub_in111. Qed. Lemma sub_in12 T1 T d1 d1' d d' (P : T1 -> T -> T -> Prop) : sub_mem d1 d1' -> sub_mem d d' -> forall Ph : ph {all3 P}, prop_in12 d1' d' Ph -> prop_in12 d1 d Ph. Proof. by move=> /= sub1 sub; apply: sub_in111. Qed. Lemma sub_in21 T T3 d d' d3 d3' (P : T -> T -> T3 -> Prop) : sub_mem d d' -> sub_mem d3 d3' -> forall Ph : ph {all3 P}, prop_in21 d' d3' Ph -> prop_in21 d d3 Ph. Proof. by move=> /= sub sub3; apply: sub_in111. Qed. Lemma equivalence_relP_in T (R : rel T) (A : pred T) : {in A & &, equivalence_rel R} <-> {in A, reflexive R} /\ {in A &, forall x y, R x y -> {in A, R x =1 R y}}. Proof. split=> [eqiR | [Rxx trR] x y z *]; last by split=> [|/trR-> //]; apply: Rxx. by split=> [x Ax|x y Ax Ay Rxy z Az]; [rewrite (eqiR x x) | rewrite (eqiR x y)]. Qed. Section MonoHomoMorphismTheory. Variables (aT rT sT : Type) (f : aT -> rT) (g : rT -> aT). Variables (aP : pred aT) (rP : pred rT) (aR : rel aT) (rR : rel rT). Lemma monoW : {mono f : x / aP x >-> rP x} -> {homo f : x / aP x >-> rP x}. Proof. by move=> hf x ax; rewrite hf. Qed. Lemma mono2W : {mono f : x y / aR x y >-> rR x y} -> {homo f : x y / aR x y >-> rR x y}. Proof. by move=> hf x y axy; rewrite hf. Qed. Hypothesis fgK : cancel g f. Lemma homoRL : {homo f : x y / aR x y >-> rR x y} -> forall x y, aR (g x) y -> rR x (f y). Proof. by move=> Hf x y /Hf; rewrite fgK. Qed. Lemma homoLR : {homo f : x y / aR x y >-> rR x y} -> forall x y, aR x (g y) -> rR (f x) y. Proof. by move=> Hf x y /Hf; rewrite fgK. Qed. Lemma homo_mono : {homo f : x y / aR x y >-> rR x y} -> {homo g : x y / rR x y >-> aR x y} -> {mono g : x y / rR x y >-> aR x y}. Proof. move=> mf mg x y; case: (boolP (rR _ _))=> [/mg //|]. by apply: contraNF=> /mf; rewrite !fgK. Qed. Lemma monoLR : {mono f : x y / aR x y >-> rR x y} -> forall x y, rR (f x) y = aR x (g y). Proof. by move=> mf x y; rewrite -{1}[y]fgK mf. Qed. Lemma monoRL : {mono f : x y / aR x y >-> rR x y} -> forall x y, rR x (f y) = aR (g x) y. Proof. by move=> mf x y; rewrite -{1}[x]fgK mf. Qed. Lemma can_mono : {mono f : x y / aR x y >-> rR x y} -> {mono g : x y / rR x y >-> aR x y}. Proof. by move=> mf x y /=; rewrite -mf !fgK. Qed. End MonoHomoMorphismTheory. Section MonoHomoMorphismTheory_in. Variables (aT rT sT : predArgType) (f : aT -> rT) (g : rT -> aT). Variable (aD : {pred aT}). Variable (aP : pred aT) (rP : pred rT) (aR : rel aT) (rR : rel rT). Notation rD := [pred x | g x \in aD]. Lemma monoW_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in aD &, {homo f : x y / aR x y >-> rR x y}}. Proof. by move=> hf x y hx hy axy; rewrite hf. Qed. Lemma mono2W_in : {in aD, {mono f : x / aP x >-> rP x}} -> {in aD, {homo f : x / aP x >-> rP x}}. Proof. by move=> hf x hx ax; rewrite hf. Qed. Hypothesis fgK_on : {on aD, cancel g & f}. Lemma homoRL_in : {in aD &, {homo f : x y / aR x y >-> rR x y}} -> {in rD & aD, forall x y, aR (g x) y -> rR x (f y)}. Proof. by move=> Hf x y hx hy /Hf; rewrite fgK_on //; apply. Qed. Lemma homoLR_in : {in aD &, {homo f : x y / aR x y >-> rR x y}} -> {in aD & rD, forall x y, aR x (g y) -> rR (f x) y}. Proof. by move=> Hf x y hx hy /Hf; rewrite fgK_on //; apply. Qed. Lemma homo_mono_in : {in aD &, {homo f : x y / aR x y >-> rR x y}} -> {in rD &, {homo g : x y / rR x y >-> aR x y}} -> {in rD &, {mono g : x y / rR x y >-> aR x y}}. Proof. move=> mf mg x y hx hy; case: (boolP (rR _ _))=> [/mg //|]; first exact. by apply: contraNF=> /mf; rewrite !fgK_on //; apply. Qed. Lemma monoLR_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in aD & rD, forall x y, rR (f x) y = aR x (g y)}. Proof. by move=> mf x y hx hy; rewrite -{1}[y]fgK_on // mf. Qed. Lemma monoRL_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in rD & aD, forall x y, rR x (f y) = aR (g x) y}. Proof. by move=> mf x y hx hy; rewrite -{1}[x]fgK_on // mf. Qed. Lemma can_mono_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in rD &, {mono g : x y / rR x y >-> aR x y}}. Proof. by move=> mf x y hx hy /=; rewrite -mf // !fgK_on. Qed. End MonoHomoMorphismTheory_in. coq-8.11.0/plugins/ssr/ssreflect.v0000644000175000017500000007274213612664533016753 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* .doc { font-family: monospace; white-space: pre; } # **) Require Import Bool. (* For bool_scope delimiter 'bool'. *) Require Import ssrmatching. Declare ML Module "ssreflect_plugin". (** This file is the Gallina part of the ssreflect plugin implementation. Files that use the ssreflect plugin should always Require ssreflect and either Import ssreflect or Import ssreflect.SsrSyntax. Part of the contents of this file is technical and will only interest advanced developers; in addition the following are defined: #[#the str of v by f#]# == the Canonical s : str such that f s = v. #[#the str of v#]# == the Canonical s : str that coerces to v. argumentType c == the T such that c : forall x : T, P x. returnType c == the R such that c : T -> R. {type of c for s} == P s where c : forall x : T, P x. nonPropType == an interface for non-Prop Types: a nonPropType coerces to a Type, and only types that do _not_ have sort Prop are canonical nonPropType instances. This is useful for applied views (see mid-file comment). notProp T == the nonPropType instance for type T. phantom T v == singleton type with inhabitant Phantom T v. phant T == singleton type with inhabitant Phant v. =^~ r == the converse of rewriting rule r (e.g., in a rewrite multirule). unkeyed t == t, but treated as an unkeyed matching pattern by the ssreflect matching algorithm. nosimpl t == t, but on the right-hand side of Definition C := nosimpl disables expansion of C by /=. locked t == t, but locked t is not convertible to t. locked_with k t == t, but not convertible to t or locked_with k' t unless k = k' (with k : unit). Coq type-checking will be much more efficient if locked_with with a bespoke k is used for sealed definitions. unlockable v == interface for sealed constant definitions of v. Unlockable def == the unlockable that registers def : C = v. #[#unlockable of C#]# == a clone for C of the canonical unlockable for the definition of C (e.g., if it uses locked_with). #[#unlockable fun C#]# == #[#unlockable of C#]# with the expansion forced to be an explicit lambda expression. -> The usage pattern for ADT operations is: Definition foo_def x1 .. xn := big_foo_expression. Fact foo_key : unit. Proof. by #[# #]#. Qed. Definition foo := locked_with foo_key foo_def. Canonical foo_unlockable := #[#unlockable fun foo#]#. This minimizes the comparison overhead for foo, while still allowing rewrite unlock to expose big_foo_expression. More information about these definitions and their use can be found in the ssreflect manual, and in specific comments below. **) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Module SsrSyntax. (** Declare Ssr keywords: 'is' 'of' '//' '/=' and '//='. We also declare the parsing level 8, as a workaround for a notation grammar factoring problem. Arguments of application-style notations (at level 10) should be declared at level 8 rather than 9 or the camlp5 grammar will not factor properly. **) Reserved Notation "(* x 'is' y 'of' z 'isn't' // /= //= *)" (at level 8). Reserved Notation "(* 69 *)" (at level 69). (** Non ambiguous keyword to check if the SsrSyntax module is imported **) Reserved Notation "(* Use to test if 'SsrSyntax_is_Imported' *)" (at level 8). Reserved Notation "" (at level 0, n at level 0, format ""). Reserved Notation "T (* n *)" (at level 200, format "T (* n *)"). End SsrSyntax. Export SsrMatchingSyntax. Export SsrSyntax. (** Save primitive notation that will be overloaded. **) Local Notation CoqGenericIf c vT vF := (if c then vT else vF) (only parsing). Local Notation CoqGenericDependentIf c x R vT vF := (if c as x return R then vT else vF) (only parsing). Local Notation CoqCast x T := (x : T) (only parsing). (** Reserve notation that introduced in this file. **) Reserved Notation "'if' c 'then' vT 'else' vF" (at level 200, c, vT, vF at level 200, only parsing). Reserved Notation "'if' c 'return' R 'then' vT 'else' vF" (at level 200, c, R, vT, vF at level 200, only parsing). Reserved Notation "'if' c 'as' x 'return' R 'then' vT 'else' vF" (at level 200, c, R, vT, vF at level 200, x ident, only parsing). Reserved Notation "x : T" (at level 100, right associativity, format "'[hv' x '/ ' : T ']'"). Reserved Notation "T : 'Type'" (at level 100, format "T : 'Type'"). Reserved Notation "P : 'Prop'" (at level 100, format "P : 'Prop'"). Reserved Notation "[ 'the' sT 'of' v 'by' f ]" (at level 0, format "[ 'the' sT 'of' v 'by' f ]"). Reserved Notation "[ 'the' sT 'of' v ]" (at level 0, format "[ 'the' sT 'of' v ]"). Reserved Notation "{ 'type' 'of' c 'for' s }" (at level 0, format "{ 'type' 'of' c 'for' s }"). Reserved Notation "=^~ r" (at level 100, format "=^~ r"). Reserved Notation "[ 'unlockable' 'of' C ]" (at level 0, format "[ 'unlockable' 'of' C ]"). Reserved Notation "[ 'unlockable' 'fun' C ]" (at level 0, format "[ 'unlockable' 'fun' C ]"). (** To define notations for tactic in intro patterns. When "=> /t" is parsed, "t:%ssripat" is actually interpreted. **) Declare Scope ssripat_scope. Delimit Scope ssripat_scope with ssripat. (** Make the general "if" into a notation, so that we can override it below. The notations are "only parsing" because the Coq decompiler will not recognize the expansion of the boolean if; using the default printer avoids a spurious trailing %%GEN_IF. **) Declare Scope general_if_scope. Delimit Scope general_if_scope with GEN_IF. Notation "'if' c 'then' vT 'else' vF" := (CoqGenericIf c vT vF) (only parsing) : general_if_scope. Notation "'if' c 'return' R 'then' vT 'else' vF" := (CoqGenericDependentIf c c R vT vF) (only parsing) : general_if_scope. Notation "'if' c 'as' x 'return' R 'then' vT 'else' vF" := (CoqGenericDependentIf c x R vT vF) (only parsing) : general_if_scope. (** Force boolean interpretation of simple if expressions. **) Declare Scope boolean_if_scope. Delimit Scope boolean_if_scope with BOOL_IF. Notation "'if' c 'return' R 'then' vT 'else' vF" := (if c is true as c in bool return R then vT else vF) : boolean_if_scope. Notation "'if' c 'then' vT 'else' vF" := (if c%bool is true as _ in bool return _ then vT else vF) : boolean_if_scope. Notation "'if' c 'as' x 'return' R 'then' vT 'else' vF" := (if c%bool is true as x in bool return R then vT else vF) : boolean_if_scope. Open Scope boolean_if_scope. (** To allow a wider variety of notations without reserving a large number of of identifiers, the ssreflect library systematically uses "forms" to enclose complex mixfix syntax. A "form" is simply a mixfix expression enclosed in square brackets and introduced by a keyword: #[#keyword ... #]# Because the keyword follows a bracket it does not need to be reserved. Non-ssreflect libraries that do not respect the form syntax (e.g., the Coq Lists library) should be loaded before ssreflect so that their notations do not mask all ssreflect forms. **) Declare Scope form_scope. Delimit Scope form_scope with FORM. Open Scope form_scope. (** Allow overloading of the cast (x : T) syntax, put whitespace around the ":" symbol to avoid lexical clashes (and for consistency with the parsing precedence of the notation, which binds less tightly than application), and put printing boxes that print the type of a long definition on a separate line rather than force-fit it at the right margin. **) Notation "x : T" := (CoqCast x T) : core_scope. (** Allow the casual use of notations like nat * nat for explicit Type declarations. Note that (nat * nat : Type) is NOT equivalent to (nat * nat)%%type, whose inferred type is legacy type "Set". **) Notation "T : 'Type'" := (CoqCast T%type Type) (only parsing) : core_scope. (** Allow similarly Prop annotation for, e.g., rewrite multirules. **) Notation "P : 'Prop'" := (CoqCast P%type Prop) (only parsing) : core_scope. (** Constants for abstract: and #[#: name #]# intro pattern **) Definition abstract_lock := unit. Definition abstract_key := tt. Definition abstract (statement : Type) (id : nat) (lock : abstract_lock) := let: tt := lock in statement. Declare Scope ssr_scope. Notation "" := (abstract _ n _) : ssr_scope. Notation "T (* n *)" := (abstract T n abstract_key) : ssr_scope. Open Scope ssr_scope. Register abstract_lock as plugins.ssreflect.abstract_lock. Register abstract_key as plugins.ssreflect.abstract_key. Register abstract as plugins.ssreflect.abstract. (** Constants for tactic-views **) Inductive external_view : Type := tactic_view of Type. (** Syntax for referring to canonical structures: #[#the struct_type of proj_val by proj_fun#]# This form denotes the Canonical instance s of the Structure type struct_type whose proj_fun projection is proj_val, i.e., such that proj_fun s = proj_val. Typically proj_fun will be A record field accessors of struct_type, but this need not be the case; it can be, for instance, a field of a record type to which struct_type coerces; proj_val will likewise be coerced to the return type of proj_fun. In all but the simplest cases, proj_fun should be eta-expanded to allow for the insertion of implicit arguments. In the common case where proj_fun itself is a coercion, the "by" part can be omitted entirely; in this case it is inferred by casting s to the inferred type of proj_val. Obviously the latter can be fixed by using an explicit cast on proj_val, and it is highly recommended to do so when the return type intended for proj_fun is "Type", as the type inferred for proj_val may vary because of sort polymorphism (it could be Set or Prop). Note when using the #[#the _ of _ #]# form to generate a substructure from a telescopes-style canonical hierarchy (implementing inheritance with coercions), one should always project or coerce the value to the BASE structure, because Coq will only find a Canonical derived structure for the Canonical base structure -- not for a base structure that is specific to proj_value. **) Module TheCanonical. #[universes(template)] Variant put vT sT (v1 v2 : vT) (s : sT) := Put. Definition get vT sT v s (p : @put vT sT v v s) := let: Put _ _ _ := p in s. Definition get_by vT sT of sT -> vT := @get vT sT. End TheCanonical. Import TheCanonical. (* Note: no export. *) Local Arguments get_by _%type_scope _%type_scope _ _ _ _. Notation "[ 'the' sT 'of' v 'by' f ]" := (@get_by _ sT f _ _ ((fun v' (s : sT) => Put v' (f s) s) v _)) (only parsing) : form_scope. Notation "[ 'the' sT 'of' v ]" := (get ((fun s : sT => Put v (*coerce*) s s) _)) (only parsing) : form_scope. (** The following are "format only" versions of the above notations. We need to do this to prevent the formatter from being be thrown off by application collapsing, coercion insertion and beta reduction in the right hand side of the notations above. **) Notation "[ 'the' sT 'of' v 'by' f ]" := (@get_by _ sT f v _ _) (only printing) : form_scope. Notation "[ 'the' sT 'of' v ]" := (@get _ sT v _ _) (only printing) : form_scope. (** We would like to recognize Notation " #[# 'the' sT 'of' v : 'Type' #]#" := (@get Type sT v _ _) (at level 0, format " #[# 'the' sT 'of' v : 'Type' #]#") : form_scope. **) (** Helper notation for canonical structure inheritance support. This is a workaround for the poor interaction between delta reduction and canonical projections in Coq's unification algorithm, by which transparent definitions hide canonical instances, i.e., in Canonical a_type_struct := @Struct a_type ... Definition my_type := a_type. my_type doesn't effectively inherit the struct structure from a_type. Our solution is to redeclare the instance as follows Canonical my_type_struct := Eval hnf in #[#struct of my_type#]#. The special notation #[#str of _ #]# must be defined for each Strucure "str" with constructor "Str", typically as follows Definition clone_str s := let: Str _ x y ... z := s return {type of Str for s} -> str in fun k => k _ x y ... z. Notation " #[# 'str' 'of' T 'for' s #]#" := (@clone_str s (@Str T)) (at level 0, format " #[# 'str' 'of' T 'for' s #]#") : form_scope. Notation " #[# 'str' 'of' T #]#" := (repack_str (fun x => @Str T x)) (at level 0, format " #[# 'str' 'of' T #]#") : form_scope. The notation for the match return predicate is defined below; the eta expansion in the second form serves both to distinguish it from the first and to avoid the delta reduction problem. There are several variations on the notation and the definition of the the "clone" function, for telescopes, mixin classes, and join (multiple inheritance) classes. We describe a different idiom for clones in ssrfun; it uses phantom types (see below) and static unification; see fintype and ssralg for examples. **) Definition argumentType T P & forall x : T, P x := T. Definition dependentReturnType T P & forall x : T, P x := P. Definition returnType aT rT & aT -> rT := rT. Notation "{ 'type' 'of' c 'for' s }" := (dependentReturnType c s) : type_scope. (** A generic "phantom" type (actually, a unit type with a phantom parameter). This type can be used for type definitions that require some Structure on one of their parameters, to allow Coq to infer said structure so it does not have to be supplied explicitly or via the " #[#the _ of _ #]#" notation (the latter interacts poorly with other Notation). The definition of a (co)inductive type with a parameter p : p_type, that needs to use the operations of a structure Structure p_str : Type := p_Str {p_repr :> p_type; p_op : p_repr -> ...} should be given as Inductive indt_type (p : p_str) := Indt ... . Definition indt_of (p : p_str) & phantom p_type p := indt_type p. Notation "{ 'indt' p }" := (indt_of (Phantom p)). Definition indt p x y ... z : {indt p} := @Indt p x y ... z. Notation " #[# 'indt' x y ... z #]#" := (indt x y ... z). That is, the concrete type and its constructor should be shadowed by definitions that use a phantom argument to infer and display the true value of p (in practice, the "indt" constructor often performs additional functions, like "locking" the representation -- see below). We also define a simpler version ("phant" / "Phant") of phantom for the common case where p_type is Type. **) #[universes(template)] Variant phantom T (p : T) := Phantom. Arguments phantom : clear implicits. Arguments Phantom : clear implicits. #[universes(template)] Variant phant (p : Type) := Phant. (** Internal tagging used by the implementation of the ssreflect elim. **) Definition protect_term (A : Type) (x : A) : A := x. Register protect_term as plugins.ssreflect.protect_term. (** The ssreflect idiom for a non-keyed pattern: - unkeyed t will match any subterm that unifies with t, regardless of whether it displays the same head symbol as t. - unkeyed t a b will match any application of a term f unifying with t, to two arguments unifying with with a and b, respectively, regardless of apparent head symbols. - unkeyed x where x is a variable will match any subterm with the same type as x (when x would raise the 'indeterminate pattern' error). **) Notation unkeyed x := (let flex := x in flex). (** Ssreflect converse rewrite rule rule idiom. **) Definition ssr_converse R (r : R) := (Logic.I, r). Notation "=^~ r" := (ssr_converse r) : form_scope. (** Term tagging (user-level). The ssreflect library uses four strengths of term tagging to restrict convertibility during type checking: nosimpl t simplifies to t EXCEPT in a definition; more precisely, given Definition foo := nosimpl bar, foo (or foo t') will NOT be expanded by the /= and //= switches unless it is in a forcing context (e.g., in match foo t' with ... end, foo t' will be reduced if this allows the match to be reduced). Note that nosimpl bar is simply notation for a a term that beta-iota reduces to bar; hence rewrite /foo will replace foo by bar, and rewrite -/foo will replace bar by foo. CAVEAT: nosimpl should not be used inside a Section, because the end of section "cooking" removes the iota redex. locked t is provably equal to t, but is not convertible to t; 'locked' provides support for selective rewriting, via the lock t : t = locked t Lemma, and the ssreflect unlock tactic. locked_with k t is equal but not convertible to t, much like locked t, but supports explicit tagging with a value k : unit. This is used to mitigate a flaw in the term comparison heuristic of the Coq kernel, which treats all terms of the form locked t as equal and compares their arguments recursively, leading to an exponential blowup of comparison. For this reason locked_with should be used rather than locked when defining ADT operations. The unlock tactic does not support locked_with but the unlock rewrite rule does, via the unlockable interface. we also use Module Type ascription to create truly opaque constants, because simple expansion of constants to reveal an unreducible term doubles the time complexity of a negative comparison. Such opaque constants can be expanded generically with the unlock rewrite rule. See the definition of card and subset in fintype for examples of this. **) Notation nosimpl t := (let: tt := tt in t). Lemma master_key : unit. Proof. exact tt. Qed. Definition locked A := let: tt := master_key in fun x : A => x. Register master_key as plugins.ssreflect.master_key. Register locked as plugins.ssreflect.locked. Lemma lock A x : x = locked x :> A. Proof. unlock; reflexivity. Qed. (** Needed for locked predicates, in particular for eqType's. **) Lemma not_locked_false_eq_true : locked false <> true. Proof. unlock; discriminate. Qed. (** The basic closing tactic "done". **) Ltac done := trivial; hnf; intros; solve [ do ![solve [trivial | apply: sym_equal; trivial] | discriminate | contradiction | split] | case not_locked_false_eq_true; assumption | match goal with H : ~ _ |- _ => solve [case H; trivial] end ]. (** Quicker done tactic not including split, syntax: /0/ **) Ltac ssrdone0 := trivial; hnf; intros; solve [ do ![solve [trivial | apply: sym_equal; trivial] | discriminate | contradiction ] | case not_locked_false_eq_true; assumption | match goal with H : ~ _ |- _ => solve [case H; trivial] end ]. (** To unlock opaque constants. **) #[universes(template)] Structure unlockable T v := Unlockable {unlocked : T; _ : unlocked = v}. Lemma unlock T x C : @unlocked T x C = x. Proof. by case: C. Qed. Notation "[ 'unlockable' 'of' C ]" := (@Unlockable _ _ C (unlock _)) : form_scope. Notation "[ 'unlockable' 'fun' C ]" := (@Unlockable _ (fun _ => _) C (unlock _)) : form_scope. (** Generic keyed constant locking. **) (** The argument order ensures that k is always compared before T. **) Definition locked_with k := let: tt := k in fun T x => x : T. (** This can be used as a cheap alternative to cloning the unlockable instance below, but with caution as unkeyed matching can be expensive. **) Lemma locked_withE T k x : unkeyed (locked_with k x) = x :> T. Proof. by case: k. Qed. (** Intensionaly, this instance will not apply to locked u. **) Canonical locked_with_unlockable T k x := @Unlockable T x (locked_with k x) (locked_withE k x). (** More accurate variant of unlock, and safer alternative to locked_withE. **) Lemma unlock_with T k x : unlocked (locked_with_unlockable k x) = x :> T. Proof. exact: unlock. Qed. (** The internal lemmas for the have tactics. **) Definition ssr_have Plemma Pgoal (step : Plemma) rest : Pgoal := rest step. Arguments ssr_have Plemma [Pgoal]. Definition ssr_have_let Pgoal Plemma step (rest : let x : Plemma := step in Pgoal) : Pgoal := rest. Arguments ssr_have_let [Pgoal]. Register ssr_have as plugins.ssreflect.ssr_have. Register ssr_have_let as plugins.ssreflect.ssr_have_let. Definition ssr_suff Plemma Pgoal step (rest : Plemma) : Pgoal := step rest. Arguments ssr_suff Plemma [Pgoal]. Definition ssr_wlog := ssr_suff. Arguments ssr_wlog Plemma [Pgoal]. Register ssr_suff as plugins.ssreflect.ssr_suff. Register ssr_wlog as plugins.ssreflect.ssr_wlog. (** Internal N-ary congruence lemmas for the congr tactic. **) Fixpoint nary_congruence_statement (n : nat) : (forall B, (B -> B -> Prop) -> Prop) -> Prop := match n with | O => fun k => forall B, k B (fun x1 x2 : B => x1 = x2) | S n' => let k' A B e (f1 f2 : A -> B) := forall x1 x2, x1 = x2 -> (e (f1 x1) (f2 x2) : Prop) in fun k => forall A, nary_congruence_statement n' (fun B e => k _ (k' A B e)) end. Lemma nary_congruence n (k := fun B e => forall y : B, (e y y : Prop)) : nary_congruence_statement n k. Proof. have: k _ _ := _; rewrite {1}/k. elim: n k => [|n IHn] k k_P /= A; first exact: k_P. by apply: IHn => B e He; apply: k_P => f x1 x2 <-. Qed. Lemma ssr_congr_arrow Plemma Pgoal : Plemma = Pgoal -> Plemma -> Pgoal. Proof. by move->. Qed. Arguments ssr_congr_arrow : clear implicits. Register nary_congruence as plugins.ssreflect.nary_congruence. Register ssr_congr_arrow as plugins.ssreflect.ssr_congr_arrow. (** View lemmas that don't use reflection. **) Section ApplyIff. Variables P Q : Prop. Hypothesis eqPQ : P <-> Q. Lemma iffLR : P -> Q. Proof. by case: eqPQ. Qed. Lemma iffRL : Q -> P. Proof. by case: eqPQ. Qed. Lemma iffLRn : ~P -> ~Q. Proof. by move=> nP tQ; case: nP; case: eqPQ tQ. Qed. Lemma iffRLn : ~Q -> ~P. Proof. by move=> nQ tP; case: nQ; case: eqPQ tP. Qed. End ApplyIff. Hint View for move/ iffLRn|2 iffRLn|2 iffLR|2 iffRL|2. Hint View for apply/ iffRLn|2 iffLRn|2 iffRL|2 iffLR|2. (** To focus non-ssreflect tactics on a subterm, eg vm_compute. Usage: elim/abstract_context: (pattern) => G defG. vm_compute; rewrite {}defG {G}. Note that vm_cast are not stored in the proof term for reductions occurring in the context, hence set here := pattern; vm_compute in (value of here) blows up at Qed time. **) Lemma abstract_context T (P : T -> Type) x : (forall Q, Q = P -> Q x) -> P x. Proof. by move=> /(_ P); apply. Qed. (*****************************************************************************) (* Material for under/over (to rewrite under binders using "context lemmas") *) Require Export ssrunder. Hint Extern 0 (@Under_rel.Over_rel _ _ _ _) => solve [ apply: Under_rel.over_rel_done ] : core. Hint Resolve Under_rel.over_rel_done : core. Register Under_rel.Under_rel as plugins.ssreflect.Under_rel. Register Under_rel.Under_rel_from_rel as plugins.ssreflect.Under_rel_from_rel. (** Closing rewrite rule *) Definition over := over_rel. (** Closing tactic *) Ltac over := by [ apply: Under_rel.under_rel_done | rewrite over ]. (** Convenience rewrite rule to unprotect evars, e.g., to instantiate them in another way than with reflexivity. *) Definition UnderE := Under_relE. (*****************************************************************************) (** An interface for non-Prop types; used to avoid improper instantiation of polymorphic lemmas with on-demand implicits when they are used as views. For example: Some_inj {T} : forall x y : T, Some x = Some y -> x = y. Using move/Some_inj on a goal of the form Some n = Some 0 will fail: SSReflect will interpret the view as @Some_inj ?T _top_assumption_ since this is the well-typed application of the view with the minimal number of inserted evars (taking ?T := Some n = Some 0), and then will later complain that it cannot erase _top_assumption_ after having abstracted the viewed assumption. Making x and y maximal implicits would avoid this and force the intended @Some_inj nat x y _top_assumption_ interpretation, but is undesirable as it makes it harder to use Some_inj with the many SSReflect and MathComp lemmas that have an injectivity premise. Specifying {T : nonPropType} solves this more elegantly, as then (?T : Type) no longer unifies with (Some n = Some 0), which has sort Prop. **) Module NonPropType. (** Implementation notes: We rely on three interface Structures: - test_of r, the middle structure, performs the actual check: it has two canonical instances whose 'condition' projection are maybeProj (?P : Prop) and tt, and which set r := true and r := false, respectively. Unifying condition (?t : test_of ?r) with maybeProj T will thus set ?r to true if T is in Prop as the test_Prop T instance will apply, and otherwise simplify maybeProp T to tt and use the test_negative instance and set ?r to false. - call_of c r sets up a call to test_of on condition c with expected result r. It has a default instance for its 'callee' projection to Type, which sets c := maybeProj T and r := false when unifying with a type T. - type is a telescope on call_of c r, which checks that unifying test_of ?r1 with c indeed sets ?r1 := r; the type structure bundles the 'test' instance and its 'result' value along with its call_of c r projection. The default instance essentially provides eta-expansion for 'type'. This is only essential for the first 'result' projection to bool; using the instance for other projection merely avoids spurious delta expansions that would spoil the notProp T notation. In detail, unifying T =~= ?S with ?S : nonPropType, i.e., (1) T =~= @callee (@condition (result ?S) (test ?S)) (result ?S) (frame ?S) first uses the default call instance with ?T := T to reduce (1) to (2a) @condition (result ?S) (test ?S) =~= maybeProp T (3) result ?S =~= false (4) frame ?S =~= call T along with some trivial universe-related checks which are irrelevant here. Then the unification tries to use the test_Prop instance to reduce (2a) to (6a) result ?S =~= true (7a) ?P =~= T with ?P : Prop (8a) test ?S =~= test_Prop ?P Now the default 'check' instance with ?result := true resolves (6a) as (9a) ?S := @check true ?test ?frame Then (7a) can be solved precisely if T has sort at most (hence exactly) Prop, and then (8a) is solved by the check instance, yielding ?test := test_Prop T, and completing the solution of (2a), and _committing_ to it. But now (3) is inconsistent with (9a), and this makes the entire problem (1) fails. If on the othe hand T does not have sort Prop then (7a) fails and the unification resorts to delta expanding (2a), which gives (2b) @condition (result ?S) (test ?S) =~= tt which is then reduced, using the test_negative instance, to (6b) result ?S =~= false (8b) test ?S =~= test_negative Both are solved using the check default instance, as in the (2a) branch, giving (9b) ?S := @check false test_negative ?frame Then (3) and (4) are similarly soved using check, giving the final assignment (9) ?S := notProp T Observe that we _must_ perform the actual test unification on the arguments of the initial canonical instance, and not on the instance itself as we do in mathcomp/matrix and mathcomp/vector, because we want the unification to fail when T has sort Prop. If both the test_of _and_ the result check unifications were done as part of the structure telescope then the latter would be a sub-problem of the former, and thus failing the check would merely make the test_of unification backtrack and delta-expand and we would not get failure. **) Structure call_of (condition : unit) (result : bool) := Call {callee : Type}. Definition maybeProp (T : Type) := tt. Definition call T := Call (maybeProp T) false T. Structure test_of (result : bool) := Test {condition :> unit}. Definition test_Prop (P : Prop) := Test true (maybeProp P). Definition test_negative := Test false tt. Structure type := Check {result : bool; test : test_of result; frame : call_of test result}. Definition check result test frame := @Check result test frame. Module Exports. Canonical call. Canonical test_Prop. Canonical test_negative. Canonical check. Notation nonPropType := type. Coercion callee : call_of >-> Sortclass. Coercion frame : type >-> call_of. Notation notProp T := (@check false test_negative (call T)). End Exports. End NonPropType. Export NonPropType.Exports. coq-8.11.0/plugins/ssr/ssrelim.ml0000644000175000017500000006072413612664533016577 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ctx, EConstr.destRel sigma hd, not (EConstr.Vars.noccurn sigma 1 t), Array.length args, t | CastType (t, _) -> loop ctx t | ProdType (x, ty, t) -> loop (RelDecl.LocalAssum (x, ty) :: ctx) t | LetInType (x,b,ty,t) -> loop (RelDecl.LocalDef (x, b, ty) :: ctx) (EConstr.Vars.subst1 b t) | _ -> let env' = EConstr.push_rel_context ctx env in let t' = Reductionops.whd_all env' sigma t in if not (EConstr.eq_constr sigma t t') then loop ctx t' else errorstrm Pp.(str"The eliminator has the wrong shape."++spc()++ str"A (applied) bound variable was expected as the conclusion of "++ str"the eliminator's"++Pp.cut()++str"type:"++spc()++pr_econstr_env env' sigma elimty) in let ctx, pred_id, elim_is_dep, n_pred_args,concl = loop [] elimty in let n_elim_args = Context.Rel.nhyps ctx in let is_rec_elim = let count_occurn n term = let count = ref 0 in let rec occur_rec n c = match EConstr.kind sigma c with | Rel m -> if m = n then incr count | _ -> EConstr.iter_with_binders sigma succ occur_rec n c in occur_rec n term; !count in let occurr2 n t = count_occurn n t > 1 in not (List.for_all_i (fun i (_,rd) -> pred_id <= i || not (occurr2 (pred_id - i) rd)) 1 (assums_of_rel_context ctx)) in n_elim_args - pred_id, n_elim_args, is_rec_elim, elim_is_dep, n_pred_args, (ctx,concl) let subgoals_tys sigma (relctx, concl) = let rec aux cur_depth acc = function | hd :: rest -> let ty = Context.Rel.Declaration.get_type hd in if EConstr.Vars.noccurn sigma cur_depth concl && List.for_all_i (fun i -> function | Context.Rel.Declaration.LocalAssum(_, t) -> EConstr.Vars.noccurn sigma i t | Context.Rel.Declaration.LocalDef (_, b, t) -> EConstr.Vars.noccurn sigma i t && EConstr.Vars.noccurn sigma i b) 1 rest then aux (cur_depth - 1) (ty :: acc) rest else aux (cur_depth - 1) acc rest | [] -> Array.of_list (List.rev acc) in aux (List.length relctx) [] (List.rev relctx) (* A case without explicit dependent terms but with both a view and an *) (* occurrence switch and/or an equation is treated as dependent, with the *) (* viewed term as the dependent term (the occurrence switch would be *) (* meaningless otherwise). When both a view and explicit dependents are *) (* present, it is forbidden to put a (meaningless) occurrence switch on *) (* the viewed term. *) (* This is both elim and case (defaulting to the former). If ~elim is omitted * the standard eliminator is chosen. The code is made of 4 parts: * 1. find the eliminator if not given as ~elim and analyze it * 2. build the patterns to be matched against the conclusion, looking at * (occ, c), deps and the pattern inferred from the type of the eliminator * 3. build the new predicate matching the patterns, and the tactic to * generalize the equality in case eqid is not None * 4. build the tactic handle instructions and clears as required in ipats and * by eqid *) let get_eq_type gl = let eq, gl = pf_fresh_global Coqlib.(lib_ref "core.eq.type") gl in gl, EConstr.of_constr eq let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = let open Proofview.Notations in Proofview.tclEVARMAP >>= begin fun sigma -> (* some sanity checks *) match what with | `EConstr(_,_,t) when EConstr.isEvar sigma t -> anomaly "elim called on a constr evar" | `EGen (_, g) when elim = None && is_wildcard g -> errorstrm Pp.(str"Indeterminate pattern and no eliminator") | `EGen ((Some clr,occ), g) when is_wildcard g -> Proofview.tclUNIT (None, clr, occ, None) | `EGen ((None, occ), g) when is_wildcard g -> Proofview.tclUNIT (None,[],occ,None) | `EGen ((_, occ), p as gen) -> pfLIFT (pf_interp_gen true gen) >>= fun (_,c,clr) -> Proofview.tclUNIT (Some c, clr, occ, Some p) | `EConstr (clr, occ, c) -> Proofview.tclUNIT (Some c, clr, occ, None) end >>= fun (oc, orig_clr, occ, c_gen) -> pfLIFT begin fun gl -> let orig_gl, concl, env = gl, pf_concl gl, pf_env gl in ppdebug(lazy(Pp.str(if is_case then "==CASE==" else "==ELIM=="))); let fire_subst gl t = Reductionops.nf_evar (project gl) t in let is_undef_pat = function | sigma, T t -> EConstr.isEvar sigma (EConstr.of_constr t) | _ -> false in let match_pat env p occ h cl = let sigma0 = project orig_gl in ppdebug(lazy Pp.(str"matching: " ++ pr_occ occ ++ pp_pattern env p)); let (c,ucst), cl = fill_occ_pattern ~raise_NoMatch:true env sigma0 (EConstr.Unsafe.to_constr cl) p occ h in ppdebug(lazy Pp.(str" got: " ++ pr_constr_env env sigma0 c)); c, EConstr.of_constr cl, ucst in let mkTpat gl t = (* takes a term, refreshes it and makes a T pattern *) let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in let t, _, _, sigma = saturate ~beta:true env (project gl) t n in Evd.merge_universe_context sigma ucst, T (EConstr.Unsafe.to_constr t) in let unif_redex gl (sigma, r as p) t = (* t is a hint for the redex of p *) let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in let t, _, _, sigma = saturate ~beta:true env sigma t n in let sigma = Evd.merge_universe_context sigma ucst in match r with | X_In_T (e, p) -> sigma, E_As_X_In_T (EConstr.Unsafe.to_constr t, e, p) | _ -> try unify_HO env sigma t (EConstr.of_constr (fst (redex_of_pattern env p))), r with e when CErrors.noncritical e -> p in (* finds the eliminator applies it to evars and c saturated as needed *) (* obtaining "elim ??? (c ???)". pred is the higher order evar *) (* cty is None when the user writes _ (hence we can't make a pattern *) (* `seed` represents the array of types from which we derive the name seeds for the block intro patterns *) let seed, cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl = match elim with | Some elim -> let gl, elimty = pf_e_type_of gl elim in let elimty = let rename_elimty r = EConstr.of_constr (Arguments_renaming.rename_type (EConstr.to_constr ~abort_on_undefined_evars:false (project gl) elimty) r) in match EConstr.kind (project gl) elim with | Constr.Var kn -> rename_elimty (GlobRef.VarRef kn) | Constr.Const (kn,_) -> rename_elimty (GlobRef.ConstRef kn) | _ -> elimty in let pred_id, n_elim_args, is_rec, elim_is_dep, n_pred_args,ctx_concl = analyze_eliminator elimty env (project gl) in let seed = subgoals_tys (project gl) ctx_concl in let elim, elimty, elim_args, gl = pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in let pred = List.assoc pred_id elim_args in let elimty = Reductionops.whd_all env (project gl) elimty in let cty, gl = if Option.is_empty oc then None, gl else let c = Option.get oc in let gl, c_ty = pfe_type_of gl c in let pc = match c_gen with | Some p -> interp_cpattern orig_gl p None | _ -> mkTpat gl c in Some(c, c_ty, pc), gl in seed, cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl | None -> let c = Option.get oc in let gl, c_ty = pfe_type_of gl c in let ((kn, i),_ as indu), unfolded_c_ty = pf_reduce_to_quantified_ind gl c_ty in let sort = Tacticals.elimination_sort_of_goal gl in let gl, elim = if not is_case then let t,gl= pf_fresh_global (Indrec.lookup_eliminator env (kn,i) sort) gl in gl, t else Tacmach.pf_eapply (fun env sigma () -> let indu = (fst indu, EConstr.EInstance.kind sigma (snd indu)) in let (sigma, ind) = Indrec.build_case_analysis_scheme env sigma indu true sort in (sigma, ind)) gl () in let elim = EConstr.of_constr elim in let gl, elimty = pfe_type_of gl elim in let pred_id,n_elim_args,is_rec,elim_is_dep,n_pred_args,ctx_concl = analyze_eliminator elimty env (project gl) in let seed = if is_case then let mind,indb = Inductive.lookup_mind_specif env (kn,i) in let tys = indb.Declarations.mind_nf_lc in let renamed_tys = Array.mapi (fun j (ctx, cty) -> let t = Term.it_mkProd_or_LetIn cty ctx in ppdebug(lazy Pp.(str "Search" ++ Printer.pr_constr_env env (project gl) t)); let t = Arguments_renaming.rename_type t (GlobRef.ConstructRef((kn,i),j+1)) in ppdebug(lazy Pp.(str"Done Search " ++ Printer.pr_constr_env env (project gl) t)); t) tys in let drop_params x = snd @@ EConstr.decompose_prod_n_assum (project gl) mind.Declarations.mind_nparams (EConstr.of_constr x) in Array.map drop_params renamed_tys else subgoals_tys (project gl) ctx_concl in let rctx = fst (EConstr.decompose_prod_assum (project gl) unfolded_c_ty) in let n_c_args = Context.Rel.length rctx in let c, c_ty, t_args, gl = pf_saturate gl c ~ty:c_ty n_c_args in let elim, elimty, elim_args, gl = pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in let pred = List.assoc pred_id elim_args in let pc = match n_c_args, c_gen with | 0, Some p -> interp_cpattern orig_gl p None | _ -> mkTpat gl c in let cty = Some (c, c_ty, pc) in let elimty = Reductionops.whd_all env (project gl) elimty in seed, cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl in let () = let sigma = project gl in ppdebug(lazy Pp.(str"elim= "++ pr_econstr_pat env sigma elim)); ppdebug(lazy Pp.(str"elimty= "++ pr_econstr_pat env sigma elimty)) in let inf_deps_r = match EConstr.kind_of_type (project gl) elimty with | AtomicType (_, args) -> List.rev (Array.to_list args) | _ -> assert false in let saturate_until gl c c_ty f = let rec loop n = try let c, c_ty, _, gl = pf_saturate gl c ~ty:c_ty n in let gl' = f c c_ty gl in Some (c, c_ty, gl, gl') with | NotEnoughProducts -> None | e when CErrors.noncritical e -> loop (n+1) in loop 0 in (* Here we try to understand if the main pattern/term the user gave is * the first pattern to be matched (i.e. if elimty ends in P t1 .. tn, * weather tn is the t the user wrote in 'elim: t' *) let c_is_head_p, gl = match cty with | None -> true, gl (* The user wrote elim: _ *) | Some (c, c_ty, _) -> let rec first = function | [] -> errorstrm Pp.(str"Unable to apply the eliminator to the term"++ spc()++pr_econstr_env env (project gl) c++spc()) | x :: rest -> match x () with | None -> first rest | Some (b,gl) -> b, gl in (* Unify two terms if their heads are not applied unif variables, eg * not (?P x). The idea is to rule out cases where the problem is too * vague to drive the current heuristics. *) let pf_unify_HO_rigid gl a b = let is_applied_evar x = match EConstr.kind (project gl) x with | App(x,_) -> EConstr.isEvar (project gl) x | _ -> false in if is_applied_evar a || is_applied_evar b then raise Evarconv.(UnableToUnify(project gl, Pretype_errors.ProblemBeyondCapabilities)) else pf_unify_HO gl a b in let try_c_last_arg () = (* we try to see if c unifies with the last arg of elim *) if elim_is_dep then None else let arg = List.assoc (n_elim_args - 1) elim_args in let gl, arg_ty = pfe_type_of gl arg in match saturate_until gl c c_ty (fun c c_ty gl -> pf_unify_HO (pf_unify_HO_rigid gl c_ty arg_ty) arg c) with | Some (c, _, _, gl) -> Some (false, gl) | None -> None in let try_c_last_pattern () = (* we try to see if c unifies with the last inferred pattern *) if inf_deps_r = [] then None else let inf_arg = List.hd inf_deps_r in let gl, inf_arg_ty = pfe_type_of gl inf_arg in match saturate_until gl c c_ty (fun _ c_ty gl -> pf_unify_HO_rigid gl c_ty inf_arg_ty) with | Some (c, _, _,gl) -> Some(true, gl) | None -> None in first [try_c_last_arg;try_c_last_pattern] in ppdebug(lazy Pp.(str"c_is_head_p= " ++ bool c_is_head_p)); let gl, predty = pfe_type_of gl pred in (* Patterns for the inductive types indexes to be bound in pred are computed * looking at the ones provided by the user and the inferred ones looking at * the type of the elimination principle *) let pp_pat (_,p,_,occ) = Pp.(pr_occ occ ++ pp_pattern env p) in let pp_inf_pat gl (_,_,t,_) = pr_econstr_pat env (project gl) (fire_subst gl t) in let patterns, clr, gl = let rec loop patterns clr i = function | [],[] -> patterns, clr, gl | ((oclr, occ), t):: deps, inf_t :: inf_deps -> let p = interp_cpattern orig_gl t None in let clr_t = interp_clr (project gl) (oclr,(tag_of_cpattern t,EConstr.of_constr (fst (redex_of_pattern env p)))) in (* if we are the index for the equation we do not clear *) let clr_t = if deps = [] && eqid <> None then [] else clr_t in let p = if is_undef_pat p then mkTpat gl inf_t else p in loop (patterns @ [i, p, inf_t, occ]) (clr_t @ clr) (i+1) (deps, inf_deps) | [], c :: inf_deps -> ppdebug(lazy Pp.(str"adding inf pattern " ++ pr_econstr_pat env (project gl) c)); loop (patterns @ [i, mkTpat gl c, c, allocc]) clr (i+1) ([], inf_deps) | _::_, [] -> errorstrm Pp.(str "Too many dependent abstractions") in let deps, head_p, inf_deps_r = match what, c_is_head_p, cty with | `EConstr _, _, None -> anomaly "Simple elim with no term" | _, false, _ -> deps, [], inf_deps_r | `EGen gen, true, None -> deps @ [gen], [], inf_deps_r | _, true, Some (c, _, pc) -> let occ = if occ = None then allocc else occ in let inf_p, inf_deps_r = List.hd inf_deps_r, List.tl inf_deps_r in deps, [1, pc, inf_p, occ], inf_deps_r in let patterns, clr, gl = loop [] orig_clr (List.length head_p+1) (List.rev deps, inf_deps_r) in head_p @ patterns, Util.List.uniquize clr, gl in ppdebug(lazy Pp.(pp_concat (str"patterns=") (List.map pp_pat patterns))); ppdebug(lazy Pp.(pp_concat (str"inf. patterns=") (List.map (pp_inf_pat gl) patterns))); (* Predicate generation, and (if necessary) tactic to generalize the * equation asked by the user *) let elim_pred, gen_eq_tac, clr, gl = let error gl t inf_t = errorstrm Pp.(str"The given pattern matches the term"++ spc()++pp_term gl t++spc()++str"while the inferred pattern"++ spc()++pr_econstr_pat env (project gl) (fire_subst gl inf_t)++spc()++ str"doesn't") in let match_or_postpone (cl, gl, post) (h, p, inf_t, occ) = let p = unif_redex gl p inf_t in if is_undef_pat p then let () = ppdebug(lazy Pp.(str"postponing " ++ pp_pattern env p)) in cl, gl, post @ [h, p, inf_t, occ] else try let c, cl, ucst = match_pat env p occ h cl in let gl = pf_merge_uc ucst gl in let c = EConstr.of_constr c in let gl = try pf_unify_HO gl inf_t c with exn when CErrors.noncritical exn -> error gl c inf_t in cl, gl, post with | NoMatch | NoProgress -> let e, ucst = redex_of_pattern env p in let gl = pf_merge_uc ucst gl in let e = EConstr.of_constr e in let n, e, _, _ucst = pf_abs_evars gl (fst p, e) in let e, _, _, gl = pf_saturate ~beta:true gl e n in let gl = try pf_unify_HO gl inf_t e with exn when CErrors.noncritical exn -> error gl e inf_t in cl, gl, post in let rec match_all concl gl patterns = let concl, gl, postponed = List.fold_left match_or_postpone (concl, gl, []) patterns in if postponed = [] then concl, gl else if List.length postponed = List.length patterns then errorstrm Pp.(str "Some patterns are undefined even after all"++spc()++ str"the defined ones matched") else match_all concl gl postponed in let concl, gl = match_all concl gl patterns in let pred_rctx, _ = EConstr.decompose_prod_assum (project gl) (fire_subst gl predty) in let concl, gen_eq_tac, clr, gl = match eqid with | Some (IPatId _) when not is_rec -> let k = List.length deps in let c = fire_subst gl (List.assoc (n_elim_args - k - 1) elim_args) in let gl, t = pfe_type_of gl c in let gl, eq = get_eq_type gl in let gen_eq_tac, eq_ty, gl = let refl = EConstr.mkApp (eq, [|t; c; c|]) in let new_concl = EConstr.mkArrow refl Sorts.Relevant (EConstr.Vars.lift 1 (pf_concl orig_gl)) in let new_concl = fire_subst gl new_concl in let erefl, gl = mkRefl t c gl in let erefl = fire_subst gl erefl in let erefl_ty = Retyping.get_type_of (pf_env gl) (project gl) erefl in let eq_ty = Retyping.get_type_of (pf_env gl) (project gl) erefl_ty in let gen_eq_tac s = let open Evd in let sigma = merge_universe_context s.sigma (evar_universe_context (project gl)) in apply_type new_concl [erefl] { s with sigma } in gen_eq_tac, eq_ty, gl in let rel = k + if c_is_head_p then 1 else 0 in let src, gl = mkProt eq_ty EConstr.(mkApp (eq,[|t; c; mkRel rel|])) gl in let concl = EConstr.mkArrow src Sorts.Relevant (EConstr.Vars.lift 1 concl) in let clr = if deps <> [] then clr else [] in concl, gen_eq_tac, clr, gl | _ -> concl, Tacticals.tclIDTAC, clr, gl in let mk_lam t r = EConstr.mkLambda_or_LetIn r t in let concl = List.fold_left mk_lam concl pred_rctx in let gl, concl = if eqid <> None && is_rec then let gl, concls = pfe_type_of gl concl in let concl, gl = mkProt concls concl gl in let gl, _ = pfe_type_of gl concl in gl, concl else gl, concl in concl, gen_eq_tac, clr, gl in let gl, pty = pf_e_type_of gl elim_pred in ppdebug(lazy Pp.(str"elim_pred=" ++ pp_term gl elim_pred)); ppdebug(lazy Pp.(str"elim_pred_ty=" ++ pp_term gl pty)); let gl = pf_unify_HO gl pred elim_pred in let elim = fire_subst gl elim in let gl = pf_resolve_typeclasses ~where:elim ~fail:false gl in let gl, _ = pf_e_type_of gl elim in (* check that the patterns do not contain non instantiated dependent metas *) let () = let evars_of_term = Evarutil.undefined_evars_of_term (project gl) in let patterns = List.map (fun (_,_,t,_) -> fire_subst gl t) patterns in let patterns_ev = List.map evars_of_term patterns in let ev = List.fold_left Evar.Set.union Evar.Set.empty patterns_ev in let ty_ev = Evar.Set.fold (fun i e -> let ex = i in let i_ty = Evd.evar_concl (Evd.find (project gl) ex) in Evar.Set.union e (evars_of_term i_ty)) ev Evar.Set.empty in let inter = Evar.Set.inter ev ty_ev in if not (Evar.Set.is_empty inter) then begin let i = Evar.Set.choose inter in let pat = List.find (fun t -> Evar.Set.mem i (evars_of_term t)) patterns in errorstrm Pp.(str"Pattern"++spc()++pr_econstr_pat env (project gl) pat++spc()++ str"was not completely instantiated and one of its variables"++spc()++ str"occurs in the type of another non-instantiated pattern variable"); end in (* the elim tactic, with the eliminator and the predicated we computed *) let elim = project gl, elim in let seed = Array.map (fun ty -> let ctx,_ = EConstr.decompose_prod_assum (project gl) ty in CList.rev_map Context.Rel.Declaration.get_name ctx) seed in (elim,seed,clr,is_rec,gen_eq_tac), orig_gl end >>= fun (elim, seed,clr,is_rec,gen_eq_tac) -> let elim_tac = Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (refine_with ~with_evars:false elim); cleartac clr] in let gen_eq_tac = Proofview.V82.tactic gen_eq_tac in Tacticals.New.tclTHENLIST [gen_eq_tac; elim_intro_tac ?seed:(Some seed) what eqid elim_tac is_rec clr] ;; let elimtac x = let k ?seed:_ _what _eqid elim_tac _is_rec _clr = elim_tac in ssrelim ~is_case:false [] (`EConstr ([],None,x)) None k let casetac x k = let k ?seed _what _eqid elim_tac _is_rec _clr = k ?seed elim_tac in ssrelim ~is_case:true [] (`EConstr ([],None,x)) None k let pf_nb_prod gl = nb_prod (project gl) (pf_concl gl) let rev_id = mk_internal_id "rev concl" let injecteq_id = mk_internal_id "injection equation" let revtoptac n0 gl = let n = pf_nb_prod gl - n0 in let dc, cl = EConstr.decompose_prod_n_assum (project gl) n (pf_concl gl) in let dc' = dc @ [Context.Rel.Declaration.LocalAssum(make_annot (Name rev_id) Sorts.Relevant, EConstr.it_mkProd_or_LetIn cl (List.rev dc))] in let f = EConstr.it_mkLambda_or_LetIn (mkEtaApp (EConstr.mkRel (n + 1)) (-n) 1) dc' in Refiner.refiner ~check:true EConstr.Unsafe.(to_constr (EConstr.mkApp (f, [|Evarutil.mk_new_meta ()|]))) gl let equality_inj l b id c gl = let msg = ref "" in try Proofview.V82.of_tactic (Equality.inj None l b None c) gl with | Gramlib.Ploc.Exc(_,CErrors.UserError (_,s)) | CErrors.UserError (_,s) when msg := Pp.string_of_ppcmds s; !msg = "Not a projectable equality but a discriminable one." || !msg = "Nothing to inject." -> Feedback.msg_warning (Pp.str !msg); discharge_hyp (id, (id, "")) gl let injectidl2rtac id c gl = Tacticals.tclTHEN (equality_inj None true id c) (revtoptac (pf_nb_prod gl)) gl let injectl2rtac sigma c = match EConstr.kind sigma c with | Var id -> injectidl2rtac id (EConstr.mkVar id, NoBindings) | _ -> let id = injecteq_id in let xhavetac id c = Proofview.V82.of_tactic (Tactics.pose_proof (Name id) c) in Tacticals.tclTHENLIST [xhavetac id c; injectidl2rtac id (EConstr.mkVar id, NoBindings); Proofview.V82.of_tactic (Tactics.clear [id])] let is_injection_case c gl = let gl, cty = pfe_type_of gl c in let (mind,_), _ = pf_reduce_to_quantified_ind gl cty in Coqlib.check_ind_ref "core.eq.type" mind let perform_injection c gl = let gl, cty = pfe_type_of gl c in let mind, t = pf_reduce_to_quantified_ind gl cty in let dc, eqt = EConstr.decompose_prod (project gl) t in if dc = [] then injectl2rtac (project gl) c gl else if not (EConstr.Vars.closed0 (project gl) eqt) then CErrors.user_err (Pp.str "can't decompose a quantified equality") else let cl = pf_concl gl in let n = List.length dc in let c_eq = mkEtaApp c n 2 in let cl1 = EConstr.mkLambda EConstr.(make_annot Anonymous Sorts.Relevant, mkArrow eqt Sorts.Relevant cl, mkApp (mkRel 1, [|c_eq|])) in let id = injecteq_id in let id_with_ebind = (EConstr.mkVar id, NoBindings) in let injtac = Tacticals.tclTHEN (introid id) (injectidl2rtac id id_with_ebind) in Tacticals.tclTHENLAST (Proofview.V82.of_tactic (Tactics.apply (EConstr.compose_lam dc cl1))) injtac gl let ssrscase_or_inj_tac c = Proofview.V82.tactic ~nf_evars:false (fun gl -> if is_injection_case c gl then perform_injection c gl else Proofview.V82.of_tactic (casetac c (fun ?seed:_ k -> k)) gl) coq-8.11.0/plugins/ssr/ssrequality.ml0000644000175000017500000007377613612664533017521 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* !ssroldreworder); optdepr = false; optwrite = (fun b -> ssroldreworder := b) }) (** The "simpl" tactic *) (* We must avoid zeta-converting any "let"s created by the "in" tactical. *) let tacred_simpl gl = let simpl_expr = Genredexpr.( Simpl(Redops.make_red_flag[FBeta;FMatch;FZeta;FDeltaBut []],None)) in let esimpl, _ = Redexpr.reduction_of_red_expr (pf_env gl) simpl_expr in let esimpl e sigma c = let (_,t) = esimpl e sigma c in t in let simpl env sigma c = (esimpl env sigma c) in simpl let safe_simpltac n gl = if n = ~-1 then let cl= red_safe (tacred_simpl gl) (pf_env gl) (project gl) (pf_concl gl) in Proofview.V82.of_tactic (convert_concl_no_check cl) gl else ssr_n_tac "simpl" n gl let simpltac = function | Simpl n -> safe_simpltac n | Cut n -> tclTRY (donetac n) | SimplCut (n,m) -> tclTHEN (safe_simpltac m) (tclTRY (donetac n)) | Nop -> tclIDTAC (** The "congr" tactic *) let interp_congrarg_at ist gl n rf ty m = ppdebug(lazy Pp.(str"===interp_congrarg_at===")); let congrn, _ = mkSsrRRef "nary_congruence" in let args1 = mkRnat n :: mkRHoles n @ [ty] in let args2 = mkRHoles (3 * n) in let rec loop i = if i + n > m then None else try let rt = mkRApp congrn (args1 @ mkRApp rf (mkRHoles i) :: args2) in ppdebug(lazy Pp.(str"rt=" ++ Printer.pr_glob_constr_env (pf_env gl) rt)); Some (interp_refine ist gl rt) with _ -> loop (i + 1) in loop 0 let pattern_id = mk_internal_id "pattern value" let congrtac ((n, t), ty) ist gl = ppdebug(lazy (Pp.str"===congr===")); ppdebug(lazy Pp.(str"concl=" ++ Printer.pr_econstr_env (pf_env gl) (project gl) (Tacmach.pf_concl gl))); let sigma, _ as it = interp_term ist gl t in let gl = pf_merge_uc_of sigma gl in let _, f, _, _ucst = pf_abs_evars gl it in let ist' = {ist with lfun = Id.Map.add pattern_id (Tacinterp.Value.of_constr f) Id.Map.empty } in let rf = mkRltacVar pattern_id in let m = pf_nbargs gl f in let _, cf = if n > 0 then match interp_congrarg_at ist' gl n rf ty m with | Some cf -> cf | None -> errorstrm Pp.(str "No " ++ int n ++ str "-congruence with " ++ pr_term t) else let rec loop i = if i > m then errorstrm Pp.(str "No congruence with " ++ pr_term t) else match interp_congrarg_at ist' gl i rf ty m with | Some cf -> cf | None -> loop (i + 1) in loop 1 in tclTHEN (refine_with cf) (tclTRY (Proofview.V82.of_tactic Tactics.reflexivity)) gl let pf_typecheck t gl = let it = sig_it gl in let sigma,_ = pf_type_of gl t in re_sig [it] sigma let newssrcongrtac arg ist gl = ppdebug(lazy Pp.(str"===newcongr===")); ppdebug(lazy Pp.(str"concl=" ++ Printer.pr_econstr_env (pf_env gl) (project gl) (pf_concl gl))); (* utils *) let fs gl t = Reductionops.nf_evar (project gl) t in let tclMATCH_GOAL (c, gl_c) proj t_ok t_fail gl = match try Some (pf_unify_HO gl_c (pf_concl gl) c) with exn when CErrors.noncritical exn -> None with | Some gl_c -> tclTHEN (Proofview.V82.of_tactic (convert_concl ~check:true (fs gl_c c))) (t_ok (proj gl_c)) gl | None -> t_fail () gl in let mk_evar gl ty = let env, sigma, si = pf_env gl, project gl, sig_it gl in let sigma = Evd.create_evar_defs sigma in let (sigma, x) = Evarutil.new_evar env sigma ty in x, re_sig si sigma in let arr, gl = pf_mkSsrConst "ssr_congr_arrow" gl in let ssr_congr lr = EConstr.mkApp (arr, lr) in let eq, gl = pf_fresh_global Coqlib.(lib_ref "core.eq.type") gl in (* here the two cases: simple equality or arrow *) let equality, _, eq_args, gl' = pf_saturate gl (EConstr.of_constr eq) 3 in tclMATCH_GOAL (equality, gl') (fun gl' -> fs gl' (List.assoc 0 eq_args)) (fun ty -> congrtac (arg, Detyping.detype Detyping.Now false Id.Set.empty (pf_env gl) (project gl) ty) ist) (fun () -> let gl', t_lhs = pfe_new_type gl in let gl', t_rhs = pfe_new_type gl' in let lhs, gl' = mk_evar gl' t_lhs in let rhs, gl' = mk_evar gl' t_rhs in let arrow = EConstr.mkArrow lhs Sorts.Relevant (EConstr.Vars.lift 1 rhs) in tclMATCH_GOAL (arrow, gl') (fun gl' -> [|fs gl' lhs;fs gl' rhs|]) (fun lr -> let a = ssr_congr lr in tclTHENLIST [ pf_typecheck a ; Proofview.V82.of_tactic (Tactics.apply a) ; congrtac (arg, mkRType) ist ]) (fun _ _ -> errorstrm Pp.(str"Conclusion is not an equality nor an arrow"))) gl (** 7. Rewriting tactics (rewrite, unlock) *) (** Rewrite rules *) type ssrwkind = RWred of ssrsimpl | RWdef | RWeq type ssrrule = ssrwkind * ssrterm (** Rewrite arguments *) type ssrrwarg = (ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule) let notimes = 0 let nomult = 1, Once let mkocc occ = None, occ let noclr = mkocc None let mkclr clr = Some clr, None let nodocc = mkclr [] let is_rw_cut = function RWred (Cut _) -> true | _ -> false let mk_rwarg (d, (n, _ as m)) ((clr, occ as docc), rx) (rt, _ as r) : ssrrwarg = if rt <> RWeq then begin if rt = RWred Nop && not (m = nomult && occ = None && rx = None) && (clr = None || clr = Some []) then anomaly "Improper rewrite clear switch"; if d = R2L && rt <> RWdef then CErrors.user_err (Pp.str "Right-to-left switch on simplification"); if n <> 1 && is_rw_cut rt then CErrors.user_err (Pp.str "Bad or useless multiplier"); if occ <> None && rx = None && rt <> RWdef then CErrors.user_err (Pp.str "Missing redex for simplification occurrence") end; (d, m), ((docc, rx), r) let norwmult = L2R, nomult let norwocc = noclr, None let simplintac occ rdx sim gl = let simptac m gl = if m <> ~-1 then begin if rdx <> None then CErrors.user_err (Pp.str "Custom simpl tactic does not support patterns"); if occ <> None then CErrors.user_err (Pp.str "Custom simpl tactic does not support occurrence numbers"); simpltac (Simpl m) gl end else let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let simp env c _ _ = EConstr.Unsafe.to_constr (red_safe Tacred.simpl env sigma0 (EConstr.of_constr c)) in Proofview.V82.of_tactic (convert_concl_no_check (EConstr.of_constr (eval_pattern env0 sigma0 (EConstr.Unsafe.to_constr concl0) rdx occ simp))) gl in match sim with | Simpl m -> simptac m gl | SimplCut (n,m) -> tclTHEN (simptac m) (tclTRY (donetac n)) gl | _ -> simpltac sim gl let rec get_evalref env sigma c = match EConstr.kind sigma c with | Var id -> EvalVarRef id | Const (k,_) -> EvalConstRef k | App (c', _) -> get_evalref env sigma c' | Cast (c', _, _) -> get_evalref env sigma c' | Proj(c,_) -> EvalConstRef(Projection.constant c) | _ -> errorstrm Pp.(str "The term " ++ pr_econstr_pat (Global.env ()) sigma c ++ str " is not unfoldable") (* Strip a pattern generated by a prenex implicit to its constant. *) let strip_unfold_term _ ((sigma, t) as p) kt = match EConstr.kind sigma t with | App (f, a) when kt = xNoFlag && Array.for_all (EConstr.isEvar sigma) a && EConstr.isConst sigma f -> (sigma, f), true | Const _ | Var _ -> p, true | Proj _ -> p, true | _ -> p, false let same_proj sigma t1 t2 = match EConstr.kind sigma t1, EConstr.kind sigma t2 with | Proj(c1,_), Proj(c2, _) -> Projection.equal c1 c2 | _ -> false let all_ok _ _ = true let fake_pmatcher_end () = mkProp, L2R, (Evd.empty, UState.empty, mkProp) let unfoldintac occ rdx t (kt,_) gl = let fs sigma x = Reductionops.nf_evar sigma x in let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let (sigma, t), const = strip_unfold_term env0 t kt in let body env t c = Tacred.unfoldn [AllOccurrences, get_evalref env sigma t] env sigma0 c in let easy = occ = None && rdx = None in let red_flags = if easy then CClosure.betaiotazeta else CClosure.betaiota in let beta env = Reductionops.clos_norm_flags red_flags env sigma0 in let unfold, conclude = match rdx with | Some (_, (In_T _ | In_X_In_T _)) | None -> let ise = Evd.create_evar_defs sigma in let ise, u = mk_tpattern env0 sigma0 (ise,EConstr.Unsafe.to_constr t) all_ok L2R (EConstr.Unsafe.to_constr t) in let find_T, end_T = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in (fun env c _ h -> try find_T env c h ~k:(fun env c _ _ -> EConstr.Unsafe.to_constr (body env t (EConstr.of_constr c))) with NoMatch when easy -> c | NoMatch | NoProgress -> errorstrm Pp.(str"No occurrence of " ++ pr_econstr_pat env sigma0 t ++ spc() ++ str "in " ++ Printer.pr_constr_env env sigma c)), (fun () -> try end_T () with | NoMatch when easy -> fake_pmatcher_end () | NoMatch -> anomaly "unfoldintac") | _ -> (fun env (c as orig_c) _ h -> if const then let rec aux c = match EConstr.kind sigma0 c with | Const _ when EConstr.eq_constr sigma0 c t -> body env t t | App (f,a) when EConstr.eq_constr sigma0 f t -> EConstr.mkApp (body env f f,a) | Proj _ when same_proj sigma0 c t -> body env t c | _ -> let c = Reductionops.whd_betaiotazeta sigma0 c in match EConstr.kind sigma0 c with | Const _ when EConstr.eq_constr sigma0 c t -> body env t t | App (f,a) when EConstr.eq_constr sigma0 f t -> EConstr.mkApp (body env f f,a) | Proj _ when same_proj sigma0 c t -> body env t c | Const f -> aux (body env c c) | App (f, a) -> aux (EConstr.mkApp (body env f f, a)) | _ -> errorstrm Pp.(str "The term "++ pr_constr_env env sigma orig_c++ str" contains no " ++ pr_econstr_env env sigma t ++ str" even after unfolding") in EConstr.Unsafe.to_constr @@ aux (EConstr.of_constr c) else try EConstr.Unsafe.to_constr @@ body env t (fs (unify_HO env sigma (EConstr.of_constr c) t) t) with _ -> errorstrm Pp.(str "The term " ++ pr_constr_env env sigma c ++spc()++ str "does not unify with " ++ pr_econstr_pat env sigma t)), fake_pmatcher_end in let concl = let concl0 = EConstr.Unsafe.to_constr concl0 in try beta env0 (EConstr.of_constr (eval_pattern env0 sigma0 concl0 rdx occ unfold)) with Option.IsNone -> errorstrm Pp.(str"Failed to unfold " ++ pr_econstr_pat env0 sigma t) in let _ = conclude () in Proofview.V82.of_tactic (convert_concl ~check:true concl) gl ;; let foldtac occ rdx ft gl = let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let sigma, t = ft in let t = EConstr.to_constr ~abort_on_undefined_evars:false sigma t in let fold, conclude = match rdx with | Some (_, (In_T _ | In_X_In_T _)) | None -> let ise = Evd.create_evar_defs sigma in let ut = EConstr.Unsafe.to_constr (red_product_skip_id env0 sigma (EConstr.of_constr t)) in let ise, ut = mk_tpattern env0 sigma0 (ise,t) all_ok L2R ut in let find_T, end_T = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[ut]) in (fun env c _ h -> try find_T env c h ~k:(fun env t _ _ -> t) with NoMatch ->c), (fun () -> try end_T () with NoMatch -> fake_pmatcher_end ()) | _ -> (fun env c _ h -> try let sigma = unify_HO env sigma (EConstr.of_constr c) (EConstr.of_constr t) in EConstr.to_constr ~abort_on_undefined_evars:false sigma (EConstr.of_constr t) with _ -> errorstrm Pp.(str "fold pattern " ++ pr_constr_pat env sigma t ++ spc () ++ str "does not match redex " ++ pr_constr_pat env sigma c)), fake_pmatcher_end in let concl0 = EConstr.Unsafe.to_constr concl0 in let concl = eval_pattern env0 sigma0 concl0 rdx occ fold in let _ = conclude () in Proofview.V82.of_tactic (convert_concl ~check:true (EConstr.of_constr concl)) gl ;; let converse_dir = function L2R -> R2L | R2L -> L2R let rw_progress rhs lhs ise = not (EConstr.eq_constr ise lhs (Evarutil.nf_evar ise rhs)) (* Coq has a more general form of "equation" (any type with a single *) (* constructor with no arguments with_rect_r elimination lemmas). *) (* However there is no clear way of determining the LHS and RHS of *) (* such a generic Leibniz equation -- short of inspecting the type *) (* of the elimination lemmas. *) let rec strip_prod_assum c = match Constr.kind c with | Prod (_, _, c') -> strip_prod_assum c' | LetIn (_, v, _, c') -> strip_prod_assum (subst1 v c) | Cast (c', _, _) -> strip_prod_assum c' | _ -> c let rule_id = mk_internal_id "rewrite rule" exception PRtype_error of (Environ.env * Evd.evar_map * Pretype_errors.pretype_error) option let id_map_redex _ sigma ~before:_ ~after = sigma, after let pirrel_rewrite ?(under=false) ?(map_redex=id_map_redex) pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl = (* ppdebug(lazy(str"sigma@pirrel_rewrite=" ++ pr_evar_map None sigma)); *) let env = pf_env gl in let beta = Reductionops.clos_norm_flags CClosure.beta env sigma in let sigma, new_rdx = map_redex env sigma ~before:rdx ~after:new_rdx in let sigma, p = (* The resulting goal *) Evarutil.new_evar env sigma (beta (EConstr.Vars.subst1 new_rdx pred)) in let pred = EConstr.mkNamedLambda (make_annot pattern_id Sorts.Relevant) rdx_ty pred in let sigma, elim = let sort = elimination_sort_of_goal gl in match Equality.eq_elimination_ref (dir = L2R) sort with | Some r -> Evd.fresh_global env sigma r | None -> let ((kn, i) as ind, _), unfolded_c_ty = Tacred.reduce_to_quantified_ind env sigma c_ty in let sort = elimination_sort_of_goal gl in let sigma, elim = Evd.fresh_global env sigma (Indrec.lookup_eliminator env ind sort) in if dir = R2L then sigma, elim else let elim, _ = EConstr.destConst sigma elim in let mp,l = Constant.repr2 (Constant.make1 (Constant.canonical elim)) in let l' = Label.of_id (Nameops.add_suffix (Label.to_id l) "_r") in let c1' = Global.constant_of_delta_kn (Constant.canonical (Constant.make2 mp l')) in sigma, EConstr.of_constr (mkConst c1') in let proof = EConstr.mkApp (elim, [| rdx_ty; new_rdx; pred; p; rdx; c |]) in (* We check the proof is well typed *) let sigma, proof_ty = try Typing.type_of env sigma proof with | Pretype_errors.PretypeError (env, sigma, te) -> raise (PRtype_error (Some (env, sigma, te))) | e when CErrors.noncritical e -> raise (PRtype_error None) in ppdebug(lazy Pp.(str"pirrel_rewrite: proof term: " ++ pr_econstr_env env sigma proof)); ppdebug(lazy Pp.(str"pirrel_rewrite of type: " ++ pr_econstr_env env sigma proof_ty)); try refine_with ~first_goes_last:(not !ssroldreworder || under) ~with_evars:under (sigma, proof) gl with _ -> (* we generate a msg like: "Unable to find an instance for the variable" *) let hd_ty, miss = match EConstr.kind sigma c with | App (hd, args) -> let hd_ty = Retyping.get_type_of env sigma hd in let names = let rec aux t = function 0 -> [] | n -> let t = Reductionops.whd_all env sigma t in match EConstr.kind_of_type sigma t with | ProdType (name, _, t) -> name.binder_name :: aux t (n-1) | _ -> assert false in aux hd_ty (Array.length args) in hd_ty, Util.List.map_filter (fun (t, name) -> let evs = Evar.Set.elements (Evarutil.undefined_evars_of_term sigma t) in let open_evs = List.filter (fun k -> Sorts.InProp <> Retyping.get_sort_family_of env sigma (Evd.evar_concl (Evd.find sigma k))) evs in if open_evs <> [] then Some name else None) (List.combine (Array.to_list args) names) | _ -> anomaly "rewrite rule not an application" in errorstrm Pp.(Himsg.explain_refiner_error env sigma (Logic.UnresolvedBindings miss)++ (Pp.fnl()++str"Rule's type:" ++ spc() ++ pr_econstr_env env sigma hd_ty)) ;; let rwcltac ?under ?map_redex cl rdx dir sr gl = let sr = let sigma, r = sr in let sigma = resolve_typeclasses ~where:r ~fail:false (pf_env gl) sigma in sigma, r in let n, r_n,_, ucst = pf_abs_evars gl sr in let r_n' = pf_abs_cterm gl n r_n in let r' = EConstr.Vars.subst_var pattern_id r_n' in let gl = pf_unsafe_merge_uc ucst gl in let rdxt = Retyping.get_type_of (pf_env gl) (fst sr) rdx in (* ppdebug(lazy(str"sigma@rwcltac=" ++ pr_evar_map None (fst sr))); *) ppdebug(lazy Pp.(str"r@rwcltac=" ++ pr_econstr_env (pf_env gl) (project gl) (snd sr))); let cvtac, rwtac, gl = if EConstr.Vars.closed0 (project gl) r' then let env, sigma, c, c_eq = pf_env gl, fst sr, snd sr, Coqlib.(lib_ref "core.eq.type") in let sigma, c_ty = Typing.type_of env sigma c in ppdebug(lazy Pp.(str"c_ty@rwcltac=" ++ pr_econstr_env env sigma c_ty)); match EConstr.kind_of_type sigma (Reductionops.whd_all env sigma c_ty) with | AtomicType(e, a) when Ssrcommon.is_ind_ref sigma e c_eq -> let new_rdx = if dir = L2R then a.(2) else a.(1) in pirrel_rewrite ?under ?map_redex cl rdx rdxt new_rdx dir (sigma,c) c_ty, tclIDTAC, gl | _ -> let cl' = EConstr.mkApp (EConstr.mkNamedLambda (make_annot pattern_id Sorts.Relevant) rdxt cl, [|rdx|]) in let sigma, _ = Typing.type_of env sigma cl' in let gl = pf_merge_uc_of sigma gl in Proofview.V82.of_tactic (convert_concl ~check:true cl'), rewritetac ?under dir r', gl else let dc, r2 = EConstr.decompose_lam_n_assum (project gl) n r' in let r3, _, r3t = try EConstr.destCast (project gl) r2 with _ -> errorstrm Pp.(str "no cast from " ++ pr_econstr_pat (pf_env gl) (project gl) (snd sr) ++ str " to " ++ pr_econstr_env (pf_env gl) (project gl) r2) in let cl' = EConstr.mkNamedProd (make_annot rule_id Sorts.Relevant) (EConstr.it_mkProd_or_LetIn r3t dc) (EConstr.Vars.lift 1 cl) in let cl'' = EConstr.mkNamedProd (make_annot pattern_id Sorts.Relevant) rdxt cl' in let itacs = [introid pattern_id; introid rule_id] in let cltac = Proofview.V82.of_tactic (Tactics.clear [pattern_id; rule_id]) in let rwtacs = [rewritetac ?under dir (EConstr.mkVar rule_id); cltac] in apply_type cl'' [rdx; EConstr.it_mkLambda_or_LetIn r3 dc], tclTHENLIST (itacs @ rwtacs), gl in let cvtac' _ = try cvtac gl with | PRtype_error e -> let error = Option.cata (fun (env, sigma, te) -> Pp.(fnl () ++ str "Type error was: " ++ Himsg.explain_pretype_error env sigma te)) (Pp.mt ()) e in if occur_existential (project gl) (Tacmach.pf_concl gl) then errorstrm Pp.(str "Rewriting impacts evars" ++ error) else errorstrm Pp.(str "Dependent type error in rewrite of " ++ pr_econstr_env (pf_env gl) (project gl) (EConstr.mkNamedLambda (make_annot pattern_id Sorts.Relevant) rdxt cl) ++ error) in tclTHEN cvtac' rwtac gl [@@@ocaml.warning "-3"] let lz_coq_prod = let prod = lazy (Coqlib.build_prod ()) in fun () -> Lazy.force prod let lz_setoid_relation = let sdir = ["Classes"; "RelationClasses"] in let last_srel = ref None in fun env -> match !last_srel with | Some (env', srel) when env' == env -> srel | _ -> let srel = try Some (UnivGen.constr_of_monomorphic_global @@ Coqlib.find_reference "Class_setoid" ("Coq"::sdir) "RewriteRelation" [@ocaml.warning "-3"]) with _ -> None in last_srel := Some (env, srel); srel let ssr_is_setoid env = match lz_setoid_relation env with | None -> fun _ _ _ -> false | Some srel -> fun sigma r args -> Rewrite.is_applied_rewrite_relation env sigma [] (EConstr.mkApp (r, args)) <> None let closed0_check cl p gl = if closed0 cl then errorstrm Pp.(str"No occurrence of redex "++ pr_constr_env (pf_env gl) (project gl) p) let dir_org = function L2R -> 1 | R2L -> 2 let rwprocess_rule dir rule gl = let env = pf_env gl in let coq_prod = lz_coq_prod () in let is_setoid = ssr_is_setoid env in let r_sigma, rules = let rec loop d sigma r t0 rs red = let t = if red = 1 then Tacred.hnf_constr env sigma t0 else Reductionops.whd_betaiotazeta sigma t0 in ppdebug(lazy Pp.(str"rewrule="++pr_econstr_pat env sigma t)); match EConstr.kind sigma t with | Prod (_, xt, at) -> let sigma = Evd.create_evar_defs sigma in let (sigma, x) = Evarutil.new_evar env sigma xt in loop d sigma EConstr.(mkApp (r, [|x|])) (EConstr.Vars.subst1 x at) rs 0 | App (pr, a) when is_ind_ref sigma pr coq_prod.Coqlib.typ -> let sr sigma = match EConstr.kind sigma (Tacred.hnf_constr env sigma r) with | App (c, ra) when is_construct_ref sigma c coq_prod.Coqlib.intro -> fun i -> ra.(i + 1), sigma | _ -> let ra = Array.append a [|r|] in function 1 -> let sigma, pi1 = Evd.fresh_global env sigma coq_prod.Coqlib.proj1 in EConstr.mkApp (pi1, ra), sigma | _ -> let sigma, pi2 = Evd.fresh_global env sigma coq_prod.Coqlib.proj2 in EConstr.mkApp (pi2, ra), sigma in let sigma,trty = Evd.fresh_global env sigma Coqlib.(lib_ref "core.True.type") in if EConstr.eq_constr sigma a.(0) trty then let s, sigma = sr sigma 2 in loop (converse_dir d) sigma s a.(1) rs 0 else let s, sigma = sr sigma 2 in let sigma, rs2 = loop d sigma s a.(1) rs 0 in let s, sigma = sr sigma 1 in loop d sigma s a.(0) rs2 0 | App (r_eq, a) when Hipattern.match_with_equality_type env sigma t != None -> let (ind, u) = EConstr.destInd sigma r_eq and rhs = Array.last a in let np = Inductiveops.inductive_nparamdecls env ind in let indu = (ind, EConstr.EInstance.kind sigma u) in let ind_ct = Inductiveops.type_of_constructors env indu in let lhs0 = last_arg sigma (EConstr.of_constr (strip_prod_assum ind_ct.(0))) in let rdesc = match EConstr.kind sigma lhs0 with | Rel i -> let lhs = a.(np - i) in let lhs, rhs = if d = L2R then lhs, rhs else rhs, lhs in (* msgnl (str "RW: " ++ pr_rwdir d ++ str " " ++ pr_constr_pat r ++ str " : " ++ pr_constr_pat lhs ++ str " ~> " ++ pr_constr_pat rhs); *) d, r, lhs, rhs (* let l_i, r_i = if d = L2R then i, 1 - ndep else 1 - ndep, i in let lhs = a.(np - l_i) and rhs = a.(np - r_i) in let a' = Array.copy a in let _ = a'.(np - l_i) <- mkVar pattern_id in let r' = mkCast (r, DEFAULTcast, mkApp (r_eq, a')) in (d, r', lhs, rhs) *) | _ -> let lhs = EConstr.Vars.substl (array_list_of_tl (Array.sub a 0 np)) lhs0 in let lhs, rhs = if d = R2L then lhs, rhs else rhs, lhs in let d' = if Array.length a = 1 then d else converse_dir d in d', r, lhs, rhs in sigma, rdesc :: rs | App (s_eq, a) when is_setoid sigma s_eq a -> let np = Array.length a and i = 3 - dir_org d in let lhs = a.(np - i) and rhs = a.(np + i - 3) in let a' = Array.copy a in let _ = a'.(np - i) <- EConstr.mkVar pattern_id in let r' = EConstr.mkCast (r, DEFAULTcast, EConstr.mkApp (s_eq, a')) in sigma, (d, r', lhs, rhs) :: rs | _ -> if red = 0 then loop d sigma r t rs 1 else errorstrm Pp.(str "not a rewritable relation: " ++ pr_econstr_pat env sigma t ++ spc() ++ str "in rule " ++ pr_econstr_pat env sigma (snd rule)) in let sigma, r = rule in let t = Retyping.get_type_of env sigma r in loop dir sigma r t [] 0 in r_sigma, rules let rwrxtac ?under ?map_redex occ rdx_pat dir rule gl = let env = pf_env gl in let r_sigma, rules = rwprocess_rule dir rule gl in let find_rule rdx = let rec rwtac = function | [] -> errorstrm Pp.(str "pattern " ++ pr_econstr_pat env (project gl) rdx ++ str " does not match " ++ pr_dir_side dir ++ str " of " ++ pr_econstr_pat env (project gl) (snd rule)) | (d, r, lhs, rhs) :: rs -> try let ise = unify_HO env (Evd.create_evar_defs r_sigma) lhs rdx in if not (rw_progress rhs rdx ise) then raise NoMatch else d, (ise, Evd.evar_universe_context ise, Reductionops.nf_evar ise r) with _ -> rwtac rs in rwtac rules in let sigma0, env0, concl0 = project gl, pf_env gl, pf_concl gl in let find_R, conclude = match rdx_pat with | Some (_, (In_T _ | In_X_In_T _)) | None -> let upats_origin = dir, EConstr.Unsafe.to_constr (snd rule) in let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) = let sigma, pat = let rw_progress rhs t evd = rw_progress rhs (EConstr.of_constr t) evd in mk_tpattern env sigma0 (sigma, EConstr.to_constr ~abort_on_undefined_evars:false sigma r) (rw_progress rhs) d (EConstr.to_constr ~abort_on_undefined_evars:false sigma lhs) in sigma, pats @ [pat] in let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in let find_R, end_R = mk_tpattern_matcher sigma0 occ ~upats_origin rpats in (fun e c _ i -> find_R ~k:(fun _ _ _ h -> mkRel h) e c i), fun cl -> let rdx,d,r = end_R () in closed0_check cl rdx gl; (d,r),rdx | Some(_, (T e | X_In_T (_,e) | E_As_X_In_T (e,_,_) | E_In_X_In_T (e,_,_))) -> let r = ref None in (fun env c _ h -> do_once r (fun () -> find_rule (EConstr.of_constr c), c); mkRel h), (fun concl -> closed0_check concl e gl; let (d,(ev,ctx,c)) , x = assert_done r in (d,(ev,ctx, EConstr.to_constr ~abort_on_undefined_evars:false ev c)) , x) in let concl0 = EConstr.Unsafe.to_constr concl0 in let concl = eval_pattern env0 sigma0 concl0 rdx_pat occ find_R in let (d, r), rdx = conclude concl in let r = Evd.merge_universe_context (pi1 r) (pi2 r), EConstr.of_constr (pi3 r) in rwcltac ?under ?map_redex (EConstr.of_constr concl) (EConstr.of_constr rdx) d r gl ;; let ssrinstancesofrule ist dir arg gl = let sigma0, env0, concl0 = project gl, pf_env gl, pf_concl gl in let rule = interp_term ist gl arg in let r_sigma, rules = rwprocess_rule dir rule gl in let find, conclude = let upats_origin = dir, EConstr.Unsafe.to_constr (snd rule) in let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) = let sigma, pat = let rw_progress rhs t evd = rw_progress rhs (EConstr.of_constr t) evd in mk_tpattern env sigma0 (sigma,EConstr.to_constr ~abort_on_undefined_evars:false sigma r) (rw_progress rhs) d (EConstr.to_constr ~abort_on_undefined_evars:false sigma lhs) in sigma, pats @ [pat] in let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in mk_tpattern_matcher ~all_instances:true ~raise_NoMatch:true sigma0 None ~upats_origin rpats in let print env p c _ = Feedback.msg_info Pp.(hov 1 (str"instance:" ++ spc() ++ pr_constr_env env r_sigma p ++ spc() ++ str "matches:" ++ spc() ++ pr_constr_env env r_sigma c)); c in Feedback.msg_info Pp.(str"BEGIN INSTANCES"); try while true do ignore(find env0 (EConstr.Unsafe.to_constr concl0) 1 ~k:print) done; raise NoMatch with NoMatch -> Feedback.msg_info Pp.(str"END INSTANCES"); tclIDTAC gl let ipat_rewrite occ dir c gl = rwrxtac occ None dir (project gl, c) gl let rwargtac ?under ?map_redex ist ((dir, mult), (((oclr, occ), grx), (kind, gt))) gl = let fail = ref false in let interp_rpattern gl gc = try interp_rpattern gl gc with _ when snd mult = May -> fail := true; project gl, T mkProp in let interp gc gl = try interp_term ist gl gc with _ when snd mult = May -> fail := true; (project gl, EConstr.mkProp) in let rwtac gl = let rx = Option.map (interp_rpattern gl) grx in let gl = match rx with | None -> gl | Some (s,_) -> pf_merge_uc_of s gl in let t = interp gt gl in let gl = pf_merge_uc_of (fst t) gl in (match kind with | RWred sim -> simplintac occ rx sim | RWdef -> if dir = R2L then foldtac occ rx t else unfoldintac occ rx t gt | RWeq -> rwrxtac ?under ?map_redex occ rx dir t) gl in let ctac = old_cleartac (interp_clr (project gl) (oclr, (fst gt, snd (interp gt gl)))) in if !fail then ctac gl else tclTHEN (tclMULT mult rwtac) ctac gl (** Rewrite argument sequence *) (* type ssrrwargs = ssrrwarg list *) (** The "rewrite" tactic *) let ssrrewritetac ?under ?map_redex ist rwargs = tclTHENLIST (List.map (rwargtac ?under ?map_redex ist) rwargs) (** The "unlock" tactic *) let unfoldtac occ ko t kt gl = let env = pf_env gl in let cl, c = pf_fill_occ_term gl occ (fst (strip_unfold_term env t kt)) in let cl' = EConstr.Vars.subst1 (pf_unfoldn [OnlyOccurrences [1], get_evalref env (project gl) c] gl c) cl in let f = if ko = None then CClosure.betaiotazeta else CClosure.betaiota in Proofview.V82.of_tactic (convert_concl ~check:true (pf_reduce (Reductionops.clos_norm_flags f) gl cl')) gl let unlocktac ist args gl = let utac (occ, gt) gl = unfoldtac occ occ (interp_term ist gl gt) (fst gt) gl in let locked, gl = pf_mkSsrConst "locked" gl in let key, gl = pf_mkSsrConst "master_key" gl in let ktacs = [ (fun gl -> unfoldtac None None (project gl,locked) xInParens gl); Proofview.V82.of_tactic (Ssrelim.casetac key (fun ?seed:_ k -> k)) ] in tclTHENLIST (List.map utac args @ ktacs) gl coq-8.11.0/plugins/ssr/ssrtacticals.mli0000644000175000017500000000320613612664533017761 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Tacinterp.Value.t -> Ssrast.ssrdir -> int Locus.or_var * (('a * Tacinterp.Value.t option list) * Tacinterp.Value.t option) -> Tacmach.tactic val tclCLAUSES : unit Proofview.tactic -> (Ssrast.ssrhyps * ((Ssrast.ssrhyp_or_id * string) * Ssrmatching.cpattern option) option) list * Ssrast.ssrclseq -> unit Proofview.tactic val hinttac : Tacinterp.interp_sign -> bool -> bool * Tacinterp.Value.t option list -> Ssrast.v82tac val ssrdotac : Tacinterp.interp_sign -> ((int Locus.or_var * Ssrast.ssrmmod) * (bool * Tacinterp.Value.t option list)) * ((Ssrast.ssrhyps * ((Ssrast.ssrhyp_or_id * string) * Ssrmatching.cpattern option) option) list * Ssrast.ssrclseq) -> Goal.goal Evd.sigma -> Goal.goal list Evd.sigma coq-8.11.0/plugins/ssr/ssrcommon.ml0000644000175000017500000016757513612664533017155 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* hyp_err ?loc "Duplicate assumption " id | SsrHyp (_, id) :: hyps -> check_hyps_uniq (id :: ids) hyps | [] -> () let check_hyp_exists hyps (SsrHyp(_, id)) = try ignore(Context.Named.lookup id hyps) with Not_found -> errorstrm Pp.(str"No assumption is named " ++ Id.print id) let test_hyp_exists hyps (SsrHyp(_, id)) = try ignore(Context.Named.lookup id hyps); true with Not_found -> false let hoik f = function Hyp x -> f x | Id x -> f x let hoi_id = hoik hyp_id let mk_hint tac = false, [Some tac] let mk_orhint tacs = true, tacs let nullhint = true, [] let nohint = false, [] type 'a tac_a = (goal * 'a) sigma -> (goal * 'a) list sigma let push_ctx a gl = re_sig (sig_it gl, a) (project gl) let push_ctxs a gl = re_sig (List.map (fun x -> x,a) (sig_it gl)) (project gl) let pull_ctx gl = let g, a = sig_it gl in re_sig g (project gl), a let pull_ctxs gl = let g, a = List.split (sig_it gl) in re_sig g (project gl), a let with_ctx f gl = let gl, ctx = pull_ctx gl in let rc, ctx = f ctx in rc, push_ctx ctx gl let without_ctx f gl = let gl, _ctx = pull_ctx gl in f gl let tac_ctx t gl = let gl, a = pull_ctx gl in let gl = t gl in push_ctxs a gl let tclTHEN_ia t1 t2 gl = let gal = t1 gl in let goals, sigma = sig_it gal, project gal in let _, opened, sigma = List.fold_left (fun (i,opened,sigma) g -> let gl = t2 i (re_sig g sigma) in i+1, sig_it gl :: opened, project gl) (1,[],sigma) goals in re_sig (List.flatten (List.rev opened)) sigma let tclTHEN_a t1 t2 gl = tclTHEN_ia t1 (fun _ -> t2) gl let tclTHENS_a t1 tl gl = tclTHEN_ia t1 (fun i -> List.nth tl (i-1)) gl let rec tclTHENLIST_a = function | [] -> tac_ctx tclIDTAC | t1::tacl -> tclTHEN_a t1 (tclTHENLIST_a tacl) (* like tclTHEN_i but passes to the tac "i of n" and not just i *) let tclTHEN_i_max tac taci gl = let maxi = ref 0 in tclTHEN_ia (tclTHEN_ia tac (fun i -> maxi := max i !maxi; tac_ctx tclIDTAC)) (fun i gl -> taci i !maxi gl) gl let tac_on_all gl tac = let goals = sig_it gl in let opened, sigma = List.fold_left (fun (opened,sigma) g -> let gl = tac (re_sig g sigma) in sig_it gl :: opened, project gl) ([],project gl) goals in re_sig (List.flatten (List.rev opened)) sigma (* Used to thread data between intro patterns at run time *) type tac_ctx = { tmp_ids : (Id.t * Name.t ref) list; wild_ids : Id.t list; delayed_clears : Id.t list; } let new_ctx () = { tmp_ids = []; wild_ids = []; delayed_clears = [] } let with_fresh_ctx t gl = let gl = push_ctx (new_ctx()) gl in let gl = t gl in fst (pull_ctxs gl) open Genarg open Stdarg open Pp let errorstrm x = CErrors.user_err ~hdr:"ssreflect" x let anomaly s = CErrors.anomaly (str s) (* Tentative patch from util.ml *) let array_fold_right_from n f v a = let rec fold n = if n >= Array.length v then a else f v.(n) (fold (succ n)) in fold n let array_app_tl v l = if Array.length v = 0 then invalid_arg "array_app_tl"; array_fold_right_from 1 (fun e l -> e::l) v l let array_list_of_tl v = if Array.length v = 0 then invalid_arg "array_list_of_tl"; array_fold_right_from 1 (fun e l -> e::l) v [] (* end patch *) let option_assert_get o msg = match o with | None -> CErrors.anomaly msg | Some x -> x (** Constructors for rawconstr *) open Glob_term let mkRHole = DAst.make @@ GHole (Evar_kinds.InternalHole, Namegen.IntroAnonymous, None) let rec mkRHoles n = if n > 0 then mkRHole :: mkRHoles (n - 1) else [] let rec isRHoles cl = match cl with | [] -> true | c :: l -> match DAst.get c with GHole _ -> isRHoles l | _ -> false let mkRApp f args = if args = [] then f else DAst.make @@ GApp (f, args) let mkRVar id = DAst.make @@ GRef (GlobRef.VarRef id,None) let mkRltacVar id = DAst.make @@ GVar (id) let mkRCast rc rt = DAst.make @@ GCast (rc, CastConv rt) let mkRType = DAst.make @@ GSort (UAnonymous {rigid=true}) let mkRProp = DAst.make @@ GSort (UNamed [GProp,0]) let mkRArrow rt1 rt2 = DAst.make @@ GProd (Anonymous, Explicit, rt1, rt2) let mkRConstruct c = DAst.make @@ GRef (GlobRef.ConstructRef c,None) let mkRInd mind = DAst.make @@ GRef (GlobRef.IndRef mind,None) let mkRLambda n s t = DAst.make @@ GLambda (n, Explicit, s, t) let rec mkRnat n = if n <= 0 then DAst.make @@ GRef (Coqlib.lib_ref "num.nat.O", None) else mkRApp (DAst.make @@ GRef (Coqlib.lib_ref "num.nat.S", None)) [mkRnat (n - 1)] let glob_constr ist genv = function | _, Some ce -> let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.Tacinterp.lfun Id.Set.empty in let ltacvars = { Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in Constrintern.intern_gen Pretyping.WithoutTypeConstraint ~ltacvars genv Evd.(from_env genv) ce | rc, None -> rc let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c let intern_term ist env (_, c) = glob_constr ist env c (* Estimate a bound on the number of arguments of a raw constr. *) (* This is not perfect, because the unifier may fail to *) (* typecheck the partial application, so we use a minimum of 5. *) (* Also, we don't handle delayed or iterated coercions to *) (* FUNCLASS, which is probably just as well since these can *) (* lead to infinite arities. *) let splay_open_constr gl (sigma, c) = let env = pf_env gl in let t = Retyping.get_type_of env sigma c in Reductionops.splay_prod env sigma t let isAppInd env sigma c = let c = Reductionops.clos_whd_flags CClosure.all env sigma c in let c, _ = decompose_app_vect sigma c in EConstr.isInd sigma c (** Generic argument-based globbing/typing utilities *) let interp_refine ist gl rc = let constrvars = Tacinterp.extract_ltac_constr_values ist (pf_env gl) in let vars = { Glob_ops.empty_lvar with Ltac_pretype.ltac_constrs = constrvars; ltac_genargs = ist.Tacinterp.lfun } in let kind = Pretyping.OfType (pf_concl gl) in let flags = { Pretyping.use_typeclasses = true; solve_unification_constraints = true; fail_evar = false; expand_evars = true; program_mode = false; polymorphic = false; } in let sigma, c = Pretyping.understand_ltac flags (pf_env gl) (project gl) vars kind rc in (* ppdebug(lazy(str"sigma@interp_refine=" ++ pr_evar_map None sigma)); *) ppdebug(lazy(str"c@interp_refine=" ++ Printer.pr_econstr_env (pf_env gl) sigma c)); (sigma, (sigma, c)) let interp_open_constr ist gl gc = let (sigma, (c, _)) = Tacinterp.interp_open_constr_with_bindings ist (pf_env gl) (project gl) (gc, Tactypes.NoBindings) in (project gl, (sigma, c)) let interp_term ist gl (_, c) = snd (interp_open_constr ist gl c) let of_ftactic ftac gl = let r = ref None in let tac = Ftactic.run ftac (fun ans -> r := Some ans; Proofview.tclUNIT ()) in let tac = Proofview.V82.of_tactic tac in let { sigma = sigma } = tac gl in let ans = match !r with | None -> assert false (* If the tactic failed we should not reach this point *) | Some ans -> ans in (sigma, ans) let interp_wit wit ist gl x = let globarg = in_gen (glbwit wit) x in let arg = Tacinterp.interp_genarg ist globarg in let (sigma, arg) = of_ftactic arg gl in sigma, Tacinterp.Value.cast (topwit wit) arg let interp_hyp ist gl (SsrHyp (loc, id)) = let s, id' = interp_wit wit_var ist gl CAst.(make ?loc id) in if not_section_id id' then s, SsrHyp (loc, id') else hyp_err ?loc "Can't clear section hypothesis " id' let interp_hyps ist gl ghyps = let hyps = List.map snd (List.map (interp_hyp ist gl) ghyps) in check_hyps_uniq [] hyps; Tacmach.project gl, hyps (* Old terms *) let mk_term k c = k, (mkRHole, Some c) let mk_lterm c = mk_term xNoFlag c (* New terms *) let mk_ast_closure_term a t = { annotation = a; body = t; interp_env = None; glob_env = None; } let glob_ast_closure_term (ist : Genintern.glob_sign) t = { t with glob_env = Some ist } let subst_ast_closure_term (_s : Mod_subst.substitution) t = (* _s makes sense only for glob constr *) t let interp_ast_closure_term (ist : Geninterp.interp_sign) (gl : 'goal Evd.sigma) t = (* gl is only useful if we want to interp *now*, later we have * a potentially different gl.sigma *) Tacmach.project gl, { t with interp_env = Some ist } let ssrterm_of_ast_closure_term { body; annotation } = let c = match annotation with | `Parens -> xInParens | `At -> xWithAt | _ -> xNoFlag in mk_term c body let ssrdgens_of_parsed_dgens = function | [], clr -> { dgens = []; gens = []; clr } | [gens], clr -> { dgens = []; gens; clr } | [dgens;gens], clr -> { dgens; gens; clr } | _ -> assert false let nbargs_open_constr gl oc = let pl, _ = splay_open_constr gl oc in List.length pl let pf_nbargs gl c = nbargs_open_constr gl (project gl, c) let internal_names = ref [] let add_internal_name pt = internal_names := pt :: !internal_names let is_internal_name s = List.exists (fun p -> p s) !internal_names let tmp_tag = "_the_" let tmp_post = "_tmp_" let mk_tmp_id i = Id.of_string (Printf.sprintf "%s%s%s" tmp_tag (CString.ordinal i) tmp_post) let new_tmp_id ctx = let id = mk_tmp_id (1 + List.length ctx.tmp_ids) in let orig = ref Anonymous in (id, orig), { ctx with tmp_ids = (id, orig) :: ctx.tmp_ids } ;; let mk_internal_id s = let s' = Printf.sprintf "_%s_" s in let s' = String.map (fun c -> if c = ' ' then '_' else c) s' in add_internal_name ((=) s'); Id.of_string s' let same_prefix s t n = let rec loop i = i = n || s.[i] = t.[i] && loop (i + 1) in loop 0 let skip_digits s = let n = String.length s in let rec loop i = if i < n && is_digit s.[i] then loop (i + 1) else i in loop let mk_tagged_id t i = Id.of_string (Printf.sprintf "%s%d_" t i) let is_tagged t s = let n = String.length s - 1 and m = String.length t in m < n && s.[n] = '_' && same_prefix s t m && skip_digits s m = n let evar_tag = "_evar_" let _ = add_internal_name (is_tagged evar_tag) let mk_evar_name n = Name (mk_tagged_id evar_tag n) let ssr_anon_hyp = "Hyp" let wildcard_tag = "_the_" let wildcard_post = "_wildcard_" let mk_wildcard_id i = Id.of_string (Printf.sprintf "%s%s%s" wildcard_tag (CString.ordinal i) wildcard_post) let has_wildcard_tag s = let n = String.length s in let m = String.length wildcard_tag in let m' = String.length wildcard_post in n < m + m' + 2 && same_prefix s wildcard_tag m && String.sub s (n - m') m' = wildcard_post && skip_digits s m = n - m' - 2 let _ = add_internal_name has_wildcard_tag let new_wild_id ctx = let i = 1 + List.length ctx.wild_ids in let id = mk_wildcard_id i in id, { ctx with wild_ids = id :: ctx.wild_ids } let discharged_tag = "_discharged_" let mk_discharged_id id = Id.of_string (Printf.sprintf "%s%s_" discharged_tag (Id.to_string id)) let has_discharged_tag s = let m = String.length discharged_tag and n = String.length s - 1 in m < n && s.[n] = '_' && same_prefix s discharged_tag m let _ = add_internal_name has_discharged_tag let is_discharged_id id = has_discharged_tag (Id.to_string id) let max_suffix m (t, j0 as tj0) id = let s = Id.to_string id in let n = String.length s - 1 in let dn = String.length t - 1 - n in let i0 = j0 - dn in if not (i0 >= m && s.[n] = '_' && same_prefix s t m) then tj0 else let rec loop i = if i < i0 && s.[i] = '0' then loop (i + 1) else if (if i < i0 then skip_digits s i = n else le_s_t i) then s, i else tj0 and le_s_t i = let ds = s.[i] and dt = t.[i + dn] in if ds = dt then i = n || le_s_t (i + 1) else dt < ds && skip_digits s i = n in loop m (** creates a fresh (w.r.t. `gl_ids` and internal names) inaccessible name of the form _tXX_ *) let mk_anon_id t gl_ids = let m, si0, id0 = let s = ref (Printf.sprintf "_%s_" t) in if is_internal_name !s then s := "_" ^ !s; let n = String.length !s - 1 in let rec loop i j = let d = !s.[i] in if not (is_digit d) then i + 1, j else loop (i - 1) (if d = '0' then j else i) in let m, j = loop (n - 1) n in m, (!s, j), Id.of_string_soft !s in if not (List.mem id0 gl_ids) then id0 else let s, i = List.fold_left (max_suffix m) si0 gl_ids in let open Bytes in let s = of_string s in let n = length s - 1 in let rec loop i = if get s i = '9' then (set s i '0'; loop (i - 1)) else if i < m then (set s n '0'; set s m '1'; cat s (of_string "_")) else (set s i (Char.chr (Char.code (get s i) + 1)); s) in Id.of_string_soft (Bytes.to_string (loop (n - 1))) let convert_concl_no_check t = Tactics.convert_concl ~check:false t DEFAULTcast let convert_concl ~check t = Tactics.convert_concl ~check t DEFAULTcast let rename_hd_prod orig_name_ref gl = match EConstr.kind (project gl) (pf_concl gl) with | Prod(x,src,tgt) -> let x = {x with binder_name = !orig_name_ref} in Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkProd (x,src,tgt))) gl | _ -> CErrors.anomaly (str "gentac creates no product") (* Reduction that preserves the Prod/Let spine of the "in" tactical. *) let inc_safe n = if n = 0 then n else n + 1 let rec safe_depth s c = match EConstr.kind s c with | LetIn ({binder_name=Name x}, _, _, c') when is_discharged_id x -> safe_depth s c' + 1 | LetIn (_, _, _, c') | Prod (_, _, c') -> inc_safe (safe_depth s c') | _ -> 0 let red_safe (r : Reductionops.reduction_function) e s c0 = let rec red_to e c n = match EConstr.kind s c with | Prod (x, t, c') when n > 0 -> let t' = r e s t in let e' = EConstr.push_rel (RelDecl.LocalAssum (x, t')) e in EConstr.mkProd (x, t', red_to e' c' (n - 1)) | LetIn (x, b, t, c') when n > 0 -> let t' = r e s t in let e' = EConstr.push_rel (RelDecl.LocalAssum (x, t')) e in EConstr.mkLetIn (x, r e s b, t', red_to e' c' (n - 1)) | _ -> r e s c in red_to e c0 (safe_depth s c0) let is_id_constr sigma c = match EConstr.kind sigma c with | Lambda(_,_,c) when EConstr.isRel sigma c -> 1 = EConstr.destRel sigma c | _ -> false let red_product_skip_id env sigma c = match EConstr.kind sigma c with | App(hd,args) when Array.length args = 1 && is_id_constr sigma hd -> args.(0) | _ -> try Tacred.red_product env sigma c with _ -> c let ssrevaltac ist gtac = Tacinterp.tactic_of_value ist gtac (** Open term to lambda-term coercion *)(* {{{ ************************************) (* This operation takes a goal gl and an open term (sigma, t), and *) (* returns a term t' where all the new evars in sigma are abstracted *) (* with the mkAbs argument, i.e., for mkAbs = mkLambda then there is *) (* some duplicate-free array args of evars of sigma such that the *) (* term mkApp (t', args) is convertible to t. *) (* This makes a useful shorthand for local definitions in proofs, *) (* i.e., pose succ := _ + 1 means pose succ := fun n : nat => n + 1, *) (* and, in context of the 4CT library, pose mid := maps id means *) (* pose mid := fun d : detaSet => @maps d d (@id (datum d)) *) (* Note that this facility does not extend to set, which tries *) (* instead to fill holes by matching a goal subterm. *) (* The argument to "have" et al. uses product abstraction, e.g. *) (* have Hmid: forall s, (maps id s) = s. *) (* stands for *) (* have Hmid: forall (d : dataSet) (s : seq d), (maps id s) = s. *) (* We also use this feature for rewrite rules, so that, e.g., *) (* rewrite: (plus_assoc _ 3). *) (* will execute as *) (* rewrite (fun n => plus_assoc n 3) *) (* i.e., it will rewrite some subterm .. + (3 + ..) to .. + 3 + ... *) (* The convention is also used for the argument of the congr tactic, *) (* e.g., congr (x + _ * 1). *) (* Replace new evars with lambda variables, retaining local dependencies *) (* but stripping global ones. We use the variable names to encode the *) (* the number of dependencies, so that the transformation is reversible. *) let env_size env = List.length (Environ.named_context env) let pf_concl gl = EConstr.Unsafe.to_constr (pf_concl gl) let pf_get_hyp gl x = EConstr.Unsafe.to_named_decl (pf_get_hyp gl x) let pf_e_type_of gl t = let sigma, env, it = project gl, pf_env gl, sig_it gl in let sigma, ty = Typing.type_of env sigma t in re_sig it sigma, ty let pf_resolve_typeclasses ~where ~fail gl = let sigma, env, it = project gl, pf_env gl, sig_it gl in let filter = let evset = Evarutil.undefined_evars_of_term sigma where in fun k _ -> Evar.Set.mem k evset in let sigma = Typeclasses.resolve_typeclasses ~filter ~fail env sigma in re_sig it sigma let resolve_typeclasses ~where ~fail env sigma = let filter = let evset = Evarutil.undefined_evars_of_term sigma where in fun k _ -> Evar.Set.mem k evset in let sigma = Typeclasses.resolve_typeclasses ~filter ~fail env sigma in sigma let nf_evar sigma t = EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr t)) let pf_abs_evars2 gl rigid (sigma, c0) = let c0 = EConstr.to_constr ~abort_on_undefined_evars:false sigma c0 in let sigma0, ucst = project gl, Evd.evar_universe_context sigma in let nenv = env_size (pf_env gl) in let abs_evar n k = let evi = Evd.find sigma k in let concl = EConstr.Unsafe.to_constr evi.evar_concl in let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in let abs_dc c = function | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t x.binder_relevance c) | NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in let t = Context.Named.fold_inside abs_dc ~init:concl dc in nf_evar sigma t in let rec put evlist c = match Constr.kind c with | Evar (k, a) -> if List.mem_assoc k evlist || Evd.mem sigma0 k || List.mem k rigid then evlist else let n = max 0 (Array.length a - nenv) in let t = abs_evar n k in (k, (n, t)) :: put evlist t | _ -> Constr.fold put evlist c in let evlist = put [] c0 in if evlist = [] then 0, EConstr.of_constr c0,[], ucst else let rec lookup k i = function | [] -> 0, 0 | (k', (n, _)) :: evl -> if k = k' then i, n else lookup k (i + 1) evl in let rec get i c = match Constr.kind c with | Evar (ev, a) -> let j, n = lookup ev i evlist in if j = 0 then Constr.map (get i) c else if n = 0 then mkRel j else mkApp (mkRel j, Array.init n (fun k -> get i a.(n - 1 - k))) | _ -> Constr.map_with_binders ((+) 1) get i c in let rec loop c i = function | (_, (n, t)) :: evl -> loop (mkLambda (make_annot (mk_evar_name n) Sorts.Relevant, get (i - 1) t, c)) (i - 1) evl | [] -> c in List.length evlist, EConstr.of_constr (loop (get 1 c0) 1 evlist), List.map fst evlist, ucst let pf_abs_evars gl t = pf_abs_evars2 gl [] t (* As before but if (?i : T(?j)) and (?j : P : Prop), then the lambda for i * looks like (fun evar_i : (forall pi : P. T(pi))) thanks to "loopP" and all * occurrences of evar_i are replaced by (evar_i evar_j) thanks to "app". * * If P can be solved by ssrautoprop (that defaults to trivial), then * the corresponding lambda looks like (fun evar_i : T(c)) where c is * the solution found by ssrautoprop. *) let ssrautoprop_tac = ref (fun gl -> assert false) (* Thanks to Arnaud Spiwack for this snippet *) let call_on_evar tac e s = let { it = gs ; sigma = s } = tac { it = e ; sigma = s; } in gs, s open Pp let pp _ = () (* FIXME *) module Intset = Evar.Set let pf_abs_evars_pirrel gl (sigma, c0) = pp(lazy(str"==PF_ABS_EVARS_PIRREL==")); pp(lazy(str"c0= " ++ Printer.pr_constr_env (pf_env gl) sigma c0)); let sigma0 = project gl in let c0 = nf_evar sigma0 (nf_evar sigma c0) in let nenv = env_size (pf_env gl) in let abs_evar n k = let evi = Evd.find sigma k in let concl = EConstr.Unsafe.to_constr evi.evar_concl in let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in let abs_dc c = function | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t x.binder_relevance c) | NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in let t = Context.Named.fold_inside abs_dc ~init:concl dc in nf_evar sigma0 (nf_evar sigma t) in let rec put evlist c = match Constr.kind c with | Evar (k, a) -> if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else let n = max 0 (Array.length a - nenv) in let k_ty = Retyping.get_sort_family_of (pf_env gl) sigma (Evd.evar_concl (Evd.find sigma k)) in let is_prop = k_ty = InProp in let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t | _ -> Constr.fold put evlist c in let evlist = put [] c0 in if evlist = [] then 0, c0 else let pr_constr t = Printer.pr_econstr_env (pf_env gl) sigma (Reductionops.nf_beta (pf_env gl) (project gl) (EConstr.of_constr t)) in pp(lazy(str"evlist=" ++ pr_list (fun () -> str";") (fun (k,_) -> Evar.print k) evlist)); let evplist = let depev = List.fold_left (fun evs (_,(_,t,_)) -> let t = EConstr.of_constr t in Intset.union evs (Evarutil.undefined_evars_of_term sigma t)) Intset.empty evlist in List.filter (fun (i,(_,_,b)) -> b && Intset.mem i depev) evlist in let evlist, evplist, sigma = if evplist = [] then evlist, [], sigma else List.fold_left (fun (ev, evp, sigma) (i, (_,t,_) as p) -> try let ng, sigma = call_on_evar !ssrautoprop_tac i sigma in if (ng <> []) then errorstrm (str "Should we tell the user?"); List.filter (fun (j,_) -> j <> i) ev, evp, sigma with _ -> ev, p::evp, sigma) (evlist, [], sigma) (List.rev evplist) in let c0 = nf_evar sigma c0 in let evlist = List.map (fun (x,(y,t,z)) -> x,(y,nf_evar sigma t,z)) evlist in let evplist = List.map (fun (x,(y,t,z)) -> x,(y,nf_evar sigma t,z)) evplist in pp(lazy(str"c0= " ++ pr_constr c0)); let rec lookup k i = function | [] -> 0, 0 | (k', (n,_,_)) :: evl -> if k = k' then i,n else lookup k (i + 1) evl in let rec get evlist i c = match Constr.kind c with | Evar (ev, a) -> let j, n = lookup ev i evlist in if j = 0 then Constr.map (get evlist i) c else if n = 0 then mkRel j else mkApp (mkRel j, Array.init n (fun k -> get evlist i a.(n - 1 - k))) | _ -> Constr.map_with_binders ((+) 1) (get evlist) i c in let rec app extra_args i c = match decompose_app c with | hd, args when isRel hd && destRel hd = i -> let j = destRel hd in mkApp (mkRel j, Array.of_list (List.map (Vars.lift (i-1)) extra_args @ args)) | _ -> Constr.map_with_binders ((+) 1) (app extra_args) i c in let rec loopP evlist c i = function | (_, (n, t, _)) :: evl -> let t = get evlist (i - 1) t in let n = Name (Id.of_string (ssr_anon_hyp ^ string_of_int n)) in loopP evlist (mkProd (make_annot n Sorts.Relevant, t, c)) (i - 1) evl | [] -> c in let rec loop c i = function | (_, (n, t, _)) :: evl -> let evs = Evarutil.undefined_evars_of_term sigma (EConstr.of_constr t) in let t_evplist = List.filter (fun (k,_) -> Intset.mem k evs) evplist in let t = loopP t_evplist (get t_evplist 1 t) 1 t_evplist in let t = get evlist (i - 1) t in let extra_args = List.map (fun (k,_) -> mkRel (fst (lookup k i evlist))) (List.rev t_evplist) in let c = if extra_args = [] then c else app extra_args 1 c in loop (mkLambda (make_annot (mk_evar_name n) Sorts.Relevant, t, c)) (i - 1) evl | [] -> c in let res = loop (get evlist 1 c0) 1 evlist in pp(lazy(str"res= " ++ pr_constr res)); List.length evlist, res (* Strip all non-essential dependencies from an abstracted term, generating *) (* standard names for the abstracted holes. *) let nb_evar_deps = function | Name id -> let s = Id.to_string id in if not (is_tagged evar_tag s) then 0 else let m = String.length evar_tag in (try int_of_string (String.sub s m (String.length s - 1 - m)) with _ -> 0) | _ -> 0 let pf_type_id gl t = Id.of_string (Namegen.hdchar (pf_env gl) (project gl) t) let pfe_type_of gl t = let sigma, ty = pf_type_of gl t in re_sig (sig_it gl) sigma, ty let pfe_new_type gl = let sigma, env, it = project gl, pf_env gl, sig_it gl in let sigma,t = Evarutil.new_Type sigma in re_sig it sigma, t let pfe_type_relevance_of gl t = let gl, ty = pfe_type_of gl t in gl, ty, pf_apply Retyping.relevance_of_term gl t let pf_type_of gl t = let sigma, ty = pf_type_of gl (EConstr.of_constr t) in re_sig (sig_it gl) sigma, EConstr.Unsafe.to_constr ty let pf_abs_cterm gl n c0 = if n <= 0 then c0 else let c0 = EConstr.Unsafe.to_constr c0 in let noargs = [|0|] in let eva = Array.make n noargs in let rec strip i c = match Constr.kind c with | App (f, a) when isRel f -> let j = i - destRel f in if j >= n || eva.(j) = noargs then mkApp (f, Array.map (strip i) a) else let dp = eva.(j) in let nd = Array.length dp - 1 in let mkarg k = strip i a.(if k < nd then dp.(k + 1) - j else k + dp.(0)) in mkApp (f, Array.init (Array.length a - dp.(0)) mkarg) | _ -> Constr.map_with_binders ((+) 1) strip i c in let rec strip_ndeps j i c = match Constr.kind c with | Prod (x, t, c1) when i < j -> let dl, c2 = strip_ndeps j (i + 1) c1 in if Vars.noccurn 1 c2 then dl, Vars.lift (-1) c2 else i :: dl, mkProd (x, strip i t, c2) | LetIn (x, b, t, c1) when i < j -> let _, _, c1' = destProd c1 in let dl, c2 = strip_ndeps j (i + 1) c1' in if Vars.noccurn 1 c2 then dl, Vars.lift (-1) c2 else i :: dl, mkLetIn (x, strip i b, strip i t, c2) | _ -> [], strip i c in let rec strip_evars i c = match Constr.kind c with | Lambda (x, t1, c1) when i < n -> let na = nb_evar_deps x.binder_name in let dl, t2 = strip_ndeps (i + na) i t1 in let na' = List.length dl in eva.(i) <- Array.of_list (na - na' :: dl); let x' = if na' = 0 then Name (pf_type_id gl (EConstr.of_constr t2)) else mk_evar_name na' in mkLambda ({x with binder_name=x'}, t2, strip_evars (i + 1) c1) (* if noccurn 1 c2 then lift (-1) c2 else mkLambda (Name (pf_type_id gl t2), t2, c2) *) | _ -> strip i c in EConstr.of_constr (strip_evars 0 c0) (* }}} *) let pf_merge_uc uc gl = re_sig (sig_it gl) (Evd.merge_universe_context (Refiner.project gl) uc) let pf_merge_uc_of sigma gl = let ucst = Evd.evar_universe_context sigma in pf_merge_uc ucst gl let rec constr_name sigma c = match EConstr.kind sigma c with | Var id -> Name id | Cast (c', _, _) -> constr_name sigma c' | Const (cn,_) -> Name (Label.to_id (Constant.label cn)) | App (c', _) -> constr_name sigma c' | _ -> Anonymous let pf_mkprod gl c ?(name=constr_name (project gl) c) cl = let gl, t, r = pfe_type_relevance_of gl c in if name <> Anonymous || EConstr.Vars.noccurn (project gl) 1 cl then gl, EConstr.mkProd (make_annot name r, t, cl) else gl, EConstr.mkProd (make_annot (Name (pf_type_id gl t)) r, t, cl) let pf_abs_prod name gl c cl = pf_mkprod gl c ~name (Termops.subst_term (project gl) c cl) (** look up a name in the ssreflect internals module *) let ssrdirpath = DirPath.make [Id.of_string "ssreflect"] let ssrqid name = Libnames.make_qualid ssrdirpath (Id.of_string name) let mkSsrRef name = let qn = Format.sprintf "plugins.ssreflect.%s" name in if Coqlib.has_ref qn then Coqlib.lib_ref qn else CErrors.user_err Pp.(str "Small scale reflection library not loaded (" ++ str name ++ str ")") let mkSsrRRef name = (DAst.make @@ GRef (mkSsrRef name,None)), None let mkSsrConst name env sigma = EConstr.fresh_global env sigma (mkSsrRef name) let pf_mkSsrConst name gl = let sigma, env, it = project gl, pf_env gl, sig_it gl in let (sigma, t) = mkSsrConst name env sigma in t, re_sig it sigma let pf_fresh_global name gl = let sigma, env, it = project gl, pf_env gl, sig_it gl in let sigma,t = Evd.fresh_global env sigma name in EConstr.Unsafe.to_constr t, re_sig it sigma let mkProt t c gl = let prot, gl = pf_mkSsrConst "protect_term" gl in EConstr.mkApp (prot, [|t; c|]), gl let mkEtaApp c n imin = let open EConstr in if n = 0 then c else let nargs, mkarg = if n < 0 then -n, (fun i -> mkRel (imin + i)) else let imax = imin + n - 1 in n, (fun i -> mkRel (imax - i)) in mkApp (c, Array.init nargs mkarg) let mkRefl t c gl = let sigma = project gl in let (sigma, refl) = EConstr.fresh_global (pf_env gl) sigma Coqlib.(lib_ref "core.eq.refl") in EConstr.mkApp (refl, [|t; c|]), { gl with sigma } let discharge_hyp (id', (id, mode)) gl = let cl' = Vars.subst_var id (pf_concl gl) in let decl = pf_get_hyp gl id in match decl, mode with | NamedDecl.LocalAssum _, _ | NamedDecl.LocalDef _, "(" -> let id' = {(NamedDecl.get_annot decl) with binder_name = Name id'} in Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:true (EConstr.of_constr (mkProd (id', NamedDecl.get_type decl, cl'))) [EConstr.of_constr (mkVar id)]) gl | NamedDecl.LocalDef (_, v, t), _ -> let id' = {(NamedDecl.get_annot decl) with binder_name = Name id'} in Proofview.V82.of_tactic (convert_concl ~check:true (EConstr.of_constr (mkLetIn (id', v, t, cl')))) gl (* wildcard names *) let clear_wilds wilds gl = Proofview.V82.of_tactic (Tactics.clear (List.filter (fun id -> List.mem id wilds) (pf_ids_of_hyps gl))) gl let clear_with_wilds wilds clr0 gl = let extend_clr clr nd = let id = NamedDecl.get_id nd in if List.mem id clr || not (List.mem id wilds) then clr else let vars = Termops.global_vars_set_of_decl (pf_env gl) (project gl) nd in let occurs id' = Id.Set.mem id' vars in if List.exists occurs clr then id :: clr else clr in Proofview.V82.of_tactic (Tactics.clear (Context.Named.fold_inside extend_clr ~init:clr0 (Tacmach.pf_hyps gl))) gl let clear_wilds_and_tmp_and_delayed_ids gl = let _, ctx = pull_ctx gl in tac_ctx (tclTHEN (clear_with_wilds ctx.wild_ids ctx.delayed_clears) (clear_wilds (List.map fst ctx.tmp_ids @ ctx.wild_ids))) gl let view_error s gv = errorstrm (str ("Cannot " ^ s ^ " view ") ++ pr_term gv) open Locus (****************************** tactics ***********************************) let rewritetac ?(under=false) dir c = (* Due to the new optional arg ?tac, application shouldn't be too partial *) let open Proofview.Notations in Proofview.V82.of_tactic begin Equality.general_rewrite (dir = L2R) AllOccurrences true false c <*> if under then Proofview.cycle 1 else Proofview.tclUNIT () end (**********************`:********* hooks ************************************) type name_hint = (int * EConstr.types array) option ref let pf_abs_ssrterm ?(resolve_typeclasses=false) ist gl t = let sigma, ct as t = interp_term ist gl t in let sigma, _ as t = let env = pf_env gl in if not resolve_typeclasses then t else let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in sigma, Evarutil.nf_evar sigma ct in let n, c, abstracted_away, ucst = pf_abs_evars gl t in List.fold_left Evd.remove sigma abstracted_away, pf_abs_cterm gl n c, ucst, n let top_id = mk_internal_id "top assumption" let ssr_n_tac seed n gl = let name = if n = -1 then seed else ("ssr" ^ seed ^ string_of_int n) in let fail msg = CErrors.user_err (Pp.str msg) in let tacname = try Tacenv.locate_tactic (Libnames.qualid_of_ident (Id.of_string name)) with Not_found -> try Tacenv.locate_tactic (ssrqid name) with Not_found -> if n = -1 then fail "The ssreflect library was not loaded" else fail ("The tactic "^name^" was not found") in let tacexpr = CAst.make @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in Proofview.V82.of_tactic (Tacinterp.eval_tactic (Tacexpr.TacArg tacexpr)) gl let donetac n gl = ssr_n_tac "done" n gl open Constrexpr open Util (** Constructors for constr_expr *) let mkCProp loc = CAst.make ?loc @@ CSort (UNamed [GProp,0]) let mkCType loc = CAst.make ?loc @@ CSort (UAnonymous {rigid=true}) let mkCVar ?loc id = CAst.make ?loc @@ CRef (qualid_of_ident ?loc id, None) let rec mkCHoles ?loc n = if n <= 0 then [] else (CAst.make ?loc @@ CHole (None, Namegen.IntroAnonymous, None)) :: mkCHoles ?loc (n - 1) let mkCHole loc = CAst.make ?loc @@ CHole (None, Namegen.IntroAnonymous, None) let mkCLambda ?loc name ty t = CAst.make ?loc @@ CLambdaN ([CLocalAssum([CAst.make ?loc name], Default Explicit, ty)], t) let mkCArrow ?loc ty t = CAst.make ?loc @@ CProdN ([CLocalAssum([CAst.make Anonymous], Default Explicit, ty)], t) let mkCCast ?loc t ty = CAst.make ?loc @@ CCast (t, CastConv ty) let rec isCHoles = function { CAst.v = CHole _ } :: cl -> isCHoles cl | cl -> cl = [] let rec isCxHoles = function ({ CAst.v = CHole _ }, None) :: ch -> isCxHoles ch | _ -> false let pf_interp_ty ?(resolve_typeclasses=false) ist gl ty = let n_binders = ref 0 in let ty = match ty with | a, (t, None) -> let rec force_type ty = DAst.(map (function | GProd (x, k, s, t) -> incr n_binders; GProd (x, k, s, force_type t) | GLetIn (x, v, oty, t) -> incr n_binders; GLetIn (x, v, oty, force_type t) | _ -> DAst.get (mkRCast ty mkRType))) ty in a, (force_type t, None) | _, (_, Some ty) -> let rec force_type ty = CAst.(map (function | CProdN (abs, t) -> n_binders := !n_binders + List.length (List.flatten (List.map (function CLocalAssum (nal,_,_) -> nal | CLocalDef (na,_,_) -> [na] | CLocalPattern _ -> (* We count a 'pat for 1; TO BE CHECKED *) [CAst.make Name.Anonymous]) abs)); CProdN (abs, force_type t) | CLetIn (n, v, oty, t) -> incr n_binders; CLetIn (n, v, oty, force_type t) | _ -> (mkCCast ty (mkCType None)).v)) ty in mk_term ' ' (force_type ty) in let strip_cast (sigma, t) = let rec aux t = match EConstr.kind_of_type sigma t with | CastType (t, ty) when !n_binders = 0 && EConstr.isSort sigma ty -> t | ProdType(n,s,t) -> decr n_binders; EConstr.mkProd (n, s, aux t) | LetInType(n,v,ty,t) -> decr n_binders; EConstr.mkLetIn (n, v, ty, aux t) | _ -> anomaly "pf_interp_ty: ssr Type cast deleted by typecheck" in sigma, aux t in let sigma, cty as ty = strip_cast (interp_term ist gl ty) in let ty = let env = pf_env gl in if not resolve_typeclasses then ty else let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in sigma, Evarutil.nf_evar sigma cty in let n, c, _, ucst = pf_abs_evars gl ty in let lam_c = pf_abs_cterm gl n c in let ctx, c = EConstr.decompose_lam_n_assum sigma n lam_c in n, EConstr.it_mkProd_or_LetIn c ctx, lam_c, ucst ;; (* TASSI: given (c : ty), generates (c ??? : ty[???/...]) with m evars *) exception NotEnoughProducts let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_of env sigma c) m = let rec loop ty args sigma n = if n = 0 then let args = List.rev args in (if beta then Reductionops.whd_beta sigma else fun x -> x) (EConstr.mkApp (c, Array.of_list (List.map snd args))), ty, args, sigma else match EConstr.kind_of_type sigma ty with | ProdType (_, src, tgt) -> let sigma = create_evar_defs sigma in let (sigma, x) = Evarutil.new_evar env sigma (if bi_types then Reductionops.nf_betaiota env sigma src else src) in loop (EConstr.Vars.subst1 x tgt) ((m - n,x) :: args) sigma (n-1) | CastType (t, _) -> loop t args sigma n | LetInType (_, v, _, t) -> loop (EConstr.Vars.subst1 v t) args sigma n | SortType _ -> assert false | AtomicType _ -> let ty = (* FIXME *) (Reductionops.whd_all env sigma) ty in match EConstr.kind_of_type sigma ty with | ProdType _ -> loop ty args sigma n | _ -> raise NotEnoughProducts in loop ty [] sigma m let pf_saturate ?beta ?bi_types gl c ?ty m = let env, sigma, si = pf_env gl, project gl, sig_it gl in let t, ty, args, sigma = saturate ?beta ?bi_types env sigma c ?ty m in t, ty, args, re_sig si sigma let pf_partial_solution gl t evl = let sigma, g = project gl, sig_it gl in let sigma = Goal.V82.partial_solution (pf_env gl) sigma g t in re_sig (List.map (fun x -> (fst (EConstr.destEvar sigma x))) evl) sigma let dependent_apply_error = try CErrors.user_err (Pp.str "Could not fill dependent hole in \"apply\"") with err -> err (* TASSI: Sometimes Coq's apply fails. According to my experience it may be * related to goals that are products and with beta redexes. In that case it * guesses the wrong number of implicit arguments for your lemma. What follows * is just like apply, but with a user-provided number n of implicits. * * Refine.refine function that handles type classes and evars but fails to * handle "dependently typed higher order evars". * * Refiner.refiner that does not handle metas with a non ground type but works * with dependently typed higher order metas. *) let applyn ~with_evars ?beta ?(with_shelve=false) ?(first_goes_last=false) n t gl = if with_evars then let refine gl = let t, ty, args, gl = pf_saturate ?beta ~bi_types:true gl t n in (* pp(lazy(str"sigma@saturate=" ++ pr_evar_map None (project gl))); *) let gl = pf_unify_HO gl ty (Tacmach.pf_concl gl) in let gs = CList.map_filter (fun (_, e) -> if EConstr.isEvar (project gl) e then Some e else None) args in pf_partial_solution gl t gs in Proofview.(V82.of_tactic (Tacticals.New.tclTHENLIST [ V82.tactic refine; (if with_shelve then shelve_unifiable else tclUNIT ()); (if first_goes_last then cycle 1 else tclUNIT ()) ])) gl else let t, gl = if n = 0 then t, gl else let sigma, si = project gl, sig_it gl in let rec loop sigma bo args = function (* saturate with metas *) | 0 -> EConstr.mkApp (t, Array.of_list (List.rev args)), re_sig si sigma | n -> match EConstr.kind sigma bo with | Lambda (_, ty, bo) -> if not (EConstr.Vars.closed0 sigma ty) then raise dependent_apply_error; let m = Evarutil.new_meta () in loop (meta_declare m ty sigma) bo ((EConstr.mkMeta m)::args) (n-1) | _ -> assert false in loop sigma t [] n in pp(lazy(str"Refiner.refiner " ++ Printer.pr_econstr_env (pf_env gl) (project gl) t)); Proofview.(V82.of_tactic (Tacticals.New.tclTHENLIST [ V82.tactic (Refiner.refiner ~check:false EConstr.Unsafe.(to_constr t)); (if first_goes_last then cycle 1 else tclUNIT ()) ])) gl let refine_with ?(first_goes_last=false) ?beta ?(with_evars=true) oc gl = let uct = Evd.evar_universe_context (fst oc) in let n, oc = pf_abs_evars_pirrel gl (fst oc, EConstr.to_constr ~abort_on_undefined_evars:false (fst oc) (snd oc)) in let gl = pf_unsafe_merge_uc uct gl in try applyn ~with_evars ~first_goes_last ~with_shelve:true ?beta n (EConstr.of_constr oc) gl with e when CErrors.noncritical e -> raise dependent_apply_error (* We wipe out all the keywords generated by the grammar rules we defined. *) (* The user is supposed to Require Import ssreflect or Require ssreflect *) (* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) (* consequence the extended ssreflect grammar. *) let () = CLexer.set_keyword_state frozen_lexer ;; (** Basic tactics *) let rec fst_prod red tac = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in match EConstr.kind (Proofview.Goal.sigma gl) concl with | Prod (id,_,tgt) | LetIn(id,_,_,tgt) -> tac id.binder_name | _ -> if red then Tacticals.New.tclZEROMSG (str"No product even after head-reduction.") else Tacticals.New.tclTHEN Tactics.hnf_in_concl (fst_prod true tac) end let introid ?(orig=ref Anonymous) name = tclTHEN (fun gl -> let g, env = Tacmach.pf_concl gl, pf_env gl in let sigma = project gl in match EConstr.kind sigma g with | App (hd, _) when EConstr.isLambda sigma hd -> Proofview.V82.of_tactic (convert_concl_no_check (Reductionops.whd_beta sigma g)) gl | _ -> tclIDTAC gl) (Proofview.V82.of_tactic (fst_prod false (fun id -> orig := id; Tactics.intro_mustbe_force name))) ;; let anontac decl gl = let id = match RelDecl.get_name decl with | Name id -> if is_discharged_id id then id else mk_anon_id (Id.to_string id) (Tacmach.pf_ids_of_hyps gl) | _ -> mk_anon_id ssr_anon_hyp (Tacmach.pf_ids_of_hyps gl) in introid id gl let rec intro_anon gl = try anontac (List.hd (fst (EConstr.decompose_prod_n_assum (project gl) 1 (Tacmach.pf_concl gl)))) gl with err0 -> try tclTHEN (Proofview.V82.of_tactic Tactics.red_in_concl) intro_anon gl with e when CErrors.noncritical e -> raise err0 (* with _ -> CErrors.error "No product even after reduction" *) let is_pf_var sigma c = EConstr.isVar sigma c && not_section_id (EConstr.destVar sigma c) let hyp_of_var sigma v = SsrHyp (Loc.tag @@ EConstr.destVar sigma v) let interp_clr sigma = function | Some clr, (k, c) when (k = xNoFlag || k = xWithAt) && is_pf_var sigma c -> hyp_of_var sigma c :: clr | Some clr, _ -> clr | None, _ -> [] (** Basic tacticals *) (** Multipliers *)(* {{{ ***********************************************************) (* tactical *) let tclID tac = tac let tclDOTRY n tac = if n <= 0 then tclIDTAC else let rec loop i gl = if i = n then tclTRY tac gl else tclTRY (tclTHEN tac (loop (i + 1))) gl in loop 1 let tclDO n tac = let prefix i = str"At iteration " ++ int i ++ str": " in let tac_err_at i gl = try tac gl with | CErrors.UserError (l, s) as e -> let _, info = CErrors.push e in let e' = CErrors.UserError (l, prefix i ++ s) in Util.iraise (e', info) | Gramlib.Ploc.Exc(loc, CErrors.UserError (l, s)) -> raise (Gramlib.Ploc.Exc(loc, CErrors.UserError (l, prefix i ++ s))) in let rec loop i gl = if i = n then tac_err_at i gl else (tclTHEN (tac_err_at i) (loop (i + 1))) gl in loop 1 let tclMULT = function | 0, May -> tclREPEAT | 1, May -> tclTRY | n, May -> tclDOTRY n | 0, Must -> tclAT_LEAST_ONCE | n, Must when n > 1 -> tclDO n | _ -> tclID let old_cleartac clr = check_hyps_uniq [] clr; Proofview.V82.of_tactic (Tactics.clear (hyps_ids clr)) let cleartac clr = check_hyps_uniq [] clr; Tactics.clear (hyps_ids clr) (* }}} *) (** Generalize tactic *) (* XXX the k of the redex should percolate out *) let pf_interp_gen_aux gl to_ind ((oclr, occ), t) = let pat = interp_cpattern gl t None in (* UGLY API *) let gl = pf_merge_uc_of (fst pat) gl in let cl, env, sigma = Tacmach.pf_concl gl, pf_env gl, project gl in let (c, ucst), cl = try fill_occ_pattern ~raise_NoMatch:true env sigma (EConstr.Unsafe.to_constr cl) pat occ 1 with NoMatch -> redex_of_pattern env pat, (EConstr.Unsafe.to_constr cl) in let gl = pf_merge_uc ucst gl in let c = EConstr.of_constr c in let cl = EConstr.of_constr cl in let clr = interp_clr sigma (oclr, (tag_of_cpattern t, c)) in if not(occur_existential sigma c) then if tag_of_cpattern t = xWithAt then if not (EConstr.isVar sigma c) then errorstrm (str "@ can be used with variables only") else match Tacmach.pf_get_hyp gl (EConstr.destVar sigma c) with | NamedDecl.LocalAssum _ -> errorstrm (str "@ can be used with let-ins only") | NamedDecl.LocalDef (name, b, ty) -> true, pat, EConstr.mkLetIn (map_annot Name.mk_name name,b,ty,cl),c,clr,ucst,gl else let gl, ccl = pf_mkprod gl c cl in false, pat, ccl, c, clr,ucst,gl else if to_ind && occ = None then let nv, p, _, ucst' = pf_abs_evars gl (fst pat, c) in let ucst = UState.union ucst ucst' in if nv = 0 then anomaly "occur_existential but no evars" else let gl, pty, rp = pfe_type_relevance_of gl p in false, pat, EConstr.mkProd (make_annot (constr_name (project gl) c) rp, pty, Tacmach.pf_concl gl), p, clr,ucst,gl else CErrors.user_err ?loc:(loc_of_cpattern t) (str "generalized term didn't match") let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:true x xs) let genclrtac cl cs clr = let tclmyORELSE tac1 tac2 gl = try tac1 gl with e when CErrors.noncritical e -> tac2 e gl in (* apply_type may give a type error, but the useful message is * the one of clear. You type "move: x" and you get * "x is used in hyp H" instead of * "The term H has type T x but is expected to have type T x0". *) tclTHEN (tclmyORELSE (apply_type cl cs) (fun type_err gl -> tclTHEN (tclTHEN (Proofview.V82.of_tactic (Tactics.elim_type (EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.(lib_ref "core.False.type"))))) (old_cleartac clr)) (fun gl -> raise type_err) gl)) (old_cleartac clr) let gentac gen gl = (* ppdebug(lazy(str"sigma@gentac=" ++ pr_evar_map None (project gl))); *) let conv, _, cl, c, clr, ucst,gl = pf_interp_gen_aux gl false gen in ppdebug(lazy(str"c@gentac=" ++ pr_econstr_env (pf_env gl) (project gl) c)); let gl = pf_merge_uc ucst gl in if conv then tclTHEN (Proofview.V82.of_tactic (convert_concl ~check:true cl)) (old_cleartac clr) gl else genclrtac cl [c] clr gl let genstac (gens, clr) = tclTHENLIST (old_cleartac clr :: List.rev_map gentac gens) let gen_tmp_ids ?(ist=Geninterp.({ lfun = Id.Map.empty; poly = false; extra = Tacinterp.TacStore.empty })) gl = let gl, ctx = pull_ctx gl in push_ctxs ctx (tclTHENLIST (List.map (fun (id,orig_ref) -> tclTHEN (gentac ((None,Some(false,[])),cpattern_of_id id)) (rename_hd_prod orig_ref)) ctx.tmp_ids) gl) ;; let pf_interp_gen to_ind gen gl = let _, _, a, b, c, ucst,gl = pf_interp_gen_aux gl to_ind gen in (a, b ,c), pf_merge_uc ucst gl let pfLIFT f = let open Proofview.Notations in let hack = ref None in Proofview.V82.tactic (fun gl -> let g = sig_it gl in let x, gl = f gl in hack := Some (x,project gl); re_sig [g] (project gl)) >>= fun () -> let x, sigma = option_assert_get !hack (Pp.str"pfLIFT") in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT x ;; (* TASSI: This version of unprotects inlines the unfold tactic definition, * since we don't want to wipe out let-ins, and it seems there is no flag * to change that behaviour in the standard unfold code *) let unprotecttac gl = let c, gl = pf_mkSsrConst "protect_term" gl in let prot, _ = EConstr.destConst (project gl) c in Tacticals.onClause (fun idopt -> let hyploc = Option.map (fun id -> id, InHyp) idopt in Proofview.V82.of_tactic (Tactics.reduct_option ~check:false (Reductionops.clos_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fBETA; CClosure.RedFlags.fCONST prot; CClosure.RedFlags.fMATCH; CClosure.RedFlags.fFIX; CClosure.RedFlags.fCOFIX]), DEFAULTcast) hyploc)) allHypsAndConcl gl let is_protect hd env sigma = let _, protectC = mkSsrConst "protect_term" env sigma in EConstr.eq_constr_nounivs sigma hd protectC let abs_wgen keep_let f gen (gl,args,c) = let sigma, env = project gl, pf_env gl in let evar_closed t p = if occur_existential sigma t then CErrors.user_err ?loc:(loc_of_cpattern p) ~hdr:"ssreflect" (pr_econstr_pat env sigma t ++ str" contains holes and matches no subterm of the goal") in match gen with | _, Some ((x, mode), None) when mode = "@" || (mode = " " && keep_let) -> let x = hoi_id x in let decl = Tacmach.pf_get_hyp gl x in gl, (if NamedDecl.is_local_def decl then args else EConstr.mkVar x :: args), EConstr.mkProd_or_LetIn (decl |> NamedDecl.to_rel_decl |> RelDecl.set_name (Name (f x))) (EConstr.Vars.subst_var x c) | _, Some ((x, _), None) -> let x = hoi_id x in let hyp = Tacmach.pf_get_hyp gl x in let x' = make_annot (Name (f x)) (NamedDecl.get_relevance hyp) in let prod = EConstr.mkProd (x', NamedDecl.get_type hyp, EConstr.Vars.subst_var x c) in gl, EConstr.mkVar x :: args, prod | _, Some ((x, "@"), Some p) -> let x = hoi_id x in let cp = interp_cpattern gl p None in let gl = pf_merge_uc_of (fst cp) gl in let (t, ucst), c = try fill_occ_pattern ~raise_NoMatch:true env sigma (EConstr.Unsafe.to_constr c) cp None 1 with NoMatch -> redex_of_pattern env cp, (EConstr.Unsafe.to_constr c) in let c = EConstr.of_constr c in let t = EConstr.of_constr t in evar_closed t p; let ut = red_product_skip_id env sigma t in let gl, ty, r = pfe_type_relevance_of gl t in pf_merge_uc ucst gl, args, EConstr.mkLetIn(make_annot (Name (f x)) r, ut, ty, c) | _, Some ((x, _), Some p) -> let x = hoi_id x in let cp = interp_cpattern gl p None in let gl = pf_merge_uc_of (fst cp) gl in let (t, ucst), c = try fill_occ_pattern ~raise_NoMatch:true env sigma (EConstr.Unsafe.to_constr c) cp None 1 with NoMatch -> redex_of_pattern env cp, (EConstr.Unsafe.to_constr c) in let c = EConstr.of_constr c in let t = EConstr.of_constr t in evar_closed t p; let gl, ty, r = pfe_type_relevance_of gl t in pf_merge_uc ucst gl, t :: args, EConstr.mkProd(make_annot (Name (f x)) r, ty, c) | _ -> gl, args, c let clr_of_wgen gen clrs = match gen with | clr, Some ((x, _), None) -> let x = hoi_id x in old_cleartac clr :: old_cleartac [SsrHyp(Loc.tag x)] :: clrs | clr, _ -> old_cleartac clr :: clrs let reduct_in_concl ~check t = Tactics.reduct_in_concl ~check (t, DEFAULTcast) let unfold cl = let module R = Reductionops in let module F = CClosure.RedFlags in reduct_in_concl ~check:false (R.clos_norm_flags (F.mkflags (List.map (fun c -> F.fCONST (fst (destConst (EConstr.Unsafe.to_constr c)))) cl @ [F.fBETA; F.fMATCH; F.fFIX; F.fCOFIX]))) open Proofview open Notations let tacSIGMA = Goal.enter_one ~__LOC__ begin fun g -> let k = Goal.goal g in let sigma = Goal.sigma g in tclUNIT (Tacmach.re_sig k sigma) end let tclINTERP_AST_CLOSURE_TERM_AS_CONSTR c = tclINDEPENDENTL begin tacSIGMA >>= fun gl -> let old_ssrterm = mkRHole, Some c.Ssrast.body in let ist = option_assert_get c.Ssrast.interp_env Pp.(str "tclINTERP_AST_CLOSURE_TERM_AS_CONSTR: term with no ist") in let sigma, t = interp_wit Stdarg.wit_constr ist gl old_ssrterm in Unsafe.tclEVARS sigma <*> tclUNIT t end let tacREDUCE_TO_QUANTIFIED_IND ty = tacSIGMA >>= fun gl -> tclUNIT (Tacmach.pf_reduce_to_quantified_ind gl ty) let tacTYPEOF c = Goal.enter_one ~__LOC__ (fun g -> let sigma, env = Goal.sigma g, Goal.env g in let sigma, ty = Typing.type_of env sigma c in Unsafe.tclEVARS sigma <*> tclUNIT ty) (** This tactic creates a partial proof realizing the introduction rule, but does not check anything. *) let unsafe_intro env decl b = let open Context.Named.Declaration in Refine.refine ~typecheck:false begin fun sigma -> let ctx = Environ.named_context_val env in let nctx = EConstr.push_named_context_val decl ctx in let inst = List.map (get_id %> EConstr.mkVar) (Environ.named_context env) in let ninst = EConstr.mkRel 1 :: inst in let nb = EConstr.Vars.subst1 (EConstr.mkVar (get_id decl)) b in let sigma, ev = Evarutil.new_evar_instance nctx sigma nb ~principal:true ninst in sigma, EConstr.mkNamedLambda_or_LetIn decl ev end let set_decl_id id = let open Context in function | Rel.Declaration.LocalAssum(name,ty) -> Named.Declaration.LocalAssum({name with binder_name=id},ty) | Rel.Declaration.LocalDef(name,ty,t) -> Named.Declaration.LocalDef({name with binder_name=id},ty,t) let rec decompose_assum env sigma orig_goal = let open Context in match EConstr.kind sigma orig_goal with | Prod(name,ty,t) -> Rel.Declaration.LocalAssum(name,ty), t, true | LetIn(name,ty,t1,t2) -> Rel.Declaration.LocalDef(name, ty, t1), t2, true | _ -> let goal = Reductionops.whd_allnolet env sigma orig_goal in match EConstr.kind sigma goal with | Prod(name,ty,t) -> Rel.Declaration.LocalAssum(name,ty), t, false | LetIn(name,ty,t1,t2) -> Rel.Declaration.LocalDef(name,ty,t1), t2, false | App(hd,args) when EConstr.isLetIn sigma hd -> (* hack *) let _,v,_,b = EConstr.destLetIn sigma hd in let ctx, t, _ = decompose_assum env sigma (EConstr.mkApp (EConstr.Vars.subst1 v b, args)) in ctx, t, false | _ -> CErrors.user_err Pp.(str "No assumption in " ++ Printer.pr_econstr_env env sigma goal) let tclFULL_BETAIOTA = Goal.enter begin fun gl -> let r, _ = Redexpr.reduction_of_red_expr (Goal.env gl) Genredexpr.(Lazy { rBeta=true; rMatch=true; rFix=true; rCofix=true; rZeta=false; rDelta=false; rConst=[]}) in Tactics.e_reduct_in_concl ~check:false (r,Constr.DEFAULTcast) end type intro_id = | Anon | Id of Id.t | Seed of string (** [intro id k] introduces the first premise (product or let-in) of the goal under the name [id], reducing the head of the goal (using beta, iota, delta but not zeta) if necessary. If [id] is None, a name is generated, that will not be user accessible. If the goal does not start with a product or a let-in even after reduction, it fails. In case of success, the original name and final id are passed to the continuation [k] which gets evaluated. *) let tclINTRO ~id ~conclusion:k = Goal.enter begin fun gl -> let open Context in let env, sigma, g = Goal.(env gl, sigma gl, concl gl) in let decl, t, no_red = decompose_assum env sigma g in let original_name = Rel.Declaration.get_name decl in let already_used = Tacmach.New.pf_ids_of_hyps gl in let id = match id, original_name with | Id id, _ -> id | Seed id, _ -> mk_anon_id id already_used | Anon, Name id -> if is_discharged_id id then id else mk_anon_id (Id.to_string id) already_used | Anon, Anonymous -> let ids = Tacmach.New.pf_ids_of_hyps gl in mk_anon_id ssr_anon_hyp ids in if List.mem id already_used then errorstrm Pp.(Id.print id ++ str" already used"); unsafe_intro env (set_decl_id id decl) t <*> (if no_red then tclUNIT () else tclFULL_BETAIOTA) <*> k ~orig_name:original_name ~new_name:id end let return ~orig_name:_ ~new_name:_ = tclUNIT () let tclINTRO_ID id = tclINTRO ~id:(Id id) ~conclusion:return let tclINTRO_ANON ?seed () = match seed with | None -> tclINTRO ~id:Anon ~conclusion:return | Some seed -> tclINTRO ~id:(Seed seed) ~conclusion:return let tclRENAME_HD_PROD name = Goal.enter begin fun gl -> let concl = Goal.concl gl in let sigma = Goal.sigma gl in match EConstr.kind sigma concl with | Prod(x,src,tgt) -> convert_concl_no_check EConstr.(mkProd ({x with binder_name = name},src,tgt)) | _ -> CErrors.anomaly (Pp.str "rename_hd_prod: no head product") end let tcl0G ~default tac = numgoals >>= fun ng -> if ng = 0 then tclUNIT default else tac let rec tclFIRSTa = function | [] -> Tacticals.New.tclZEROMSG Pp.(str"No applicable tactic.") | tac :: rest -> tclORELSE tac (fun _ -> tclFIRSTa rest) let rec tclFIRSTi tac n = if n < 0 then Tacticals.New.tclZEROMSG Pp.(str "tclFIRSTi") else tclORELSE (tclFIRSTi tac (n-1)) (fun _ -> tac n) let tacCONSTR_NAME ?name c = match name with | Some n -> tclUNIT n | None -> Goal.enter_one ~__LOC__ (fun g -> let sigma = Goal.sigma g in tclUNIT (constr_name sigma c)) let tacMKPROD c ?name cl = tacTYPEOF c >>= fun t -> tacCONSTR_NAME ?name c >>= fun name -> Goal.enter_one ~__LOC__ begin fun g -> let sigma, env = Goal.sigma g, Goal.env g in let r = Retyping.relevance_of_term env sigma c in if name <> Names.Name.Anonymous || EConstr.Vars.noccurn sigma 1 cl then tclUNIT (EConstr.mkProd (make_annot name r, t, cl)) else let name = Names.Id.of_string (Namegen.hdchar env sigma t) in tclUNIT (EConstr.mkProd (make_annot (Name.Name name) r, t, cl)) end let tacINTERP_CPATTERN cp = tacSIGMA >>= begin fun gl -> tclUNIT (Ssrmatching.interp_cpattern gl cp None) end let tacUNIFY a b = tacSIGMA >>= begin fun gl -> let gl = Ssrmatching.pf_unify_HO gl a b in Unsafe.tclEVARS (Tacmach.project gl) end let tclOPTION o d = match o with | None -> d >>= tclUNIT | Some x -> tclUNIT x let tacIS_INJECTION_CASE ?ty t = begin tclOPTION ty (tacTYPEOF t) >>= fun ty -> tacREDUCE_TO_QUANTIFIED_IND ty >>= fun ((mind,_),_) -> tclUNIT (Coqlib.check_ind_ref "core.eq.type" mind) end let tclWITHTOP tac = Goal.enter begin fun gl -> let top = mk_anon_id "top_assumption" (Tacmach.New.pf_ids_of_hyps gl) in tclINTRO_ID top <*> tac (EConstr.mkVar top) <*> Tactics.clear [top] end let tacMK_SSR_CONST name = Goal.enter_one ~__LOC__ begin fun g -> let sigma, env = Goal.(sigma g, env g) in let sigma, c = mkSsrConst name env sigma in Unsafe.tclEVARS sigma <*> tclUNIT c end module type StateType = sig type state val init : state end module MakeState(S : StateType) = struct let state_field : S.state Proofview_monad.StateStore.field = Proofview_monad.StateStore.field () (* FIXME: should not inject fresh_state, but initialize it at the beginning *) let lift_upd_state upd s = let open Proofview_monad.StateStore in let old_state = Option.default S.init (get s state_field) in upd old_state >>= fun new_state -> tclUNIT (set s state_field new_state) let tacUPDATE upd = Goal.enter begin fun gl -> let s0 = Goal.state gl in Goal.enter_one ~__LOC__ (fun _ -> lift_upd_state upd s0) >>= fun s -> Unsafe.tclGETGOALS >>= fun gls -> let gls = List.map (fun gs -> let g = Proofview_monad.drop_state gs in Proofview_monad.goal_with_state g s) gls in Unsafe.tclSETGOALS gls end let tclGET k = Goal.enter begin fun gl -> let open Proofview_monad.StateStore in k (Option.default S.init (get (Goal.state gl) state_field)) end let tclGET1 k = Goal.enter_one begin fun gl -> let open Proofview_monad.StateStore in k (Option.default S.init (get (Goal.state gl) state_field)) end let tclSET new_s = let open Proofview_monad.StateStore in Unsafe.tclGETGOALS >>= fun gls -> let gls = List.map (fun gs -> let g = Proofview_monad.drop_state gs in let s = Proofview_monad.get_state gs in Proofview_monad.goal_with_state g (set s state_field new_s)) gls in Unsafe.tclSETGOALS gls let get g = Option.default S.init (Proofview_monad.StateStore.get (Goal.state g) state_field) end let is_construct_ref sigma c r = EConstr.isConstruct sigma c && GlobRef.equal (GlobRef.ConstructRef (fst(EConstr.destConstruct sigma c))) r let is_ind_ref sigma c r = EConstr.isInd sigma c && GlobRef.equal (GlobRef.IndRef (fst(EConstr.destInd sigma c))) r let is_const_ref sigma c r = EConstr.isConst sigma c && GlobRef.equal (GlobRef.ConstRef (fst(EConstr.destConst sigma c))) r (* vim: set filetype=ocaml foldmethod=marker: *) coq-8.11.0/plugins/ssr/ssrvernac.mlg0000644000175000017500000005740613612664533017301 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* is then ... else ... *) (* if is [in ..] return ... then ... else ... *) (* let: := in ... *) (* let: [in ...] := return ... in ... *) (* The scope of a top-level 'as' in the pattern extends over the *) (* 'return' type (dependent if/let). *) (* Note that the optional "in ..." appears next to the *) (* rather than the in then "let:" syntax. The alternative *) (* would lead to ambiguities in, e.g., *) (* let: p1 := (*v---INNER LET:---v *) *) (* let: p2 := let: p3 := e3 in k return t in k2 in k1 return t' *) (* in b (*^--ALTERNATIVE INNER LET--------^ *) *) (* Caveat : There is no pretty-printing support, since this would *) (* require a modification to the Coq kernel (adding a new match *) (* display style -- why aren't these strings?); also, the v8.1 *) (* pretty-printer only allows extension hooks for printing *) (* integer or string literals. *) (* Also note that in the v8 grammar "is" needs to be a keyword; *) (* as this can't be done from an ML extension file, the new *) (* syntax will only work when ssreflect.v is imported. *) let no_ct = None, None and no_rt = None let aliasvar = function | [[{ CAst.v = CPatAlias (_, na); loc }]] -> Some na | _ -> None let mk_cnotype mp = aliasvar mp, None let mk_ctype mp t = aliasvar mp, Some t let mk_rtype t = Some t let mk_dthen ?loc (mp, ct, rt) c = (CAst.make ?loc (mp, c)), ct, rt let mk_let ?loc rt ct mp c1 = CAst.make ?loc @@ CCases (LetPatternStyle, rt, ct, [CAst.make ?loc (mp, c1)]) let mk_pat c (na, t) = (c, na, t) } GRAMMAR EXTEND Gram GLOBAL: binder_constr; ssr_rtype: [[ "return"; t = operconstr LEVEL "100" -> { mk_rtype t } ]]; ssr_mpat: [[ p = pattern -> { [[p]] } ]]; ssr_dpat: [ [ mp = ssr_mpat; "in"; t = pattern; rt = ssr_rtype -> { mp, mk_ctype mp t, rt } | mp = ssr_mpat; rt = ssr_rtype -> { mp, mk_cnotype mp, rt } | mp = ssr_mpat -> { mp, no_ct, no_rt } ] ]; ssr_dthen: [[ dp = ssr_dpat; "then"; c = lconstr -> { mk_dthen ~loc dp c } ]]; ssr_elsepat: [[ "else" -> { [[CAst.make ~loc @@ CPatAtom None]] } ]]; ssr_else: [[ mp = ssr_elsepat; c = lconstr -> { CAst.make ~loc (mp, c) } ]]; binder_constr: [ [ "if"; c = operconstr LEVEL "200"; "is"; db1 = ssr_dthen; b2 = ssr_else -> { let b1, ct, rt = db1 in CAst.make ~loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2]) } | "if"; c = operconstr LEVEL "200";"isn't";db1 = ssr_dthen; b2 = ssr_else -> { let b1, ct, rt = db1 in let b1, b2 = let open CAst in let {loc=l1; v=(p1, r1)}, {loc=l2; v=(p2, r2)} = b1, b2 in (make ?loc:l1 (p1, r2), make ?loc:l2 (p2, r1)) in CAst.make ~loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2]) } | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; "in"; c1 = lconstr -> { mk_let ~loc no_rt [mk_pat c no_ct] mp c1 } | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; rt = ssr_rtype; "in"; c1 = lconstr -> { mk_let ~loc rt [mk_pat c (mk_cnotype mp)] mp c1 } | "let"; ":"; mp = ssr_mpat; "in"; t = pattern; ":="; c = lconstr; rt = ssr_rtype; "in"; c1 = lconstr -> { mk_let ~loc rt [mk_pat c (mk_ctype mp t)] mp c1 } ] ]; END GRAMMAR EXTEND Gram GLOBAL: closed_binder; closed_binder: [ [ ["of" -> { () } | "&" -> { () } ]; c = operconstr LEVEL "99" -> { [CLocalAssum ([CAst.make ~loc Anonymous], Default Explicit, c)] } ] ]; END (** Vernacular commands: Prenex Implicits and Search *)(***********************) (* This should really be implemented as an extension to the implicit *) (* arguments feature, but unfortuately that API is sealed. The current *) (* workaround uses a combination of notations that works reasonably, *) (* with the following caveats: *) (* - The pretty-printing always elides prenex implicits, even when *) (* they are obviously needed. *) (* - Prenex Implicits are NEVER exported from a module, because this *) (* would lead to faulty pretty-printing and scoping errors. *) (* - The command "Import Prenex Implicits" can be used to reassert *) (* Prenex Implicits for all the visible constants that had been *) (* declared as Prenex Implicits. *) { let declare_one_prenex_implicit locality f = let fref = try Smartlocate.global_with_alias f with _ -> errorstrm (pr_qualid f ++ str " is not declared") in let rec loop = function | a :: args' when Impargs.is_status_implicit a -> Impargs.MaximallyImplicit :: loop args' | args' when List.exists Impargs.is_status_implicit args' -> errorstrm (str "Expected prenex implicits for " ++ pr_qualid f) | _ -> [] in let impls = match Impargs.implicits_of_global fref with | [cond,impls] -> impls | [] -> errorstrm (str "Expected some implicits for " ++ pr_qualid f) | _ -> errorstrm (str "Multiple implicits not supported") in match loop impls with | [] -> errorstrm (str "Expected some implicits for " ++ pr_qualid f) | impls -> Impargs.set_implicits locality fref [impls] } VERNAC COMMAND EXTEND Ssrpreneximplicits CLASSIFIED AS SIDEFF | #[ locality = Attributes.locality; ] [ "Prenex" "Implicits" ne_global_list(fl) ] -> { let locality = Locality.make_section_locality locality in List.iter (declare_one_prenex_implicit locality) fl; } END (* Vernac grammar visibility patch *) GRAMMAR EXTEND Gram GLOBAL: gallina_ext; gallina_ext: [ [ IDENT "Import"; IDENT "Prenex"; IDENT "Implicits" -> { Vernacexpr.VernacSetOption (false, ["Printing"; "Implicit"; "Defensive"], Vernacexpr.OptionUnset) } ] ] ; END (** Extend Search to subsume SearchAbout, also adding hidden Type coercions. *) (* Main prefilter *) { type raw_glob_search_about_item = | RGlobSearchSubPattern of constr_expr | RGlobSearchString of Loc.t * string * string option let pr_search_item env sigma = function | RGlobSearchString (_,s,_) -> str s | RGlobSearchSubPattern p -> pr_constr_expr env sigma p let wit_ssr_searchitem = add_genarg "ssr_searchitem" pr_search_item let pr_ssr_search_item env sigma _ _ _ = pr_search_item env sigma (* Workaround the notation API that can only print notations *) let is_ident s = try CLexer.check_ident s; true with _ -> false let is_ident_part s = is_ident ("H" ^ s) let interp_search_notation ?loc tag okey = let err msg = CErrors.user_err ?loc ~hdr:"interp_search_notation" msg in let mk_pntn s for_key = let n = String.length s in let s' = Bytes.make (n + 2) ' ' in let rec loop i i' = if i >= n then s', i' - 2 else if s.[i] = ' ' then loop (i + 1) i' else let j = try String.index_from s (i + 1) ' ' with _ -> n in let m = j - i in if s.[i] = '\'' && i < j - 2 && s.[j - 1] = '\'' then (String.blit s (i + 1) s' i' (m - 2); loop (j + 1) (i' + m - 1)) else if for_key && is_ident (String.sub s i m) then (Bytes.set s' i' '_'; loop (j + 1) (i' + 2)) else (String.blit s i s' i' m; loop (j + 1) (i' + m + 1)) in loop 0 1 in let trim_ntn (pntn, m) = (InConstrEntrySomeLevel,Bytes.sub_string pntn 1 (max 0 m)) in let pr_ntn ntn = str "(" ++ Notation.pr_notation ntn ++ str ")" in let pr_and_list pr = function | [x] -> pr x | x :: lx -> pr_list pr_comma pr lx ++ pr_comma () ++ str "and " ++ pr x | [] -> mt () in let pr_sc sc = str (if sc = "" then "independently" else sc) in let pr_scs = function | [""] -> pr_sc "" | scs -> str "in " ++ pr_and_list pr_sc scs in let generator, pr_tag_sc = let ign _ = mt () in match okey with | Some key -> let sc = Notation.find_delimiters_scope ?loc key in let pr_sc s_in = str s_in ++ spc() ++ str sc ++ pr_comma() in Notation.pr_scope ign sc, pr_sc | None -> Notation.pr_scopes ign, ign in let qtag s_in = pr_tag_sc s_in ++ qstring tag ++ spc()in let ptag, ttag = let ptag, m = mk_pntn tag false in if m <= 0 then err (str "empty notation fragment"); ptag, trim_ntn (ptag, m) in let last = ref "" and last_sc = ref "" in let scs = ref [] and ntns = ref [] in let push_sc sc = match !scs with | "" :: scs' -> scs := "" :: sc :: scs' | scs' -> scs := sc :: scs' in let get s _ _ = match !last with | "Scope " -> last_sc := s; last := "" | "Lonely notation" -> last_sc := ""; last := "" | "\"" -> let pntn, m = mk_pntn s true in if String.string_contains ~where:(Bytes.to_string pntn) ~what:(Bytes.to_string ptag) then begin let ntn = trim_ntn (pntn, m) in match !ntns with | [] -> ntns := [ntn]; scs := [!last_sc] | ntn' :: _ when ntn' = ntn -> push_sc !last_sc | _ when ntn = ttag -> ntns := ntn :: !ntns; scs := [!last_sc] | _ :: ntns' when List.mem ntn ntns' -> () | ntn' :: ntns' -> ntns := ntn' :: ntn :: ntns' end; last := "" | _ -> last := s in pp_with (Format.make_formatter get (fun _ -> ())) generator; let ntn = match !ntns with | [] -> err (hov 0 (qtag "in" ++ str "does not occur in any notation")) | ntn :: ntns' when ntn = ttag -> if ntns' <> [] then begin let pr_ntns' = pr_and_list pr_ntn ntns' in Feedback.msg_warning (hov 4 (qtag "In" ++ str "also occurs in " ++ pr_ntns')) end; ntn | [ntn] -> Feedback.msg_notice (hov 4 (qtag "In" ++ str "is part of notation " ++ pr_ntn ntn)); ntn | ntns' -> let e = str "occurs in" ++ spc() ++ pr_and_list pr_ntn ntns' in err (hov 4 (str "ambiguous: " ++ qtag "in" ++ e)) in let (nvars, body), ((_, pat), osc) = match !scs with | [sc] -> Notation.interp_notation ?loc ntn (None, [sc]) | scs' -> try Notation.interp_notation ?loc ntn (None, []) with _ -> let e = pr_ntn ntn ++ spc() ++ str "is defined " ++ pr_scs scs' in err (hov 4 (str "ambiguous: " ++ pr_tag_sc "in" ++ e)) in let sc = Option.default "" osc in let _ = let m_sc = if osc <> None then str "In " ++ str sc ++ pr_comma() else mt() in let ntn_pat = trim_ntn (mk_pntn pat false) in let rbody = glob_constr_of_notation_constr ?loc body in let m_body = hov 0 (Constrextern.without_symbols prl_glob_constr rbody) in let m = m_sc ++ pr_ntn ntn_pat ++ spc () ++ str "denotes " ++ m_body in Feedback.msg_notice (hov 0 m) in if List.length !scs > 1 then let scs' = List.remove (=) sc !scs in let w = pr_ntn ntn ++ str " is also defined " ++ pr_scs scs' in Feedback.msg_warning (hov 4 w) else if String.string_contains ~where:(snd ntn) ~what:" .. " then err (pr_ntn ntn ++ str " is an n-ary notation"); let nvars = List.filter (fun (_,(_,typ)) -> typ = NtnTypeConstr) nvars in let rec sub () = function | NVar x when List.mem_assoc x nvars -> DAst.make ?loc @@ GPatVar (FirstOrderPatVar x) | c -> glob_constr_of_notation_constr_with_binders ?loc (fun _ x -> (), None, x) sub () c in let _, npat = Patternops.pattern_of_glob_constr (sub () body) in Search.GlobSearchSubPattern npat } ARGUMENT EXTEND ssr_search_item TYPED AS ssr_searchitem PRINTED BY { pr_ssr_search_item env sigma } | [ string(s) ] -> { RGlobSearchString (loc,s,None) } | [ string(s) "%" preident(key) ] -> { RGlobSearchString (loc,s,Some key) } | [ constr_pattern(p) ] -> { RGlobSearchSubPattern p } END { let pr_ssr_search_arg env sigma _ _ _ = let pr_item (b, p) = str (if b then "-" else "") ++ pr_search_item env sigma p in pr_list spc pr_item } ARGUMENT EXTEND ssr_search_arg TYPED AS (bool * ssr_searchitem) list PRINTED BY { pr_ssr_search_arg env sigma } | [ "-" ssr_search_item(p) ssr_search_arg(a) ] -> { (false, p) :: a } | [ ssr_search_item(p) ssr_search_arg(a) ] -> { (true, p) :: a } | [ ] -> { [] } END { (* Main type conclusion pattern filter *) let rec splay_search_pattern na = function | Pattern.PApp (fp, args) -> splay_search_pattern (na + Array.length args) fp | Pattern.PLetIn (_, _, _, bp) -> splay_search_pattern na bp | Pattern.PRef hr -> hr, na | _ -> CErrors.user_err (Pp.str "no head constant in head search pattern") let push_rels_assum l e = let l = List.map (fun (n,t) -> n, EConstr.Unsafe.to_constr t) l in push_rels_assum l e let coerce_search_pattern_to_sort hpat = let env = Global.env () in let sigma = Evd.(from_env env) in let mkPApp fp n_imps args = let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in Pattern.PApp (fp, args') in let hr, na = splay_search_pattern 0 hpat in let dc, ht = let hr, _ = Typeops.type_of_global_in_context env hr (* FIXME *) in Reductionops.splay_prod env sigma (EConstr.of_constr hr) in let np = List.length dc in if np < na then CErrors.user_err (Pp.str "too many arguments in head search pattern") else let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in let warn () = Feedback.msg_warning (str "Listing only lemmas with conclusion matching " ++ pr_constr_pattern_env env sigma hpat') in if EConstr.isSort sigma ht then begin warn (); true, hpat' end else let filter_head, coe_path = try let _, cp = Classops.lookup_path_to_sort_from (push_rels_assum dc env) sigma ht in warn (); true, cp with _ -> false, [] in let coerce hp coe_index = let coe_ref = coe_index.Classops.coe_value in try let n_imps = Option.get (Classops.hide_coercion coe_ref) in mkPApp (Pattern.PRef coe_ref) n_imps [|hp|] with Not_found | Option.IsNone -> errorstrm (str "need explicit coercion " ++ pr_global coe_ref ++ spc () ++ str "to interpret head search pattern as type") in filter_head, List.fold_left coerce hpat' coe_path let interp_head_pat hpat = let filter_head, p = coerce_search_pattern_to_sort hpat in let rec loop c = match CoqConstr.kind c with | Cast (c', _, _) -> loop c' | Prod (_, _, c') -> loop c' | LetIn (_, _, _, c') -> loop c' | _ -> let env = Global.env () in let sigma = Evd.from_env env in Constr_matching.is_matching env sigma p (EConstr.of_constr c) in filter_head, loop let all_true _ = true let rec interp_search_about args accu = match args with | [] -> accu | (flag, arg) :: rem -> fun gr env typ -> let ans = Search.search_about_filter arg gr env typ in (if flag then ans else not ans) && interp_search_about rem accu gr env typ let interp_search_arg arg = let arg = List.map (fun (x,arg) -> x, match arg with | RGlobSearchString (loc,s,key) -> if is_ident_part s then Search.GlobSearchString s else interp_search_notation ~loc s key | RGlobSearchSubPattern p -> let env = Global.env () in let _, p = Constrintern.intern_constr_pattern env (Evd.from_env env) p in Search.GlobSearchSubPattern p) arg in let hpat, a1 = match arg with | (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a' | (true, Search.GlobSearchSubPattern p) :: a' -> let filter_head, p = interp_head_pat p in if filter_head then p, a' else all_true, arg | _ -> all_true, arg in let is_string = function (_, Search.GlobSearchString _) -> true | _ -> false in let a2, a3 = List.partition is_string a1 in interp_search_about (a2 @ a3) (fun gr env typ -> hpat typ) (* Module path postfilter *) let pr_modloc (b, m) = if b then str "-" ++ pr_qualid m else pr_qualid m let wit_ssrmodloc = add_genarg "ssrmodloc" (fun env sigma -> pr_modloc) let pr_ssr_modlocs _ _ _ ml = if ml = [] then str "" else spc () ++ str "in " ++ pr_list spc pr_modloc ml } ARGUMENT EXTEND ssr_modlocs TYPED AS ssrmodloc list PRINTED BY { pr_ssr_modlocs } | [ ] -> { [] } END GRAMMAR EXTEND Gram GLOBAL: ssr_modlocs; modloc: [[ "-"; m = global -> { true, m } | m = global -> { false, m } ]]; ssr_modlocs: [[ "in"; ml = LIST1 modloc -> { ml } ]]; END { let interp_modloc mr = let interp_mod (_, qid) = try Nametab.full_name_module qid with Not_found -> CErrors.user_err ?loc:qid.CAst.loc (str "No Module " ++ pr_qualid qid) in let mr_out, mr_in = List.partition fst mr in let interp_bmod b = function | [] -> fun _ _ _ -> true | rmods -> Search.module_filter (List.map interp_mod rmods, b) in let is_in = interp_bmod false mr_in and is_out = interp_bmod true mr_out in fun gr env typ -> is_in gr env typ && is_out gr env typ (* The unified, extended vernacular "Search" command *) let ssrdisplaysearch gr env t = let pr_res = pr_global gr ++ str ":" ++ spc () ++ pr_lconstr_env env Evd.empty t in Feedback.msg_notice (hov 2 pr_res ++ fnl ()) } VERNAC COMMAND EXTEND SsrSearchPattern CLASSIFIED AS QUERY | [ "Search" ssr_search_arg(a) ssr_modlocs(mr) ] -> { let hpat = interp_search_arg a in let in_mod = interp_modloc mr in let post_filter gr env typ = in_mod gr env typ && hpat gr env typ in let display gr env typ = if post_filter gr env typ then ssrdisplaysearch gr env typ in Search.generic_search None display } END (** View hint database and View application. *)(* ******************************) (* There are three databases of lemmas used to mediate the application *) (* of reflection lemmas: one for forward chaining, one for backward *) (* chaining, and one for secondary backward chaining. *) (* View hints *) { let pr_raw_ssrhintref env sigma prc _ _ = let open CAst in function | { v = CAppExpl ((None, r,x), args) } when isCHoles args -> prc env sigma (CAst.make @@ CRef (r,x)) ++ str "|" ++ int (List.length args) | { v = CApp ((_, { v = CRef _ }), _) } as c -> prc env sigma c | { v = CApp ((_, c), args) } when isCxHoles args -> prc env sigma c ++ str "|" ++ int (List.length args) | c -> prc env sigma c let pr_rawhintref env sigma c = match DAst.get c with | GApp (f, args) when isRHoles args -> pr_glob_constr_env env f ++ str "|" ++ int (List.length args) | _ -> pr_glob_constr_env env c let pr_glob_ssrhintref env sigma _ _ _ (c, _) = pr_rawhintref env sigma c let pr_ssrhintref env sigma prc _ _ = prc env sigma let mkhintref ?loc c n = match c.CAst.v with | CRef (r,x) -> CAst.make ?loc @@ CAppExpl ((None, r, x), mkCHoles ?loc n) | _ -> mkAppC (c, mkCHoles ?loc n) } ARGUMENT EXTEND ssrhintref TYPED AS constr PRINTED BY { pr_ssrhintref env sigma } RAW_PRINTED BY { pr_raw_ssrhintref env sigma } GLOB_PRINTED BY { pr_glob_ssrhintref env sigma } | [ constr(c) ] -> { c } | [ constr(c) "|" natural(n) ] -> { mkhintref ~loc c n } END { (* View purpose *) let pr_viewpos = function | Some Ssrview.AdaptorDb.Forward -> str " for move/" | Some Ssrview.AdaptorDb.Backward -> str " for apply/" | Some Ssrview.AdaptorDb.Equivalence -> str " for apply//" | None -> mt () let pr_ssrviewpos _ _ _ = pr_viewpos } ARGUMENT EXTEND ssrviewpos PRINTED BY { pr_ssrviewpos } | [ "for" "move" "/" ] -> { Some Ssrview.AdaptorDb.Forward } | [ "for" "apply" "/" ] -> { Some Ssrview.AdaptorDb.Backward } | [ "for" "apply" "/" "/" ] -> { Some Ssrview.AdaptorDb.Equivalence } | [ "for" "apply" "//" ] -> { Some Ssrview.AdaptorDb.Equivalence } | [ ] -> { None } END { let pr_ssrviewposspc _ _ _ i = pr_viewpos i ++ spc () } ARGUMENT EXTEND ssrviewposspc TYPED AS ssrviewpos PRINTED BY { pr_ssrviewposspc } | [ ssrviewpos(i) ] -> { i } END { let print_view_hints env sigma kind l = let pp_viewname = str "Hint View" ++ pr_viewpos (Some kind) ++ str " " in let pp_hints = pr_list spc (pr_rawhintref env sigma) l in Feedback.msg_notice (pp_viewname ++ hov 0 pp_hints ++ Pp.cut ()) } VERNAC COMMAND EXTEND PrintView CLASSIFIED AS QUERY | [ "Print" "Hint" "View" ssrviewpos(i) ] -> { let env = Global.env () in let sigma = Evd.from_env env in (match i with | Some k -> print_view_hints env sigma k (Ssrview.AdaptorDb.get k) | None -> List.iter (fun k -> print_view_hints env sigma k (Ssrview.AdaptorDb.get k)) [ Ssrview.AdaptorDb.Forward; Ssrview.AdaptorDb.Backward; Ssrview.AdaptorDb.Equivalence ]) } END { let glob_view_hints lvh = List.map (Constrintern.intern_constr (Global.env ()) (Evd.from_env (Global.env ()))) lvh } VERNAC COMMAND EXTEND HintView CLASSIFIED AS SIDEFF | [ "Hint" "View" ssrviewposspc(n) ne_ssrhintref_list(lvh) ] -> { let hints = glob_view_hints lvh in match n with | None -> Ssrview.AdaptorDb.declare Ssrview.AdaptorDb.Forward hints; Ssrview.AdaptorDb.declare Ssrview.AdaptorDb.Backward hints | Some k -> Ssrview.AdaptorDb.declare k hints } END (** Keyword compatibility fixes. *) (* Coq v8.1 notation uses "by" and "of" quasi-keywords, i.e., reserved *) (* identifiers used as keywords. This is incompatible with ssreflect.v *) (* which makes "by" and "of" true keywords, because of technicalities *) (* in the internal lexer-parser API of Coq. We patch this here by *) (* adding new parsing rules that recognize the new keywords. *) (* To make matters worse, the Coq grammar for tactics fails to *) (* export the non-terminals we need to patch. Fortunately, the CamlP5 *) (* API provides a backdoor access (with loads of Obj.magic trickery). *) (* Coq v8.3 defines "by" as a keyword, some hacks are not needed any *) (* longer and thus comment out. Such comments are marked with v8.3 *) { open Pltac } GRAMMAR EXTEND Gram GLOBAL: hypident; hypident: [ [ "("; IDENT "type"; "of"; id = Prim.identref; ")" -> { id, Locus.InHypTypeOnly } | "("; IDENT "value"; "of"; id = Prim.identref; ")" -> { id, Locus.InHypValueOnly } ] ]; END GRAMMAR EXTEND Gram GLOBAL: hloc; hloc: [ [ "in"; "("; "Type"; "of"; id = ident; ")" -> { Tacexpr.HypLocation (CAst.make id, Locus.InHypTypeOnly) } | "in"; "("; IDENT "Value"; "of"; id = ident; ")" -> { Tacexpr.HypLocation (CAst.make id, Locus.InHypValueOnly) } ] ]; END GRAMMAR EXTEND Gram GLOBAL: constr_eval; constr_eval: [ [ IDENT "type"; "of"; c = Constr.constr -> { Genredexpr.ConstrTypeOf c }] ]; END (* We wipe out all the keywords generated by the grammar rules we defined. *) (* The user is supposed to Require Import ssreflect or Require ssreflect *) (* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) (* consequence the extended ssreflect grammar. *) { let () = CLexer.set_keyword_state frozen_lexer ;; } (* vim: set filetype=ocaml foldmethod=marker: *) coq-8.11.0/plugins/ssr/ssrfun.v0000644000175000017500000010173513612664533016274 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* .doc { font-family: monospace; white-space: pre; } # **) Require Import ssreflect. (** This file contains the basic definitions and notations for working with functions. The definitions provide for: - Pair projections: p.1 == first element of a pair p.2 == second element of a pair These notations also apply to p : P /\ Q, via an and >-> pair coercion. - Simplifying functions, beta-reduced by /= and simpl: #[#fun : T => E#]# == constant function from type T that returns E #[#fun x => E#]# == unary function #[#fun x : T => E#]# == unary function with explicit domain type #[#fun x y => E#]# == binary function #[#fun x y : T => E#]# == binary function with common domain type #[#fun (x : T) y => E#]# \ #[#fun (x : xT) (y : yT) => E#]# | == binary function with (some) explicit, #[#fun x (y : T) => E#]# / independent domain types for each argument - Partial functions using option type: oapp f d ox == if ox is Some x returns f x, d otherwise odflt d ox == if ox is Some x returns x, d otherwise obind f ox == if ox is Some x returns f x, None otherwise omap f ox == if ox is Some x returns Some (f x), None otherwise - Singleton types: all_equal_to x0 == x0 is the only value in its type, so any such value can be rewritten to x0. - A generic wrapper type: wrapped T == the inductive type with values Wrap x for x : T. unwrap w == the projection of w : wrapped T on T. wrap x == the canonical injection of x : T into wrapped T; it is equivalent to Wrap x, but is declared as a (default) Canonical Structure, which lets the Coq HO unification automatically expand x into unwrap (wrap x). The delta reduction of wrap x to Wrap can be exploited to introduce controlled nondeterminism in Canonical Structure inference, as in the implementation of the mxdirect predicate in matrix.v. - The empty type: void == a notation for the Empty_set type of the standard library. of_void T == the canonical injection void -> T. - Sigma types: tag w == the i of w : {i : I & T i}. tagged w == the T i component of w : {i : I & T i}. Tagged T x == the {i : I & T i} with component x : T i. tag2 w == the i of w : {i : I & T i & U i}. tagged2 w == the T i component of w : {i : I & T i & U i}. tagged2' w == the U i component of w : {i : I & T i & U i}. Tagged2 T U x y == the {i : I & T i} with components x : T i and y : U i. sval u == the x of u : {x : T | P x}. s2val u == the x of u : {x : T | P x & Q x}. The properties of sval u, s2val u are given by lemmas svalP, s2valP, and s2valP'. We provide coercions sigT2 >-> sigT and sig2 >-> sig >-> sigT. A suite of lemmas (all_sig, ...) let us skolemize sig, sig2, sigT, sigT2 and pair, e.g., have /all_sig#[#f fP#]# (x : T): {y : U | P y} by ... yields an f : T -> U such that fP : forall x, P (f x). - Identity functions: id == NOTATION for the explicit identity function fun x => x. @id T == notation for the explicit identity at type T. idfun == an expression with a head constant, convertible to id; idfun x simplifies to x. @idfun T == the expression above, specialized to type T. phant_id x y == the function type phantom _ x -> phantom _ y. *** In addition to their casual use in functional programming, identity functions are often used to trigger static unification as part of the construction of dependent Records and Structures. For example, if we need a structure sT over a type T, we take as arguments T, sT, and a "dummy" function T -> sort sT: Definition foo T sT & T -> sort sT := ... We can avoid specifying sT directly by calling foo (@id T), or specify the call completely while still ensuring the consistency of T and sT, by calling @foo T sT idfun. The phant_id type allows us to extend this trick to non-Type canonical projections. It also allows us to sidestep dependent type constraints when building explicit records, e.g., given Record r := R {x; y : T(x)}. if we need to build an r from a given y0 while inferring some x0, such that y0 : T(x0), we pose Definition mk_r .. y .. (x := ...) y' & phant_id y y' := R x y'. Calling @mk_r .. y0 .. id will cause Coq to use y' := y0, while checking the dependent type constraint y0 : T(x0). - Extensional equality for functions and relations (i.e. functions of two arguments): f1 =1 f2 == f1 x is equal to f2 x for all x. f1 =1 f2 :> A == ... and f2 is explicitly typed. f1 =2 f2 == f1 x y is equal to f2 x y for all x y. f1 =2 f2 :> A == ... and f2 is explicitly typed. - Composition for total and partial functions: f^~ y == function f with second argument specialised to y, i.e., fun x => f x y CAVEAT: conditional (non-maximal) implicit arguments of f are NOT inserted in this context @^~ x == application at x, i.e., fun f => f x #[#eta f#]# == the explicit eta-expansion of f, i.e., fun x => f x CAVEAT: conditional (non-maximal) implicit arguments of f are NOT inserted in this context. fun=> v := the constant function fun _ => v. f1 \o f2 == composition of f1 and f2. Note: (f1 \o f2) x simplifies to f1 (f2 x). f1 \; f2 == categorical composition of f1 and f2. This expands to to f2 \o f1 and (f1 \; f2) x simplifies to f2 (f1 x). pcomp f1 f2 == composition of partial functions f1 and f2. - Properties of functions: injective f <-> f is injective. cancel f g <-> g is a left inverse of f / f is a right inverse of g. pcancel f g <-> g is a left inverse of f where g is partial. ocancel f g <-> g is a left inverse of f where f is partial. bijective f <-> f is bijective (has a left and right inverse). involutive f <-> f is involutive. - Properties for operations. left_id e op <-> e is a left identity for op (e op x = x). right_id e op <-> e is a right identity for op (x op e = x). left_inverse e inv op <-> inv is a left inverse for op wrt identity e, i.e., (inv x) op x = e. right_inverse e inv op <-> inv is a right inverse for op wrt identity e i.e., x op (i x) = e. self_inverse e op <-> each x is its own op-inverse (x op x = e). idempotent op <-> op is idempotent for op (x op x = x). associative op <-> op is associative, i.e., x op (y op z) = (x op y) op z. commutative op <-> op is commutative (x op y = y op x). left_commutative op <-> op is left commutative, i.e., x op (y op z) = y op (x op z). right_commutative op <-> op is right commutative, i.e., (x op y) op z = (x op z) op y. left_zero z op <-> z is a left zero for op (z op x = z). right_zero z op <-> z is a right zero for op (x op z = z). left_distributive op1 op2 <-> op1 distributes over op2 to the left: (x op2 y) op1 z = (x op1 z) op2 (y op1 z). right_distributive op1 op2 <-> op distributes over add to the right: x op1 (y op2 z) = (x op1 z) op2 (x op1 z). interchange op1 op2 <-> op1 and op2 satisfy an interchange law: (x op2 y) op1 (z op2 t) = (x op1 z) op2 (y op1 t). Note that interchange op op is a commutativity property. left_injective op <-> op is injective in its left argument: x op y = z op y -> x = z. right_injective op <-> op is injective in its right argument: x op y = x op z -> y = z. left_loop inv op <-> op, inv obey the inverse loop left axiom: (inv x) op (x op y) = y for all x, y, i.e., op (inv x) is always a left inverse of op x rev_left_loop inv op <-> op, inv obey the inverse loop reverse left axiom: x op ((inv x) op y) = y, for all x, y. right_loop inv op <-> op, inv obey the inverse loop right axiom: (x op y) op (inv y) = x for all x, y. rev_right_loop inv op <-> op, inv obey the inverse loop reverse right axiom: (x op (inv y)) op y = x for all x, y. Note that familiar "cancellation" identities like x + y - y = x or x - y + y = x are respectively instances of right_loop and rev_right_loop The corresponding lemmas will use the K and NK/VK suffixes, respectively. - Morphisms for functions and relations: {morph f : x / a >-> r} <-> f is a morphism with respect to functions (fun x => a) and (fun x => r); if r == R#[#x#]#, this states that f a = R#[#f x#]# for all x. {morph f : x / a} <-> f is a morphism with respect to the function expression (fun x => a). This is shorthand for {morph f : x / a >-> a}; note that the two instances of a are often interpreted at different types. {morph f : x y / a >-> r} <-> f is a morphism with respect to functions (fun x y => a) and (fun x y => r). {morph f : x y / a} <-> f is a morphism with respect to the function expression (fun x y => a). {homo f : x / a >-> r} <-> f is a homomorphism with respect to the predicates (fun x => a) and (fun x => r); if r == R#[#x#]#, this states that a -> R#[#f x#]# for all x. {homo f : x / a} <-> f is a homomorphism with respect to the predicate expression (fun x => a). {homo f : x y / a >-> r} <-> f is a homomorphism with respect to the relations (fun x y => a) and (fun x y => r). {homo f : x y / a} <-> f is a homomorphism with respect to the relation expression (fun x y => a). {mono f : x / a >-> r} <-> f is monotone with respect to projectors (fun x => a) and (fun x => r); if r == R#[#x#]#, this states that R#[#f x#]# = a for all x. {mono f : x / a} <-> f is monotone with respect to the projector expression (fun x => a). {mono f : x y / a >-> r} <-> f is monotone with respect to relators (fun x y => a) and (fun x y => r). {mono f : x y / a} <-> f is monotone with respect to the relator expression (fun x y => a). The file also contains some basic lemmas for the above concepts. Lemmas relative to cancellation laws use some abbreviated suffixes: K - a cancellation rule like esymK : cancel (@esym T x y) (@esym T y x). LR - a lemma moving an operation from the left hand side of a relation to the right hand side, like canLR: cancel g f -> x = g y -> f x = y. RL - a lemma moving an operation from the right to the left, e.g., canRL. Beware that the LR and RL orientations refer to an "apply" (back chaining) usage; when using the same lemmas with "have" or "move" (forward chaining) the directions will be reversed!. **) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. (** Parsing / printing declarations. *) Reserved Notation "p .1" (at level 2, left associativity, format "p .1"). Reserved Notation "p .2" (at level 2, left associativity, format "p .2"). Reserved Notation "f ^~ y" (at level 10, y at level 8, no associativity, format "f ^~ y"). Reserved Notation "@^~ x" (at level 10, x at level 8, no associativity, format "@^~ x"). Reserved Notation "[ 'eta' f ]" (at level 0, format "[ 'eta' f ]"). Reserved Notation "'fun' => E" (at level 200, format "'fun' => E"). Reserved Notation "[ 'fun' : T => E ]" (at level 0, format "'[hv' [ 'fun' : T => '/ ' E ] ']'"). Reserved Notation "[ 'fun' x => E ]" (at level 0, x ident, format "'[hv' [ 'fun' x => '/ ' E ] ']'"). Reserved Notation "[ 'fun' x : T => E ]" (at level 0, x ident, format "'[hv' [ 'fun' x : T => '/ ' E ] ']'"). Reserved Notation "[ 'fun' x y => E ]" (at level 0, x ident, y ident, format "'[hv' [ 'fun' x y => '/ ' E ] ']'"). Reserved Notation "[ 'fun' x y : T => E ]" (at level 0, x ident, y ident, format "'[hv' [ 'fun' x y : T => '/ ' E ] ']'"). Reserved Notation "[ 'fun' ( x : T ) y => E ]" (at level 0, x ident, y ident, format "'[hv' [ 'fun' ( x : T ) y => '/ ' E ] ']'"). Reserved Notation "[ 'fun' x ( y : T ) => E ]" (at level 0, x ident, y ident, format "'[hv' [ 'fun' x ( y : T ) => '/ ' E ] ']'"). Reserved Notation "[ 'fun' ( x : T ) ( y : U ) => E ]" (at level 0, x ident, y ident, format "[ 'fun' ( x : T ) ( y : U ) => E ]" ). Reserved Notation "f =1 g" (at level 70, no associativity). Reserved Notation "f =1 g :> A" (at level 70, g at next level, A at level 90). Reserved Notation "f =2 g" (at level 70, no associativity). Reserved Notation "f =2 g :> A" (at level 70, g at next level, A at level 90). Reserved Notation "f \o g" (at level 50, format "f \o '/ ' g"). Reserved Notation "f \; g" (at level 60, right associativity, format "f \; '/ ' g"). Reserved Notation "{ 'morph' f : x / a >-> r }" (at level 0, f at level 99, x ident, format "{ 'morph' f : x / a >-> r }"). Reserved Notation "{ 'morph' f : x / a }" (at level 0, f at level 99, x ident, format "{ 'morph' f : x / a }"). Reserved Notation "{ 'morph' f : x y / a >-> r }" (at level 0, f at level 99, x ident, y ident, format "{ 'morph' f : x y / a >-> r }"). Reserved Notation "{ 'morph' f : x y / a }" (at level 0, f at level 99, x ident, y ident, format "{ 'morph' f : x y / a }"). Reserved Notation "{ 'homo' f : x / a >-> r }" (at level 0, f at level 99, x ident, format "{ 'homo' f : x / a >-> r }"). Reserved Notation "{ 'homo' f : x / a }" (at level 0, f at level 99, x ident, format "{ 'homo' f : x / a }"). Reserved Notation "{ 'homo' f : x y / a >-> r }" (at level 0, f at level 99, x ident, y ident, format "{ 'homo' f : x y / a >-> r }"). Reserved Notation "{ 'homo' f : x y / a }" (at level 0, f at level 99, x ident, y ident, format "{ 'homo' f : x y / a }"). Reserved Notation "{ 'homo' f : x y /~ a }" (at level 0, f at level 99, x ident, y ident, format "{ 'homo' f : x y /~ a }"). Reserved Notation "{ 'mono' f : x / a >-> r }" (at level 0, f at level 99, x ident, format "{ 'mono' f : x / a >-> r }"). Reserved Notation "{ 'mono' f : x / a }" (at level 0, f at level 99, x ident, format "{ 'mono' f : x / a }"). Reserved Notation "{ 'mono' f : x y / a >-> r }" (at level 0, f at level 99, x ident, y ident, format "{ 'mono' f : x y / a >-> r }"). Reserved Notation "{ 'mono' f : x y / a }" (at level 0, f at level 99, x ident, y ident, format "{ 'mono' f : x y / a }"). Reserved Notation "{ 'mono' f : x y /~ a }" (at level 0, f at level 99, x ident, y ident, format "{ 'mono' f : x y /~ a }"). Reserved Notation "@ 'id' T" (at level 10, T at level 8, format "@ 'id' T"). Reserved Notation "@ 'sval'" (at level 10, format "@ 'sval'"). (** Syntax for defining auxiliary recursive function. Usage: Section FooDefinition. Variables (g1 : T1) (g2 : T2). (globals) Fixoint foo_auxiliary (a3 : T3) ... := body, using #[#rec e3, ... #]# for recursive calls where " #[# 'rec' a3 , a4 , ... #]#" := foo_auxiliary. Definition foo x y .. := #[#rec e1, ... #]#. + proofs about foo End FooDefinition. **) Reserved Notation "[ 'rec' a ]" (at level 0, format "[ 'rec' a ]"). Reserved Notation "[ 'rec' a , b ]" (at level 0, format "[ 'rec' a , b ]"). Reserved Notation "[ 'rec' a , b , c ]" (at level 0, format "[ 'rec' a , b , c ]"). Reserved Notation "[ 'rec' a , b , c , d ]" (at level 0, format "[ 'rec' a , b , c , d ]"). Reserved Notation "[ 'rec' a , b , c , d , e ]" (at level 0, format "[ 'rec' a , b , c , d , e ]"). Reserved Notation "[ 'rec' a , b , c , d , e , f ]" (at level 0, format "[ 'rec' a , b , c , d , e , f ]"). Reserved Notation "[ 'rec' a , b , c , d , e , f , g ]" (at level 0, format "[ 'rec' a , b , c , d , e , f , g ]"). Reserved Notation "[ 'rec' a , b , c , d , e , f , g , h ]" (at level 0, format "[ 'rec' a , b , c , d , e , f , g , h ]"). Reserved Notation "[ 'rec' a , b , c , d , e , f , g , h , i ]" (at level 0, format "[ 'rec' a , b , c , d , e , f , g , h , i ]"). Reserved Notation "[ 'rec' a , b , c , d , e , f , g , h , i , j ]" (at level 0, format "[ 'rec' a , b , c , d , e , f , g , h , i , j ]"). Declare Scope pair_scope. Delimit Scope pair_scope with PAIR. Open Scope pair_scope. (** Notations for pair/conjunction projections **) Notation "p .1" := (fst p) : pair_scope. Notation "p .2" := (snd p) : pair_scope. Coercion pair_of_and P Q (PandQ : P /\ Q) := (proj1 PandQ, proj2 PandQ). Definition all_pair I T U (w : forall i : I, T i * U i) := (fun i => (w i).1, fun i => (w i).2). (** Complements on the option type constructor, used below to encode partial functions. **) Module Option. Definition apply aT rT (f : aT -> rT) x u := if u is Some y then f y else x. Definition default T := apply (fun x : T => x). Definition bind aT rT (f : aT -> option rT) := apply f None. Definition map aT rT (f : aT -> rT) := bind (fun x => Some (f x)). End Option. Notation oapp := Option.apply. Notation odflt := Option.default. Notation obind := Option.bind. Notation omap := Option.map. Notation some := (@Some _) (only parsing). (** Shorthand for some basic equality lemmas. **) Notation erefl := refl_equal. Notation ecast i T e x := (let: erefl in _ = i := e return T in x). Definition esym := sym_eq. Definition nesym := sym_not_eq. Definition etrans := trans_eq. Definition congr1 := f_equal. Definition congr2 := f_equal2. (** Force at least one implicit when used as a view. **) Prenex Implicits esym nesym. (** A predicate for singleton types. **) Definition all_equal_to T (x0 : T) := forall x, unkeyed x = x0. Lemma unitE : all_equal_to tt. Proof. by case. Qed. (** A generic wrapper type **) #[universes(template)] Structure wrapped T := Wrap {unwrap : T}. Canonical wrap T x := @Wrap T x. Prenex Implicits unwrap wrap Wrap. Declare Scope fun_scope. Delimit Scope fun_scope with FUN. Open Scope fun_scope. (** Notations for argument transpose **) Notation "f ^~ y" := (fun x => f x y) : fun_scope. Notation "@^~ x" := (fun f => f x) : fun_scope. (** Definitions and notation for explicit functions with simplification, i.e., which simpl and /= beta expand (this is complementary to nosimpl). **) #[universes(template)] Variant simpl_fun (aT rT : Type) := SimplFun of aT -> rT. Section SimplFun. Variables aT rT : Type. Definition fun_of_simpl (f : simpl_fun aT rT) := fun x => let: SimplFun lam := f in lam x. End SimplFun. Coercion fun_of_simpl : simpl_fun >-> Funclass. Notation "[ 'fun' : T => E ]" := (SimplFun (fun _ : T => E)) : fun_scope. Notation "[ 'fun' x => E ]" := (SimplFun (fun x => E)) : fun_scope. Notation "[ 'fun' x y => E ]" := (fun x => [fun y => E]) : fun_scope. Notation "[ 'fun' x : T => E ]" := (SimplFun (fun x : T => E)) (only parsing) : fun_scope. Notation "[ 'fun' x y : T => E ]" := (fun x : T => [fun y : T => E]) (only parsing) : fun_scope. Notation "[ 'fun' ( x : T ) y => E ]" := (fun x : T => [fun y => E]) (only parsing) : fun_scope. Notation "[ 'fun' x ( y : T ) => E ]" := (fun x => [fun y : T => E]) (only parsing) : fun_scope. Notation "[ 'fun' ( x : T ) ( y : U ) => E ]" := (fun x : T => [fun y : U => E]) (only parsing) : fun_scope. (** For delta functions in eqtype.v. **) Definition SimplFunDelta aT rT (f : aT -> aT -> rT) := [fun z => f z z]. (** Extensional equality, for unary and binary functions, including syntactic sugar. **) Section ExtensionalEquality. Variables A B C : Type. Definition eqfun (f g : B -> A) : Prop := forall x, f x = g x. Definition eqrel (r s : C -> B -> A) : Prop := forall x y, r x y = s x y. Lemma frefl f : eqfun f f. Proof. by []. Qed. Lemma fsym f g : eqfun f g -> eqfun g f. Proof. by move=> eq_fg x. Qed. Lemma ftrans f g h : eqfun f g -> eqfun g h -> eqfun f h. Proof. by move=> eq_fg eq_gh x; rewrite eq_fg. Qed. Lemma rrefl r : eqrel r r. Proof. by []. Qed. End ExtensionalEquality. Typeclasses Opaque eqfun. Typeclasses Opaque eqrel. Hint Resolve frefl rrefl : core. Notation "f1 =1 f2" := (eqfun f1 f2) : fun_scope. Notation "f1 =1 f2 :> A" := (f1 =1 (f2 : A)) : fun_scope. Notation "f1 =2 f2" := (eqrel f1 f2) : fun_scope. Notation "f1 =2 f2 :> A" := (f1 =2 (f2 : A)) : fun_scope. Section Composition. Variables A B C : Type. Definition comp (f : B -> A) (g : C -> B) x := f (g x). Definition catcomp g f := comp f g. Definition pcomp (f : B -> option A) (g : C -> option B) x := obind f (g x). Lemma eq_comp f f' g g' : f =1 f' -> g =1 g' -> comp f g =1 comp f' g'. Proof. by move=> eq_ff' eq_gg' x; rewrite /comp eq_gg' eq_ff'. Qed. End Composition. Arguments comp {A B C} f g x /. Arguments catcomp {A B C} g f x /. Notation "f1 \o f2" := (comp f1 f2) : fun_scope. Notation "f1 \; f2" := (catcomp f1 f2) : fun_scope. Notation "[ 'eta' f ]" := (fun x => f x) : fun_scope. Notation "'fun' => E" := (fun _ => E) : fun_scope. Notation id := (fun x => x). Notation "@ 'id' T" := (fun x : T => x) (only parsing) : fun_scope. Definition idfun T x : T := x. Arguments idfun {T} x /. Definition phant_id T1 T2 v1 v2 := phantom T1 v1 -> phantom T2 v2. (** The empty type. **) Notation void := Empty_set. Definition of_void T (x : void) : T := match x with end. (** Strong sigma types. **) Section Tag. Variables (I : Type) (i : I) (T_ U_ : I -> Type). Definition tag := projT1. Definition tagged : forall w, T_(tag w) := @projT2 I [eta T_]. Definition Tagged x := @existT I [eta T_] i x. Definition tag2 (w : @sigT2 I T_ U_) := let: existT2 _ _ i _ _ := w in i. Definition tagged2 w : T_(tag2 w) := let: existT2 _ _ _ x _ := w in x. Definition tagged2' w : U_(tag2 w) := let: existT2 _ _ _ _ y := w in y. Definition Tagged2 x y := @existT2 I [eta T_] [eta U_] i x y. End Tag. Arguments Tagged [I i]. Arguments Tagged2 [I i]. Prenex Implicits tag tagged Tagged tag2 tagged2 tagged2' Tagged2. Coercion tag_of_tag2 I T_ U_ (w : @sigT2 I T_ U_) := Tagged (fun i => T_ i * U_ i)%type (tagged2 w, tagged2' w). Lemma all_tag I T U : (forall x : I, {y : T x & U x y}) -> {f : forall x, T x & forall x, U x (f x)}. Proof. by move=> fP; exists (fun x => tag (fP x)) => x; case: (fP x). Qed. Lemma all_tag2 I T U V : (forall i : I, {y : T i & U i y & V i y}) -> {f : forall i, T i & forall i, U i (f i) & forall i, V i (f i)}. Proof. by case/all_tag=> f /all_pair[]; exists f. Qed. (** Refinement types. **) (** Prenex Implicits and renaming. **) Notation sval := (@proj1_sig _ _). Notation "@ 'sval'" := (@proj1_sig) (at level 10, format "@ 'sval'"). Section Sig. Variables (T : Type) (P Q : T -> Prop). Lemma svalP (u : sig P) : P (sval u). Proof. by case: u. Qed. Definition s2val (u : sig2 P Q) := let: exist2 _ _ x _ _ := u in x. Lemma s2valP u : P (s2val u). Proof. by case: u. Qed. Lemma s2valP' u : Q (s2val u). Proof. by case: u. Qed. End Sig. Prenex Implicits svalP s2val s2valP s2valP'. Coercion tag_of_sig I P (u : @sig I P) := Tagged P (svalP u). Coercion sig_of_sig2 I P Q (u : @sig2 I P Q) := exist (fun i => P i /\ Q i) (s2val u) (conj (s2valP u) (s2valP' u)). Lemma all_sig I T P : (forall x : I, {y : T x | P x y}) -> {f : forall x, T x | forall x, P x (f x)}. Proof. by case/all_tag=> f; exists f. Qed. Lemma all_sig2 I T P Q : (forall x : I, {y : T x | P x y & Q x y}) -> {f : forall x, T x | forall x, P x (f x) & forall x, Q x (f x)}. Proof. by case/all_sig=> f /all_pair[]; exists f. Qed. Section Morphism. Variables (aT rT sT : Type) (f : aT -> rT). (** Morphism property for unary and binary functions **) Definition morphism_1 aF rF := forall x, f (aF x) = rF (f x). Definition morphism_2 aOp rOp := forall x y, f (aOp x y) = rOp (f x) (f y). (** Homomorphism property for unary and binary relations **) Definition homomorphism_1 (aP rP : _ -> Prop) := forall x, aP x -> rP (f x). Definition homomorphism_2 (aR rR : _ -> _ -> Prop) := forall x y, aR x y -> rR (f x) (f y). (** Stability property for unary and binary relations **) Definition monomorphism_1 (aP rP : _ -> sT) := forall x, rP (f x) = aP x. Definition monomorphism_2 (aR rR : _ -> _ -> sT) := forall x y, rR (f x) (f y) = aR x y. End Morphism. Notation "{ 'morph' f : x / a >-> r }" := (morphism_1 f (fun x => a) (fun x => r)) : type_scope. Notation "{ 'morph' f : x / a }" := (morphism_1 f (fun x => a) (fun x => a)) : type_scope. Notation "{ 'morph' f : x y / a >-> r }" := (morphism_2 f (fun x y => a) (fun x y => r)) : type_scope. Notation "{ 'morph' f : x y / a }" := (morphism_2 f (fun x y => a) (fun x y => a)) : type_scope. Notation "{ 'homo' f : x / a >-> r }" := (homomorphism_1 f (fun x => a) (fun x => r)) : type_scope. Notation "{ 'homo' f : x / a }" := (homomorphism_1 f (fun x => a) (fun x => a)) : type_scope. Notation "{ 'homo' f : x y / a >-> r }" := (homomorphism_2 f (fun x y => a) (fun x y => r)) : type_scope. Notation "{ 'homo' f : x y / a }" := (homomorphism_2 f (fun x y => a) (fun x y => a)) : type_scope. Notation "{ 'homo' f : x y /~ a }" := (homomorphism_2 f (fun y x => a) (fun x y => a)) : type_scope. Notation "{ 'mono' f : x / a >-> r }" := (monomorphism_1 f (fun x => a) (fun x => r)) : type_scope. Notation "{ 'mono' f : x / a }" := (monomorphism_1 f (fun x => a) (fun x => a)) : type_scope. Notation "{ 'mono' f : x y / a >-> r }" := (monomorphism_2 f (fun x y => a) (fun x y => r)) : type_scope. Notation "{ 'mono' f : x y / a }" := (monomorphism_2 f (fun x y => a) (fun x y => a)) : type_scope. Notation "{ 'mono' f : x y /~ a }" := (monomorphism_2 f (fun y x => a) (fun x y => a)) : type_scope. (** In an intuitionistic setting, we have two degrees of injectivity. The weaker one gives only simplification, and the strong one provides a left inverse (we show in `fintype' that they coincide for finite types). We also define an intermediate version where the left inverse is only a partial function. **) Section Injections. Variables (rT aT : Type) (f : aT -> rT). Definition injective := forall x1 x2, f x1 = f x2 -> x1 = x2. Definition cancel g := forall x, g (f x) = x. Definition pcancel g := forall x, g (f x) = Some x. Definition ocancel (g : aT -> option rT) h := forall x, oapp h x (g x) = x. Lemma can_pcan g : cancel g -> pcancel (fun y => Some (g y)). Proof. by move=> fK x; congr (Some _). Qed. Lemma pcan_inj g : pcancel g -> injective. Proof. by move=> fK x y /(congr1 g); rewrite !fK => [[]]. Qed. Lemma can_inj g : cancel g -> injective. Proof. by move/can_pcan; apply: pcan_inj. Qed. Lemma canLR g x y : cancel g -> x = f y -> g x = y. Proof. by move=> fK ->. Qed. Lemma canRL g x y : cancel g -> f x = y -> x = g y. Proof. by move=> fK <-. Qed. End Injections. Lemma Some_inj {T : nonPropType} : injective (@Some T). Proof. by move=> x y []. Qed. Lemma of_voidK T : pcancel (of_void T) [fun _ => None]. Proof. by case. Qed. (** cancellation lemmas for dependent type casts. **) Lemma esymK T x y : cancel (@esym T x y) (@esym T y x). Proof. by case: y /. Qed. Lemma etrans_id T x y (eqxy : x = y :> T) : etrans (erefl x) eqxy = eqxy. Proof. by case: y / eqxy. Qed. Section InjectionsTheory. Variables (A B C : Type) (f g : B -> A) (h : C -> B). Lemma inj_id : injective (@id A). Proof. by []. Qed. Lemma inj_can_sym f' : cancel f f' -> injective f' -> cancel f' f. Proof. by move=> fK injf' x; apply: injf'. Qed. Lemma inj_comp : injective f -> injective h -> injective (f \o h). Proof. by move=> injf injh x y /injf; apply: injh. Qed. Lemma inj_compr : injective (f \o h) -> injective h. Proof. by move=> injfh x y /(congr1 f) /injfh. Qed. Lemma can_comp f' h' : cancel f f' -> cancel h h' -> cancel (f \o h) (h' \o f'). Proof. by move=> fK hK x; rewrite /= fK hK. Qed. Lemma pcan_pcomp f' h' : pcancel f f' -> pcancel h h' -> pcancel (f \o h) (pcomp h' f'). Proof. by move=> fK hK x; rewrite /pcomp fK /= hK. Qed. Lemma eq_inj : injective f -> f =1 g -> injective g. Proof. by move=> injf eqfg x y; rewrite -2!eqfg; apply: injf. Qed. Lemma eq_can f' g' : cancel f f' -> f =1 g -> f' =1 g' -> cancel g g'. Proof. by move=> fK eqfg eqfg' x; rewrite -eqfg -eqfg'. Qed. Lemma inj_can_eq f' : cancel f f' -> injective f' -> cancel g f' -> f =1 g. Proof. by move=> fK injf' gK x; apply: injf'; rewrite fK. Qed. End InjectionsTheory. Section Bijections. Variables (A B : Type) (f : B -> A). Variant bijective : Prop := Bijective g of cancel f g & cancel g f. Hypothesis bijf : bijective. Lemma bij_inj : injective f. Proof. by case: bijf => g fK _; apply: can_inj fK. Qed. Lemma bij_can_sym f' : cancel f' f <-> cancel f f'. Proof. split=> fK; first exact: inj_can_sym fK bij_inj. by case: bijf => h _ hK x; rewrite -[x]hK fK. Qed. Lemma bij_can_eq f' f'' : cancel f f' -> cancel f f'' -> f' =1 f''. Proof. by move=> fK fK'; apply: (inj_can_eq _ bij_inj); apply/bij_can_sym. Qed. End Bijections. Section BijectionsTheory. Variables (A B C : Type) (f : B -> A) (h : C -> B). Lemma eq_bij : bijective f -> forall g, f =1 g -> bijective g. Proof. by case=> f' fK f'K g eqfg; exists f'; eapply eq_can; eauto. Qed. Lemma bij_comp : bijective f -> bijective h -> bijective (f \o h). Proof. by move=> [f' fK f'K] [h' hK h'K]; exists (h' \o f'); apply: can_comp; auto. Qed. Lemma bij_can_bij : bijective f -> forall f', cancel f f' -> bijective f'. Proof. by move=> bijf; exists f; first by apply/(bij_can_sym bijf). Qed. End BijectionsTheory. Section Involutions. Variables (A : Type) (f : A -> A). Definition involutive := cancel f f. Hypothesis Hf : involutive. Lemma inv_inj : injective f. Proof. exact: can_inj Hf. Qed. Lemma inv_bij : bijective f. Proof. by exists f. Qed. End Involutions. Section OperationProperties. Variables S T R : Type. Section SopTisR. Implicit Type op : S -> T -> R. Definition left_inverse e inv op := forall x, op (inv x) x = e. Definition right_inverse e inv op := forall x, op x (inv x) = e. Definition left_injective op := forall x, injective (op^~ x). Definition right_injective op := forall y, injective (op y). End SopTisR. Section SopTisS. Implicit Type op : S -> T -> S. Definition right_id e op := forall x, op x e = x. Definition left_zero z op := forall x, op z x = z. Definition right_commutative op := forall x y z, op (op x y) z = op (op x z) y. Definition left_distributive op add := forall x y z, op (add x y) z = add (op x z) (op y z). Definition right_loop inv op := forall y, cancel (op^~ y) (op^~ (inv y)). Definition rev_right_loop inv op := forall y, cancel (op^~ (inv y)) (op^~ y). End SopTisS. Section SopTisT. Implicit Type op : S -> T -> T. Definition left_id e op := forall x, op e x = x. Definition right_zero z op := forall x, op x z = z. Definition left_commutative op := forall x y z, op x (op y z) = op y (op x z). Definition right_distributive op add := forall x y z, op x (add y z) = add (op x y) (op x z). Definition left_loop inv op := forall x, cancel (op x) (op (inv x)). Definition rev_left_loop inv op := forall x, cancel (op (inv x)) (op x). End SopTisT. Section SopSisT. Implicit Type op : S -> S -> T. Definition self_inverse e op := forall x, op x x = e. Definition commutative op := forall x y, op x y = op y x. End SopSisT. Section SopSisS. Implicit Type op : S -> S -> S. Definition idempotent op := forall x, op x x = x. Definition associative op := forall x y z, op x (op y z) = op (op x y) z. Definition interchange op1 op2 := forall x y z t, op1 (op2 x y) (op2 z t) = op2 (op1 x z) (op1 y t). End SopSisS. End OperationProperties. coq-8.11.0/plugins/ssr/ssrcommon.mli0000644000175000017500000004034413612664533017306 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Id.t val hyps_ids : ssrhyps -> Id.t list val check_hyp_exists : ('a, 'b) Context.Named.pt -> ssrhyp -> unit val test_hyp_exists : ('a, 'b) Context.Named.pt -> ssrhyp -> bool val check_hyps_uniq : Id.t list -> ssrhyps -> unit val not_section_id : Id.t -> bool val hyp_err : ?loc:Loc.t -> string -> Id.t -> 'a val hoik : (ssrhyp -> 'a) -> ssrhyp_or_id -> 'a val hoi_id : ssrhyp_or_id -> Id.t (******************************* hints ***********************************) val mk_hint : 'a -> 'a ssrhint val mk_orhint : 'a -> bool * 'a val nullhint : bool * 'a list val nohint : 'a ssrhint (******************************** misc ************************************) val errorstrm : Pp.t -> 'a val anomaly : string -> 'a val array_app_tl : 'a array -> 'a list -> 'a list val array_list_of_tl : 'a array -> 'a list val array_fold_right_from : int -> ('a -> 'b -> 'b) -> 'a array -> 'b -> 'b val option_assert_get : 'a option -> Pp.t -> 'a (**************************** lifted tactics ******************************) (* tactics with extra data attached to each goals, e.g. the list of * temporary variables to be cleared *) type 'a tac_a = (goal * 'a) sigma -> (goal * 'a) list sigma (* Thread around names to be cleared or generalized back, and the speed *) type tac_ctx = { tmp_ids : (Id.t * Name.t ref) list; wild_ids : Id.t list; (* List of variables to be cleared at the end of the sentence *) delayed_clears : Id.t list; } val new_ctx : unit -> tac_ctx (* REMOVE *) val pull_ctxs : ('a * tac_ctx) list sigma -> 'a list sigma * tac_ctx list (* REMOVE *) val with_fresh_ctx : tac_ctx tac_a -> tactic val pull_ctx : ('a * tac_ctx) sigma -> 'a sigma * tac_ctx val push_ctx : tac_ctx -> 'a sigma -> ('a * tac_ctx) sigma val push_ctxs : tac_ctx -> 'a list sigma -> ('a * tac_ctx) list sigma val tac_ctx : tactic -> tac_ctx tac_a val with_ctx : (tac_ctx -> 'b * tac_ctx) -> ('a * tac_ctx) sigma -> 'b * ('a * tac_ctx) sigma val without_ctx : ('a sigma -> 'b) -> ('a * tac_ctx) sigma -> 'b (* Standard tacticals lifted to the tac_a type *) val tclTHENLIST_a : tac_ctx tac_a list -> tac_ctx tac_a val tclTHEN_i_max : tac_ctx tac_a -> (int -> int -> tac_ctx tac_a) -> tac_ctx tac_a val tclTHEN_a : tac_ctx tac_a -> tac_ctx tac_a -> tac_ctx tac_a val tclTHENS_a : tac_ctx tac_a -> tac_ctx tac_a list -> tac_ctx tac_a val tac_on_all : (goal * tac_ctx) list sigma -> tac_ctx tac_a -> (goal * tac_ctx) list sigma (************************ ssr tactic arguments ******************************) (*********************** Misc helpers *****************************) val mkRHole : Glob_term.glob_constr val mkRHoles : int -> Glob_term.glob_constr list val isRHoles : Glob_term.glob_constr list -> bool val mkRApp : Glob_term.glob_constr -> Glob_term.glob_constr list -> Glob_term.glob_constr val mkRVar : Id.t -> Glob_term.glob_constr val mkRltacVar : Id.t -> Glob_term.glob_constr val mkRCast : Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr val mkRType : Glob_term.glob_constr val mkRProp : Glob_term.glob_constr val mkRArrow : Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr val mkRConstruct : Names.constructor -> Glob_term.glob_constr val mkRInd : Names.inductive -> Glob_term.glob_constr val mkRLambda : Name.t -> Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr val mkRnat : int -> Glob_term.glob_constr val mkCHole : Loc.t option -> constr_expr val mkCHoles : ?loc:Loc.t -> int -> constr_expr list val mkCVar : ?loc:Loc.t -> Id.t -> constr_expr val mkCCast : ?loc:Loc.t -> constr_expr -> constr_expr -> constr_expr val mkCType : Loc.t option -> constr_expr val mkCProp : Loc.t option -> constr_expr val mkCArrow : ?loc:Loc.t -> constr_expr -> constr_expr -> constr_expr val mkCLambda : ?loc:Loc.t -> Name.t -> constr_expr -> constr_expr -> constr_expr val isCHoles : constr_expr list -> bool val isCxHoles : (constr_expr * 'a option) list -> bool val intern_term : Tacinterp.interp_sign -> env -> ssrterm -> Glob_term.glob_constr val pf_intern_term : Tacinterp.interp_sign -> Goal.goal Evd.sigma -> ssrterm -> Glob_term.glob_constr val interp_term : Tacinterp.interp_sign -> Goal.goal Evd.sigma -> ssrterm -> evar_map * EConstr.t val interp_wit : ('a, 'b, 'c) genarg_type -> ist -> goal sigma -> 'b -> evar_map * 'c val interp_hyp : ist -> goal sigma -> ssrhyp -> evar_map * ssrhyp val interp_hyps : ist -> goal sigma -> ssrhyps -> evar_map * ssrhyps val interp_refine : Tacinterp.interp_sign -> Goal.goal Evd.sigma -> Glob_term.glob_constr -> evar_map * (evar_map * EConstr.constr) val interp_open_constr : Tacinterp.interp_sign -> Goal.goal Evd.sigma -> Genintern.glob_constr_and_expr -> evar_map * (evar_map * EConstr.t) val pf_e_type_of : Goal.goal Evd.sigma -> EConstr.constr -> Goal.goal Evd.sigma * EConstr.types val splay_open_constr : Goal.goal Evd.sigma -> evar_map * EConstr.t -> (Names.Name.t Context.binder_annot * EConstr.t) list * EConstr.t val isAppInd : Environ.env -> Evd.evar_map -> EConstr.types -> bool val mk_term : ssrtermkind -> constr_expr -> ssrterm val mk_lterm : constr_expr -> ssrterm val mk_ast_closure_term : [ `None | `Parens | `DoubleParens | `At ] -> Constrexpr.constr_expr -> ast_closure_term val interp_ast_closure_term : Geninterp.interp_sign -> Goal.goal Evd.sigma -> ast_closure_term -> Evd.evar_map * ast_closure_term val subst_ast_closure_term : Mod_subst.substitution -> ast_closure_term -> ast_closure_term val glob_ast_closure_term : Genintern.glob_sign -> ast_closure_term -> ast_closure_term val ssrterm_of_ast_closure_term : ast_closure_term -> ssrterm val ssrdgens_of_parsed_dgens : (ssrdocc * Ssrmatching.cpattern) list list * ssrclear -> ssrdgens val is_internal_name : string -> bool val add_internal_name : (string -> bool) -> unit val mk_internal_id : string -> Id.t val mk_tagged_id : string -> int -> Id.t val mk_evar_name : int -> Name.t val ssr_anon_hyp : string val pf_type_id : Goal.goal Evd.sigma -> EConstr.types -> Id.t val pf_abs_evars : Goal.goal Evd.sigma -> evar_map * EConstr.t -> int * EConstr.t * Evar.t list * UState.t val pf_abs_evars2 : (* ssr2 *) Goal.goal Evd.sigma -> Evar.t list -> evar_map * EConstr.t -> int * EConstr.t * Evar.t list * UState.t val pf_abs_cterm : Goal.goal Evd.sigma -> int -> EConstr.t -> EConstr.t val pf_merge_uc : UState.t -> 'a Evd.sigma -> 'a Evd.sigma val pf_merge_uc_of : evar_map -> 'a Evd.sigma -> 'a Evd.sigma val constr_name : evar_map -> EConstr.t -> Name.t val pf_type_of : Goal.goal Evd.sigma -> Constr.constr -> Goal.goal Evd.sigma * Constr.types val pfe_type_of : Goal.goal Evd.sigma -> EConstr.t -> Goal.goal Evd.sigma * EConstr.types val pfe_new_type : Goal.goal Evd.sigma -> Goal.goal Evd.sigma * EConstr.types val pfe_type_relevance_of : Goal.goal Evd.sigma -> EConstr.t -> Goal.goal Evd.sigma * EConstr.types * Sorts.relevance val pf_abs_prod : Name.t -> Goal.goal Evd.sigma -> EConstr.t -> EConstr.t -> Goal.goal Evd.sigma * EConstr.types val mkSsrRRef : string -> Glob_term.glob_constr * 'a option val mkSsrConst : string -> env -> evar_map -> evar_map * EConstr.t val pf_mkSsrConst : string -> Goal.goal Evd.sigma -> EConstr.t * Goal.goal Evd.sigma val new_wild_id : tac_ctx -> Names.Id.t * tac_ctx val pf_fresh_global : GlobRef.t -> Goal.goal Evd.sigma -> Constr.constr * Goal.goal Evd.sigma val is_discharged_id : Id.t -> bool val mk_discharged_id : Id.t -> Id.t val is_tagged : string -> string -> bool val has_discharged_tag : string -> bool val ssrqid : string -> Libnames.qualid val new_tmp_id : tac_ctx -> (Names.Id.t * Name.t ref) * tac_ctx val mk_anon_id : string -> Id.t list -> Id.t val pf_abs_evars_pirrel : Goal.goal Evd.sigma -> evar_map * Constr.constr -> int * Constr.constr val nbargs_open_constr : Goal.goal Evd.sigma -> Evd.evar_map * EConstr.t -> int val pf_nbargs : Goal.goal Evd.sigma -> EConstr.t -> int val gen_tmp_ids : ?ist:Geninterp.interp_sign -> (Goal.goal * tac_ctx) Evd.sigma -> (Goal.goal * tac_ctx) list Evd.sigma val ssrevaltac : Tacinterp.interp_sign -> Tacinterp.Value.t -> unit Proofview.tactic val convert_concl_no_check : EConstr.t -> unit Proofview.tactic val convert_concl : check:bool -> EConstr.t -> unit Proofview.tactic val red_safe : Reductionops.reduction_function -> env -> evar_map -> EConstr.t -> EConstr.t val red_product_skip_id : env -> evar_map -> EConstr.t -> EConstr.t val ssrautoprop_tac : (Evar.t Evd.sigma -> Evar.t list Evd.sigma) ref val mkProt : EConstr.t -> EConstr.t -> Goal.goal Evd.sigma -> EConstr.t * Goal.goal Evd.sigma val mkEtaApp : EConstr.t -> int -> int -> EConstr.t val mkRefl : EConstr.t -> EConstr.t -> Goal.goal Evd.sigma -> EConstr.t * Goal.goal Evd.sigma val discharge_hyp : Id.t * (Id.t * string) -> Goal.goal Evd.sigma -> Evar.t list Evd.sigma val clear_wilds_and_tmp_and_delayed_ids : (Goal.goal * tac_ctx) Evd.sigma -> (Goal.goal * tac_ctx) list Evd.sigma val view_error : string -> ssrterm -> 'a val top_id : Id.t val pf_abs_ssrterm : ?resolve_typeclasses:bool -> ist -> Goal.goal Evd.sigma -> ssrterm -> evar_map * EConstr.t * UState.t * int val pf_interp_ty : ?resolve_typeclasses:bool -> Tacinterp.interp_sign -> Goal.goal Evd.sigma -> Ssrast.ssrtermkind * (Glob_term.glob_constr * Constrexpr.constr_expr option) -> int * EConstr.t * EConstr.t * UState.t val ssr_n_tac : string -> int -> v82tac val donetac : int -> v82tac val applyn : with_evars:bool -> ?beta:bool -> ?with_shelve:bool -> ?first_goes_last:bool -> int -> EConstr.t -> v82tac exception NotEnoughProducts val pf_saturate : ?beta:bool -> ?bi_types:bool -> Goal.goal Evd.sigma -> EConstr.constr -> ?ty:EConstr.types -> int -> EConstr.constr * EConstr.types * (int * EConstr.constr) list * Goal.goal Evd.sigma val saturate : ?beta:bool -> ?bi_types:bool -> env -> evar_map -> EConstr.constr -> ?ty:EConstr.types -> int -> EConstr.constr * EConstr.types * (int * EConstr.constr) list * evar_map val refine_with : ?first_goes_last:bool -> ?beta:bool -> ?with_evars:bool -> evar_map * EConstr.t -> v82tac val pf_resolve_typeclasses : where:EConstr.t -> fail:bool -> Goal.goal Evd.sigma -> Goal.goal Evd.sigma val resolve_typeclasses : where:EConstr.t -> fail:bool -> Environ.env -> Evd.evar_map -> Evd.evar_map (*********************** Wrapped Coq tactics *****************************) val rewritetac : ?under:bool -> ssrdir -> EConstr.t -> tactic type name_hint = (int * EConstr.types array) option ref val gentac : Ssrast.ssrdocc * Ssrmatching.cpattern -> v82tac val genstac : ((Ssrast.ssrhyp list option * Ssrmatching.occ) * Ssrmatching.cpattern) list * Ssrast.ssrhyp list -> Tacmach.tactic val pf_interp_gen : bool -> (Ssrast.ssrhyp list option * Ssrmatching.occ) * Ssrmatching.cpattern -> Goal.goal Evd.sigma -> (EConstr.t * EConstr.t * Ssrast.ssrhyp list) * Goal.goal Evd.sigma (* HACK: use to put old pf_code in the tactic monad *) val pfLIFT : (Goal.goal Evd.sigma -> 'a * Goal.goal Evd.sigma) -> 'a Proofview.tactic (** Basic tactics *) val introid : ?orig:Name.t ref -> Id.t -> v82tac val intro_anon : v82tac val interp_clr : evar_map -> ssrhyps option * (ssrtermkind * EConstr.t) -> ssrhyps val genclrtac : EConstr.constr -> EConstr.constr list -> Ssrast.ssrhyp list -> Tacmach.tactic val old_cleartac : ssrhyps -> v82tac val cleartac : ssrhyps -> unit Proofview.tactic val tclMULT : int * ssrmmod -> Tacmach.tactic -> Tacmach.tactic val unprotecttac : Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val is_protect : EConstr.t -> Environ.env -> Evd.evar_map -> bool val abs_wgen : bool -> (Id.t -> Id.t) -> 'a * ((Ssrast.ssrhyp_or_id * string) * Ssrmatching.cpattern option) option -> Goal.goal Evd.sigma * EConstr.t list * EConstr.t -> Goal.goal Evd.sigma * EConstr.t list * EConstr.t val clr_of_wgen : ssrhyps * ((ssrhyp_or_id * 'a) * 'b option) option -> Proofview.V82.tac list -> Proofview.V82.tac list val unfold : EConstr.t list -> unit Proofview.tactic val apply_type : EConstr.types -> EConstr.t list -> Proofview.V82.tac (* New code ****************************************************************) (* To call old code *) val tacSIGMA : Goal.goal Evd.sigma Proofview.tactic val tclINTERP_AST_CLOSURE_TERM_AS_CONSTR : ast_closure_term -> EConstr.t list Proofview.tactic val tacREDUCE_TO_QUANTIFIED_IND : EConstr.types -> ((Names.inductive * EConstr.EInstance.t) * EConstr.types) Proofview.tactic val tacTYPEOF : EConstr.t -> EConstr.types Proofview.tactic val tclINTRO_ID : Id.t -> unit Proofview.tactic val tclINTRO_ANON : ?seed:string -> unit -> unit Proofview.tactic (* Lower level API, calls conclusion with the name taken from the prod *) type intro_id = | Anon | Id of Id.t | Seed of string val tclINTRO : id:intro_id -> conclusion:(orig_name:Name.t -> new_name:Id.t -> unit Proofview.tactic) -> unit Proofview.tactic val tclRENAME_HD_PROD : Name.t -> unit Proofview.tactic (* calls the tactic only if there are more than 0 goals *) val tcl0G : default:'a -> 'a Proofview.tactic -> 'a Proofview.tactic (* like tclFIRST but with 'a tactic *) val tclFIRSTa : 'a Proofview.tactic list -> 'a Proofview.tactic val tclFIRSTi : (int -> 'a Proofview.tactic) -> int -> 'a Proofview.tactic val tacCONSTR_NAME : ?name:Name.t -> EConstr.t -> Name.t Proofview.tactic (* [tacMKPROD t name ctx] (where ctx is a term possibly containing an unbound * Rel 1) builds [forall name : ty_t, ctx] *) val tacMKPROD : EConstr.t -> ?name:Name.t -> EConstr.types -> EConstr.types Proofview.tactic val tacINTERP_CPATTERN : Ssrmatching.cpattern -> Ssrmatching.pattern Proofview.tactic val tacUNIFY : EConstr.t -> EConstr.t -> unit Proofview.tactic (* if [(t : eq _ _ _)] then we can inject it *) val tacIS_INJECTION_CASE : ?ty:EConstr.types -> EConstr.t -> bool Proofview.tactic (** 1 shot, hands-on the top of the stack, eg for [=> ->] *) val tclWITHTOP : (EConstr.t -> unit Proofview.tactic) -> unit Proofview.tactic val tacMK_SSR_CONST : string -> EConstr.t Proofview.tactic module type StateType = sig type state val init : state end module MakeState(S : StateType) : sig val tclGET : (S.state -> unit Proofview.tactic) -> unit Proofview.tactic val tclGET1 : (S.state -> 'a Proofview.tactic) -> 'a Proofview.tactic val tclSET : S.state -> unit Proofview.tactic val tacUPDATE : (S.state -> S.state Proofview.tactic) -> unit Proofview.tactic val get : Proofview.Goal.t -> S.state end val is_ind_ref : Evd.evar_map -> EConstr.t -> Names.GlobRef.t -> bool val is_construct_ref : Evd.evar_map -> EConstr.t -> Names.GlobRef.t -> bool val is_const_ref : Evd.evar_map -> EConstr.t -> Names.GlobRef.t -> bool coq-8.11.0/plugins/ssr/ssrbwd.ml0000644000175000017500000001313513612664533016417 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* xInParens | '@' -> xWithAt | ' ' -> xNoFlag | 'x' -> xCpattern | _ -> assert false (** Backward chaining tactics: apply, exact, congr. *) (** The "apply" tactic *) let interp_agen ist gl ((goclr, _), (k, gc as c)) (clr, rcs) = (* ppdebug(lazy(str"sigma@interp_agen=" ++ pr_evar_map None (project gl))); *) let k = char_to_kind k in let rc = pf_intern_term ist gl c in let rcs' = rc :: rcs in match goclr with | None -> clr, rcs' | Some ghyps -> let clr' = snd (interp_hyps ist gl ghyps) @ clr in if k <> xNoFlag then clr', rcs' else let loc = rc.CAst.loc in match DAst.get rc with | GVar id when not_section_id id -> SsrHyp (Loc.tag ?loc id) :: clr', rcs' | GRef (Names.GlobRef.VarRef id, _) when not_section_id id -> SsrHyp (Loc.tag ?loc id) :: clr', rcs' | _ -> clr', rcs' let pf_pr_glob_constr gl = pr_glob_constr_env (pf_env gl) let interp_nbargs ist gl rc = try let rc6 = mkRApp rc (mkRHoles 6) in let sigma, t = interp_open_constr ist gl (rc6, None) in let si = sig_it gl in let gl = re_sig si sigma in 6 + Ssrcommon.nbargs_open_constr gl t with _ -> 5 let interp_view_nbimps ist gl rc = try let sigma, t = interp_open_constr ist gl (rc, None) in let si = sig_it gl in let gl = re_sig si sigma in let pl, c = splay_open_constr gl t in if Ssrcommon.isAppInd (pf_env gl) (project gl) c then List.length pl else (-(List.length pl)) with _ -> 0 let interp_agens ist gl gagens = match List.fold_right (interp_agen ist gl) gagens ([], []) with | clr, rlemma :: args -> let n = interp_nbargs ist gl rlemma - List.length args in let rec loop i = if i > n then errorstrm Pp.(str "Cannot apply lemma " ++ pf_pr_glob_constr gl rlemma) else try interp_refine ist gl (mkRApp rlemma (mkRHoles i @ args)) with _ -> loop (i + 1) in clr, loop 0 | _ -> assert false let pf_match = pf_apply (fun e s c t -> understand_tcc e s ~expected_type:t c) let apply_rconstr ?ist t gl = (* ppdebug(lazy(str"sigma@apply_rconstr=" ++ pr_evar_map None (project gl))); *) let n = match ist, DAst.get t with | None, (GVar id | GRef (Names.GlobRef.VarRef id,_)) -> pf_nbargs gl (EConstr.mkVar id) | Some ist, _ -> interp_nbargs ist gl t | _ -> anomaly "apply_rconstr without ist and not RVar" in let mkRlemma i = mkRApp t (mkRHoles i) in let cl = pf_concl gl in let rec loop i = if i > n then errorstrm Pp.(str"Cannot apply lemma "++pf_pr_glob_constr gl t) else try pf_match gl (mkRlemma i) (OfType cl) with _ -> loop (i + 1) in refine_with (loop 0) gl let mkRAppView ist gl rv gv = let nb_view_imps = interp_view_nbimps ist gl rv in mkRApp rv (mkRHoles (abs nb_view_imps)) let refine_interp_apply_view dbl ist gl gv = let pair i = List.map (fun x -> i, x) in let rv = pf_intern_term ist gl gv in let v = mkRAppView ist gl rv gv in let interp_with (dbl, hint) = let i = if dbl = Ssrview.AdaptorDb.Equivalence then 2 else 1 in interp_refine ist gl (mkRApp hint (v :: mkRHoles i)) in let rec loop = function | [] -> (try apply_rconstr ~ist rv gl with _ -> view_error "apply" gv) | h :: hs -> (try refine_with (snd (interp_with h)) gl with _ -> loop hs) in loop (pair dbl (Ssrview.AdaptorDb.get dbl) @ if dbl = Ssrview.AdaptorDb.Equivalence then pair Ssrview.AdaptorDb.Backward (Ssrview.AdaptorDb.(get Backward)) else []) let apply_top_tac = Tacticals.tclTHENLIST [ introid top_id; apply_rconstr (mkRVar top_id); old_cleartac [SsrHyp(None,top_id)] ] let inner_ssrapplytac gviews (ggenl, gclr) ist = Proofview.V82.tactic ~nf_evars:false (fun gl -> let _, clr = interp_hyps ist gl gclr in let vtac gv i gl' = refine_interp_apply_view i ist gl' gv in let ggenl, tclGENTAC = if gviews <> [] && ggenl <> [] then let ggenl= List.map (fun (x,g) -> x, cpattern_of_term g ist) (List.hd ggenl) in [], Tacticals.tclTHEN (genstac (ggenl,[])) else ggenl, Tacticals.tclTHEN Tacticals.tclIDTAC in tclGENTAC (fun gl -> match gviews, ggenl with | v :: tl, [] -> let dbl = if List.length tl = 1 then Ssrview.AdaptorDb.Equivalence else Ssrview.AdaptorDb.Backward in Tacticals.tclTHEN (List.fold_left (fun acc v -> Tacticals.tclTHENLAST acc (vtac v dbl)) (vtac v Ssrview.AdaptorDb.Backward) tl) (old_cleartac clr) gl | [], [agens] -> let clr', (sigma, lemma) = interp_agens ist gl agens in let gl = pf_merge_uc_of sigma gl in Tacticals.tclTHENLIST [old_cleartac clr; refine_with ~beta:true lemma; old_cleartac clr'] gl | _, _ -> Tacticals.tclTHENLIST [apply_top_tac; old_cleartac clr] gl) gl ) let apply_top_tac = Proofview.V82.tactic ~nf_evars:false apply_top_tac coq-8.11.0/plugins/ssr/ssrtacticals.ml0000644000175000017500000001555713612664533017624 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* , first, and last). *) let get_index = function Locus.ArgArg i -> i | _ -> anomaly "Uninterpreted index" (* Toplevel constr must be globalized twice ! *) (** The "first" and "last" tacticals. *) let tclPERM perm tac gls = let subgls = tac gls in let subgll' = perm subgls.Evd.it in re_sig subgll' subgls.Evd.sigma let rot_hyps dir i hyps = let n = List.length hyps in if i = 0 then List.rev hyps else if i > n then CErrors.user_err (Pp.str "Not enough subgoals") else let rec rot i l_hyps = function | hyp :: hyps' when i > 0 -> rot (i - 1) (hyp :: l_hyps) hyps' | hyps' -> hyps' @ (List.rev l_hyps) in rot (match dir with L2R -> i | R2L -> n - i) [] hyps let tclSEQAT ist atac1 dir (ivar, ((_, atacs2), atac3)) = let i = get_index ivar in let evtac t = Proofview.V82.of_tactic (ssrevaltac ist t) in let tac1 = evtac atac1 in if atacs2 = [] && atac3 <> None then tclPERM (rot_hyps dir i) tac1 else let evotac = function Some atac -> evtac atac | _ -> Tacticals.tclIDTAC in let tac3 = evotac atac3 in let rec mk_pad n = if n > 0 then tac3 :: mk_pad (n - 1) else [] in match dir, mk_pad (i - 1), List.map evotac atacs2 with | L2R, [], [tac2] when atac3 = None -> Tacticals.tclTHENFIRST tac1 tac2 | L2R, [], [tac2] when atac3 = None -> Tacticals.tclTHENLAST tac1 tac2 | L2R, pad, tacs2 -> Tacticals.tclTHENSFIRSTn tac1 (Array.of_list (pad @ tacs2)) tac3 | R2L, pad, tacs2 -> Tacticals.tclTHENSLASTn tac1 tac3 (Array.of_list (tacs2 @ pad)) (** The "in" pseudo-tactical *)(* {{{ **********************************************) let hidden_goal_tag = "the_hidden_goal" let check_wgen_uniq gens = let clears = List.flatten (List.map fst gens) in check_hyps_uniq [] clears; let ids = CList.map_filter (function (_,Some ((id,_),_)) -> Some (hoi_id id) | _ -> None) gens in let rec check ids = function | id :: _ when List.mem id ids -> errorstrm Pp.(str"Duplicate generalization " ++ Id.print id) | id :: hyps -> check (id :: ids) hyps | [] -> () in check [] ids let pf_clauseids gl gens clseq = let keep_clears = List.map (fun (x, _) -> x, None) in if gens <> [] then (check_wgen_uniq gens; gens) else if clseq <> InAll && clseq <> InAllHyps then keep_clears gens else CErrors.user_err (Pp.str "assumptions should be named explicitly") let hidden_clseq = function InHyps | InHypsSeq | InAllHyps -> true | _ -> false let posetac id cl = Proofview.V82.of_tactic (Tactics.pose_tac (Name id) cl) let hidetacs clseq idhide cl0 = if not (hidden_clseq clseq) then [] else [posetac idhide cl0; Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkVar idhide))] let endclausestac id_map clseq gl_id cl0 gl = let not_hyp' id = not (List.mem_assoc id id_map) in let orig_id id = try List.assoc id id_map with Not_found -> id in let dc, c = EConstr.decompose_prod_assum (project gl) (pf_concl gl) in let hide_goal = hidden_clseq clseq in let c_hidden = hide_goal && EConstr.eq_constr (project gl) c (EConstr.mkVar gl_id) in let rec fits forced = function | (id, _) :: ids, decl :: dc' when RelDecl.get_name decl = Name id -> fits true (ids, dc') | ids, dc' -> forced && ids = [] && (not hide_goal || dc' = [] && c_hidden) in let rec unmark c = match EConstr.kind (project gl) c with | Var id when hidden_clseq clseq && id = gl_id -> cl0 | Prod ({binder_name=Name id} as na, t, c') when List.mem_assoc id id_map -> EConstr.mkProd ({na with binder_name=Name (orig_id id)}, unmark t, unmark c') | LetIn ({binder_name=Name id} as na, v, t, c') when List.mem_assoc id id_map -> EConstr.mkLetIn ({na with binder_name=Name (orig_id id)}, unmark v, unmark t, unmark c') | _ -> EConstr.map (project gl) unmark c in let utac hyp = Proofview.V82.of_tactic (Tactics.convert_hyp ~check:false ~reorder:false (NamedDecl.map_constr unmark hyp)) in let utacs = List.map utac (pf_hyps gl) in let ugtac gl' = Proofview.V82.of_tactic (convert_concl_no_check (unmark (pf_concl gl'))) gl' in let ctacs = if hide_goal then [Proofview.V82.of_tactic (Tactics.clear [gl_id])] else [] in let mktac itacs = Tacticals.tclTHENLIST (itacs @ utacs @ ugtac :: ctacs) in let itac (_, id) = Proofview.V82.of_tactic (Tactics.introduction id) in if fits false (id_map, List.rev dc) then mktac (List.map itac id_map) gl else let all_ids = ids_of_rel_context dc @ pf_ids_of_hyps gl in if List.for_all not_hyp' all_ids && not c_hidden then mktac [] gl else errorstrm Pp.(str "tampering with discharged assumptions of \"in\" tactical") let tclCLAUSES tac (gens, clseq) gl = if clseq = InGoal || clseq = InSeqGoal then tac gl else let clr_gens = pf_clauseids gl gens clseq in let clear = Tacticals.tclTHENLIST (List.rev(List.fold_right clr_of_wgen clr_gens [])) in let gl_id = mk_anon_id hidden_goal_tag (Tacmach.pf_ids_of_hyps gl) in let cl0 = pf_concl gl in let dtac gl = let c = pf_concl gl in let gl, args, c = List.fold_right (abs_wgen true mk_discharged_id) gens (gl,[], c) in apply_type c args gl in let endtac = let id_map = CList.map_filter (function | _, Some ((x,_),_) -> let id = hoi_id x in Some (mk_discharged_id id, id) | _, None -> None) gens in endclausestac id_map clseq gl_id cl0 in Tacticals.tclTHENLIST (hidetacs clseq gl_id cl0 @ [dtac; clear; tac; endtac]) gl (** The "do" tactical. ********************************************************) let hinttac ist is_by (is_or, atacs) = let dtac = if is_by then donetac ~-1 else Tacticals.tclIDTAC in let mktac = function | Some atac -> Tacticals.tclTHEN (Proofview.V82.of_tactic (ssrevaltac ist atac)) dtac | _ -> dtac in match List.map mktac atacs with | [] -> if is_or then dtac else Tacticals.tclIDTAC | [tac] -> tac | tacs -> Tacticals.tclFIRST tacs let ssrdotac ist (((n, m), tac), clauses) = let mul = get_index n, m in tclCLAUSES (tclMULT mul (hinttac ist false tac)) clauses let tclCLAUSES tac g_c = Proofview.V82.(tactic (tclCLAUSES (of_tactic tac) g_c)) coq-8.11.0/plugins/ssr/ssrparser.mlg0000644000175000017500000024642413612664533017317 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { tac } ]]; END (* Copy of ssrtacarg with LEVEL "3", useful for: "under ... do ..." *) ARGUMENT EXTEND ssrtac3arg TYPED AS tactic PRINTED BY { pr_ssrtacarg env sigma } END GRAMMAR EXTEND Gram GLOBAL: ssrtac3arg; ssrtac3arg: [[ tac = tactic_expr LEVEL "3" -> { tac } ]]; END { (* Lexically closed tactic for tacticals. *) let pr_ssrtclarg env sigma _ _ prt tac = prt env sigma tacltop tac } ARGUMENT EXTEND ssrtclarg TYPED AS ssrtacarg PRINTED BY { pr_ssrtclarg env sigma } | [ ssrtacarg(tac) ] -> { tac } END { open Genarg (** Adding a new uninterpreted generic argument type *) let add_genarg tag pr = let wit = Genarg.make0 tag in let tag = Geninterp.Val.create tag in let glob ist x = (ist, x) in let subst _ x = x in let interp ist x = Ftactic.return (Geninterp.Val.Dyn (tag, x)) in let gen_pr env sigma _ _ _ = pr env sigma in let () = Genintern.register_intern0 wit glob in let () = Genintern.register_subst0 wit subst in let () = Geninterp.register_interp0 wit interp in let () = Geninterp.register_val0 wit (Some (Geninterp.Val.Base tag)) in Pptactic.declare_extra_genarg_pprule wit gen_pr gen_pr gen_pr; wit (** Primitive parsing to avoid syntax conflicts with basic tactics. *) let accept_before_syms syms strm = match Util.stream_nth 1 strm with | Tok.KEYWORD sym when List.mem sym syms -> () | _ -> raise Stream.Failure let accept_before_syms_or_any_id syms strm = match Util.stream_nth 1 strm with | Tok.KEYWORD sym when List.mem sym syms -> () | Tok.IDENT _ -> () | _ -> raise Stream.Failure let accept_before_syms_or_ids syms ids strm = match Util.stream_nth 1 strm with | Tok.KEYWORD sym when List.mem sym syms -> () | Tok.IDENT id when List.mem id ids -> () | _ -> raise Stream.Failure open Ssrast let pr_id = Ppconstr.pr_id let pr_name = function Name id -> pr_id id | Anonymous -> str "_" let pr_spc () = str " " let pr_list = prlist_with_sep (**************************** ssrhyp **************************************) let pr_ssrhyp _ _ _ = pr_hyp let wit_ssrhyprep = add_genarg "ssrhyprep" (fun env sigma -> pr_hyp) let intern_hyp ist (SsrHyp (loc, id) as hyp) = let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_var) CAst.(make ?loc id)) in if not_section_id id then hyp else hyp_err ?loc "Can't clear section hypothesis " id open Pcoq.Prim } ARGUMENT EXTEND ssrhyp TYPED AS ssrhyprep PRINTED BY { pr_ssrhyp } INTERPRETED BY { interp_hyp } GLOBALIZED BY { intern_hyp } | [ ident(id) ] -> { SsrHyp (Loc.tag ~loc id) } END { let pr_hoi = hoik pr_hyp let pr_ssrhoi _ _ _ = pr_hoi let wit_ssrhoirep = add_genarg "ssrhoirep" (fun env sigma -> pr_hoi) let intern_ssrhoi ist = function | Hyp h -> Hyp (intern_hyp ist h) | Id (SsrHyp (_, id)) as hyp -> let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_ident) id) in hyp let interp_ssrhoi ist gl = function | Hyp h -> let s, h' = interp_hyp ist gl h in s, Hyp h' | Id (SsrHyp (loc, id)) -> let s, id' = interp_wit wit_ident ist gl id in s, Id (SsrHyp (loc, id')) } ARGUMENT EXTEND ssrhoi_hyp TYPED AS ssrhoirep PRINTED BY { pr_ssrhoi } INTERPRETED BY { interp_ssrhoi } GLOBALIZED BY { intern_ssrhoi } | [ ident(id) ] -> { Hyp (SsrHyp(Loc.tag ~loc id)) } END ARGUMENT EXTEND ssrhoi_id TYPED AS ssrhoirep PRINTED BY { pr_ssrhoi } INTERPRETED BY { interp_ssrhoi } GLOBALIZED BY { intern_ssrhoi } | [ ident(id) ] -> { Id (SsrHyp(Loc.tag ~loc id)) } END (** Rewriting direction *) { let pr_rwdir = function L2R -> mt() | R2L -> str "-" let wit_ssrdir = add_genarg "ssrdir" (fun env sigma -> pr_dir) (** Simpl switch *) let pr_ssrsimpl _ _ _ = pr_simpl let wit_ssrsimplrep = add_genarg "ssrsimplrep" (fun env sigma -> pr_simpl) let test_ssrslashnum b1 b2 _ strm = match Util.stream_nth 0 strm with | Tok.KEYWORD "/" -> (match Util.stream_nth 1 strm with | Tok.NUMERAL _ when b1 -> (match Util.stream_nth 2 strm with | Tok.KEYWORD "=" | Tok.KEYWORD "/=" when not b2 -> () | Tok.KEYWORD "/" -> if not b2 then () else begin match Util.stream_nth 3 strm with | Tok.NUMERAL _ -> () | _ -> raise Stream.Failure end | _ -> raise Stream.Failure) | Tok.KEYWORD "/" when not b1 -> (match Util.stream_nth 2 strm with | Tok.KEYWORD "=" when not b2 -> () | Tok.NUMERAL _ when b2 -> (match Util.stream_nth 3 strm with | Tok.KEYWORD "=" -> () | _ -> raise Stream.Failure) | _ when not b2 -> () | _ -> raise Stream.Failure) | Tok.KEYWORD "=" when not b1 && not b2 -> () | _ -> raise Stream.Failure) | Tok.KEYWORD "//" when not b1 -> (match Util.stream_nth 1 strm with | Tok.KEYWORD "=" when not b2 -> () | Tok.NUMERAL _ when b2 -> (match Util.stream_nth 2 strm with | Tok.KEYWORD "=" -> () | _ -> raise Stream.Failure) | _ when not b2 -> () | _ -> raise Stream.Failure) | _ -> raise Stream.Failure let test_ssrslashnum10 = test_ssrslashnum true false let test_ssrslashnum11 = test_ssrslashnum true true let test_ssrslashnum01 = test_ssrslashnum false true let test_ssrslashnum00 = test_ssrslashnum false false let negate_parser f tok x = let rc = try Some (f tok x) with Stream.Failure -> None in match rc with | None -> () | Some _ -> raise Stream.Failure let test_not_ssrslashnum = Pcoq.Entry.of_parser "test_not_ssrslashnum" (negate_parser test_ssrslashnum10) let test_ssrslashnum00 = Pcoq.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum00 let test_ssrslashnum10 = Pcoq.Entry.of_parser "test_ssrslashnum10" test_ssrslashnum10 let test_ssrslashnum11 = Pcoq.Entry.of_parser "test_ssrslashnum11" test_ssrslashnum11 let test_ssrslashnum01 = Pcoq.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum01 } ARGUMENT EXTEND ssrsimpl_ne TYPED AS ssrsimplrep PRINTED BY { pr_ssrsimpl } | [ "//=" ] -> { SimplCut (~-1,~-1) } | [ "/=" ] -> { Simpl ~-1 } END (* Pcoq.Prim. *) GRAMMAR EXTEND Gram GLOBAL: ssrsimpl_ne; ssrsimpl_ne: [ [ test_ssrslashnum11; "/"; n = natural; "/"; m = natural; "=" -> { SimplCut(n,m) } | test_ssrslashnum10; "/"; n = natural; "/" -> { Cut n } | test_ssrslashnum10; "/"; n = natural; "=" -> { Simpl n } | test_ssrslashnum10; "/"; n = natural; "/=" -> { SimplCut (n,~-1) } | test_ssrslashnum10; "/"; n = natural; "/"; "=" -> { SimplCut (n,~-1) } | test_ssrslashnum01; "//"; m = natural; "=" -> { SimplCut (~-1,m) } | test_ssrslashnum00; "//" -> { Cut ~-1 } ]]; END { let pr_ssrclear _ _ _ = pr_clear mt } ARGUMENT EXTEND ssrclear_ne TYPED AS ssrhyp list PRINTED BY { pr_ssrclear } | [ "{" ne_ssrhyp_list(clr) "}" ] -> { check_hyps_uniq [] clr; clr } END ARGUMENT EXTEND ssrclear TYPED AS ssrclear_ne PRINTED BY { pr_ssrclear } | [ ssrclear_ne(clr) ] -> { clr } | [ ] -> { [] } END (** Indexes *) (* Since SSR indexes are always positive numbers, we use the 0 value *) (* to encode an omitted index. We reuse the in or_var type, but we *) (* supply our own interpretation function, which checks for non *) (* positive values, and allows the use of constr numerals, so that *) (* e.g., "let n := eval compute in (1 + 3) in (do n!clear)" works. *) { let pr_index = function | ArgVar {CAst.v=id} -> pr_id id | ArgArg n when n > 0 -> int n | _ -> mt () let pr_ssrindex _ _ _ = pr_index let noindex = ArgArg 0 let check_index ?loc i = if i > 0 then i else CErrors.user_err ?loc (str"Index not positive") let mk_index ?loc = function | ArgArg i -> ArgArg (check_index ?loc i) | iv -> iv let interp_index ist gl idx = Tacmach.project gl, match idx with | ArgArg _ -> idx | ArgVar id -> let i = try let v = Id.Map.find id.CAst.v ist.Tacinterp.lfun in begin match Tacinterp.Value.to_int v with | Some i -> i | None -> begin match Tacinterp.Value.to_constr v with | Some c -> let rc = Detyping.detype Detyping.Now false Id.Set.empty (pf_env gl) (project gl) c in begin match Notation.uninterp_prim_token rc with | _, Constrexpr.Numeral (b,{NumTok.int = s; frac = ""; exp = ""}) -> let n = int_of_string s in (match b with SPlus -> n | SMinus -> -n) | _ -> raise Not_found end | None -> raise Not_found end end with _ -> CErrors.user_err ?loc:id.CAst.loc (str"Index not a number") in ArgArg (check_index ?loc:id.CAst.loc i) open Pltac } ARGUMENT EXTEND ssrindex PRINTED BY { pr_ssrindex } INTERPRETED BY { interp_index } END (** Occurrence switch *) (* The standard syntax of complemented occurrence lists involves a single *) (* initial "-", e.g., {-1 3 5}. An initial *) (* "+" may be used to indicate positive occurrences (the default). The *) (* "+" is optional, except if the list of occurrences starts with a *) (* variable or is empty (to avoid confusion with a clear switch). The *) (* empty positive switch "{+}" selects no occurrences, while the empty *) (* negative switch "{-}" selects all occurrences explicitly; this is the *) (* default, but "{-}" prevents the implicit clear, and can be used to *) (* force dependent elimination -- see ndefectelimtac below. *) { let pr_ssrocc _ _ _ = pr_occ open Pcoq.Prim } ARGUMENT EXTEND ssrocc TYPED AS (bool * int list) option PRINTED BY { pr_ssrocc } | [ natural(n) natural_list(occ) ] -> { Some (false, List.map (check_index ~loc) (n::occ)) } | [ "-" natural_list(occ) ] -> { Some (true, occ) } | [ "+" natural_list(occ) ] -> { Some (false, occ) } END (* modality *) { let pr_mmod = function May -> str "?" | Must -> str "!" | Once -> mt () let wit_ssrmmod = add_genarg "ssrmmod" (fun env sigma -> pr_mmod) let ssrmmod = Pcoq.create_generic_entry Pcoq.utactic "ssrmmod" (Genarg.rawwit wit_ssrmmod);; } GRAMMAR EXTEND Gram GLOBAL: ssrmmod; ssrmmod: [[ "!" -> { Must } | LEFTQMARK -> { May } | "?" -> { May } ]]; END (** Rewrite multiplier: !n ?n *) { let pr_mult (n, m) = if n > 0 && m <> Once then int n ++ pr_mmod m else pr_mmod m let pr_ssrmult _ _ _ = pr_mult } ARGUMENT EXTEND ssrmult_ne TYPED AS (int * ssrmmod) PRINTED BY { pr_ssrmult } | [ natural(n) ssrmmod(m) ] -> { check_index ~loc n, m } | [ ssrmmod(m) ] -> { notimes, m } END ARGUMENT EXTEND ssrmult TYPED AS ssrmult_ne PRINTED BY { pr_ssrmult } | [ ssrmult_ne(m) ] -> { m } | [ ] -> { nomult } END { (** Discharge occ switch (combined occurrence / clear switch *) let pr_docc = function | None, occ -> pr_occ occ | Some clr, _ -> pr_clear mt clr let pr_ssrdocc _ _ _ = pr_docc } ARGUMENT EXTEND ssrdocc TYPED AS (ssrclear option * ssrocc) PRINTED BY { pr_ssrdocc } | [ "{" ssrocc(occ) "}" ] -> { mkocc occ } | [ "{" ssrhyp_list(clr) "}" ] -> { mkclr clr } END { (* Old kinds of terms *) let input_ssrtermkind _ strm = match Util.stream_nth 0 strm with | Tok.KEYWORD "(" -> xInParens | Tok.KEYWORD "@" -> xWithAt | _ -> xNoFlag let ssrtermkind = Pcoq.Entry.of_parser "ssrtermkind" input_ssrtermkind (* New kinds of terms *) let input_term_annotation _ strm = match Stream.npeek 2 strm with | Tok.KEYWORD "(" :: Tok.KEYWORD "(" :: _ -> `DoubleParens | Tok.KEYWORD "(" :: _ -> `Parens | Tok.KEYWORD "@" :: _ -> `At | _ -> `None let term_annotation = Pcoq.Entry.of_parser "term_annotation" input_term_annotation (* terms *) (** Terms parsing. ********************************************************) (* Because we allow wildcards in term references, we need to stage the *) (* interpretation of terms so that it occurs at the right time during *) (* the execution of the tactic (e.g., so that we don't report an error *) (* for a term that isn't actually used in the execution). *) (* The term representation tracks whether the concrete initial term *) (* started with an opening paren, which might avoid a conflict between *) (* the ssrreflect term syntax and Gallina notation. *) (* Old terms *) let pr_ssrterm _ _ _ = pr_term let glob_ssrterm gs = function | k, (_, Some c) -> k, Tacintern.intern_constr gs c | ct -> ct let subst_ssrterm s (k, c) = k, Tacsubst.subst_glob_constr_and_expr s c let interp_ssrterm _ gl t = Tacmach.project gl, t open Pcoq.Constr } ARGUMENT EXTEND ssrterm PRINTED BY { pr_ssrterm } INTERPRETED BY { interp_ssrterm } GLOBALIZED BY { glob_ssrterm } SUBSTITUTED BY { subst_ssrterm } RAW_PRINTED BY { pr_ssrterm } GLOB_PRINTED BY { pr_ssrterm } END GRAMMAR EXTEND Gram GLOBAL: ssrterm; ssrterm: [[ k = ssrtermkind; c = Pcoq.Constr.constr -> { mk_term k c } ]]; END (* New terms *) { let pp_ast_closure_term _ _ _ = pr_ast_closure_term } ARGUMENT EXTEND ast_closure_term PRINTED BY { pp_ast_closure_term } INTERPRETED BY { interp_ast_closure_term } GLOBALIZED BY { glob_ast_closure_term } SUBSTITUTED BY { subst_ast_closure_term } RAW_PRINTED BY { pp_ast_closure_term } GLOB_PRINTED BY { pp_ast_closure_term } | [ term_annotation(a) constr(c) ] -> { mk_ast_closure_term a c } END ARGUMENT EXTEND ast_closure_lterm PRINTED BY { pp_ast_closure_term } INTERPRETED BY { interp_ast_closure_term } GLOBALIZED BY { glob_ast_closure_term } SUBSTITUTED BY { subst_ast_closure_term } RAW_PRINTED BY { pp_ast_closure_term } GLOB_PRINTED BY { pp_ast_closure_term } | [ term_annotation(a) lconstr(c) ] -> { mk_ast_closure_term a c } END (* Old Views *) { let pr_view = pr_list mt (fun c -> str "/" ++ pr_term c) let pr_ssrbwdview _ _ _ = pr_view } ARGUMENT EXTEND ssrbwdview TYPED AS ssrterm list PRINTED BY { pr_ssrbwdview } END (* Pcoq *) GRAMMAR EXTEND Gram GLOBAL: ssrbwdview; ssrbwdview: [ [ test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr -> { [mk_term xNoFlag c] } | test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr; w = ssrbwdview -> { (mk_term xNoFlag c) :: w } ]]; END (* New Views *) { type ssrfwdview = ast_closure_term list let pr_ssrfwdview _ _ _ = pr_view2 } ARGUMENT EXTEND ssrfwdview TYPED AS ast_closure_term list PRINTED BY { pr_ssrfwdview } END (* Pcoq *) GRAMMAR EXTEND Gram GLOBAL: ssrfwdview; ssrfwdview: [ [ test_not_ssrslashnum; "/"; c = ast_closure_term -> { [c] } | test_not_ssrslashnum; "/"; c = ast_closure_term; w = ssrfwdview -> { c :: w } ]]; END (* ipats *) { let remove_loc x = x.CAst.v let ipat_of_intro_pattern p = Tactypes.( let rec ipat_of_intro_pattern = function | IntroNaming (IntroIdentifier id) -> IPatId id | IntroAction IntroWildcard -> IPatAnon Drop | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) -> IPatCase (Regular( List.map (List.map ipat_of_intro_pattern) (List.map (List.map remove_loc) iorpat))) | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) -> IPatCase (Regular [List.map ipat_of_intro_pattern (List.map remove_loc iandpat)]) | IntroNaming IntroAnonymous -> IPatAnon (One None) | IntroAction (IntroRewrite b) -> IPatRewrite (allocc, if b then L2R else R2L) | IntroNaming (IntroFresh id) -> IPatAnon (One None) | IntroAction (IntroApplyOn _) -> (* to do *) CErrors.user_err (Pp.str "TO DO") | IntroAction (IntroInjection ips) -> IPatInj [List.map ipat_of_intro_pattern (List.map remove_loc ips)] | IntroForthcoming _ -> (* Unable to determine which kind of ipat interp_introid could * return [HH] *) assert false in ipat_of_intro_pattern p ) let rec map_ipat map_id map_ssrhyp map_ast_closure_term = function | (IPatSimpl _ | IPatAnon _ | IPatRewrite _ | IPatNoop | IPatFastNondep) as x -> x | IPatId id -> IPatId (map_id id) | IPatAbstractVars l -> IPatAbstractVars (List.map map_id l) | IPatClear clr -> IPatClear (List.map map_ssrhyp clr) | IPatCase (Regular iorpat) -> IPatCase (Regular (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat)) | IPatCase (Block(hat)) -> IPatCase (Block(map_block map_id hat)) | IPatDispatch (Regular iorpat) -> IPatDispatch (Regular (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat)) | IPatDispatch (Block (hat)) -> IPatDispatch (Block(map_block map_id hat)) | IPatInj iorpat -> IPatInj (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat) | IPatView v -> IPatView (List.map map_ast_closure_term v) and map_block map_id = function | Prefix id -> Prefix (map_id id) | SuffixId id -> SuffixId (map_id id) | SuffixNum _ as x -> x type ssripatrep = ssripat let wit_ssripatrep = add_genarg "ssripatrep" (fun env sigma -> pr_ipat) let pr_ssripat _ _ _ = pr_ipat let pr_ssripats _ _ _ = pr_ipats let pr_ssriorpat _ _ _ = pr_iorpat let intern_ipat ist = map_ipat (fun id -> id) (intern_hyp ist) (glob_ast_closure_term ist) let intern_ipats ist = List.map (intern_ipat ist) let interp_intro_pattern = interp_wit wit_intro_pattern let interp_introid ist gl id = try IntroNaming (IntroIdentifier (hyp_id (snd (interp_hyp ist gl (SsrHyp (Loc.tag id)))))) with _ -> (snd (interp_intro_pattern ist gl (CAst.make @@ IntroNaming (IntroIdentifier id)))).CAst.v let get_intro_id = function | IntroNaming (IntroIdentifier id) -> id | _ -> assert false let rec add_intro_pattern_hyps ipat hyps = let {CAst.loc=loc;v=ipat} = ipat in match ipat with | IntroNaming (IntroIdentifier id) -> if not_section_id id then SsrHyp (loc, id) :: hyps else hyp_err ?loc "Can't delete section hypothesis " id | IntroAction IntroWildcard -> hyps | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) -> List.fold_right (List.fold_right add_intro_pattern_hyps) iorpat hyps | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) -> List.fold_right add_intro_pattern_hyps iandpat hyps | IntroNaming IntroAnonymous -> [] | IntroNaming (IntroFresh _) -> [] | IntroAction (IntroRewrite _) -> hyps | IntroAction (IntroInjection ips) -> List.fold_right add_intro_pattern_hyps ips hyps | IntroAction (IntroApplyOn (c,pat)) -> add_intro_pattern_hyps pat hyps | IntroForthcoming _ -> (* As in ipat_of_intro_pattern, was unable to determine which kind of ipat interp_introid could return [HH] *) assert false (* We interp the ipat using the standard ltac machinery for ids, since * we have no clue what a name could be bound to (maybe another ipat) *) let interp_ipat ist gl = let ltacvar id = Id.Map.mem id ist.Tacinterp.lfun in let interp_block = function | Prefix id when ltacvar id -> begin match interp_introid ist gl id with | IntroNaming (IntroIdentifier id) -> Prefix id | _ -> Ssrcommon.errorstrm Pp.(str"Variable " ++ Id.print id ++ str" in block intro pattern should be bound to an identifier.") end | SuffixId id when ltacvar id -> begin match interp_introid ist gl id with | IntroNaming (IntroIdentifier id) -> SuffixId id | _ -> Ssrcommon.errorstrm Pp.(str"Variable " ++ Id.print id ++ str" in block intro pattern should be bound to an identifier.") end | x -> x in let rec interp = function | IPatId id when ltacvar id -> ipat_of_intro_pattern (interp_introid ist gl id) | IPatId _ as x -> x | IPatClear clr -> let add_hyps (SsrHyp (loc, id) as hyp) hyps = if not (ltacvar id) then hyp :: hyps else add_intro_pattern_hyps CAst.(make ?loc (interp_introid ist gl id)) hyps in let clr' = List.fold_right add_hyps clr [] in check_hyps_uniq [] clr'; IPatClear clr' | IPatCase(Regular iorpat) -> IPatCase(Regular(List.map (List.map interp) iorpat)) | IPatCase(Block(hat)) -> IPatCase(Block(interp_block hat)) | IPatDispatch(Regular iorpat) -> IPatDispatch(Regular (List.map (List.map interp) iorpat)) | IPatDispatch(Block(hat)) -> IPatDispatch(Block(interp_block hat)) | IPatInj iorpat -> IPatInj (List.map (List.map interp) iorpat) | IPatAbstractVars l -> IPatAbstractVars (List.map get_intro_id (List.map (interp_introid ist gl) l)) | IPatView l -> IPatView (List.map (fun x -> snd(interp_ast_closure_term ist gl x)) l) | (IPatSimpl _ | IPatAnon _ | IPatRewrite _ | IPatNoop | IPatFastNondep) as x -> x in interp let interp_ipats ist gl l = project gl, List.map (interp_ipat ist gl) l let pushIPatRewrite = function | pats :: orpat -> (IPatRewrite (allocc, L2R) :: pats) :: orpat | [] -> [] let pushIPatNoop = function | pats :: orpat -> (IPatNoop :: pats) :: orpat | [] -> [] let test_ident_no_do _ strm = match Util.stream_nth 0 strm with | Tok.IDENT s when s <> "do" -> () | _ -> raise Stream.Failure let test_ident_no_do = Pcoq.Entry.of_parser "test_ident_no_do" test_ident_no_do } ARGUMENT EXTEND ident_no_do PRINTED BY { fun _ _ _ -> Names.Id.print } END GRAMMAR EXTEND Gram GLOBAL: ident_no_do; ident_no_do: [ [ test_ident_no_do; id = IDENT -> { Id.of_string id } ] ]; END ARGUMENT EXTEND ssripat TYPED AS ssripatrep list PRINTED BY { pr_ssripats } INTERPRETED BY { interp_ipats } GLOBALIZED BY { intern_ipats } | [ "_" ] -> { [IPatAnon Drop] } | [ "*" ] -> { [IPatAnon All] } | [ ">" ] -> { [IPatFastNondep] } | [ ident_no_do(id) ] -> { [IPatId id] } | [ "?" ] -> { [IPatAnon (One None)] } | [ "+" ] -> { [IPatAnon Temporary] } | [ "++" ] -> { [IPatAnon Temporary; IPatAnon Temporary] } | [ ssrsimpl_ne(sim) ] -> { [IPatSimpl sim] } | [ ssrdocc(occ) "->" ] -> { match occ with | Some [], _ -> CErrors.user_err ~loc (str"occ_switch expected") | None, occ -> [IPatRewrite (occ, L2R)] | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, L2R)] } | [ ssrdocc(occ) "<-" ] -> { match occ with | Some [], _ -> CErrors.user_err ~loc (str"occ_switch expected") | None, occ -> [IPatRewrite (occ, R2L)] | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, R2L)] } | [ ssrdocc(occ) ] -> { match occ with | Some cl, _ -> check_hyps_uniq [] cl; [IPatClear cl] | _ -> CErrors.user_err ~loc (str"Only identifiers are allowed here") } | [ "->" ] -> { [IPatRewrite (allocc, L2R)] } | [ "<-" ] -> { [IPatRewrite (allocc, R2L)] } | [ "-" ] -> { [IPatNoop] } | [ "-/" "=" ] -> { [IPatNoop;IPatSimpl(Simpl ~-1)] } | [ "-/=" ] -> { [IPatNoop;IPatSimpl(Simpl ~-1)] } | [ "-/" "/" ] -> { [IPatNoop;IPatSimpl(Cut ~-1)] } | [ "-//" ] -> { [IPatNoop;IPatSimpl(Cut ~-1)] } | [ "-/" integer(n) "/" ] -> { [IPatNoop;IPatSimpl(Cut n)] } | [ "-/" "/=" ] -> { [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] } | [ "-//" "=" ] -> { [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] } | [ "-//=" ] -> { [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] } | [ "-/" integer(n) "/=" ] -> { [IPatNoop;IPatSimpl(SimplCut (n,~-1))] } | [ "-/" integer(n) "/" integer (m) "=" ] -> { [IPatNoop;IPatSimpl(SimplCut(n,m))] } | [ ssrfwdview(v) ] -> { [IPatView v] } | [ "[" ":" ident_list(idl) "]" ] -> { [IPatAbstractVars idl] } | [ "[:" ident_list(idl) "]" ] -> { [IPatAbstractVars idl] } END ARGUMENT EXTEND ssripats TYPED AS ssripat PRINTED BY { pr_ssripats } | [ ssripat(i) ssripats(tl) ] -> { i @ tl } | [ ] -> { [] } END ARGUMENT EXTEND ssriorpat TYPED AS ssripat list PRINTED BY { pr_ssriorpat } | [ ssripats(pats) "|" ssriorpat(orpat) ] -> { pats :: orpat } | [ ssripats(pats) "|-" ">" ssriorpat(orpat) ] -> { pats :: pushIPatRewrite orpat } | [ ssripats(pats) "|-" ssriorpat(orpat) ] -> { pats :: pushIPatNoop orpat } | [ ssripats(pats) "|->" ssriorpat(orpat) ] -> { pats :: pushIPatRewrite orpat } | [ ssripats(pats) "||" ssriorpat(orpat) ] -> { pats :: [] :: orpat } | [ ssripats(pats) "|||" ssriorpat(orpat) ] -> { pats :: [] :: [] :: orpat } | [ ssripats(pats) "||||" ssriorpat(orpat) ] -> { [pats; []; []; []] @ orpat } | [ ssripats(pats) ] -> { [pats] } END { let reject_ssrhid _ strm = match Util.stream_nth 0 strm with | Tok.KEYWORD "[" -> (match Util.stream_nth 1 strm with | Tok.KEYWORD ":" -> raise Stream.Failure | _ -> ()) | _ -> () let test_nohidden = Pcoq.Entry.of_parser "test_ssrhid" reject_ssrhid let rec reject_binder crossed_paren k tok s = match try Some (Util.stream_nth k s) with Stream.Failure -> None with | Some (Tok.KEYWORD "(") when not crossed_paren -> reject_binder true (k+1) tok s | Some (Tok.IDENT _) when crossed_paren -> reject_binder true (k+1) tok s | Some (Tok.KEYWORD ":" | Tok.KEYWORD ":=") when crossed_paren -> raise Stream.Failure | Some (Tok.KEYWORD ")") when crossed_paren -> raise Stream.Failure | _ -> if crossed_paren then () else raise Stream.Failure let _test_nobinder = Pcoq.Entry.of_parser "test_nobinder" (reject_binder false 0) } ARGUMENT EXTEND ssrcpat TYPED AS ssripatrep PRINTED BY { pr_ssripat } END (* Pcoq *) GRAMMAR EXTEND Gram GLOBAL: ssrcpat; hat: [ [ "^"; id = ident -> { Prefix id } | "^"; "~"; id = ident -> { SuffixId id } | "^"; "~"; n = natural -> { SuffixNum n } | "^~"; id = ident -> { SuffixId id } | "^~"; n = natural -> { SuffixNum n } ]]; ssrcpat: [ [ test_nohidden; "["; hat_id = hat; "]" -> { IPatCase (Block(hat_id)) } | test_nohidden; "["; iorpat = ssriorpat; "]" -> { IPatCase (Regular iorpat) } | test_nohidden; "[="; iorpat = ssriorpat; "]" -> { IPatInj iorpat } ]]; END GRAMMAR EXTEND Gram GLOBAL: ssripat; ssripat: [[ pat = ssrcpat -> { [pat] } ]]; END ARGUMENT EXTEND ssripats_ne TYPED AS ssripat PRINTED BY { pr_ssripats } | [ ssripat(i) ssripats(tl) ] -> { i @ tl } END (* subsets of patterns *) { (* TODO: review what this function does, it looks suspicious *) let check_ssrhpats loc w_binders ipats = let err_loc s = CErrors.user_err ~loc ~hdr:"ssreflect" s in let clr, ipats = let opt_app = function None -> fun l -> Some l | Some l1 -> fun l2 -> Some (l1 @ l2) in let rec aux clr = function | IPatClear cl :: tl -> aux (opt_app clr cl) tl | tl -> clr, tl in aux None ipats in let simpl, ipats = match List.rev ipats with | IPatSimpl _ as s :: tl -> [s], List.rev tl | _ -> [], ipats in if simpl <> [] && not w_binders then err_loc (str "No s-item allowed here: " ++ pr_ipats simpl); let ipat, binders = let rec loop ipat = function | [] -> ipat, [] | ( IPatId _| IPatAnon _| IPatCase _ | IPatDispatch _ | IPatRewrite _ as i) :: tl -> if w_binders then if simpl <> [] && tl <> [] then err_loc(str"binders XOR s-item allowed here: "++pr_ipats(tl@simpl)) else if not (List.for_all (function IPatId _ -> true | _ -> false) tl) then err_loc (str "Only binders allowed here: " ++ pr_ipats tl) else ipat @ [i], tl else if tl = [] then ipat @ [i], [] else err_loc (str "No binder or s-item allowed here: " ++ pr_ipats tl) | hd :: tl -> loop (ipat @ [hd]) tl in loop [] ipats in ((clr, ipat), binders), simpl let pr_clear_opt sep = function None -> mt () | Some x -> pr_clear sep x let pr_hpats (((clr, ipat), binders), simpl) = pr_clear_opt mt clr ++ pr_ipats ipat ++ pr_ipats binders ++ pr_ipats simpl let pr_ssrhpats _ _ _ = pr_hpats let pr_ssrhpats_wtransp _ _ _ (_, x) = pr_hpats x } ARGUMENT EXTEND ssrhpats TYPED AS (((ssrclear option * ssripat) * ssripat) * ssripat) PRINTED BY { pr_ssrhpats } | [ ssripats(i) ] -> { check_ssrhpats loc true i } END ARGUMENT EXTEND ssrhpats_wtransp TYPED AS (bool * (((ssrclear option * ssripats) * ssripats) * ssripats)) PRINTED BY { pr_ssrhpats_wtransp } | [ ssripats(i) ] -> { false,check_ssrhpats loc true i } | [ ssripats(i) "@" ssripats(j) ] -> { true,check_ssrhpats loc true (i @ j) } END ARGUMENT EXTEND ssrhpats_nobs TYPED AS (((ssrclear option * ssripats) * ssripats) * ssripats) PRINTED BY { pr_ssrhpats } | [ ssripats(i) ] -> { check_ssrhpats loc false i } END ARGUMENT EXTEND ssrrpat TYPED AS ssripatrep PRINTED BY { pr_ssripat } | [ "->" ] -> { IPatRewrite (allocc, L2R) } | [ "<-" ] -> { IPatRewrite (allocc, R2L) } END { let pr_intros sep intrs = if intrs = [] then mt() else sep () ++ str "=>" ++ pr_ipats intrs let pr_ssrintros _ _ _ = pr_intros mt } ARGUMENT EXTEND ssrintros_ne TYPED AS ssripat PRINTED BY { pr_ssrintros } | [ "=>" ssripats_ne(pats) ] -> { pats } (* TODO | [ "=>" ">" ssripats_ne(pats) ] -> { IPatFastMode :: pats } | [ "=>>" ssripats_ne(pats) ] -> [ IPatFastMode :: pats ] *) END ARGUMENT EXTEND ssrintros TYPED AS ssrintros_ne PRINTED BY { pr_ssrintros } | [ ssrintros_ne(intrs) ] -> { intrs } | [ ] -> { [] } END { let pr_ssrintrosarg env sigma _ _ prt (tac, ipats) = prt env sigma tacltop tac ++ pr_intros spc ipats } ARGUMENT EXTEND ssrintrosarg TYPED AS (tactic * ssrintros) PRINTED BY { pr_ssrintrosarg env sigma } END { let () = register_ssrtac "tclintros" begin fun args ist -> match args with | [arg] -> let arg = cast_arg wit_ssrintrosarg arg in let tac, intros = arg in ssrevaltac ist tac <*> tclIPATssr intros | _ -> assert false end (** Defined identifier *) let pr_ssrfwdid id = pr_spc () ++ pr_id id let pr_ssrfwdidx _ _ _ = pr_ssrfwdid } (* We use a primitive parser for the head identifier of forward *) (* tactis to avoid syntactic conflicts with basic Coq tactics. *) ARGUMENT EXTEND ssrfwdid TYPED AS ident PRINTED BY { pr_ssrfwdidx } END { let accept_ssrfwdid _ strm = match stream_nth 0 strm with | Tok.IDENT id -> accept_before_syms_or_any_id [":"; ":="; "("] strm | _ -> raise Stream.Failure let test_ssrfwdid = Pcoq.Entry.of_parser "test_ssrfwdid" accept_ssrfwdid } GRAMMAR EXTEND Gram GLOBAL: ssrfwdid; ssrfwdid: [[ test_ssrfwdid; id = Prim.ident -> { id } ]]; END (* by *) (** Tactical arguments. *) (* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) (* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) (* and subgoal reordering tacticals (; first & ; last), respectively. *) { let pr_ortacs env sigma prt = let rec pr_rec = function | [None] -> spc() ++ str "|" ++ spc() | None :: tacs -> spc() ++ str "|" ++ pr_rec tacs | Some tac :: tacs -> spc() ++ str "| " ++ prt env sigma tacltop tac ++ pr_rec tacs | [] -> mt() in function | [None] -> spc() | None :: tacs -> pr_rec tacs | Some tac :: tacs -> prt env sigma tacltop tac ++ pr_rec tacs | [] -> mt() let pr_ssrortacs env sigma _ _ = pr_ortacs env sigma } ARGUMENT EXTEND ssrortacs TYPED AS tactic option list PRINTED BY { pr_ssrortacs env sigma } | [ ssrtacarg(tac) "|" ssrortacs(tacs) ] -> { Some tac :: tacs } | [ ssrtacarg(tac) "|" ] -> { [Some tac; None] } | [ ssrtacarg(tac) ] -> { [Some tac] } | [ "|" ssrortacs(tacs) ] -> { None :: tacs } | [ "|" ] -> { [None; None] } END { let pr_hintarg env sigma prt = function | true, tacs -> hv 0 (str "[ " ++ pr_ortacs env sigma prt tacs ++ str " ]") | false, [Some tac] -> prt env sigma tacltop tac | _, _ -> mt() let pr_ssrhintarg env sigma _ _ = pr_hintarg env sigma } ARGUMENT EXTEND ssrhintarg TYPED AS (bool * ssrortacs) PRINTED BY { pr_ssrhintarg env sigma } | [ "[" "]" ] -> { nullhint } | [ "[" ssrortacs(tacs) "]" ] -> { mk_orhint tacs } | [ ssrtacarg(arg) ] -> { mk_hint arg } END (* Copy of ssrhintarg with LEVEL "3", useful for: "under ... do ..." *) ARGUMENT EXTEND ssrhint3arg TYPED AS (bool * ssrortacs) PRINTED BY { pr_ssrhintarg env sigma } | [ "[" "]" ] -> { nullhint } | [ "[" ssrortacs(tacs) "]" ] -> { mk_orhint tacs } | [ ssrtac3arg(arg) ] -> { mk_hint arg } END ARGUMENT EXTEND ssrortacarg TYPED AS ssrhintarg PRINTED BY { pr_ssrhintarg env sigma } | [ "[" ssrortacs(tacs) "]" ] -> { mk_orhint tacs } END { let pr_hint env sigma prt arg = if arg = nohint then mt() else str "by " ++ pr_hintarg env sigma prt arg let pr_ssrhint env sigma _ _ = pr_hint env sigma } ARGUMENT EXTEND ssrhint TYPED AS ssrhintarg PRINTED BY { pr_ssrhint env sigma } | [ ] -> { nohint } END (** The "in" pseudo-tactical *) (* We can't make "in" into a general tactical because this would create a *) (* crippling conflict with the ltac let .. in construct. Hence, we add *) (* explicitly an "in" suffix to all the extended tactics for which it is *) (* relevant (including move, case, elim) and to the extended do tactical *) (* below, which yields a general-purpose "in" of the form do [...] in ... *) (* This tactical needs to come before the intro tactics because the latter *) (* must take precautions in order not to interfere with the discharged *) (* assumptions. This is especially difficult for discharged "let"s, which *) (* the default simpl and unfold tactics would erase blindly. *) { open Ssrmatching_plugin.Ssrmatching open Ssrmatching_plugin.G_ssrmatching let pr_wgen = function | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hoi id | (clr, Some((id,k),Some p)) -> spc() ++ pr_clear mt clr ++ str"(" ++ str k ++ pr_hoi id ++ str ":=" ++ pr_cpattern p ++ str ")" | (clr, None) -> spc () ++ pr_clear mt clr let pr_ssrwgen _ _ _ = pr_wgen } (* no globwith for char *) ARGUMENT EXTEND ssrwgen TYPED AS (ssrclear * ((ssrhoi_hyp * string) * cpattern option) option) PRINTED BY { pr_ssrwgen } | [ ssrclear_ne(clr) ] -> { clr, None } | [ ssrhoi_hyp(hyp) ] -> { [], Some((hyp, " "), None) } | [ "@" ssrhoi_hyp(hyp) ] -> { [], Some((hyp, "@"), None) } | [ "(" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> { [], Some ((id," "),Some p) } | [ "(" ssrhoi_id(id) ")" ] -> { [], Some ((id,"("), None) } | [ "(@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> { [], Some ((id,"@"),Some p) } | [ "(" "@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> { [], Some ((id,"@"),Some p) } END { let pr_clseq = function | InGoal | InHyps -> mt () | InSeqGoal -> str "|- *" | InHypsSeqGoal -> str " |- *" | InHypsGoal -> str " *" | InAll -> str "*" | InHypsSeq -> str " |-" | InAllHyps -> str "* |-" let wit_ssrclseq = add_genarg "ssrclseq" (fun env sigma -> pr_clseq) let pr_clausehyps = pr_list pr_spc pr_wgen let pr_ssrclausehyps _ _ _ = pr_clausehyps } ARGUMENT EXTEND ssrclausehyps TYPED AS ssrwgen list PRINTED BY { pr_ssrclausehyps } | [ ssrwgen(hyp) "," ssrclausehyps(hyps) ] -> { hyp :: hyps } | [ ssrwgen(hyp) ssrclausehyps(hyps) ] -> { hyp :: hyps } | [ ssrwgen(hyp) ] -> { [hyp] } END { (* type ssrclauses = ssrahyps * ssrclseq *) let pr_clauses (hyps, clseq) = if clseq = InGoal then mt () else str "in " ++ pr_clausehyps hyps ++ pr_clseq clseq let pr_ssrclauses _ _ _ = pr_clauses } ARGUMENT EXTEND ssrclauses TYPED AS (ssrwgen list * ssrclseq) PRINTED BY { pr_ssrclauses } | [ "in" ssrclausehyps(hyps) "|-" "*" ] -> { hyps, InHypsSeqGoal } | [ "in" ssrclausehyps(hyps) "|-" ] -> { hyps, InHypsSeq } | [ "in" ssrclausehyps(hyps) "*" ] -> { hyps, InHypsGoal } | [ "in" ssrclausehyps(hyps) ] -> { hyps, InHyps } | [ "in" "|-" "*" ] -> { [], InSeqGoal } | [ "in" "*" ] -> { [], InAll } | [ "in" "*" "|-" ] -> { [], InAllHyps } | [ ] -> { [], InGoal } END { (** Definition value formatting *) (* We use an intermediate structure to correctly render the binder list *) (* abbreviations. We use a list of hints to extract the binders and *) (* base term from a term, for the two first levels of representation of *) (* of constr terms. *) let pr_binder prl = function | Bvar x -> pr_name x | Bdecl (xs, t) -> str "(" ++ pr_list pr_spc pr_name xs ++ str " : " ++ prl t ++ str ")" | Bdef (x, None, v) -> str "(" ++ pr_name x ++ str " := " ++ prl v ++ str ")" | Bdef (x, Some t, v) -> str "(" ++ pr_name x ++ str " : " ++ prl t ++ str " := " ++ prl v ++ str ")" | Bstruct x -> str "{struct " ++ pr_name x ++ str "}" | Bcast t -> str ": " ++ prl t let rec format_local_binders h0 bl0 = match h0, bl0 with | BFvar :: h, CLocalAssum ([{CAst.v=x}], _, _) :: bl -> Bvar x :: format_local_binders h bl | BFdecl _ :: h, CLocalAssum (lxs, _, t) :: bl -> Bdecl (List.map (fun x -> x.CAst.v) lxs, t) :: format_local_binders h bl | BFdef :: h, CLocalDef ({CAst.v=x}, v, oty) :: bl -> Bdef (x, oty, v) :: format_local_binders h bl | _ -> [] let rec format_constr_expr h0 c0 = let open CAst in match h0, c0 with | BFvar :: h, { v = CLambdaN ([CLocalAssum([{CAst.v=x}], _, _)], c) } -> let bs, c' = format_constr_expr h c in Bvar x :: bs, c' | BFdecl _:: h, { v = CLambdaN ([CLocalAssum(lxs, _, t)], c) } -> let bs, c' = format_constr_expr h c in Bdecl (List.map (fun x -> x.CAst.v) lxs, t) :: bs, c' | BFdef :: h, { v = CLetIn({CAst.v=x}, v, oty, c) } -> let bs, c' = format_constr_expr h c in Bdef (x, oty, v) :: bs, c' | [BFcast], { v = CCast (c, Glob_term.CastConv t) } -> [Bcast t], c | BFrec (has_str, has_cast) :: h, { v = CFix ( _, [_, Some {CAst.v = CStructRec locn}, bl, t, c]) } -> let bs = format_local_binders h bl in let bstr = if has_str then [Bstruct (Name locn.CAst.v)] else [] in bs @ bstr @ (if has_cast then [Bcast t] else []), c | BFrec (_, has_cast) :: h, { v = CCoFix ( _, [_, bl, t, c]) } -> format_local_binders h bl @ (if has_cast then [Bcast t] else []), c | _, c -> [], c (** Forward chaining argument *) (* There are three kinds of forward definitions: *) (* - Hint: type only, cast to Type, may have proof hint. *) (* - Have: type option + value, no space before type *) (* - Pose: binders + value, space before binders. *) let pr_fwdkind = function | FwdHint (s,_) -> str (s ^ " ") | _ -> str " :=" ++ spc () let pr_fwdfmt (fk, _ : ssrfwdfmt) = pr_fwdkind fk let wit_ssrfwdfmt = add_genarg "ssrfwdfmt" (fun env sigma -> pr_fwdfmt) (* type ssrfwd = ssrfwdfmt * ssrterm *) let mkFwdVal fk c = ((fk, []), c) let mkssrFwdVal fk c = ((fk, []), (c,None)) let dC t = Glob_term.CastConv t let same_ist { interp_env = x } { interp_env = y } = match x,y with | None, None -> true | Some a, Some b -> a == b | _ -> false let mkFwdCast fk ?loc ?c t = let c = match c with | None -> mkCHole loc | Some c -> assert (same_ist t c); c.body in ((fk, [BFcast]), { t with annotation = `None; body = (CAst.make ?loc @@ CCast (c, dC t.body)) }) let mkssrFwdCast fk loc t c = ((fk, [BFcast]), (c, Some t)) let mkFwdHint s t = let loc = Constrexpr_ops.constr_loc t.body in mkFwdCast (FwdHint (s,false)) ?loc t let mkFwdHintNoTC s t = let loc = Constrexpr_ops.constr_loc t.body in mkFwdCast (FwdHint (s,true)) ?loc t let pr_gen_fwd prval prc prlc fk (bs, c) = let prc s = str s ++ spc () ++ prval prc prlc c in match fk, bs with | FwdHint (s,_), [Bcast t] -> str s ++ spc () ++ prlc t | FwdHint (s,_), _ -> prc (s ^ "(* typeof *)") | FwdHave, [Bcast t] -> str ":" ++ spc () ++ prlc t ++ prc " :=" | _, [] -> prc " :=" | _, _ -> spc () ++ pr_list spc (pr_binder prlc) bs ++ prc " :=" let pr_fwd_guarded prval prval' = function | (fk, h), c -> pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c.body) let pr_unguarded prc prlc = prlc let pr_fwd = pr_fwd_guarded pr_unguarded pr_unguarded let pr_ssrfwd _ _ _ = pr_fwd } ARGUMENT EXTEND ssrfwd TYPED AS (ssrfwdfmt * ast_closure_lterm) PRINTED BY { pr_ssrfwd } | [ ":=" ast_closure_lterm(c) ] -> { mkFwdVal FwdPose c } | [ ":" ast_closure_lterm (t) ":=" ast_closure_lterm(c) ] -> { mkFwdCast FwdPose ~loc t ~c } END (** Independent parsing for binders *) (* The pose, pose fix, and pose cofix tactics use these internally to *) (* parse argument fragments. *) { let pr_ssrbvar env sigma prc _ _ v = prc env sigma v } ARGUMENT EXTEND ssrbvar TYPED AS constr PRINTED BY { pr_ssrbvar env sigma } | [ ident(id) ] -> { mkCVar ~loc id } | [ "_" ] -> { mkCHole (Some loc) } END { let bvar_lname = let open CAst in function | { v = CRef (qid, _) } when qualid_is_ident qid -> CAst.make ?loc:qid.CAst.loc @@ Name (qualid_basename qid) | { loc = loc } -> CAst.make ?loc Anonymous let pr_ssrbinder env sigma prc _ _ (_, c) = prc env sigma c } ARGUMENT EXTEND ssrbinder TYPED AS (ssrfwdfmt * constr) PRINTED BY { pr_ssrbinder env sigma } | [ ssrbvar(bv) ] -> { let { CAst.loc=xloc } as x = bvar_lname bv in (FwdPose, [BFvar]), CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Glob_term.Explicit,mkCHole xloc)],mkCHole (Some loc)) } | [ "(" ssrbvar(bv) ")" ] -> { let { CAst.loc=xloc } as x = bvar_lname bv in (FwdPose, [BFvar]), CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Glob_term.Explicit,mkCHole xloc)],mkCHole (Some loc)) } | [ "(" ssrbvar(bv) ":" lconstr(t) ")" ] -> { let x = bvar_lname bv in (FwdPose, [BFdecl 1]), CAst.make ~loc @@ CLambdaN ([CLocalAssum([x], Default Glob_term.Explicit, t)], mkCHole (Some loc)) } | [ "(" ssrbvar(bv) ne_ssrbvar_list(bvs) ":" lconstr(t) ")" ] -> { let xs = List.map bvar_lname (bv :: bvs) in let n = List.length xs in (FwdPose, [BFdecl n]), CAst.make ~loc @@ CLambdaN ([CLocalAssum (xs, Default Glob_term.Explicit, t)], mkCHole (Some loc)) } | [ "(" ssrbvar(id) ":" lconstr(t) ":=" lconstr(v) ")" ] -> { (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, Some t, mkCHole (Some loc)) } | [ "(" ssrbvar(id) ":=" lconstr(v) ")" ] -> { (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, None, mkCHole (Some loc)) } END GRAMMAR EXTEND Gram GLOBAL: ssrbinder; ssrbinder: [ [ ["of" -> { () } | "&" -> { () } ]; c = operconstr LEVEL "99" -> { (FwdPose, [BFvar]), CAst.make ~loc @@ CLambdaN ([CLocalAssum ([CAst.make ~loc Anonymous],Default Glob_term.Explicit,c)],mkCHole (Some loc)) } ] ]; END { let rec binders_fmts = function | ((_, h), _) :: bs -> h @ binders_fmts bs | _ -> [] let push_binders c2 bs = let loc2 = constr_loc c2 in let mkloc loc1 = Loc.merge_opt loc1 loc2 in let open CAst in let rec loop ty c = function | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs when ty -> CAst.make ?loc:(mkloc loc1) @@ CProdN (b, loop ty c bs) | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs -> CAst.make ?loc:(mkloc loc1) @@ CLambdaN (b, loop ty c bs) | (_, { loc = loc1; v = CLetIn (x, v, oty, _) } ) :: bs -> CAst.make ?loc:(mkloc loc1) @@ CLetIn (x, v, oty, loop ty c bs) | [] -> c | _ -> anomaly "binder not a lambda nor a let in" in match c2 with | { loc; v = CCast (ct, Glob_term.CastConv cty) } -> CAst.make ?loc @@ (CCast (loop false ct bs, Glob_term.CastConv (loop true cty bs))) | ct -> loop false ct bs let rec fix_binders = let open CAst in function | (_, { v = CLambdaN ([CLocalAssum(xs, _, t)], _) } ) :: bs -> CLocalAssum (xs, Default Glob_term.Explicit, t) :: fix_binders bs | (_, { v = CLetIn (x, v, oty, _) } ) :: bs -> CLocalDef (x, v, oty) :: fix_binders bs | _ -> [] let pr_ssrstruct _ _ _ = function | Some id -> str "{struct " ++ pr_id id ++ str "}" | None -> mt () } ARGUMENT EXTEND ssrstruct TYPED AS ident option PRINTED BY { pr_ssrstruct } | [ "{" "struct" ident(id) "}" ] -> { Some id } | [ ] -> { None } END (** The "pose" tactic *) (* The plain pose form. *) { let bind_fwd bs ((fk, h), c) = (fk,binders_fmts bs @ h), { c with body = push_binders c.body bs } } ARGUMENT EXTEND ssrposefwd TYPED AS ssrfwd PRINTED BY { pr_ssrfwd } | [ ssrbinder_list(bs) ssrfwd(fwd) ] -> { bind_fwd bs fwd } END (* The pose fix form. *) { let pr_ssrfixfwd _ _ _ (id, fwd) = str " fix " ++ pr_id id ++ pr_fwd fwd let bvar_locid = function | { CAst.v = CRef (qid, _) } when qualid_is_ident qid -> CAst.make ?loc:qid.CAst.loc (qualid_basename qid) | _ -> CErrors.user_err (Pp.str "Missing identifier after \"(co)fix\"") } ARGUMENT EXTEND ssrfixfwd TYPED AS (ident * ssrfwd) PRINTED BY { pr_ssrfixfwd } | [ "fix" ssrbvar(bv) ssrbinder_list(bs) ssrstruct(sid) ssrfwd(fwd) ] -> { let { CAst.v=id } as lid = bvar_locid bv in let (fk, h), ac = fwd in let c = ac.body in let has_cast, t', c' = match format_constr_expr h c with | [Bcast t'], c' -> true, t', c' | _ -> false, mkCHole (constr_loc c), c in let lb = fix_binders bs in let has_struct, i = let rec loop = function | {CAst.loc=l'; v=Name id'} :: _ when Option.equal Id.equal sid (Some id') -> true, CAst.make ?loc:l' id' | [{CAst.loc=l';v=Name id'}] when sid = None -> false, CAst.make ?loc:l' id' | _ :: bn -> loop bn | [] -> CErrors.user_err (Pp.str "Bad structural argument") in loop (names_of_local_assums lb) in let h' = BFrec (has_struct, has_cast) :: binders_fmts bs in let fix = CAst.make ~loc @@ CFix (lid, [lid, (Some (CAst.make (CStructRec i))), lb, t', c']) in id, ((fk, h'), { ac with body = fix }) } END (* The pose cofix form. *) { let pr_ssrcofixfwd _ _ _ (id, fwd) = str " cofix " ++ pr_id id ++ pr_fwd fwd } ARGUMENT EXTEND ssrcofixfwd TYPED AS ssrfixfwd PRINTED BY { pr_ssrcofixfwd } | [ "cofix" ssrbvar(bv) ssrbinder_list(bs) ssrfwd(fwd) ] -> { let { CAst.v=id } as lid = bvar_locid bv in let (fk, h), ac = fwd in let c = ac.body in let has_cast, t', c' = match format_constr_expr h c with | [Bcast t'], c' -> true, t', c' | _ -> false, mkCHole (constr_loc c), c in let h' = BFrec (false, has_cast) :: binders_fmts bs in let cofix = CAst.make ~loc @@ CCoFix (lid, [lid, fix_binders bs, t', c']) in id, ((fk, h'), { ac with body = cofix }) } END { (* This does not print the type, it should be fixed... *) let pr_ssrsetfwd _ _ _ (((fk,_),(t,_)), docc) = pr_gen_fwd (fun _ _ -> pr_cpattern) (fun _ -> mt()) (fun _ -> mt()) fk ([Bcast ()],t) } ARGUMENT EXTEND ssrsetfwd TYPED AS ((ssrfwdfmt * (lcpattern * ast_closure_lterm option)) * ssrdocc) PRINTED BY { pr_ssrsetfwd } | [ ":" ast_closure_lterm(t) ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> { mkssrFwdCast FwdPose loc t c, mkocc occ } | [ ":" ast_closure_lterm(t) ":=" lcpattern(c) ] -> { mkssrFwdCast FwdPose loc t c, nodocc } | [ ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> { mkssrFwdVal FwdPose c, mkocc occ } | [ ":=" lcpattern(c) ] -> { mkssrFwdVal FwdPose c, nodocc } END { let pr_ssrhavefwd env sigma _ _ prt (fwd, hint) = pr_fwd fwd ++ pr_hint env sigma prt hint } ARGUMENT EXTEND ssrhavefwd TYPED AS (ssrfwd * ssrhint) PRINTED BY { pr_ssrhavefwd env sigma } | [ ":" ast_closure_lterm(t) ssrhint(hint) ] -> { mkFwdHint ":" t, hint } | [ ":" ast_closure_lterm(t) ":=" ast_closure_lterm(c) ] -> { mkFwdCast FwdHave ~loc t ~c, nohint } | [ ":" ast_closure_lterm(t) ":=" ] -> { mkFwdHintNoTC ":" t, nohint } | [ ":=" ast_closure_lterm(c) ] -> { mkFwdVal FwdHave c, nohint } END { let intro_id_to_binder = List.map (function | IPatId id -> let { CAst.loc=xloc } as x = bvar_lname (mkCVar id) in (FwdPose, [BFvar]), CAst.make @@ CLambdaN ([CLocalAssum([x], Default Glob_term.Explicit, mkCHole xloc)], mkCHole None) | _ -> anomaly "non-id accepted as binder") let binder_to_intro_id = CAst.(List.map (function | (FwdPose, [BFvar]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) } | (FwdPose, [BFdecl _]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) } -> List.map (function {v=Name id} -> IPatId id | _ -> IPatAnon (One None)) ids | (FwdPose, [BFdef]), { v = CLetIn ({v=Name id},_,_,_) } -> [IPatId id] | (FwdPose, [BFdef]), { v = CLetIn ({v=Anonymous},_,_,_) } -> [IPatAnon (One None)] | _ -> anomaly "ssrbinder is not a binder")) let pr_ssrhavefwdwbinders env sigma _ _ prt (tr,((hpats, (fwd, hint)))) = pr_hpats hpats ++ pr_fwd fwd ++ pr_hint env sigma prt hint } ARGUMENT EXTEND ssrhavefwdwbinders TYPED AS (bool * (ssrhpats * (ssrfwd * ssrhint))) PRINTED BY { pr_ssrhavefwdwbinders env sigma } | [ ssrhpats_wtransp(trpats) ssrbinder_list(bs) ssrhavefwd(fwd) ] -> { let tr, pats = trpats in let ((clr, pats), binders), simpl = pats in let allbs = intro_id_to_binder binders @ bs in let allbinders = binders @ List.flatten (binder_to_intro_id bs) in let hint = bind_fwd allbs (fst fwd), snd fwd in tr, ((((clr, pats), allbinders), simpl), hint) } END { let pr_ssrdoarg env sigma prc _ prt (((n, m), tac), clauses) = pr_index n ++ pr_mmod m ++ pr_hintarg env sigma prt tac ++ pr_clauses clauses } ARGUMENT EXTEND ssrdoarg TYPED AS (((ssrindex * ssrmmod) * ssrhintarg) * ssrclauses) PRINTED BY { pr_ssrdoarg env sigma } END { (* type ssrseqarg = ssrindex * (ssrtacarg * ssrtac option) *) let pr_seqtacarg env sigma prt = function | (is_first, []), _ -> str (if is_first then "first" else "last") | tac, Some dtac -> hv 0 (pr_hintarg env sigma prt tac ++ spc() ++ str "|| " ++ prt env sigma tacltop dtac) | tac, _ -> pr_hintarg env sigma prt tac let pr_ssrseqarg env sigma _ _ prt = function | ArgArg 0, tac -> pr_seqtacarg env sigma prt tac | i, tac -> pr_index i ++ str " " ++ pr_seqtacarg env sigma prt tac } (* We must parse the index separately to resolve the conflict with *) (* an unindexed tactic. *) ARGUMENT EXTEND ssrseqarg TYPED AS (ssrindex * (ssrhintarg * tactic option)) PRINTED BY { pr_ssrseqarg env sigma } END { let sq_brace_tacnames = ["first"; "solve"; "do"; "rewrite"; "have"; "suffices"; "wlog"] (* "by" is a keyword *) let accept_ssrseqvar _ strm = match stream_nth 0 strm with | Tok.IDENT id when not (List.mem id sq_brace_tacnames) -> accept_before_syms_or_ids ["["] ["first";"last"] strm | _ -> raise Stream.Failure let test_ssrseqvar = Pcoq.Entry.of_parser "test_ssrseqvar" accept_ssrseqvar let swaptacarg (loc, b) = (b, []), Some (TacId []) let check_seqtacarg dir arg = match snd arg, dir with | ((true, []), Some (TacAtom { CAst.loc })), L2R -> CErrors.user_err ?loc (str "expected \"last\"") | ((false, []), Some (TacAtom { CAst.loc })), R2L -> CErrors.user_err ?loc (str "expected \"first\"") | _, _ -> arg let ssrorelse = Entry.create "ssrorelse" } GRAMMAR EXTEND Gram GLOBAL: ssrorelse ssrseqarg; ssrseqidx: [ [ test_ssrseqvar; id = Prim.ident -> { ArgVar (CAst.make ~loc id) } | n = Prim.natural -> { ArgArg (check_index ~loc n) } ] ]; ssrswap: [[ IDENT "first" -> { loc, true } | IDENT "last" -> { loc, false } ]]; ssrorelse: [[ "||"; tac = tactic_expr LEVEL "2" -> { tac } ]]; ssrseqarg: [ [ arg = ssrswap -> { noindex, swaptacarg arg } | i = ssrseqidx; tac = ssrortacarg; def = OPT ssrorelse -> { i, (tac, def) } | i = ssrseqidx; arg = ssrswap -> { i, swaptacarg arg } | tac = tactic_expr LEVEL "3" -> { noindex, (mk_hint tac, None) } ] ]; END { let tactic_expr = Pltac.tactic_expr } (** 1. Utilities *) (** Tactic-level diagnosis *) (* debug *) { (* Let's play with the new proof engine API *) let old_tac = V82.tactic } (** Name generation *) (* Since Coq now does repeated internal checks of its external lexical *) (* rules, we now need to carve ssreflect reserved identifiers out of *) (* out of the user namespace. We use identifiers of the form _id_ for *) (* this purpose, e.g., we "anonymize" an identifier id as _id_, adding *) (* an extra leading _ if this might clash with an internal identifier. *) (* We check for ssreflect identifiers in the ident grammar rule; *) (* when the ssreflect Module is present this is normally an error, *) (* but we provide a compatibility flag to reduce this to a warning. *) { let ssr_reserved_ids = Summary.ref ~name:"SSR:idents" true let () = Goptions.(declare_bool_option { optname = "ssreflect identifiers"; optkey = ["SsrIdents"]; optdepr = false; optread = (fun _ -> !ssr_reserved_ids); optwrite = (fun b -> ssr_reserved_ids := b) }) let is_ssr_reserved s = let n = String.length s in n > 2 && s.[0] = '_' && s.[n - 1] = '_' let ssr_id_of_string loc s = if is_ssr_reserved s && is_ssr_loaded () then begin if !ssr_reserved_ids then CErrors.user_err ~loc (str ("The identifier " ^ s ^ " is reserved.")) else if is_internal_name s then Feedback.msg_warning (str ("Conflict between " ^ s ^ " and ssreflect internal names.")) else Feedback.msg_warning (str ( "The name " ^ s ^ " fits the _xxx_ format used for anonymous variables.\n" ^ "Scripts with explicit references to anonymous variables are fragile.")) end; Id.of_string s let ssr_null_entry = Pcoq.Entry.of_parser "ssr_null" (fun _ _ -> ()) } GRAMMAR EXTEND Gram GLOBAL: Prim.ident; Prim.ident: [[ s = IDENT; ssr_null_entry -> { ssr_id_of_string loc s } ]]; END { let perm_tag = "_perm_Hyp_" let _ = add_internal_name (is_tagged perm_tag) } (* We must not anonymize context names discharged by the "in" tactical. *) (** Tactical extensions. *) { type ssrargfmt = ArgSsr of string | ArgSep of string let set_pr_ssrtac name prec afmt = (* FIXME *) () (* let fmt = List.map (function | ArgSep s -> Egramml.GramTerminal s | ArgSsr s -> Egramml.GramTerminal s | ArgCoq at -> Egramml.GramTerminal "COQ_ARG") afmt in let tacname = ssrtac_name name in () *) let ssrtac_expr ?loc name args = TacML (CAst.make ?loc (ssrtac_entry name, args)) let tclintros_expr ?loc tac ipats = let args = [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrintrosarg) (tac, ipats))] in ssrtac_expr ?loc "tclintros" args } GRAMMAR EXTEND Gram GLOBAL: tactic_expr; tactic_expr: LEVEL "1" [ RIGHTA [ tac = tactic_expr; intros = ssrintros_ne -> { tclintros_expr ~loc tac intros } ] ]; END (** Bracketing tactical *) (* The tactic pretty-printer doesn't know that some extended tactics *) (* are actually tacticals. To prevent it from improperly removing *) (* parentheses we override the parsing rule for bracketed tactic *) (* expressions so that the pretty-print always reflects the input. *) (* (Removing user-specified parentheses is dubious anyway). *) GRAMMAR EXTEND Gram GLOBAL: tactic_expr; ssrparentacarg: [[ "("; tac = tactic_expr; ")" -> { CAst.make ~loc (Tacexp tac) } ]]; tactic_expr: LEVEL "0" [[ arg = ssrparentacarg -> { TacArg arg } ]]; END (** The internal "done" and "ssrautoprop" tactics. *) (* For additional flexibility, "done" and "ssrautoprop" are *) (* defined in Ltac. *) (* Although we provide a default definition in ssreflect, *) (* we look up the definition dynamically at each call point, *) (* to allow for user extensions. "ssrautoprop" defaults to *) (* trivial. *) { let ssrautoprop gl = try let tacname = try Tacenv.locate_tactic (qualid_of_ident (Id.of_string "ssrautoprop")) with Not_found -> Tacenv.locate_tactic (ssrqid "ssrautoprop") in let tacexpr = CAst.make @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in V82.of_tactic (eval_tactic (Tacexpr.TacArg tacexpr)) gl with Not_found -> V82.of_tactic (Auto.full_trivial []) gl let () = ssrautoprop_tac := ssrautoprop let tclBY tac = Tacticals.tclTHEN tac (donetac ~-1) (** Tactical arguments. *) (* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) (* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) (* and subgoal reordering tacticals (; first & ; last), respectively. *) (* Force use of the tactic_expr parsing entry, to rule out tick marks. *) (** The "by" tactical. *) open Ssrfwd } TACTIC EXTEND ssrtclby | [ "by" ssrhintarg(tac) ] -> { V82.tactic (hinttac ist true tac) } END (* We can't parse "by" in ARGUMENT EXTEND because it will only be made *) (* into a keyword in ssreflect.v; so we anticipate this in GEXTEND. *) GRAMMAR EXTEND Gram GLOBAL: ssrhint simple_tactic; ssrhint: [[ "by"; arg = ssrhintarg -> { arg } ]]; END (** The "do" tactical. ********************************************************) { let () = register_ssrtac "tcldo" begin fun args ist -> match args with | [arg] -> let arg = cast_arg wit_ssrdoarg arg in V82.tactic (ssrdotac ist arg) | _ -> assert false end let _ = set_pr_ssrtac "tcldo" 3 [ArgSep "do "; ArgSsr "doarg"] let ssrdotac_expr ?loc n m tac clauses = let arg = ((n, m), tac), clauses in ssrtac_expr ?loc "tcldo" [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrdoarg) arg)] } GRAMMAR EXTEND Gram GLOBAL: tactic_expr; ssrdotac: [ [ tac = tactic_expr LEVEL "3" -> { mk_hint tac } | tacs = ssrortacarg -> { tacs } ] ]; tactic_expr: LEVEL "3" [ RIGHTA [ IDENT "do"; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses -> { ssrdotac_expr ~loc noindex m tac clauses } | IDENT "do"; tac = ssrortacarg; clauses = ssrclauses -> { ssrdotac_expr ~loc noindex Once tac clauses } | IDENT "do"; n = int_or_var; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses -> { ssrdotac_expr ~loc (mk_index ~loc n) m tac clauses } ] ]; END { (* We can't actually parse the direction separately because this *) (* would introduce conflicts with the basic ltac syntax. *) let pr_ssrseqdir _ _ _ = function | L2R -> str ";" ++ spc () ++ str "first " | R2L -> str ";" ++ spc () ++ str "last " } ARGUMENT EXTEND ssrseqdir TYPED AS ssrdir PRINTED BY { pr_ssrseqdir } END { let () = register_ssrtac "tclseq" begin fun args ist -> match args with | [tac; dir; arg] -> let tac = cast_arg wit_ssrtclarg tac in let dir = cast_arg wit_ssrseqdir dir in let arg = cast_arg wit_ssrseqarg arg in V82.tactic (tclSEQAT ist tac dir arg) | _ -> assert false end let _ = set_pr_ssrtac "tclseq" 5 [ArgSsr "tclarg"; ArgSsr "seqdir"; ArgSsr "seqarg"] let tclseq_expr ?loc tac dir arg = let arg1 = in_gen (rawwit wit_ssrtclarg) tac in let arg2 = in_gen (rawwit wit_ssrseqdir) dir in let arg3 = in_gen (rawwit wit_ssrseqarg) (check_seqtacarg dir arg) in ssrtac_expr ?loc "tclseq" (List.map (fun x -> Tacexpr.TacGeneric x) [arg1; arg2; arg3]) } GRAMMAR EXTEND Gram GLOBAL: tactic_expr; ssr_first: [ [ tac = ssr_first; ipats = ssrintros_ne -> { tclintros_expr ~loc tac ipats } | "["; tacl = LIST0 tactic_expr SEP "|"; "]" -> { TacFirst tacl } ] ]; ssr_first_else: [ [ tac1 = ssr_first; tac2 = ssrorelse -> { TacOrelse (tac1, tac2) } | tac = ssr_first -> { tac } ]]; tactic_expr: LEVEL "4" [ LEFTA [ tac1 = tactic_expr; ";"; IDENT "first"; tac2 = ssr_first_else -> { TacThen (tac1, tac2) } | tac = tactic_expr; ";"; IDENT "first"; arg = ssrseqarg -> { tclseq_expr ~loc tac L2R arg } | tac = tactic_expr; ";"; IDENT "last"; arg = ssrseqarg -> { tclseq_expr ~loc tac R2L arg } ] ]; END (** 5. Bookkeeping tactics (clear, move, case, elim) *) (** Generalization (discharge) item *) (* An item is a switch + term pair. *) (* type ssrgen = ssrdocc * ssrterm *) { let pr_gen (docc, dt) = pr_docc docc ++ pr_cpattern dt let pr_ssrgen _ _ _ = pr_gen } ARGUMENT EXTEND ssrgen TYPED AS (ssrdocc * cpattern) PRINTED BY { pr_ssrgen } | [ ssrdocc(docc) cpattern(dt) ] -> { match docc with | Some [], _ -> CErrors.user_err ~loc (str"Clear flag {} not allowed here") | _ -> docc, dt } | [ cpattern(dt) ] -> { nodocc, dt } END { let has_occ ((_, occ), _) = occ <> None (** Generalization (discharge) sequence *) (* A discharge sequence is represented as a list of up to two *) (* lists of d-items, plus an ident list set (the possibly empty *) (* final clear switch). The main list is empty iff the command *) (* is defective, and has length two if there is a sequence of *) (* dependent terms (and in that case it is the first of the two *) (* lists). Thus, the first of the two lists is never empty. *) (* type ssrgens = ssrgen list *) (* type ssrdgens = ssrgens list * ssrclear *) let gens_sep = function [], [] -> mt | _ -> spc let pr_dgens pr_gen (gensl, clr) = let prgens s gens = str s ++ pr_list spc pr_gen gens in let prdeps deps = prgens ": " deps ++ spc () ++ str "/" in match gensl with | [deps; []] -> prdeps deps ++ pr_clear pr_spc clr | [deps; gens] -> prdeps deps ++ prgens " " gens ++ pr_clear spc clr | [gens] -> prgens ": " gens ++ pr_clear spc clr | _ -> pr_clear pr_spc clr let pr_ssrdgens _ _ _ = pr_dgens pr_gen let cons_gen gen = function | gens :: gensl, clr -> (gen :: gens) :: gensl, clr | _ -> anomaly "missing gen list" let cons_dep (gensl, clr) = if List.length gensl = 1 then ([] :: gensl, clr) else CErrors.user_err (Pp.str "multiple dependents switches '/'") } ARGUMENT EXTEND ssrdgens_tl TYPED AS (ssrgen list list * ssrclear) PRINTED BY { pr_ssrdgens } | [ "{" ne_ssrhyp_list(clr) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> { cons_gen (mkclr clr, dt) dgens } | [ "{" ne_ssrhyp_list(clr) "}" ] -> { [[]], clr } | [ "{" ssrocc(occ) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> { cons_gen (mkocc occ, dt) dgens } | [ "/" ssrdgens_tl(dgens) ] -> { cons_dep dgens } | [ cpattern(dt) ssrdgens_tl(dgens) ] -> { cons_gen (nodocc, dt) dgens } | [ ] -> { [[]], [] } END ARGUMENT EXTEND ssrdgens TYPED AS ssrdgens_tl PRINTED BY { pr_ssrdgens } | [ ":" ssrgen(gen) ssrdgens_tl(dgens) ] -> { cons_gen gen dgens } END (** Equations *) (* argument *) { type ssreqid = ssripatrep option let pr_eqid = function Some pat -> str " " ++ pr_ipat pat | None -> mt () let pr_ssreqid _ _ _ = pr_eqid } (* We must use primitive parsing here to avoid conflicts with the *) (* basic move, case, and elim tactics. *) ARGUMENT EXTEND ssreqid TYPED AS ssripatrep option PRINTED BY { pr_ssreqid } END { let accept_ssreqid _ strm = match Util.stream_nth 0 strm with | Tok.IDENT _ -> accept_before_syms [":"] strm | Tok.KEYWORD ":" -> () | Tok.KEYWORD pat when List.mem pat ["_"; "?"; "->"; "<-"] -> accept_before_syms [":"] strm | _ -> raise Stream.Failure let test_ssreqid = Pcoq.Entry.of_parser "test_ssreqid" accept_ssreqid } GRAMMAR EXTEND Gram GLOBAL: ssreqid; ssreqpat: [ [ id = Prim.ident -> { IPatId id } | "_" -> { IPatAnon Drop } | "?" -> { IPatAnon (One None) } | "+" -> { IPatAnon Temporary } | occ = ssrdocc; "->" -> { match occ with | None, occ -> IPatRewrite (occ, L2R) | _ -> CErrors.user_err ~loc (str"Only occurrences are allowed here") } | occ = ssrdocc; "<-" -> { match occ with | None, occ -> IPatRewrite (occ, R2L) | _ -> CErrors.user_err ~loc (str "Only occurrences are allowed here") } | "->" -> { IPatRewrite (allocc, L2R) } | "<-" -> { IPatRewrite (allocc, R2L) } ]]; ssreqid: [ [ test_ssreqid; pat = ssreqpat -> { Some pat } | test_ssreqid -> { None } ]]; END (** Bookkeeping (discharge-intro) argument *) (* Since all bookkeeping ssr commands have the same discharge-intro *) (* argument format we use a single grammar entry point to parse them. *) (* the entry point parses only non-empty arguments to avoid conflicts *) (* with the basic Coq tactics. *) { type ssrarg = ssrfwdview * (ssreqid * (cpattern ssragens * ssripats)) (* type ssrarg = ssrbwdview * (ssreqid * (ssrdgens * ssripats)) *) let pr_ssrarg _ _ _ (view, (eqid, (dgens, ipats))) = let pri = pr_intros (gens_sep dgens) in pr_view2 view ++ pr_eqid eqid ++ pr_dgens pr_gen dgens ++ pri ipats } ARGUMENT EXTEND ssrarg TYPED AS (ssrfwdview * (ssreqid * (ssrdgens * ssrintros))) PRINTED BY { pr_ssrarg } | [ ssrfwdview(view) ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> { view, (eqid, (dgens, ipats)) } | [ ssrfwdview(view) ssrclear(clr) ssrintros(ipats) ] -> { view, (None, (([], clr), ipats)) } | [ ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> { [], (eqid, (dgens, ipats)) } | [ ssrclear_ne(clr) ssrintros(ipats) ] -> { [], (None, (([], clr), ipats)) } | [ ssrintros_ne(ipats) ] -> { [], (None, (([], []), ipats)) } END (** The "clear" tactic *) (* We just add a numeric version that clears the n top assumptions. *) TACTIC EXTEND ssrclear | [ "clear" natural(n) ] -> { tclIPAT (List.init n (fun _ -> IOpDrop)) } END (** The "move" tactic *) { (* TODO: review this, in particular the => _ and => [] cases *) let rec improper_intros = function | IPatSimpl _ :: ipats -> improper_intros ipats | (IPatId _ | IPatAnon _ | IPatCase _ | IPatDispatch _) :: _ -> false | _ -> true (* FIXME *) let check_movearg = function | view, (eqid, _) when view <> [] && eqid <> None -> CErrors.user_err (Pp.str "incompatible view and equation in move tactic") | view, (_, (([gen :: _], _), _)) when view <> [] && has_occ gen -> CErrors.user_err (Pp.str "incompatible view and occurrence switch in move tactic") | _, (_, ((dgens, _), _)) when List.length dgens > 1 -> CErrors.user_err (Pp.str "dependents switch `/' in move tactic") | _, (eqid, (_, ipats)) when eqid <> None && improper_intros ipats -> CErrors.user_err (Pp.str "no proper intro pattern for equation in move tactic") | arg -> arg } ARGUMENT EXTEND ssrmovearg TYPED AS ssrarg PRINTED BY { pr_ssrarg } | [ ssrarg(arg) ] -> { check_movearg arg } END { let movearg_of_parsed_movearg (v,(eq,(dg,ip))) = (v,(eq,(ssrdgens_of_parsed_dgens dg,ip))) } TACTIC EXTEND ssrmove | [ "move" ssrmovearg(arg) ssrrpat(pat) ] -> { ssrmovetac (movearg_of_parsed_movearg arg) <*> tclIPAT (tclCompileIPats [pat]) } | [ "move" ssrmovearg(arg) ssrclauses(clauses) ] -> { tclCLAUSES (ssrmovetac (movearg_of_parsed_movearg arg)) clauses } | [ "move" ssrrpat(pat) ] -> { tclIPAT (tclCompileIPats [pat]) } | [ "move" ] -> { ssrsmovetac } END { let check_casearg = function | view, (_, (([_; gen :: _], _), _)) when view <> [] && has_occ gen -> CErrors.user_err (Pp.str "incompatible view and occurrence switch in dependent case tactic") | arg -> arg } ARGUMENT EXTEND ssrcasearg TYPED AS ssrarg PRINTED BY { pr_ssrarg } | [ ssrarg(arg) ] -> { check_casearg arg } END TACTIC EXTEND ssrcase | [ "case" ssrcasearg(arg) ssrclauses(clauses) ] -> { tclCLAUSES (ssrcasetac (movearg_of_parsed_movearg arg)) clauses } | [ "case" ] -> { ssrscasetoptac } END (** The "elim" tactic *) TACTIC EXTEND ssrelim | [ "elim" ssrarg(arg) ssrclauses(clauses) ] -> { tclCLAUSES (ssrelimtac (movearg_of_parsed_movearg arg)) clauses } | [ "elim" ] -> { ssrselimtoptac } END (** 6. Backward chaining tactics: apply, exact, congr. *) (** The "apply" tactic *) { let pr_agen (docc, dt) = pr_docc docc ++ pr_term dt let pr_ssragen _ _ _ = pr_agen let pr_ssragens _ _ _ = pr_dgens pr_agen } ARGUMENT EXTEND ssragen TYPED AS (ssrdocc * ssrterm) PRINTED BY { pr_ssragen } | [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ] -> { mkclr clr, dt } | [ ssrterm(dt) ] -> { nodocc, dt } END ARGUMENT EXTEND ssragens TYPED AS (ssragen list list * ssrclear) PRINTED BY { pr_ssragens } | [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ssragens(agens) ] -> { cons_gen (mkclr clr, dt) agens } | [ "{" ne_ssrhyp_list(clr) "}" ] -> { [[]], clr} | [ ssrterm(dt) ssragens(agens) ] -> { cons_gen (nodocc, dt) agens } | [ ] -> { [[]], [] } END { let mk_applyarg views agens intros = views, (agens, intros) let pr_ssraarg _ _ _ (view, (dgens, ipats)) = let pri = pr_intros (gens_sep dgens) in pr_view view ++ pr_dgens pr_agen dgens ++ pri ipats } ARGUMENT EXTEND ssrapplyarg TYPED AS (ssrbwdview * (ssragens * ssrintros)) PRINTED BY { pr_ssraarg } | [ ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> { mk_applyarg [] (cons_gen gen dgens) intros } | [ ssrclear_ne(clr) ssrintros(intros) ] -> { mk_applyarg [] ([], clr) intros } | [ ssrintros_ne(intros) ] -> { mk_applyarg [] ([], []) intros } | [ ssrbwdview(view) ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> { mk_applyarg view (cons_gen gen dgens) intros } | [ ssrbwdview(view) ssrclear(clr) ssrintros(intros) ] -> { mk_applyarg view ([], clr) intros } END TACTIC EXTEND ssrapply | [ "apply" ssrapplyarg(arg) ] -> { let views, (gens_clr, intros) = arg in inner_ssrapplytac views gens_clr ist <*> tclIPATssr intros } | [ "apply" ] -> { apply_top_tac } END (** The "exact" tactic *) { let mk_exactarg views dgens = mk_applyarg views dgens [] } ARGUMENT EXTEND ssrexactarg TYPED AS ssrapplyarg PRINTED BY { pr_ssraarg } | [ ":" ssragen(gen) ssragens(dgens) ] -> { mk_exactarg [] (cons_gen gen dgens) } | [ ssrbwdview(view) ssrclear(clr) ] -> { mk_exactarg view ([], clr) } | [ ssrclear_ne(clr) ] -> { mk_exactarg [] ([], clr) } END { let vmexacttac pf = Goal.enter begin fun gl -> exact_no_check (EConstr.mkCast (pf, _vmcast, Tacmach.New.pf_concl gl)) end } TACTIC EXTEND ssrexact | [ "exact" ssrexactarg(arg) ] -> { let views, (gens_clr, _) = arg in V82.tactic (tclBY (V82.of_tactic (inner_ssrapplytac views gens_clr ist))) } | [ "exact" ] -> { V82.tactic (Tacticals.tclORELSE (donetac ~-1) (tclBY (V82.of_tactic apply_top_tac))) } | [ "exact" "<:" lconstr(pf) ] -> { vmexacttac pf } END (** The "congr" tactic *) { let pr_ssrcongrarg _ _ _ ((n, f), dgens) = (if n <= 0 then mt () else str " " ++ int n) ++ str " " ++ pr_term f ++ pr_dgens pr_gen dgens } ARGUMENT EXTEND ssrcongrarg TYPED AS ((int * ssrterm) * ssrdgens) PRINTED BY { pr_ssrcongrarg } | [ natural(n) constr(c) ssrdgens(dgens) ] -> { (n, mk_term xNoFlag c), dgens } | [ natural(n) constr(c) ] -> { (n, mk_term xNoFlag c),([[]],[]) } | [ constr(c) ssrdgens(dgens) ] -> { (0, mk_term xNoFlag c), dgens } | [ constr(c) ] -> { (0, mk_term xNoFlag c), ([[]],[]) } END TACTIC EXTEND ssrcongr | [ "congr" ssrcongrarg(arg) ] -> { let arg, dgens = arg in V82.tactic begin match dgens with | [gens], clr -> Tacticals.tclTHEN (genstac (gens,clr)) (newssrcongrtac arg ist) | _ -> errorstrm (str"Dependent family abstractions not allowed in congr") end } END (** 7. Rewriting tactics (rewrite, unlock) *) (** Coq rewrite compatibility flag *) (** Rewrite clear/occ switches *) { let pr_rwocc = function | None, None -> mt () | None, occ -> pr_occ occ | Some clr, _ -> pr_clear_ne clr let pr_ssrrwocc _ _ _ = pr_rwocc } ARGUMENT EXTEND ssrrwocc TYPED AS ssrdocc PRINTED BY { pr_ssrrwocc } | [ "{" ssrhyp_list(clr) "}" ] -> { mkclr clr } | [ "{" ssrocc(occ) "}" ] -> { mkocc occ } | [ ] -> { noclr } END (** Rewrite rules *) { let pr_rwkind = function | RWred s -> pr_simpl s | RWdef -> str "/" | RWeq -> mt () let wit_ssrrwkind = add_genarg "ssrrwkind" (fun env sigma -> pr_rwkind) let pr_rule = function | RWred s, _ -> pr_simpl s | RWdef, r-> str "/" ++ pr_term r | RWeq, r -> pr_term r let pr_ssrrule _ _ _ = pr_rule let noruleterm loc = mk_term xNoFlag (mkCProp loc) } ARGUMENT EXTEND ssrrule_ne TYPED AS (ssrrwkind * ssrterm) PRINTED BY { pr_ssrrule } END GRAMMAR EXTEND Gram GLOBAL: ssrrule_ne; ssrrule_ne : [ [ test_not_ssrslashnum; x = [ "/"; t = ssrterm -> { RWdef, t } | t = ssrterm -> { RWeq, t } | s = ssrsimpl_ne -> { RWred s, noruleterm (Some loc) } ] -> { x } | s = ssrsimpl_ne -> { RWred s, noruleterm (Some loc) } ]]; END ARGUMENT EXTEND ssrrule TYPED AS ssrrule_ne PRINTED BY { pr_ssrrule } | [ ssrrule_ne(r) ] -> { r } | [ ] -> { RWred Nop, noruleterm (Some loc) } END (** Rewrite arguments *) { let pr_option f = function None -> mt() | Some x -> f x let pr_pattern_squarep= pr_option (fun r -> str "[" ++ pr_rpattern r ++ str "]") let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep let pr_rwarg ((d, m), ((docc, rx), r)) = pr_rwdir d ++ pr_mult m ++ pr_rwocc docc ++ pr_pattern_squarep rx ++ pr_rule r let pr_ssrrwarg _ _ _ = pr_rwarg } ARGUMENT EXTEND ssrpattern_squarep TYPED AS rpattern option PRINTED BY { pr_ssrpattern_squarep } | [ "[" rpattern(rdx) "]" ] -> { Some rdx } | [ ] -> { None } END ARGUMENT EXTEND ssrpattern_ne_squarep TYPED AS rpattern option PRINTED BY { pr_ssrpattern_squarep } | [ "[" rpattern(rdx) "]" ] -> { Some rdx } END ARGUMENT EXTEND ssrrwarg TYPED AS ((ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule)) PRINTED BY { pr_ssrrwarg } | [ "-" ssrmult(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> { mk_rwarg (R2L, m) (docc, rx) r } | [ "-/" ssrterm(t) ] -> (* just in case '-/' should become a token *) { mk_rwarg (R2L, nomult) norwocc (RWdef, t) } | [ ssrmult_ne(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> { mk_rwarg (L2R, m) (docc, rx) r } | [ "{" ne_ssrhyp_list(clr) "}" ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> { mk_rwarg norwmult (mkclr clr, rx) r } | [ "{" ne_ssrhyp_list(clr) "}" ssrrule(r) ] -> { mk_rwarg norwmult (mkclr clr, None) r } | [ "{" ssrocc(occ) "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> { mk_rwarg norwmult (mkocc occ, rx) r } | [ "{" "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> { mk_rwarg norwmult (nodocc, rx) r } | [ ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> { mk_rwarg norwmult (noclr, rx) r } | [ ssrrule_ne(r) ] -> { mk_rwarg norwmult norwocc r } END TACTIC EXTEND ssrinstofruleL2R | [ "ssrinstancesofruleL2R" ssrterm(arg) ] -> { V82.tactic (ssrinstancesofrule ist L2R arg) } END TACTIC EXTEND ssrinstofruleR2L | [ "ssrinstancesofruleR2L" ssrterm(arg) ] -> { V82.tactic (ssrinstancesofrule ist R2L arg) } END (** Rewrite argument sequence *) (* type ssrrwargs = ssrrwarg list *) { let pr_ssrrwargs _ _ _ rwargs = pr_list spc pr_rwarg rwargs } ARGUMENT EXTEND ssrrwargs TYPED AS ssrrwarg list PRINTED BY { pr_ssrrwargs } END { let ssr_rw_syntax = Summary.ref ~name:"SSR:rewrite" true let () = Goptions.(declare_bool_option { optname = "ssreflect rewrite"; optkey = ["SsrRewrite"]; optread = (fun _ -> !ssr_rw_syntax); optdepr = false; optwrite = (fun b -> ssr_rw_syntax := b) }) let lbrace = Char.chr 123 (** Workaround to a limitation of coqpp *) let test_ssr_rw_syntax = let test _ strm = if not !ssr_rw_syntax then raise Stream.Failure else if is_ssr_loaded () then () else match Util.stream_nth 0 strm with | Tok.KEYWORD key when List.mem key.[0] [lbrace; '['; '/'] -> () | _ -> raise Stream.Failure in Pcoq.Entry.of_parser "test_ssr_rw_syntax" test } GRAMMAR EXTEND Gram GLOBAL: ssrrwargs; ssrrwargs: [[ test_ssr_rw_syntax; a = LIST1 ssrrwarg -> { a } ]]; END (** The "rewrite" tactic *) TACTIC EXTEND ssrrewrite | [ "rewrite" ssrrwargs(args) ssrclauses(clauses) ] -> { tclCLAUSES (old_tac (ssrrewritetac ist args)) clauses } END (** The "unlock" tactic *) { let pr_unlockarg (occ, t) = pr_occ occ ++ pr_term t let pr_ssrunlockarg _ _ _ = pr_unlockarg } ARGUMENT EXTEND ssrunlockarg TYPED AS (ssrocc * ssrterm) PRINTED BY { pr_ssrunlockarg } | [ "{" ssrocc(occ) "}" ssrterm(t) ] -> { occ, t } | [ ssrterm(t) ] -> { None, t } END { let pr_ssrunlockargs _ _ _ args = pr_list spc pr_unlockarg args } ARGUMENT EXTEND ssrunlockargs TYPED AS ssrunlockarg list PRINTED BY { pr_ssrunlockargs } | [ ssrunlockarg_list(args) ] -> { args } END TACTIC EXTEND ssrunlock | [ "unlock" ssrunlockargs(args) ssrclauses(clauses) ] -> { tclCLAUSES (old_tac (unlocktac ist args)) clauses } END (** 8. Forward chaining tactics (pose, set, have, suffice, wlog) *) TACTIC EXTEND ssrpose | [ "pose" ssrfixfwd(ffwd) ] -> { V82.tactic (ssrposetac ffwd) } | [ "pose" ssrcofixfwd(ffwd) ] -> { V82.tactic (ssrposetac ffwd) } | [ "pose" ssrfwdid(id) ssrposefwd(fwd) ] -> { V82.tactic (ssrposetac (id, fwd)) } END (** The "set" tactic *) (* type ssrsetfwd = ssrfwd * ssrdocc *) TACTIC EXTEND ssrset | [ "set" ssrfwdid(id) ssrsetfwd(fwd) ssrclauses(clauses) ] -> { tclCLAUSES (old_tac (ssrsettac id fwd)) clauses } END (** The "have" tactic *) (* type ssrhavefwd = ssrfwd * ssrhint *) (* Pltac. *) (* The standard TACTIC EXTEND does not work for abstract *) GRAMMAR EXTEND Gram GLOBAL: tactic_expr; tactic_expr: LEVEL "3" [ RIGHTA [ IDENT "abstract"; gens = ssrdgens -> { ssrtac_expr ~loc "abstract" [Tacexpr.TacGeneric (Genarg.in_gen (Genarg.rawwit wit_ssrdgens) gens)] } ]]; END TACTIC EXTEND ssrabstract | [ "abstract" ssrdgens(gens) ] -> { if List.length (fst gens) <> 1 then errorstrm (str"dependents switches '/' not allowed here"); Ssripats.ssrabstract (ssrdgens_of_parsed_dgens gens) } END TACTIC EXTEND ssrhave | [ "have" ssrhavefwdwbinders(fwd) ] -> { V82.tactic (havetac ist fwd false false) } END TACTIC EXTEND ssrhavesuff | [ "have" "suff" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> { V82.tactic (havetac ist (false,(pats,fwd)) true false) } END TACTIC EXTEND ssrhavesuffices | [ "have" "suffices" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> { V82.tactic (havetac ist (false,(pats,fwd)) true false) } END TACTIC EXTEND ssrsuffhave | [ "suff" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> { V82.tactic (havetac ist (false,(pats,fwd)) true true) } END TACTIC EXTEND ssrsufficeshave | [ "suffices" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> { V82.tactic (havetac ist (false,(pats,fwd)) true true) } END (** The "suffice" tactic *) { let pr_ssrsufffwdwbinders env sigma _ _ prt (hpats, (fwd, hint)) = pr_hpats hpats ++ pr_fwd fwd ++ pr_hint env sigma prt hint } ARGUMENT EXTEND ssrsufffwd TYPED AS (ssrhpats * (ssrfwd * ssrhint)) PRINTED BY { pr_ssrsufffwdwbinders env sigma } | [ ssrhpats(pats) ssrbinder_list(bs) ":" ast_closure_lterm(t) ssrhint(hint) ] -> { let ((clr, pats), binders), simpl = pats in let allbs = intro_id_to_binder binders @ bs in let allbinders = binders @ List.flatten (binder_to_intro_id bs) in let fwd = mkFwdHint ":" t in (((clr, pats), allbinders), simpl), (bind_fwd allbs fwd, hint) } END TACTIC EXTEND ssrsuff | [ "suff" ssrsufffwd(fwd) ] -> { V82.tactic (sufftac ist fwd) } END TACTIC EXTEND ssrsuffices | [ "suffices" ssrsufffwd(fwd) ] -> { V82.tactic (sufftac ist fwd) } END (** The "wlog" (Without Loss Of Generality) tactic *) (* type ssrwlogfwd = ssrwgen list * ssrfwd *) { let pr_ssrwlogfwd _ _ _ (gens, t) = str ":" ++ pr_list mt pr_wgen gens ++ spc() ++ pr_fwd t } ARGUMENT EXTEND ssrwlogfwd TYPED AS (ssrwgen list * ssrfwd) PRINTED BY { pr_ssrwlogfwd } | [ ":" ssrwgen_list(gens) "/" ast_closure_lterm(t) ] -> { gens, mkFwdHint "/" t} END TACTIC EXTEND ssrwlog | [ "wlog" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { V82.tactic (wlogtac ist pats fwd hint false `NoGen) } END TACTIC EXTEND ssrwlogs | [ "wlog" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } END TACTIC EXTEND ssrwlogss | [ "wlog" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } END TACTIC EXTEND ssrwithoutloss | [ "without" "loss" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { V82.tactic (wlogtac ist pats fwd hint false `NoGen) } END TACTIC EXTEND ssrwithoutlosss | [ "without" "loss" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } END TACTIC EXTEND ssrwithoutlossss | [ "without" "loss" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } END { (* Generally have *) let pr_idcomma _ _ _ = function | None -> mt() | Some None -> str"_, " | Some (Some id) -> pr_id id ++ str", " } ARGUMENT EXTEND ssr_idcomma TYPED AS ident option option PRINTED BY { pr_idcomma } | [ ] -> { None } END { let accept_idcomma _ strm = match stream_nth 0 strm with | Tok.IDENT _ | Tok.KEYWORD "_" -> accept_before_syms [","] strm | _ -> raise Stream.Failure let test_idcomma = Pcoq.Entry.of_parser "test_idcomma" accept_idcomma } GRAMMAR EXTEND Gram GLOBAL: ssr_idcomma; ssr_idcomma: [ [ test_idcomma; ip = [ id = IDENT -> { Some (Id.of_string id) } | "_" -> { None } ]; "," -> { Some ip } ] ]; END { let augment_preclr clr1 (((clr0, x),y),z) = let cl = match clr0 with | None -> if clr1 = [] then None else Some clr1 | Some clr0 -> Some (clr1 @ clr0) in (((cl, x),y),z) } TACTIC EXTEND ssrgenhave | [ "gen" "have" ssrclear(clr) ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { let pats = augment_preclr clr pats in V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) } END TACTIC EXTEND ssrgenhave2 | [ "generally" "have" ssrclear(clr) ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { let pats = augment_preclr clr pats in V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) } END { let check_under_arg ((_dir,mult),((_occ,_rpattern),_rule)) = if mult <> nomult then CErrors.user_err Pp.(str"under does not support multipliers") } TACTIC EXTEND under | [ "under" ssrrwarg(arg) ] -> { check_under_arg arg; Ssrfwd.undertac ist None arg nohint } | [ "under" ssrrwarg(arg) ssrintros_ne(ipats) ] -> { check_under_arg arg; Ssrfwd.undertac ist (Some ipats) arg nohint } | [ "under" ssrrwarg(arg) ssrintros_ne(ipats) "do" ssrhint3arg(h) ] -> { check_under_arg arg; Ssrfwd.undertac ist (Some ipats) arg h } | [ "under" ssrrwarg(arg) "do" ssrhint3arg(h) ] -> { (* implicit "=> [*|*]" *) check_under_arg arg; Ssrfwd.undertac ~pad_intro:true ist (Some [IPatAnon All]) arg h } END { (* We wipe out all the keywords generated by the grammar rules we defined. *) (* The user is supposed to Require Import ssreflect or Require ssreflect *) (* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) (* consequence the extended ssreflect grammar. *) let () = CLexer.set_keyword_state frozen_lexer ;; } (* vim: set filetype=ocaml foldmethod=marker: *) coq-8.11.0/plugins/ssr/ssrunder.v0000644000175000017500000000564713612664533016626 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* .doc { font-family: monospace; white-space: pre; } # **) (** Constants for under/over, to rewrite under binders using "context lemmas" Note: this file does not require [ssreflect]; it is both required by [ssrsetoid] and *exported* by [ssrunder]. This preserves the following feature: we can use [Setoid] without requiring [ssreflect] and use [ssreflect] without requiring [Setoid]. *) Require Import ssrclasses. Module Type UNDER_REL. Parameter Under_rel : forall (A : Type) (eqA : A -> A -> Prop), A -> A -> Prop. Parameter Under_rel_from_rel : forall (A : Type) (eqA : A -> A -> Prop) (x y : A), @Under_rel A eqA x y -> eqA x y. Parameter Under_relE : forall (A : Type) (eqA : A -> A -> Prop), @Under_rel A eqA = eqA. (** [Over_rel, over_rel, over_rel_done]: for "by rewrite over_rel" *) Parameter Over_rel : forall (A : Type) (eqA : A -> A -> Prop), A -> A -> Prop. Parameter over_rel : forall (A : Type) (eqA : A -> A -> Prop) (x y : A), @Under_rel A eqA x y = @Over_rel A eqA x y. Parameter over_rel_done : forall (A : Type) (eqA : A -> A -> Prop) (EeqA : Reflexive eqA) (x : A), @Over_rel A eqA x x. (** [under_rel_done]: for Ltac-style over *) Parameter under_rel_done : forall (A : Type) (eqA : A -> A -> Prop) (EeqA : Reflexive eqA) (x : A), @Under_rel A eqA x x. Notation "''Under[' x ]" := (@Under_rel _ _ x _) (at level 8, format "''Under[' x ]", only printing). End UNDER_REL. Module Export Under_rel : UNDER_REL. Definition Under_rel (A : Type) (eqA : A -> A -> Prop) := eqA. Lemma Under_rel_from_rel : forall (A : Type) (eqA : A -> A -> Prop) (x y : A), @Under_rel A eqA x y -> eqA x y. Proof. now trivial. Qed. Lemma Under_relE (A : Type) (eqA : A -> A -> Prop) : @Under_rel A eqA = eqA. Proof. now trivial. Qed. Definition Over_rel := Under_rel. Lemma over_rel : forall (A : Type) (eqA : A -> A -> Prop) (x y : A), @Under_rel A eqA x y = @Over_rel A eqA x y. Proof. now trivial. Qed. Lemma over_rel_done : forall (A : Type) (eqA : A -> A -> Prop) (EeqA : Reflexive eqA) (x : A), @Over_rel A eqA x x. Proof. now unfold Over_rel. Qed. Lemma under_rel_done : forall (A : Type) (eqA : A -> A -> Prop) (EeqA : Reflexive eqA) (x : A), @Under_rel A eqA x x. Proof. now trivial. Qed. End Under_rel. coq-8.11.0/plugins/ssr/ssrprinters.ml0000644000175000017500000001270713612664533017515 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* hd | x :: xs -> hd ++ List.fold_left (fun acc x -> acc ++ sep ++ x) x xs let pp_term gl t = let t = Reductionops.nf_evar (project gl) t in pr_econstr_env (pf_env gl) (project gl) t (* FIXME *) (* terms are pre constr, the kind is parsing/printing flag to distinguish * between x, @x and (x). It affects automatic clear and let-in preservation. * Cpattern is a temporary flag that becomes InParens ASAP. *) (* type ssrtermkind = InParens | WithAt | NoFlag | Cpattern *) let xInParens = '(' let xWithAt = '@' let xNoFlag = ' ' let xCpattern = 'x' (* Term printing utilities functions for deciding bracketing. *) let pr_paren prx x = hov 1 (str "(" ++ prx x ++ str ")") (* String lexing utilities *) let skip_wschars s = let rec loop i = match s.[i] with '\n'..' ' -> loop (i + 1) | _ -> i in loop (* We also guard characters that might interfere with the ssreflect *) (* tactic syntax. *) let guard_term ch1 s i = match s.[i] with | '(' -> false | '{' | '/' | '=' -> true | _ -> ch1 = xInParens (* We also guard characters that might interfere with the ssreflect *) (* tactic syntax. *) let pr_guarded guard prc c = pp_with Format.str_formatter (prc c); let s = Format.flush_str_formatter () ^ "$" in if guard s (skip_wschars s 0) then pr_paren prc c else prc c let prl_constr_expr = let env = Global.env () in let sigma = Evd.from_env env in Ppconstr.pr_lconstr_expr env sigma let pr_glob_constr c = Printer.pr_glob_constr_env (Global.env ()) c let prl_glob_constr c = Printer.pr_lglob_constr_env (Global.env ()) c let pr_glob_constr_and_expr = function | _, Some c -> let env = Global.env () in let sigma = Evd.from_env env in Ppconstr.pr_constr_expr env sigma c | c, None -> pr_glob_constr c let pr_term (k, c) = pr_guarded (guard_term k) pr_glob_constr_and_expr c let pr_hyp (SsrHyp (_, id)) = Id.print id let pr_hyps = pr_list pr_spc pr_hyp let pr_occ = function | Some (true, occ) -> str "{-" ++ pr_list pr_spc int occ ++ str "}" | Some (false, occ) -> str "{+" ++ pr_list pr_spc int occ ++ str "}" | None -> str "{}" let pr_clear_ne clr = str "{" ++ pr_hyps clr ++ str "}" let pr_clear sep clr = sep () ++ pr_clear_ne clr let pr_dir = function L2R -> str "->" | R2L -> str "<-" let pr_simpl = function | Simpl -1 -> str "/=" | Cut -1 -> str "//" | Simpl n -> str "/" ++ int n ++ str "=" | Cut n -> str "/" ++ int n ++ str"/" | SimplCut (-1,-1) -> str "//=" | SimplCut (n,-1) -> str "/" ++ int n ++ str "/=" | SimplCut (-1,n) -> str "//" ++ int n ++ str "=" | SimplCut (n,m) -> str "/" ++ int n ++ str "/" ++ int m ++ str "=" | Nop -> mt () (* New terms *) let pr_ast_closure_term { body } = let env = Global.env () in let sigma = Evd.from_env env in Ppconstr.pr_constr_expr env sigma body let pr_view2 = pr_list mt (fun c -> str "/" ++ pr_ast_closure_term c) let rec pr_ipat p = match p with | IPatId id -> Id.print id | IPatSimpl sim -> pr_simpl sim | IPatClear clr -> pr_clear mt clr | IPatCase (Regular iorpat) -> hov 1 (str "[" ++ pr_iorpat iorpat ++ str "]") | IPatCase (Block m) -> hov 1 (str"[" ++ pr_block m ++ str"]") | IPatDispatch(Regular iorpat) -> hov 1 (str "(" ++ pr_iorpat iorpat ++ str ")") | IPatDispatch (Block m) -> hov 1 (str"(" ++ pr_block m ++ str")") | IPatInj iorpat -> hov 1 (str "[=" ++ pr_iorpat iorpat ++ str "]") | IPatRewrite (occ, dir) -> pr_occ occ ++ pr_dir dir | IPatAnon All -> str "*" | IPatAnon Drop -> str "_" | IPatAnon (One _) -> str "?" | IPatView v -> pr_view2 v | IPatAnon Temporary -> str "+" | IPatNoop -> str "-" | IPatAbstractVars l -> str "[:" ++ pr_list spc Id.print l ++ str "]" | IPatFastNondep -> str">" and pr_ipats ipats = pr_list spc pr_ipat ipats and pr_iorpat iorpat = pr_list pr_bar pr_ipats iorpat and pr_block = function (Prefix id) -> str"^" ++ Id.print id | (SuffixId id) -> str"^~" ++ Id.print id | (SuffixNum n) -> str"^~" ++ int n (* 0 cost pp function. Active only if Debug Ssreflect is Set *) let ppdebug_ref = ref (fun _ -> ()) let ssr_pp s = Feedback.msg_debug (str"SSR: "++Lazy.force s) let () = Goptions.(declare_bool_option { optname = "ssreflect debugging"; optkey = ["Debug";"Ssreflect"]; optdepr = false; optread = (fun _ -> !ppdebug_ref == ssr_pp); optwrite = (fun b -> Ssrmatching.debug b; if b then ppdebug_ref := ssr_pp else ppdebug_ref := fun _ -> ()) }) let ppdebug s = !ppdebug_ref s coq-8.11.0/plugins/ssr/ssreflect_plugin.mlpack0000644000175000017500000000016113612664533021315 0ustar treinentreinenSsrast Ssrprinters Ssrcommon Ssrtacticals Ssrelim Ssrview Ssrbwd Ssrequality Ssripats Ssrfwd Ssrparser Ssrvernac coq-8.11.0/plugins/ssr/ssrfwd.ml0000644000175000017500000005143113612664533016424 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ist, Ssrcommon.ssrterm_of_ast_closure_term t | None -> assert false in let sigma, t, ucst, _ = pf_abs_ssrterm ist gl t in posetac id t (pf_merge_uc ucst gl) let ssrsettac id ((_, (pat, pty)), (_, occ)) gl = let pty = Option.map (fun { Ssrast.body; interp_env } -> let ist = Option.get interp_env in (mkRHole, Some body), ist) pty in let pat = interp_cpattern gl pat pty in let cl, sigma, env = pf_concl gl, project gl, pf_env gl in let (c, ucst), cl = let cl = EConstr.Unsafe.to_constr cl in try fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1 with NoMatch -> redex_of_pattern ~resolve_typeclasses:true env pat, cl in let gl = pf_merge_uc ucst gl in let c = EConstr.of_constr c in let cl = EConstr.of_constr cl in if Termops.occur_existential sigma c then errorstrm(str"The pattern"++spc()++ pr_econstr_pat env sigma c++spc()++str"did not match and has holes."++spc()++ str"Did you mean pose?") else let c, (gl, cty) = match EConstr.kind sigma c with | Cast(t, DEFAULTcast, ty) -> t, (gl, ty) | _ -> c, pfe_type_of gl c in let cl' = EConstr.mkLetIn (make_annot (Name id) Sorts.Relevant, c, cty, cl) in Tacticals.tclTHEN (Proofview.V82.of_tactic (convert_concl ~check:true cl')) (introid id) gl open Util open Printer open Ssrast open Ssripats let ssrhaveNOtcresolution = Summary.ref ~name:"SSR:havenotcresolution" false let () = Goptions.(declare_bool_option { optname = "have type classes"; optkey = ["SsrHave";"NoTCResolution"]; optread = (fun _ -> !ssrhaveNOtcresolution); optdepr = false; optwrite = (fun b -> ssrhaveNOtcresolution := b); }) open Constrexpr open Glob_term let combineCG t1 t2 f g = match t1, t2 with | (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None) | (x, (_, Some t1)), (_, (_, Some t2)) -> x, (mkRHole, Some (f t1 t2)) | _, (_, (_, None)) -> anomaly "have: mixed C-G constr" | _ -> anomaly "have: mixed G-C constr" let basecuttac name c gl = let hd, gl = pf_mkSsrConst name gl in let t = EConstr.mkApp (hd, [|c|]) in let gl, _ = pf_e_type_of gl t in Proofview.V82.of_tactic (Tactics.apply t) gl let introstac ipats = Proofview.V82.of_tactic (tclIPAT ipats) let havetac ist (transp,((((clr, orig_pats), binders), simpl), (((fk, _), t), hint))) suff namefst gl = let concl = pf_concl gl in let pats = tclCompileIPats orig_pats in let binders = tclCompileIPats binders in let simpl = tclCompileIPats simpl in let skols, pats = List.partition (function IOpAbstractVars _ -> true | _ -> false) pats in let itac_mkabs = introstac skols in let itac_c, clr = match clr with | None -> introstac pats, [] | Some clr -> introstac (tclCompileIPats (IPatClear clr :: orig_pats)), clr in let itac, id, clr = introstac pats, Tacticals.tclIDTAC, old_cleartac clr in let binderstac n = let rec aux = function 0 -> [] | n -> IOpInaccessible None :: aux (n-1) in Tacticals.tclTHEN (if binders <> [] then introstac (aux n) else Tacticals.tclIDTAC) (introstac binders) in let simpltac = introstac simpl in let fixtc = not !ssrhaveNOtcresolution && match fk with FwdHint(_,true) -> false | _ -> true in let hint = hinttac ist true hint in let cuttac t gl = if transp then let have_let, gl = pf_mkSsrConst "ssr_have_let" gl in let step = EConstr.mkApp (have_let, [|concl;t|]) in let gl, _ = pf_e_type_of gl step in applyn ~with_evars:true ~with_shelve:false 2 step gl else basecuttac "ssr_have" t gl in (* Introduce now abstract constants, so that everything sees them *) let abstract_key, gl = pf_mkSsrConst "abstract_key" gl in let unlock_abs (idty,args_id) gl = let gl, _ = pf_e_type_of gl idty in pf_unify_HO gl args_id.(2) abstract_key in Tacticals.tclTHENFIRST itac_mkabs (fun gl -> let mkt t = mk_term xNoFlag t in let mkl t = (xNoFlag, (t, None)) in let interp gl rtc t = pf_abs_ssrterm ~resolve_typeclasses:rtc ist gl t in let interp_ty gl rtc t = let a,b,_,u = pf_interp_ty ~resolve_typeclasses:rtc ist gl t in a,b,u in let open CAst in let ct, cty, hole, loc = match Ssrcommon.ssrterm_of_ast_closure_term t with | _, (_, Some { loc; v = CCast (ct, CastConv cty)}) -> mkt ct, mkt cty, mkt (mkCHole None), loc | _, (_, Some ct) -> mkt ct, mkt (mkCHole None), mkt (mkCHole None), None | _, (t, None) -> begin match DAst.get t with | GCast (ct, CastConv cty) -> mkl ct, mkl cty, mkl mkRHole, t.CAst.loc | _ -> mkl t, mkl mkRHole, mkl mkRHole, None end in let gl, cut, sol, itac1, itac2 = match fk, namefst, suff with | FwdHave, true, true -> errorstrm (str"Suff have does not accept a proof term") | FwdHave, false, true -> let cty = combineCG cty hole (mkCArrow ?loc) mkRArrow in let _,t,uc,_ = interp gl false (combineCG ct cty (mkCCast ?loc) mkRCast) in let gl = pf_merge_uc uc gl in let gl, ty = pfe_type_of gl t in let ctx, _ = EConstr.decompose_prod_n_assum (project gl) 1 ty in let assert_is_conv gl = try Proofview.V82.of_tactic (convert_concl ~check:true (EConstr.it_mkProd_or_LetIn concl ctx)) gl with _ -> errorstrm (str "Given proof term is not of type " ++ pr_econstr_env (pf_env gl) (project gl) (EConstr.mkArrow (EConstr.mkVar (Id.of_string "_")) Sorts.Relevant concl)) in gl, ty, Tacticals.tclTHEN assert_is_conv (Proofview.V82.of_tactic (Tactics.apply t)), id, itac_c | FwdHave, false, false -> let skols = List.flatten (List.map (function | IOpAbstractVars ids -> ids | _ -> assert false) skols) in let skols_args = List.map (fun id -> Ssripats.Internal.examine_abstract (EConstr.mkVar id) gl) skols in let gl = List.fold_right unlock_abs skols_args gl in let sigma, t, uc, n_evars = interp gl false (combineCG ct cty (mkCCast ?loc) mkRCast) in if skols <> [] && n_evars <> 0 then CErrors.user_err (Pp.strbrk @@ "Automatic generalization of unresolved implicit "^ "arguments together with abstract variables is "^ "not supported"); let gl = re_sig (sig_it gl) (Evd.merge_universe_context sigma uc) in let gs = List.map (fun (_,a) -> Ssripats.Internal.pf_find_abstract_proof false gl (EConstr.Unsafe.to_constr a.(1))) skols_args in let tacopen_skols gl = re_sig (gs @ [gl.Evd.it]) gl.Evd.sigma in let gl, ty = pf_e_type_of gl t in gl, ty, Proofview.V82.of_tactic (Tactics.apply t), id, Tacticals.tclTHEN (Tacticals.tclTHEN itac_c simpltac) (Tacticals.tclTHEN tacopen_skols (fun gl -> let abstract, gl = pf_mkSsrConst "abstract" gl in Proofview.V82.of_tactic (unfold [abstract; abstract_key]) gl)) | _,true,true -> let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in gl, EConstr.mkArrow ty Sorts.Relevant concl, hint, itac, clr | _,false,true -> let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in gl, EConstr.mkArrow ty Sorts.Relevant concl, hint, id, itac_c | _, false, false -> let n, cty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in gl, cty, Tacticals.tclTHEN (binderstac n) hint, id, Tacticals.tclTHEN itac_c simpltac | _, true, false -> assert false in Tacticals.tclTHENS (cuttac cut) [ Tacticals.tclTHEN sol itac1; itac2 ] gl) gl ;; let destProd_or_LetIn sigma c = match EConstr.kind sigma c with | Prod (n,ty,c) -> RelDecl.LocalAssum (n, ty), c | LetIn (n,bo,ty,c) -> RelDecl.LocalDef (n, bo, ty), c | _ -> raise DestKO let wlogtac ist (((clr0, pats),_),_) (gens, ((_, ct))) hint suff ghave gl = let clr0 = Option.default [] clr0 in let pats = tclCompileIPats pats in let mkabs gen = abs_wgen false (fun x -> x) gen in let mkclr gen clrs = clr_of_wgen gen clrs in let mkpats = function | _, Some ((x, _), _) -> fun pats -> IOpId (hoi_id x) :: pats | _ -> fun x -> x in let ct = match Ssrcommon.ssrterm_of_ast_closure_term ct with | (a, (b, Some ct)) -> begin match ct.CAst.v with | CCast (_, CastConv cty) -> a, (b, Some cty) | _ -> anomaly "wlog: ssr cast hole deleted by typecheck" end | (a, (t, None)) -> begin match DAst.get t with | GCast (_, CastConv cty) -> a, (cty, None) | _ -> anomaly "wlog: ssr cast hole deleted by typecheck" end in let cut_implies_goal = not (suff || ghave <> `NoGen) in let c, args, ct, gl = let gens = List.filter (function _, Some _ -> true | _ -> false) gens in let concl = pf_concl gl in let c = EConstr.mkProp in let c = if cut_implies_goal then EConstr.mkArrow c Sorts.Relevant concl else c in let gl, args, c = List.fold_right mkabs gens (gl,[],c) in let env, _ = List.fold_left (fun (env, c) _ -> let rd, c = destProd_or_LetIn (project gl) c in EConstr.push_rel rd env, c) (pf_env gl, c) gens in let sigma = project gl in let (sigma, ev) = Evarutil.new_evar env sigma EConstr.mkProp in let k, _ = EConstr.destEvar sigma ev in let fake_gl = {Evd.it = k; Evd.sigma = sigma} in let _, ct, _, uc = pf_interp_ty ist fake_gl ct in let rec var2rel c g s = match EConstr.kind sigma c, g with | Prod({binder_name=Anonymous} as x,_,c), [] -> EConstr.mkProd(x, EConstr.Vars.subst_vars s ct, c) | Sort _, [] -> EConstr.Vars.subst_vars s ct | LetIn({binder_name=Name id} as n,b,ty,c), _::g -> EConstr.mkLetIn (n,b,ty,var2rel c g (id::s)) | Prod({binder_name=Name id} as n,ty,c), _::g -> EConstr.mkProd (n,ty,var2rel c g (id::s)) | _ -> CErrors.anomaly(str"SSR: wlog: var2rel: " ++ pr_econstr_env env sigma c) in let c = var2rel c gens [] in let rec pired c = function | [] -> c | t::ts as args -> match EConstr.kind sigma c with | Prod(_,_,c) -> pired (EConstr.Vars.subst1 t c) ts | LetIn(id,b,ty,c) -> EConstr.mkLetIn (id,b,ty,pired c args) | _ -> CErrors.anomaly(str"SSR: wlog: pired: " ++ pr_econstr_env env sigma c) in c, args, pired c args, pf_merge_uc uc gl in let tacipat pats = introstac pats in let tacigens = Tacticals.tclTHEN (Tacticals.tclTHENLIST(List.rev(List.fold_right mkclr gens [old_cleartac clr0]))) (introstac (List.fold_right mkpats gens [])) in let hinttac = hinttac ist true hint in let cut_kind, fst_goal_tac, snd_goal_tac = match suff, ghave with | true, `NoGen -> "ssr_wlog", Tacticals.tclTHEN hinttac (tacipat pats), tacigens | false, `NoGen -> "ssr_wlog", hinttac, Tacticals.tclTHEN tacigens (tacipat pats) | true, `Gen _ -> assert false | false, `Gen id -> if gens = [] then errorstrm(str"gen have requires some generalizations"); let clear0 = old_cleartac clr0 in let id, name_general_hyp, cleanup, pats = match id, pats with | None, (IOpId id as ip)::pats -> Some id, tacipat [ip], clear0, pats | None, _ -> None, Tacticals.tclIDTAC, clear0, pats | Some (Some id),_ -> Some id, introid id, clear0, pats | Some _,_ -> let id = mk_anon_id "tmp" (Tacmach.pf_ids_of_hyps gl) in Some id, introid id, Tacticals.tclTHEN clear0 (Proofview.V82.of_tactic (Tactics.clear [id])), pats in let tac_specialize = match id with | None -> Tacticals.tclIDTAC | Some id -> if pats = [] then Tacticals.tclIDTAC else let args = Array.of_list args in ppdebug(lazy(str"specialized="++ pr_econstr_env (pf_env gl) (project gl) EConstr.(mkApp (mkVar id,args)))); ppdebug(lazy(str"specialized_ty="++ pr_econstr_env (pf_env gl) (project gl) ct)); Tacticals.tclTHENS (basecuttac "ssr_have" ct) [Proofview.V82.of_tactic (Tactics.apply EConstr.(mkApp (mkVar id,args))); Tacticals.tclIDTAC] in "ssr_have", (if hint = nohint then tacigens else hinttac), Tacticals.tclTHENLIST [name_general_hyp; tac_specialize; tacipat pats; cleanup] in Tacticals.tclTHENS (basecuttac cut_kind c) [fst_goal_tac; snd_goal_tac] gl (** The "suffice" tactic *) let sufftac ist ((((clr, pats),binders),simpl), ((_, c), hint)) = let clr = Option.default [] clr in let pats = tclCompileIPats pats in let binders = tclCompileIPats binders in let simpl = tclCompileIPats simpl in let htac = Tacticals.tclTHEN (introstac pats) (hinttac ist true hint) in let c = match Ssrcommon.ssrterm_of_ast_closure_term c with | (a, (b, Some ct)) -> begin match ct.CAst.v with | CCast (_, CastConv cty) -> a, (b, Some cty) | _ -> anomaly "suff: ssr cast hole deleted by typecheck" end | (a, (t, None)) -> begin match DAst.get t with | GCast (_, CastConv cty) -> a, (cty, None) | _ -> anomaly "suff: ssr cast hole deleted by typecheck" end in let ctac gl = let _,ty,_,uc = pf_interp_ty ist gl c in let gl = pf_merge_uc uc gl in basecuttac "ssr_suff" ty gl in Tacticals.tclTHENS ctac [htac; Tacticals.tclTHEN (old_cleartac clr) (introstac (binders@simpl))] open Proofview.Notations let is_app_evar sigma t = match EConstr.kind sigma t with | Constr.Evar _ -> true | Constr.App(t,_) -> begin match EConstr.kind sigma t with | Constr.Evar _ -> true | _ -> false end | _ -> false let rec ncons n e = match n with | 0 -> [] | n when n > 0 -> e :: ncons (n - 1) e | _ -> failwith "ncons" let intro_lock ipats = let hnf' = Proofview.numgoals >>= fun ng -> Proofview.tclDISPATCH (ncons (ng - 1) ssrsmovetac @ [Proofview.tclUNIT ()]) in let protect_subgoal env sigma hd args = Tactics.New.refine ~typecheck:true (fun sigma -> let lm2 = Array.length args - 2 in let sigma, carrier = Typing.type_of env sigma args.(lm2) in let rel = EConstr.mkApp (hd, Array.sub args 0 lm2) in let rel_args = Array.sub args lm2 2 in let sigma, under_rel = Ssrcommon.mkSsrConst "Under_rel" env sigma in let sigma, under_from_rel = Ssrcommon.mkSsrConst "Under_rel_from_rel" env sigma in let under_rel_args = Array.append [|carrier; rel|] rel_args in let ty = EConstr.mkApp (under_rel, under_rel_args) in let sigma, t = Evarutil.new_evar env sigma ty in sigma, EConstr.mkApp(under_from_rel,Array.append under_rel_args [|t|])) in let rec lock_eq () : unit Proofview.tactic = Proofview.Goal.enter begin fun _ -> Proofview.tclORELSE (Ssripats.tclIPAT [Ssripats.IOpTemporay; Ssripats.IOpEqGen (lock_eq ())]) (fun _exn -> Proofview.Goal.enter begin fun gl -> let c = Proofview.Goal.concl gl in let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in match EConstr.kind_of_type sigma c with | Term.AtomicType(hd, args) when Array.length args >= 2 && is_app_evar sigma (Array.last args) && Ssrequality.ssr_is_setoid env sigma hd args (* if the last condition above [ssr_is_setoid ...] holds then [Coq.Classes.RelationClasses] has been required *) || (* if this is not the case, the tactic can still succeed when the considered relation is [Coq.Init.Logic.iff] *) Ssrcommon.is_const_ref sigma hd (Coqlib.lib_ref "core.iff.type") && Array.length args = 2 && is_app_evar sigma args.(1) -> protect_subgoal env sigma hd args | _ -> let t = Reductionops.whd_all env sigma c in match EConstr.kind_of_type sigma t with | Term.AtomicType(hd, args) when Ssrcommon.is_ind_ref sigma hd (Coqlib.lib_ref "core.eq.type") && Array.length args = 3 && is_app_evar sigma args.(2) -> protect_subgoal env sigma hd args | _ -> ppdebug(lazy Pp.(str"under: stop:" ++ pr_econstr_env env sigma t)); Proofview.tclUNIT () end) end in hnf' <*> Ssripats.tclIPATssr ipats <*> lock_eq () let pretty_rename evar_map term varnames = let rec aux term vars = try match vars with | [] -> term | Names.Name.Anonymous :: varnames -> let name, types, body = EConstr.destLambda evar_map term in let res = aux body varnames in EConstr.mkLambda (name, types, res) | Names.Name.Name _ as name :: varnames -> let { Context.binder_relevance = r }, types, body = EConstr.destLambda evar_map term in let res = aux body varnames in EConstr.mkLambda (Context.make_annot name r, types, res) with DestKO -> term in aux term varnames let overtac = Proofview.V82.tactic (ssr_n_tac "over" ~-1) let check_numgoals ?(minus = 0) nh = Proofview.numgoals >>= fun ng -> if nh <> ng then let errmsg = str"Incorrect number of tactics" ++ spc() ++ str"(expected "++int (ng - minus)++str(String.plural ng " tactic") ++ str", was given "++ int (nh - minus)++str")." in CErrors.user_err errmsg else Proofview.tclUNIT () let undertac ?(pad_intro = false) ist ipats ((dir,_),_ as rule) hint = (* total number of implied hints *) let nh = List.length (snd hint) + (if hint = nullhint then 2 else 1) in let varnames = let rec aux acc = function | IPatId id :: rest -> aux (Names.Name.Name id :: acc) rest | IPatClear _ :: rest -> aux acc rest | IPatSimpl _ :: rest -> aux acc rest | IPatAnon (One _ | Drop) :: rest -> aux (Names.Name.Anonymous :: acc) rest | _ -> List.rev acc in aux [] @@ match ipats with | None -> [] | Some (IPatCase(Regular (l :: _)) :: _) -> l | Some l -> l in (* If we find a "=> [|]" we add 1 | to get "=> [||]" for the extra * goal (the one that is left once we run over) *) let ipats = match ipats with | None -> [IPatNoop] | Some l when pad_intro -> (* typically, ipats = Some [IPatAnon All] *) let new_l = ncons (nh - 1) l in [IPatCase(Regular (new_l @ [[]]))] | Some (IPatCase(Regular []) :: _ as ipats) -> ipats (* Erik: is the previous line correct/useful? *) | Some (IPatCase(Regular l) :: rest) -> IPatCase(Regular(l @ [[]])) :: rest | Some (IPatCase(Block _) :: _ as l) -> l | Some l -> [IPatCase(Regular [l;[]])] in let map_redex env evar_map ~before:_ ~after:t = ppdebug(lazy Pp.(str"under vars: " ++ prlist Names.Name.print varnames)); let evar_map, ty = Typing.type_of env evar_map t in let new_t = (* pretty-rename the bound variables *) try begin match EConstr.destApp evar_map t with (f, ar) -> let lam = Array.last ar in ppdebug(lazy Pp.(str"under: mapping:" ++ pr_econstr_env env evar_map lam)); let new_lam = pretty_rename evar_map lam varnames in let new_ar, len1 = Array.copy ar, pred (Array.length ar) in new_ar.(len1) <- new_lam; EConstr.mkApp (f, new_ar) end with | DestKO -> ppdebug(lazy Pp.(str"under: cannot pretty-rename bound variables with destApp")); t in ppdebug(lazy Pp.(str"under: to:" ++ pr_econstr_env env evar_map new_t)); evar_map, new_t in let undertacs = if hint = nohint then Proofview.tclUNIT () else let betaiota = Tactics.reduct_in_concl ~check:false (Reductionops.nf_betaiota, DEFAULTcast) in (* Usefulness of check_numgoals: tclDISPATCH would be enough, except for the error message w.r.t. the number of provided/expected tactics, as the last one is implied *) check_numgoals ~minus:1 nh <*> Proofview.tclDISPATCH ((List.map (function None -> overtac | Some e -> ssrevaltac ist e <*> overtac) (if hint = nullhint then [None] else snd hint)) @ [betaiota]) in let rew = Proofview.V82.tactic (Ssrequality.ssrrewritetac ~under:true ~map_redex ist [rule]) in rew <*> intro_lock ipats <*> undertacs coq-8.11.0/plugins/ssr/ssrview.mli0000644000175000017500000000347613612664533016775 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Glob_term.glob_constr list val declare : kind -> Glob_term.glob_constr list -> unit end (* Apply views to the top of the stack (intro pattern). If clear_if_id is * true (default false) then views that happen to be a variable are considered * as to be cleared (see the to_clear argument to the continuation) * * returns true if the last view was a tactic *) val tclIPAT_VIEWS : views:ast_closure_term list -> ?clear_if_id:bool -> conclusion:(to_clear:Names.Id.t list -> unit Proofview.tactic) -> bool Proofview.tactic (* Apply views to a given subject (as if was the top of the stack), then call conclusion on the obtained term (something like [v2 (v1 subject)]). The term being passed to conclusion is abstracted over non-resolved evars: if [simple_types] then all unnecessary dependencies among the abstracted evars are pruned *) val tclWITH_FWD_VIEWS : simple_types:bool -> subject:EConstr.t -> views:ast_closure_term list -> conclusion:(EConstr.t -> unit Proofview.tactic) -> unit Proofview.tactic coq-8.11.0/plugins/ssr/ssrbwd.mli0000644000175000017500000000146013612664533016566 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ssrterm ssragens -> ist -> unit tactic coq-8.11.0/plugins/ssr/ssrprinters.mli0000644000175000017500000000354013612664533017661 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* EConstr.constr -> Pp.t val pr_spc : unit -> Pp.t val pr_bar : unit -> Pp.t val pr_list : (unit -> Pp.t) -> ('a -> Pp.t) -> 'a list -> Pp.t val pp_concat : Pp.t -> ?sep:Pp.t -> Pp.t list -> Pp.t val xInParens : ssrtermkind val xWithAt : ssrtermkind val xNoFlag : ssrtermkind val xCpattern : ssrtermkind val pr_clear : (unit -> Pp.t) -> ssrclear -> Pp.t val pr_clear_ne : ssrclear -> Pp.t val pr_dir : ssrdir -> Pp.t val pr_simpl : ssrsimpl -> Pp.t val pr_term : ssrtermkind * (Glob_term.glob_constr * Constrexpr.constr_expr option) -> Pp.t val pr_ast_closure_term : ast_closure_term -> Pp.t val pr_view2 : ast_closure_term list -> Pp.t val pr_ipat : ssripat -> Pp.t val pr_ipats : ssripats -> Pp.t val pr_iorpat : ssripatss -> Pp.t val pr_block : id_block -> Pp.t val pr_hyp : ssrhyp -> Pp.t val pr_hyps : ssrhyps -> Pp.t val prl_constr_expr : Constrexpr.constr_expr -> Pp.t val prl_glob_constr : Glob_term.glob_constr -> Pp.t val pr_guarded : (string -> int -> bool) -> ('a -> Pp.t) -> 'a -> Pp.t val pr_occ : ssrocc -> Pp.t val ppdebug : Pp.t Lazy.t -> unit coq-8.11.0/plugins/ssr/ssrview.ml0000644000175000017500000003635713612664533016630 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* [] let cache_adaptor (_, (k, t)) = let lk = get k in if not (List.exists (Glob_ops.glob_constr_eq t) lk) then term_view_adaptor_db := AdaptorMap.add k (t :: lk) !term_view_adaptor_db let subst_adaptor ( subst, (k, t as a)) = let t' = Detyping.subst_glob_constr (Global.env()) subst t in if t' == t then a else k, t' let in_db = let open Libobject in declare_object @@ global_object_nodischarge "VIEW_ADAPTOR_DB" ~cache:cache_adaptor ~subst:(Some subst_adaptor) let declare kind terms = List.iter (fun term -> Lib.add_anonymous_leaf (in_db (kind,term))) (List.rev terms) end (* Forward View application code *****************************************) let reduce_or l = tclUNIT (List.fold_left (||) false l) module State : sig (* View storage API *) val vsINIT : view:EConstr.t -> subject_name:Id.t list -> to_clear:Id.t list -> unit tactic val vsPUSH : (EConstr.t -> (EConstr.t * Id.t list) tactic) -> unit tactic val vsCONSUME : (names:Id.t list -> EConstr.t -> to_clear:Id.t list -> unit tactic) -> unit tactic (* The bool is the || of the bool returned by the continuations *) val vsCONSUME_IF_PRESENT : (names:Id.t list -> EConstr.t option -> to_clear:Id.t list -> bool tactic) -> bool tactic val vsASSERT_EMPTY : unit tactic end = struct (* {{{ *) type vstate = { subject_name : Id.t list; (* top *) (* None if views are being applied to a term *) view : EConstr.t; (* v2 (v1 top) *) to_clear : Id.t list; } include Ssrcommon.MakeState(struct type state = vstate option let init = None end) let vsINIT ~view ~subject_name ~to_clear = tclSET (Some { subject_name; view; to_clear }) (** Initializes the state in which view data is accumulated when views are applied to the first assumption in the goal *) let vsBOOTSTRAP = Goal.enter_one ~__LOC__ begin fun gl -> let concl = Goal.concl gl in let id = (* We keep the orig name for checks in "in" tcl *) match EConstr.kind_of_type (Goal.sigma gl) concl with | Term.ProdType({binder_name=Name.Name id}, _, _) when Ssrcommon.is_discharged_id id -> id | _ -> mk_anon_id "view_subject" (Tacmach.New.pf_ids_of_hyps gl) in let view = EConstr.mkVar id in Ssrcommon.tclINTRO_ID id <*> tclSET (Some { subject_name = [id]; view; to_clear = [] }) end let rec vsPUSH k = tclINDEPENDENT (tclGET (function | Some { subject_name; view; to_clear } -> k view >>= fun (view, clr) -> tclSET (Some { subject_name; view; to_clear = to_clear @ clr }) | None -> vsBOOTSTRAP <*> vsPUSH k)) let rec vsCONSUME k = tclINDEPENDENT (tclGET (function | Some { subject_name; view; to_clear } -> tclSET None <*> k ~names:subject_name view ~to_clear | None -> vsBOOTSTRAP <*> vsCONSUME k)) let vsCONSUME_IF_PRESENT k = tclINDEPENDENTL (tclGET1 (function | Some { subject_name; view; to_clear } -> tclSET None <*> k ~names:subject_name (Some view) ~to_clear | None -> k ~names:[] None ~to_clear:[])) >>= reduce_or let vsASSERT_EMPTY = tclGET (function | Some _ -> anomaly ("vsASSERT_EMPTY: not empty") | _ -> tclUNIT ()) end (* }}} *) let intern_constr_expr { Genintern.genv; ltacvars = vars } sigma ce = let ltacvars = { Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in Constrintern.intern_gen Pretyping.WithoutTypeConstraint ~ltacvars genv sigma ce (* Disambiguation of /t - t is ltac:(tactic args) - t is a term To allow for t being a notation, like "Notation foo x := ltac:(foo x)", we need to internalize t. *) let is_tac_in_term ?extra_scope { annotation; body; glob_env; interp_env } = Goal.(enter_one ~__LOC__ begin fun goal -> let genv = env goal in let sigma = sigma goal in let ist = Ssrcommon.option_assert_get glob_env (Pp.str"not a term") in (* We use the env of the goal, not the global one *) let ist = { ist with Genintern.genv } in (* We open extra_scope *) let body = match extra_scope with | None -> body | Some s -> CAst.make (Constrexpr.CDelimiters(s,body)) in (* We unravel notations *) let g = intern_constr_expr ist sigma body in match DAst.get g with | Glob_term.GHole (_,_, Some x) when Genarg.has_type x (Genarg.glbwit Tacarg.wit_tactic) -> tclUNIT (`Tac (Genarg.out_gen (Genarg.glbwit Tacarg.wit_tactic) x)) | _ -> tclUNIT (`Term (annotation, interp_env, g)) end) (* To inject a constr into a glob_constr we use an Ltac variable *) let tclINJ_CONSTR_IST ist p = let fresh_id = Ssrcommon.mk_internal_id "ssr_inj_constr_in_glob" in let ist = { ist with Geninterp.lfun = Id.Map.add fresh_id (Taccoerce.Value.of_constr p) ist.Geninterp.lfun} in tclUNIT (ist,Glob_term.GVar fresh_id) let mkGHole = DAst.make (Glob_term.GHole(Evar_kinds.InternalHole, Namegen.IntroAnonymous, None)) let rec mkGHoles n = if n > 0 then mkGHole :: mkGHoles (n - 1) else [] let mkGApp f args = if args = [] then f else DAst.make (Glob_term.GApp (f, args)) (* From glob_constr to open_constr === (env,sigma,constr) *) let interp_glob ist glob = Goal.enter_one ~__LOC__ begin fun goal -> let env = Goal.env goal in let sigma = Goal.sigma goal in Ssrprinters.ppdebug (lazy Pp.(str"interp-in: " ++ Printer.pr_glob_constr_env env glob)); try let sigma,term = Tacinterp.interp_open_constr ist env sigma (glob,None) in Ssrprinters.ppdebug (lazy Pp.(str"interp-out: " ++ Printer.pr_econstr_env env sigma term)); tclUNIT (env,sigma,term) with e -> Ssrprinters.ppdebug (lazy Pp.(str"interp-err: " ++ Printer.pr_glob_constr_env env glob)); tclZERO e end (* Commits the term to the monad *) (* I think we should make the API safe by storing here the original evar map, * so that one cannot commit it wrongly. * We could also commit the term automatically, but this makes the code less * modular, see the 2 functions below that would need to "uncommit" *) let tclKeepOpenConstr (_env, sigma, t) = Unsafe.tclEVARS sigma <*> tclUNIT t let tclADD_CLEAR_IF_ID (env, ist, t) x = Ssrprinters.ppdebug (lazy Pp.(str"tclADD_CLEAR_IF_ID: " ++ Printer.pr_econstr_env env ist t)); let hd, args = EConstr.decompose_app ist t in match EConstr.kind ist hd with | Constr.Var id when Ssrcommon.not_section_id id -> tclUNIT (x, [id]) | _ -> tclUNIT (x,[]) let tclPAIR p x = tclUNIT (x, p) (* The ssr heuristic : *) (* Estimate a bound on the number of arguments of a raw constr. *) (* This is not perfect, because the unifier may fail to *) (* typecheck the partial application, so we use a minimum of 5. *) (* Also, we don't handle delayed or iterated coercions to *) (* FUNCLASS, which is probably just as well since these can *) (* lead to infinite arities. *) let guess_max_implicits ist glob = Proofview.tclORELSE (interp_glob ist (mkGApp glob (mkGHoles 6)) >>= fun (env,sigma,term) -> let term_ty = Retyping.get_type_of env sigma term in let ctx, _ = Reductionops.splay_prod env sigma term_ty in tclUNIT (List.length ctx + 6)) (fun _ -> tclUNIT 5) let pad_to_inductive ist glob = Goal.enter_one ~__LOC__ begin fun goal -> interp_glob ist glob >>= fun (env, sigma, term as ot) -> let term_ty = Retyping.get_type_of env sigma term in let ctx, i = Reductionops.splay_prod env sigma term_ty in let rel_ctx = List.map (fun (a,b) -> Context.Rel.Declaration.LocalAssum(a,b)) ctx in if not (Ssrcommon.isAppInd (EConstr.push_rel_context rel_ctx env) sigma i) then Tacticals.New.tclZEROMSG Pp.(str"not an inductive") else tclUNIT (mkGApp glob (mkGHoles (List.length ctx))) >>= tclADD_CLEAR_IF_ID ot end (* There are two ways of "applying" a view to term: *) (* 1- using a view hint if the view is an instance of some *) (* (reflection) inductive predicate. *) (* 2- applying the view if it coerces to a function, adding *) (* implicit arguments. *) (* They require guessing the view hints and the number of *) (* implicits, respectively, which we do by brute force. *) (* Builds v p *) let interp_view ~clear_if_id ist v p = let is_specialize hd = match DAst.get hd with Glob_term.GHole _ -> true | _ -> false in (* We cast the pile of views p into a term p_id *) tclINJ_CONSTR_IST ist p >>= fun (ist, p_id) -> let p_id = DAst.make p_id in match DAst.get v with | Glob_term.GApp (hd, rargs) when is_specialize hd -> Ssrprinters.ppdebug (lazy Pp.(str "specialize")); interp_glob ist (mkGApp p_id rargs) >>= tclKeepOpenConstr >>= tclPAIR [] | _ -> Ssrprinters.ppdebug (lazy Pp.(str "view")); (* We find out how to build (v p) eventually using an adaptor *) let adaptors = AdaptorDb.(get Forward) in Proofview.tclORELSE (pad_to_inductive ist v >>= fun (vpad,clr) -> Ssrcommon.tclFIRSTa (List.map (fun a -> interp_glob ist (mkGApp a [vpad; p_id])) adaptors) >>= tclPAIR clr) (fun _ -> guess_max_implicits ist v >>= fun n -> Ssrcommon.tclFIRSTi (fun n -> interp_glob ist (mkGApp v (mkGHoles n @ [p_id]))) n >>= fun x -> tclADD_CLEAR_IF_ID x x) >>= fun (ot,clr) -> if clear_if_id then tclKeepOpenConstr ot >>= tclPAIR clr else tclKeepOpenConstr ot >>= tclPAIR [] (* we store in the state (v top), then (v1 (v2 top))... *) let pile_up_view ~clear_if_id (annotation, ist, v) = let ist = Ssrcommon.option_assert_get ist (Pp.str"not a term") in let clear_if_id = clear_if_id && annotation <> `Parens in State.vsPUSH (fun p -> interp_view ~clear_if_id ist v p) let finalize_view s0 ?(simple_types=true) p = Goal.enter_one ~__LOC__ begin fun g -> let env = Goal.env g in let sigma = Goal.sigma g in let evars_of_p = Evd.evars_of_term sigma p in let filter x _ = Evar.Set.mem x evars_of_p in let sigma = Typeclasses.resolve_typeclasses ~fail:false ~filter env sigma in let p = Reductionops.nf_evar sigma p in let get_body = function Evd.Evar_defined x -> x | _ -> assert false in let evars_of_econstr sigma t = Evarutil.undefined_evars_of_term sigma (EConstr.of_constr t) in let rigid_of s = List.fold_left (fun l k -> if Evd.is_defined sigma k then let bo = get_body Evd.(evar_body (find sigma k)) in k :: l @ Evar.Set.elements (evars_of_econstr sigma (EConstr.Unsafe.to_constr bo)) else l ) [] s in let und0 = (* Unassigned evars in the initial goal *) let sigma0 = Tacmach.project s0 in let g0info = Evd.find sigma0 (Tacmach.sig_it s0) in let g0 = Evd.evars_of_filtered_evar_info sigma0 g0info in List.filter (fun k -> Evar.Set.mem k g0) (List.map fst (Evar.Map.bindings (Evd.undefined_map sigma0))) in let rigid = rigid_of und0 in let n, p, to_prune, _ucst = pf_abs_evars2 s0 rigid (sigma, p) in let p = if simple_types then pf_abs_cterm s0 n p else p in Ssrprinters.ppdebug (lazy Pp.(str"view@finalized: " ++ Printer.pr_econstr_env env sigma p)); let sigma = List.fold_left Evd.remove sigma to_prune in Unsafe.tclEVARS sigma <*> tclUNIT p end let pose_proof subject_name p = Tactics.generalize [p] <*> begin match subject_name with | id :: _ -> Ssrcommon.tclRENAME_HD_PROD (Name.Name id) | _ -> tclUNIT() end <*> Tactics.New.reduce_after_refine (* returns true if the last item was a tactic *) let rec apply_all_views_aux ~clear_if_id vs finalization conclusion s0 = match vs with | [] -> finalization s0 (fun name p -> (match p with | None -> conclusion ~to_clear:[] | Some p -> pose_proof name p <*> conclusion ~to_clear:name) <*> tclUNIT false) | v :: vs -> Ssrprinters.ppdebug (lazy Pp.(str"piling...")); is_tac_in_term ~extra_scope:"ssripat" v >>= function | `Term v -> Ssrprinters.ppdebug (lazy Pp.(str"..a term")); pile_up_view ~clear_if_id v <*> apply_all_views_aux ~clear_if_id vs finalization conclusion s0 | `Tac tac -> Ssrprinters.ppdebug (lazy Pp.(str"..a tactic")); finalization s0 (fun name p -> (match p with | None -> tclUNIT () | Some p -> pose_proof name p) <*> Tacinterp.eval_tactic tac <*> if vs = [] then begin Ssrprinters.ppdebug (lazy Pp.(str"..was the last view")); conclusion ~to_clear:name <*> tclUNIT true end else Tactics.clear name <*> tclINDEPENDENTL begin Ssrprinters.ppdebug (lazy Pp.(str"..was NOT the last view")); Ssrcommon.tacSIGMA >>= apply_all_views_aux ~clear_if_id vs finalization conclusion end >>= reduce_or) let apply_all_views vs ~conclusion ~clear_if_id = let finalization s0 k = State.vsCONSUME_IF_PRESENT (fun ~names t ~to_clear -> match t with | None -> k [] None | Some t -> finalize_view s0 t >>= fun p -> k (names @ to_clear) (Some p)) in Ssrcommon.tacSIGMA >>= apply_all_views_aux ~clear_if_id vs finalization conclusion (* We apply a view to a term given by the user, e.g. `case/V: x`. `x` is `subject` *) let apply_all_views_to subject ~simple_types vs ~conclusion = begin let rec process_all_vs = function | [] -> tclUNIT () | v :: vs -> is_tac_in_term v >>= function | `Tac _ -> Ssrcommon.errorstrm Pp.(str"tactic view not supported") | `Term v -> pile_up_view ~clear_if_id:false v <*> process_all_vs vs in State.vsASSERT_EMPTY <*> State.vsINIT ~subject_name:[] ~to_clear:[] ~view:subject <*> Ssrcommon.tacSIGMA >>= fun s0 -> process_all_vs vs <*> State.vsCONSUME (fun ~names:_ t ~to_clear:_ -> finalize_view s0 ~simple_types t >>= conclusion) end (* Entry points *********************************************************) let tclIPAT_VIEWS ~views:vs ?(clear_if_id=false) ~conclusion = tclINDEPENDENTL begin State.vsASSERT_EMPTY <*> apply_all_views vs ~conclusion ~clear_if_id >>= fun was_tac -> State.vsASSERT_EMPTY <*> tclUNIT was_tac end >>= reduce_or let tclWITH_FWD_VIEWS ~simple_types ~subject ~views:vs ~conclusion = tclINDEPENDENT (apply_all_views_to subject ~simple_types vs ~conclusion) (* vim: set filetype=ocaml foldmethod=marker: *) coq-8.11.0/plugins/ssr/ssrast.mli0000644000175000017500000001367713612664533016616 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* or rewrite flag *) type ssrdir = Ssrmatching_plugin.Ssrmatching.ssrdir = L2R | R2L (* simpl: "/=", cut: "//", simplcut: "//=" nop: commodity placeholder *) type ssrsimpl = Simpl of int | Cut of int | SimplCut of int * int | Nop (* modality for rewrite and do: ! ? *) type ssrmmod = May | Must | Once (* modality with a bound for rewrite and do: !n ?n *) type ssrmult = int * ssrmmod (** Occurrence switch {1 2}, all is Some(false,[]) *) type ssrocc = (bool * int list) option (* index MAYBE REMOVE ONLY INTERNAL stuff between {} *) type ssrindex = int Locus.or_var (* clear switch {H G} *) type ssrclear = ssrhyps (* Discharge occ switch (combined occurrence / clear switch) *) type ssrdocc = ssrclear option * ssrocc (* OLD ssr terms *) type ssrtermkind = char (* FIXME, make algebraic *) type ssrterm = ssrtermkind * Genintern.glob_constr_and_expr (* NEW ssr term *) (* These terms are raw but closed with the intenalization/interpretation * context. It is up to the tactic receiving it to decide if such contexts * are useful or not, and eventually manipulate the term before turning it * into a constr *) type ast_closure_term = { body : Constrexpr.constr_expr; glob_env : Genintern.glob_sign option; (* for Tacintern.intern_constr *) interp_env : Geninterp.interp_sign option; (* for Tacinterp.interp_open_constr_with_bindings *) annotation : [ `None | `Parens | `DoubleParens | `At ]; } type ssrview = ast_closure_term list type id_block = Prefix of Id.t | SuffixId of Id.t | SuffixNum of int (* Only [One] forces an introduction, possibly reducing the goal. *) type anon_kind = | One of string option (* name hint *) | Drop | All | Temporary type ssripat = | IPatNoop | IPatId of Id.t | IPatAnon of anon_kind (* inaccessible name *) | IPatDispatch of ssripatss_or_block (* (..|..) *) | IPatCase of ssripatss_or_block (* [..|..] *) | IPatInj of ssripatss | IPatRewrite of (*occurrence option * rewrite_pattern **) ssrocc * ssrdir | IPatView of ssrview (* /view *) | IPatClear of ssrclear (* {H1 H2} *) | IPatSimpl of ssrsimpl | IPatAbstractVars of Id.t list | IPatFastNondep and ssripats = ssripat list and ssripatss = ssripats list and ssripatss_or_block = | Block of id_block | Regular of ssripats list type ssrhpats = ((ssrclear option * ssripats) * ssripats) * ssripats type ssrhpats_wtransp = bool * ssrhpats (* tac => inpats *) type ssrintrosarg = Tacexpr.raw_tactic_expr * ssripats type ssrfwdid = Id.t (** Binders (for fwd tactics) *) type 'term ssrbind = | Bvar of Name.t | Bdecl of Name.t list * 'term | Bdef of Name.t * 'term option * 'term | Bstruct of Name.t | Bcast of 'term (* We use an intermediate structure to correctly render the binder list *) (* abbreviations. We use a list of hints to extract the binders and *) (* base term from a term, for the two first levels of representation of *) (* of constr terms. *) type ssrbindfmt = | BFvar | BFdecl of int (* #xs *) | BFcast (* final cast *) | BFdef (* has cast? *) | BFrec of bool * bool (* has struct? * has cast? *) type 'term ssrbindval = 'term ssrbind list * 'term (** Forward chaining argument *) (* There are three kinds of forward definitions: *) (* - Hint: type only, cast to Type, may have proof hint. *) (* - Have: type option + value, no space before type *) (* - Pose: binders + value, space before binders. *) type ssrfwdkind = FwdHint of string * bool | FwdHave | FwdPose type ssrfwdfmt = ssrfwdkind * ssrbindfmt list (* in *) type ssrclseq = InGoal | InHyps | InHypsGoal | InHypsSeqGoal | InSeqGoal | InHypsSeq | InAll | InAllHyps type 'tac ssrhint = bool * 'tac option list type 'tac fwdbinders = bool * (ssrhpats * ((ssrfwdfmt * ast_closure_term) * 'tac ssrhint)) type 'tac ffwbinders = (ssrhpats * ((ssrfwdfmt * ast_closure_term) * 'tac ssrhint)) type clause = (ssrclear * ((ssrhyp_or_id * string) * Ssrmatching_plugin.Ssrmatching.cpattern option) option) type clauses = clause list * ssrclseq type wgen = (ssrclear * ((ssrhyp_or_id * string) * Ssrmatching_plugin.Ssrmatching.cpattern option) option) type 'a ssrdoarg = ((ssrindex * ssrmmod) * 'a ssrhint) * clauses type 'a ssrseqarg = ssrindex * ('a ssrhint * 'a option) open Ssrmatching_plugin open Ssrmatching type 'a ssrcasearg = ssripat option * ('a * ssripats) type 'a ssrmovearg = ssrview * 'a ssrcasearg type ssrdgens = { dgens : (ssrdocc * cpattern) list; gens : (ssrdocc * cpattern) list; clr : ssrclear } type 'a ssragens = (ssrdocc * 'a) list list * ssrclear type ssrapplyarg = ssrterm list * (ssrterm ssragens * ssripats) (* OOP : these are general shortcuts *) type gist = Tacintern.glob_sign type ist = Tacinterp.interp_sign type goal = Goal.goal type 'a sigma = 'a Evd.sigma type v82tac = Tacmach.tactic coq-8.11.0/plugins/ssr/ssrelim.mli0000644000175000017500000000337613612664533016750 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ((Ssrast.ssrhyps option * Ssrast.ssrocc) * Ssrmatching.cpattern) list -> ([< `EConstr of Ssrast.ssrhyp list * Ssrmatching.occ * EConstr.constr & 'b | `EGen of (Ssrast.ssrhyp list option * Ssrmatching.occ) * Ssrmatching.cpattern ] as 'a) -> ?elim:EConstr.constr -> Ssrast.ssripat option -> (?seed:Names.Name.t list array -> 'a -> Ssrast.ssripat option -> unit Proofview.tactic -> bool -> Ssrast.ssrhyp list -> unit Proofview.tactic) -> unit Proofview.tactic val elimtac : EConstr.constr -> unit Proofview.tactic val casetac : EConstr.constr -> (?seed:Names.Name.t list array -> unit Proofview.tactic -> unit Proofview.tactic) -> unit Proofview.tactic val is_injection_case : EConstr.t -> Goal.goal Evd.sigma -> bool val perform_injection : EConstr.constr -> Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val ssrscase_or_inj_tac : EConstr.constr -> unit Proofview.tactic coq-8.11.0/plugins/ssr/ssrsetoid.v0000644000175000017500000000303413612664533016764 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* .doc { font-family: monospace; white-space: pre; } # **) (** Compatibility layer for [under] and [setoid_rewrite]. This file is intended to be required by [Require Import Setoid]. In particular, we can use the [under] tactic with other relations than [eq] or [iff], e.g. a [RewriteRelation], by doing: [Require Import ssreflect. Require Setoid.] This file's instances have priority 12 > other stdlib instances. (Note: this file could be skipped when porting [under] to stdlib2.) *) Require Import ssrclasses. Require Import ssrunder. Require Import RelationClasses. Require Import Relation_Definitions. (** Reconcile [Coq.Classes.RelationClasses.Reflexive] with [Coq.ssr.ssrclasses.Reflexive] *) Instance compat_Reflexive : forall {A} {R : relation A}, RelationClasses.Reflexive R -> ssrclasses.Reflexive R | 12. Proof. now trivial. Qed. coq-8.11.0/plugins/ssr/ssripats.ml0000644000175000017500000011212113612664533016756 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Names.Id.print id | IOpDrop -> Pp.str "_" | IOpTemporay -> Pp.str "+" | IOpInaccessible None -> Pp.str "?" | IOpInaccessible (Some s) -> Pp.str ("?«"^s^"»") | IOpInaccessibleAll -> Pp.str "*" | IOpAbstractVars l -> Pp.str ("[:"^String.concat " " (List.map Names.Id.to_string l)^"]") | IOpFastNondep -> Pp.str ">" | IOpInj l -> Pp.(str "[=" ++ ppl l ++ str "]") | IOpDispatchBlock b -> Pp.(str"(" ++ Ssrprinters.pr_block b ++ str")") | IOpDispatchBranches l -> Pp.(str "(" ++ ppl l ++ str ")") | IOpCaseBlock b -> Pp.(str"[" ++ Ssrprinters.pr_block b ++ str"]") | IOpCaseBranches l -> Pp.(str "[" ++ ppl l ++ str "]") | IOpRewrite (occ,dir) -> Pp.(Ssrprinters.(pr_occ occ ++ pr_dir dir)) | IOpView (None,vs) -> Pp.(prlist_with_sep mt (fun c -> str "/" ++ Ssrprinters.pr_ast_closure_term c) vs) | IOpView (Some cl,vs) -> Pp.(Ssrprinters.pr_clear Pp.spc cl ++ prlist_with_sep mt (fun c -> str "/" ++ Ssrprinters.pr_ast_closure_term c) vs) | IOpClear (clmust,clmay) -> Pp.(Ssrprinters.pr_clear spc clmust ++ match clmay with | Some cl -> str "(try " ++ Ssrprinters.pr_clear spc [cl] ++ str")" | None -> mt ()) | IOpSimpl s -> Ssrprinters.pr_simpl s | IOpEqGen _ -> Pp.str "E:" | IOpNoop -> Pp.str"-" and ppl x = Pp.(prlist_with_sep (fun () -> str"|") (prlist_with_sep spc pr_ipatop)) x module IpatMachine : sig (* the => tactical. ?eqtac is a tactic to be eventually run * after the first [..] block. first_case_is_dispatch is the * ssr exception to elim: and case: *) val main : ?eqtac:unit tactic -> first_case_is_dispatch:bool -> ssriops -> unit tactic val tclCompileIPats : ssripats -> ssriops val tclSEED_SUBGOALS : Names.Name.t list array -> unit tactic -> unit tactic end = struct (* {{{ *) module State : sig type delayed_gen = { tmp_id : Id.t; (* Temporary name *) orig_name : Name.t (* Old name *) } (* to_clear API *) val isCLR_PUSH : Id.t -> unit tactic val isCLR_PUSHL : Id.t list -> unit tactic val isCLR_CONSUME : unit tactic (* to_generalize API *) val isGEN_PUSH : delayed_gen -> unit tactic val isGEN_CONSUME : unit tactic (* name_seed API *) val isNSEED_SET : Names.Name.t list -> unit tactic val isNSEED_CONSUME : (Names.Name.t list option -> unit tactic) -> unit tactic (* Some data may expire *) val isTICK : ssriop -> unit tactic val isPRINT : Proofview.Goal.t -> Pp.t end = struct (* {{{ *) type istate = { (* Delayed clear *) to_clear : Id.t list; (* Temporary intros, to be generalized back *) to_generalize : delayed_gen list; (* The type of the inductive constructor corresponding to the current proof * branch: name seeds are taken from that in an intro block *) name_seed : Names.Name.t list option; } and delayed_gen = { tmp_id : Id.t; (* Temporary name *) orig_name : Name.t (* Old name *) } let empty_state = { to_clear = []; to_generalize = []; name_seed = None; } include Ssrcommon.MakeState(struct type state = istate let init = empty_state end) let print_name_seed env sigma = function | None -> Pp.str "-" | Some nl -> Pp.prlist Names.Name.print nl let print_delayed_gen { tmp_id; orig_name } = Pp.(Id.print tmp_id ++ str"->" ++ Name.print orig_name) let isPRINT g = let env, sigma = Goal.env g, Goal.sigma g in let state = get g in Pp.(str"{{ to_clear: " ++ prlist_with_sep spc Id.print state.to_clear ++ spc () ++ str"to_generalize: " ++ prlist_with_sep spc print_delayed_gen state.to_generalize ++ spc () ++ str"name_seed: " ++ print_name_seed env sigma state.name_seed ++ str" }}") let isCLR_PUSH id = tclGET (fun ({ to_clear = ids } as s) -> tclSET { s with to_clear = id :: ids }) let isCLR_PUSHL more_ids = tclGET (fun ({ to_clear = ids } as s) -> tclSET { s with to_clear = more_ids @ ids }) let isCLR_CONSUME = tclGET (fun ({ to_clear = ids } as s) -> tclSET { s with to_clear = [] } <*> Tactics.clear ids) let isGEN_PUSH dg = tclGET (fun s -> tclSET { s with to_generalize = dg :: s.to_generalize }) (* generalize `id` as `new_name` *) let gen_astac id new_name = let gen = ((None,Some(false,[])),Ssrmatching.cpattern_of_id id) in V82.tactic (Ssrcommon.gentac gen) <*> Ssrcommon.tclRENAME_HD_PROD new_name (* performs and resets all delayed generalizations *) let isGEN_CONSUME = tclGET (fun ({ to_generalize = dgs } as s) -> tclSET { s with to_generalize = [] } <*> Tacticals.New.tclTHENLIST (List.map (fun { tmp_id; orig_name } -> gen_astac tmp_id orig_name) dgs) <*> Tactics.clear (List.map (fun gen -> gen.tmp_id) dgs)) let isNSEED_SET ty = tclGET (fun s -> tclSET { s with name_seed = Some ty }) let isNSEED_CONSUME k = tclGET (fun ({ name_seed = x } as s) -> tclSET { s with name_seed = None } <*> k x) let isTICK = function | IOpSimpl _ | IOpClear _ -> tclUNIT () | _ -> tclGET (fun s -> tclSET { s with name_seed = None }) end (* }}} *************************************************************** *) open State (***[=> *] ****************************************************************) (** [nb_assums] returns the number of dependent premises Warning: unlike [nb_deps_assums], it does not perform reduction *) let rec nb_assums cur env sigma t = match EConstr.kind sigma t with | Prod(name,ty,body) -> nb_assums (cur+1) env sigma body | LetIn(name,ty,t1,t2) -> nb_assums (cur+1) env sigma t2 | Cast(t,_,_) -> nb_assums cur env sigma t | _ -> cur let nb_assums = nb_assums 0 let intro_anon_all = Goal.enter begin fun gl -> let env = Goal.env gl in let sigma = Goal.sigma gl in let g = Goal.concl gl in let n = nb_assums env sigma g in Tacticals.New.tclDO n (Ssrcommon.tclINTRO_ANON ()) end (*** [=> >*] **************************************************************) (** [nb_deps_assums] returns the number of dependent premises *) let rec nb_deps_assums cur env sigma t = let t' = Reductionops.whd_allnolet env sigma t in match EConstr.kind sigma t' with | Constr.Prod(name,ty,body) -> if EConstr.Vars.noccurn sigma 1 body && not (Typeclasses.is_class_type sigma ty) then cur else nb_deps_assums (cur+1) env sigma body | Constr.LetIn(name,ty,t1,t2) -> nb_deps_assums (cur+1) env sigma t2 | Constr.Cast(t,_,_) -> nb_deps_assums cur env sigma t | _ -> cur let nb_deps_assums = nb_deps_assums 0 let intro_anon_deps = Goal.enter begin fun gl -> let env = Goal.env gl in let sigma = Goal.sigma gl in let g = Goal.concl gl in let n = nb_deps_assums env sigma g in Tacticals.New.tclDO n (Ssrcommon.tclINTRO_ANON ()) end (** [intro_drop] behaves like [intro_anon] but registers the id of the introduced assumption for a delayed clear. *) let intro_drop = Ssrcommon.tclINTRO ~id:Ssrcommon.Anon ~conclusion:(fun ~orig_name:_ ~new_name -> isCLR_PUSH new_name) (** [intro_temp] behaves like [intro_anon] but registers the id of the introduced assumption for a regeneralization. *) let intro_anon_temp = Ssrcommon.tclINTRO ~id:Ssrcommon.Anon ~conclusion:(fun ~orig_name ~new_name -> isGEN_PUSH { tmp_id = new_name; orig_name }) (** [intro_end] performs the actions that have been delayed. *) let intro_end = Ssrcommon.tcl0G ~default:() (isCLR_CONSUME <*> isGEN_CONSUME) (** [=> _] *****************************************************************) let intro_clear ids = Goal.enter begin fun gl -> let _, clear_ids, ren = List.fold_left (fun (used_ids, clear_ids, ren) id -> let new_id = Ssrcommon.mk_anon_id (Id.to_string id) used_ids in (new_id :: used_ids, new_id :: clear_ids, (id, new_id) :: ren)) (Tacmach.New.pf_ids_of_hyps gl, [], []) ids in Tactics.rename_hyp ren <*> isCLR_PUSHL clear_ids end let tacCHECK_HYPS_EXIST hyps = Goal.enter begin fun gl -> let ctx = Goal.hyps gl in List.iter (Ssrcommon.check_hyp_exists ctx) hyps; tclUNIT () end let tacFILTER_HYP_EXIST hyps k = Goal.enter begin fun gl -> let ctx = Goal.hyps gl in k (Option.bind hyps (fun h -> if Ssrcommon.test_hyp_exists ctx h && Ssrcommon.(not_section_id (hyp_id h)) then Some h else None)) end (** [=> []] *****************************************************************) (* calls t1 then t2 on each subgoal passing to t2 the index of the current * subgoal (starting from 0) as well as the number of subgoals *) let tclTHENin t1 t2 = tclUNIT () >>= begin fun () -> let i = ref (-1) in t1 <*> numgoals >>= fun n -> Goal.enter begin fun g -> incr i; t2 !i n end end (* Attaches one element of `seeds` to each of the last k goals generated by `tac`, where k is the size of `seeds` *) let tclSEED_SUBGOALS seeds tac = tclTHENin tac (fun i n -> Ssrprinters.ppdebug (lazy Pp.(str"seeding")); (* eg [case: (H _ : nat)] generates 3 goals: - 1 for _ - 2 for the nat constructors *) let extra_goals = n - Array.length seeds in if i < extra_goals then tclUNIT () else isNSEED_SET seeds.(i - extra_goals)) let tac_case t = Goal.enter begin fun _ -> Ssrcommon.tacTYPEOF t >>= fun ty -> Ssrcommon.tacIS_INJECTION_CASE ~ty t >>= fun is_inj -> if is_inj then V82.tactic ~nf_evars:false (Ssrelim.perform_injection t) else Goal.enter begin fun g -> (Ssrelim.casetac t (fun ?seed k -> match seed with | None -> k | Some seed -> tclSEED_SUBGOALS seed k)) end end (** [=> [^ seed ]] *********************************************************) let tac_intro_seed interp_ipats fix = Goal.enter begin fun gl -> isNSEED_CONSUME begin fun seeds -> let seeds = Ssrcommon.option_assert_get seeds Pp.(str"tac_intro_seed: no seed") in let ipats = List.map (function | Anonymous -> let s = match fix with | Prefix id -> Id.to_string id ^ "?" | SuffixNum n -> "?" ^ string_of_int n | SuffixId id -> "?" ^ Id.to_string id in IOpInaccessible (Some s) | Name id -> let s = match fix with | Prefix fix -> Id.to_string fix ^ Id.to_string id | SuffixNum n -> Id.to_string id ^ string_of_int n | SuffixId fix -> Id.to_string id ^ Id.to_string fix in IOpId (Id.of_string s)) seeds in interp_ipats ipats end end (*** [=> [: id]] ************************************************************) let mk_abstract_id = let open Coqlib in let ssr_abstract_id = Summary.ref ~name:"SSR:abstractid" 0 in begin fun env sigma -> let sigma, zero = EConstr.fresh_global env sigma (lib_ref "num.nat.O") in let sigma, succ = EConstr.fresh_global env sigma (lib_ref "num.nat.S") in let rec nat_of_n n = if n = 0 then zero else EConstr.mkApp (succ, [|nat_of_n (n-1)|]) in incr ssr_abstract_id; sigma, nat_of_n !ssr_abstract_id end let tclMK_ABSTRACT_VAR id = Goal.enter begin fun gl -> let env, concl = Goal.(env gl, concl gl) in let step = begin fun sigma -> let (sigma, (abstract_proof, abstract_ty)) = let (sigma, (ty, _)) = Evarutil.new_type_evar env sigma Evd.univ_flexible_alg in let (sigma, ablock) = Ssrcommon.mkSsrConst "abstract_lock" env sigma in let (sigma, lock) = Evarutil.new_evar env sigma ablock in let (sigma, abstract) = Ssrcommon.mkSsrConst "abstract" env sigma in let (sigma, abstract_id) = mk_abstract_id env sigma in let abstract_ty = EConstr.mkApp(abstract, [|ty; abstract_id; lock|]) in let sigma, m = Evarutil.new_evar env sigma abstract_ty in sigma, (m, abstract_ty) in let sigma, kont = let rd = Context.Rel.Declaration.LocalAssum (make_annot (Name id) Sorts.Relevant, abstract_ty) in let sigma, ev = Evarutil.new_evar (EConstr.push_rel rd env) sigma concl in sigma, ev in let term = EConstr.(mkApp (mkLambda(make_annot (Name id) Sorts.Relevant,abstract_ty,kont),[|abstract_proof|])) in let sigma, _ = Typing.type_of env sigma term in sigma, term end in Tactics.New.refine ~typecheck:false step <*> tclFOCUS 1 3 Proofview.shelve end let tclMK_ABSTRACT_VARS ids = List.fold_right (fun id tac -> Tacticals.New.tclTHENFIRST (tclMK_ABSTRACT_VAR id) tac) ids (tclUNIT ()) (* Debugging *) let tclLOG p t = tclUNIT () >>= begin fun () -> Ssrprinters.ppdebug (lazy Pp.(str "exec: " ++ pr_ipatop p)); tclUNIT () end <*> Goal.enter begin fun g -> Ssrprinters.ppdebug (lazy Pp.(str" on state:" ++ spc () ++ isPRINT g ++ str" goal:" ++ spc () ++ Printer.pr_goal (Goal.print g))); tclUNIT () end <*> t p >>= fun ret -> Goal.enter begin fun g -> Ssrprinters.ppdebug (lazy Pp.(str "done: " ++ isPRINT g)); tclUNIT () end >>= fun () -> tclUNIT ret let notTAC = tclUNIT false let duplicate_clear = CWarnings.create ~name:"duplicate-clear" ~category:"ssr" (fun id -> Pp.(str "Duplicate clear of " ++ Id.print id)) (* returns true if it was a tactic (eg /ltac:tactic) *) let rec ipat_tac1 ipat : bool tactic = match ipat with | IOpView (glued_clear,l) -> let clear_if_id, extra_clear = match glued_clear with | None -> false, [] | Some x -> true, List.map Ssrcommon.hyp_id x in Ssrview.tclIPAT_VIEWS ~views:l ~clear_if_id ~conclusion:(fun ~to_clear:clr -> let inter = CList.intersect Id.equal clr extra_clear in List.iter duplicate_clear inter; let cl = CList.union Id.equal clr extra_clear in intro_clear cl) | IOpDispatchBranches ipatss -> tclDISPATCH (List.map ipat_tac ipatss) <*> notTAC | IOpDispatchBlock id_block -> tac_intro_seed ipat_tac id_block <*> notTAC | IOpCaseBlock id_block -> Ssrcommon.tclWITHTOP tac_case <*> tac_intro_seed ipat_tac id_block <*> notTAC | IOpCaseBranches ipatss -> tclIORPAT (Ssrcommon.tclWITHTOP tac_case) ipatss <*> notTAC | IOpId id -> Ssrcommon.tclINTRO_ID id <*> notTAC | IOpFastNondep -> intro_anon_deps <*> notTAC | IOpDrop -> intro_drop <*> notTAC | IOpInaccessible seed -> Ssrcommon.tclINTRO_ANON ?seed () <*> notTAC | IOpInaccessibleAll -> intro_anon_all <*> notTAC | IOpTemporay -> intro_anon_temp <*> notTAC | IOpSimpl Nop -> assert false | IOpInj ipatss -> tclIORPAT (Ssrcommon.tclWITHTOP (fun t -> V82.tactic ~nf_evars:false (Ssrelim.perform_injection t))) ipatss <*> notTAC | IOpClear (must,may) -> tacCHECK_HYPS_EXIST must <*> tacFILTER_HYP_EXIST may (fun may -> let must = List.map Ssrcommon.hyp_id must in let cl = Option.fold_left (fun cls (SsrHyp(_,id)) -> if CList.mem_f Id.equal id cls then begin duplicate_clear id; cls end else id :: cls) must may in intro_clear cl) <*> notTAC | IOpSimpl x -> V82.tactic ~nf_evars:false (Ssrequality.simpltac x) <*> notTAC | IOpRewrite (occ,dir) -> Ssrcommon.tclWITHTOP (fun x -> V82.tactic ~nf_evars:false (Ssrequality.ipat_rewrite occ dir x)) <*> notTAC | IOpAbstractVars ids -> tclMK_ABSTRACT_VARS ids <*> notTAC | IOpEqGen t -> t <*> notTAC | IOpNoop -> notTAC and ipat_tac pl : unit tactic = match pl with | [] -> tclUNIT () | pat :: pl -> Ssrcommon.tcl0G ~default:false (tclLOG pat ipat_tac1) >>= fun was_tac -> isTICK pat (* drops expired seeds *) >>= fun () -> if was_tac then (* exception *) let ip_before, case, ip_after = split_at_first_case pl in let case = ssr_exception true case in let case = option_to_list case in ipat_tac (ip_before @ case @ ip_after) else ipat_tac pl and tclIORPAT tac = function | [[]] -> tac | p -> Tacticals.New.tclTHENS tac (List.map ipat_tac p) and ssr_exception is_on = function | Some (IOpCaseBranches [[]]) when is_on -> Some IOpNoop | Some (IOpCaseBranches l) when is_on -> Some (IOpDispatchBranches l) | Some (IOpCaseBlock s) when is_on -> Some (IOpDispatchBlock s) | x -> x and option_to_list = function None -> [] | Some x -> [x] and split_at_first_case ipats = let rec loop acc = function | (IOpSimpl _ | IOpClear _) as x :: rest -> loop (x :: acc) rest | (IOpCaseBlock _ | IOpCaseBranches _ | IOpDispatchBlock _ | IOpDispatchBranches _) as x :: xs -> CList.rev acc, Some x, xs | pats -> CList.rev acc, None, pats in loop [] ipats ;; (* Simple pass doing {x}/v -> /v{x} *) let tclCompileIPats l = let rec elab = function | (IPatClear cl) :: (IPatView v) :: rest -> (IOpView(Some cl,v)) :: elab rest | (IPatClear cl) :: (IPatId id) :: rest -> (IOpClear (cl,Some (SsrHyp(None,id)))) :: IOpId id :: elab rest (* boring code *) | [] -> [] | IPatId id :: rest -> IOpId id :: elab rest | IPatAnon (One hint) ::rest -> IOpInaccessible hint :: elab rest | IPatAnon Drop :: rest -> IOpDrop :: elab rest | IPatAnon All :: rest -> IOpInaccessibleAll :: elab rest | IPatAnon Temporary :: rest -> IOpTemporay :: elab rest | IPatAbstractVars vs :: rest -> IOpAbstractVars vs :: elab rest | IPatFastNondep :: rest -> IOpFastNondep :: elab rest | IPatInj pats :: rest -> IOpInj (List.map elab pats) :: elab rest | IPatRewrite(occ,dir) :: rest -> IOpRewrite(occ,dir) :: elab rest | IPatView vs :: rest -> IOpView (None,vs) :: elab rest | IPatSimpl s :: rest -> IOpSimpl s :: elab rest | IPatClear cl :: rest -> IOpClear (cl,None) :: elab rest | IPatCase (Block seed) :: rest -> IOpCaseBlock seed :: elab rest | IPatCase (Regular bs) :: rest -> IOpCaseBranches (List.map elab bs) :: elab rest | IPatDispatch (Block seed) :: rest -> IOpDispatchBlock seed :: elab rest | IPatDispatch (Regular bs) :: rest -> IOpDispatchBranches (List.map elab bs) :: elab rest | IPatNoop :: rest -> IOpNoop :: elab rest in elab l ;; let tclCompileIPats l = Ssrprinters.ppdebug (lazy Pp.(str "tclCompileIPats input: " ++ prlist_with_sep spc Ssrprinters.pr_ipat l)); let ops = tclCompileIPats l in Ssrprinters.ppdebug (lazy Pp.(str "tclCompileIPats output: " ++ prlist_with_sep spc pr_ipatop ops)); ops let main ?eqtac ~first_case_is_dispatch iops = let ip_before, case, ip_after = split_at_first_case iops in let case = ssr_exception first_case_is_dispatch case in let case = option_to_list case in let eqtac = option_to_list (Option.map (fun x -> IOpEqGen x) eqtac) in let ipats = ip_before @ case @ eqtac @ ip_after in Ssrcommon.tcl0G ~default:() (ipat_tac ipats <*> intro_end) end (* }}} *) let tclIPAT_EQ eqtac ip = Ssrprinters.ppdebug (lazy Pp.(str "ipat@run: " ++ Ssrprinters.pr_ipats ip)); IpatMachine.(main ~eqtac ~first_case_is_dispatch:true (tclCompileIPats ip)) let tclIPATssr ip = Ssrprinters.ppdebug (lazy Pp.(str "ipat@run: " ++ Ssrprinters.pr_ipats ip)); IpatMachine.(main ~first_case_is_dispatch:true (tclCompileIPats ip)) let tclCompileIPats = IpatMachine.tclCompileIPats (* Common code to handle generalization lists along with the defective case *) let with_defective maintac deps clr = Goal.enter begin fun g -> let sigma, concl = Goal.(sigma g, concl g) in let top_id = match EConstr.kind_of_type sigma concl with | Term.ProdType ({binder_name=Name id}, _, _) when Ssrcommon.is_discharged_id id -> id | _ -> Ssrcommon.top_id in let top_gen = Ssrequality.mkclr clr, Ssrmatching.cpattern_of_id top_id in Ssrcommon.tclINTRO_ID top_id <*> maintac deps top_gen end let with_dgens { dgens; gens; clr } maintac = match gens with | [] -> with_defective maintac dgens clr | gen :: gens -> V82.tactic ~nf_evars:false (Ssrcommon.genstac (gens, clr)) <*> maintac dgens gen let mkCoqEq env sigma = let eq = Coqlib.((build_coq_eq_data ()).eq) in let sigma, eq = EConstr.fresh_global env sigma eq in eq, sigma let mkCoqRefl t c env sigma = let refl = Coqlib.((build_coq_eq_data()).refl) in let sigma, refl = EConstr.fresh_global env sigma refl in EConstr.mkApp (refl, [|t; c|]), sigma (** Intro patterns processing for elim tactic, in particular when used in conjunction with equation generation as in [elim E: x] *) let elim_intro_tac ipats ?seed what eqid ssrelim is_rec clr = let intro_eq = match eqid with | Some (IPatId ipat) when not is_rec -> let rec intro_eq () = Goal.enter begin fun g -> let sigma, env, concl = Goal.(sigma g, env g, concl g) in match EConstr.kind_of_type sigma concl with | Term.ProdType (_, src, tgt) -> begin match EConstr.kind_of_type sigma src with | Term.AtomicType (hd, _) when Ssrcommon.is_protect hd env sigma -> V82.tactic ~nf_evars:false Ssrcommon.unprotecttac <*> Ssrcommon.tclINTRO_ID ipat | _ -> Ssrcommon.tclINTRO_ANON () <*> intro_eq () end |_ -> Ssrcommon.errorstrm (Pp.str "Too many names in intro pattern") end in intro_eq () | Some (IPatId ipat) -> let intro_lhs = Goal.enter begin fun g -> let sigma = Goal.sigma g in let elim_name = match clr, what with | [SsrHyp(_, x)], _ -> x | _, `EConstr(_,_,t) when EConstr.isVar sigma t -> EConstr.destVar sigma t | _ -> Ssrcommon.mk_anon_id "K" (Tacmach.New.pf_ids_of_hyps g) in Tacticals.New.tclFIRST [ Ssrcommon.tclINTRO_ID elim_name ; Ssrcommon.tclINTRO_ANON ~seed:"K" ()] end in let rec gen_eq_tac () = Goal.enter begin fun g -> let sigma, env, concl = Goal.(sigma g, env g, concl g) in let sigma, eq = EConstr.fresh_global env sigma (Coqlib.lib_ref "core.eq.type") in let ctx, last = EConstr.decompose_prod_assum sigma concl in let args = match EConstr.kind_of_type sigma last with | Term.AtomicType (hd, args) -> if Ssrcommon.is_protect hd env sigma then args else Ssrcommon.errorstrm (Pp.str "Too many names in intro pattern") | _ -> assert false in let case = args.(Array.length args-1) in if not(EConstr.Vars.closed0 sigma case) then Ssrcommon.tclINTRO_ANON () <*> gen_eq_tac () else Ssrcommon.tacTYPEOF case >>= fun case_ty -> let open EConstr in let refl = mkApp (eq, [|Vars.lift 1 case_ty; mkRel 1; Vars.lift 1 case|]) in let name = Ssrcommon.mk_anon_id "K" (Tacmach.New.pf_ids_of_hyps g) in let new_concl = mkProd (make_annot (Name name) Sorts.Relevant, case_ty, mkArrow refl Sorts.Relevant (Vars.lift 2 concl)) in let erefl, sigma = mkCoqRefl case_ty case env sigma in Proofview.Unsafe.tclEVARS sigma <*> Tactics.apply_type ~typecheck:true new_concl [case;erefl] end in gen_eq_tac () <*> intro_lhs <*> Ssrcommon.tclINTRO_ID ipat | _ -> tclUNIT () in let unprotect = if eqid <> None && is_rec then V82.tactic ~nf_evars:false Ssrcommon.unprotecttac else tclUNIT () in begin match seed with | None -> ssrelim | Some s -> IpatMachine.tclSEED_SUBGOALS s ssrelim end <*> tclIPAT_EQ (intro_eq <*> unprotect) ipats ;; let mkEq dir cl c t n env sigma = let open EConstr in let eqargs = [|t; c; c|] in eqargs.(Ssrequality.dir_org dir) <- mkRel n; let eq, sigma = mkCoqEq env sigma in let refl, sigma = mkCoqRefl t c env sigma in mkArrow (mkApp (eq, eqargs)) Sorts.Relevant (Vars.lift 1 cl), refl, sigma (** in [tac/v: last gens..] the first (last to be run) generalization is "special" in that is it also the main argument of [tac] and is eventually to be processed forward with view [v]. The behavior implemented is very close to [tac: (v last) gens..] but: - [v last] may use a view adaptor - eventually clear for [last] is taken into account - [tac/v {clr}] is also supported, and [{clr}] is to be run later The code here does not "grab" [v last] nor apply [v] to [last], see the [tacVIEW_THEN_GRAB] combinator. *) let tclLAST_GEN ~to_ind ((oclr, occ), t) conclusion = tclINDEPENDENTL begin Ssrcommon.tacSIGMA >>= fun sigma0 -> Goal.enter_one begin fun g -> let pat = Ssrmatching.interp_cpattern sigma0 t None in let cl0, env, sigma, hyps = Goal.(concl g, env g, sigma g, hyps g) in let cl = EConstr.to_constr ~abort_on_undefined_evars:false sigma cl0 in let (c, ucst), cl = try Ssrmatching.fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1 with Ssrmatching.NoMatch -> Ssrmatching.redex_of_pattern env pat, cl in let sigma = Evd.merge_universe_context sigma ucst in let c, cl = EConstr.of_constr c, EConstr.of_constr cl in let clr = Ssrcommon.interp_clr sigma (oclr, (Ssrmatching.tag_of_cpattern t,c)) in (* Historically in Coq, and hence in ssr, [case t] accepts [t] of type [A.. -> Ind] and opens new goals for [A..] as well as for the branches of [Ind], see the [~to_ind] argument *) if not(Termops.occur_existential sigma c) then if Ssrmatching.tag_of_cpattern t = Ssrprinters.xWithAt then if not (EConstr.isVar sigma c) then Ssrcommon.errorstrm Pp.(str "@ can be used with variables only") else match Context.Named.lookup (EConstr.destVar sigma c) hyps with | Context.Named.Declaration.LocalAssum _ -> Ssrcommon.errorstrm Pp.(str "@ can be used with let-ins only") | Context.Named.Declaration.LocalDef (name, b, ty) -> Unsafe.tclEVARS sigma <*> tclUNIT (true, EConstr.mkLetIn (map_annot Name.mk_name name,b,ty,cl), c, clr) else Unsafe.tclEVARS sigma <*> Ssrcommon.tacMKPROD c cl >>= fun ccl -> tclUNIT (false, ccl, c, clr) else if to_ind && occ = None then let _, p, _, ucst' = (* TODO: use abs_evars2 *) Ssrcommon.pf_abs_evars sigma0 (fst pat, c) in let sigma = Evd.merge_universe_context sigma ucst' in Unsafe.tclEVARS sigma <*> Ssrcommon.tacTYPEOF p >>= fun pty -> (* TODO: check bug: cl0 no lift? *) let ccl = EConstr.mkProd (make_annot (Ssrcommon.constr_name sigma c) Sorts.Relevant, pty, cl0) in tclUNIT (false, ccl, p, clr) else Ssrcommon.errorstrm Pp.(str "generalized term didn't match") end end >>= begin fun infos -> tclDISPATCH (infos |> List.map conclusion) end (** a typical mate of [tclLAST_GEN] doing the job of applying the views [cs] to [c] and generalizing the resulting term *) let tacVIEW_THEN_GRAB ?(simple_types=true) vs ~conclusion (is_letin, new_concl, c, clear) = Ssrview.tclWITH_FWD_VIEWS ~simple_types ~subject:c ~views:vs ~conclusion:(fun t -> Ssrcommon.tacCONSTR_NAME c >>= fun name -> Goal.enter_one ~__LOC__ begin fun g -> let sigma, env = Goal.sigma g, Goal.env g in Ssrcommon.tacMKPROD t ~name (Termops.subst_term sigma t (* NOTE: we grab t here *) (Termops.prod_applist sigma new_concl [c])) >>= conclusion is_letin t clear end) (* Elim views are elimination lemmas, so the eliminated term is not added *) (* to the dependent terms as for "case", unless it actually occurs in the *) (* goal, the "all occurrences" {+} switch is used, or the equation switch *) (* is used and there are no dependents. *) let ssrelimtac (view, (eqid, (dgens, ipats))) = let ndefectelimtac view eqid ipats deps gen = match view with | [v] -> Ssrcommon.tclINTERP_AST_CLOSURE_TERM_AS_CONSTR v >>= fun cs -> tclDISPATCH (List.map (fun elim -> (Ssrelim.ssrelim deps (`EGen gen) ~elim eqid (elim_intro_tac ipats))) cs) | [] -> tclINDEPENDENT (Ssrelim.ssrelim deps (`EGen gen) eqid (elim_intro_tac ipats)) | _ -> Ssrcommon.errorstrm Pp.(str "elim: only one elimination lemma can be provided") in with_dgens dgens (ndefectelimtac view eqid ipats) let ssrcasetac (view, (eqid, (dgens, ipats))) = let ndefectcasetac view eqid ipats deps ((_, occ), _ as gen) = tclLAST_GEN ~to_ind:true gen (fun (_, cl, c, clear as info) -> let conclusion _ vc _clear _cl = Ssrcommon.tacIS_INJECTION_CASE vc >>= fun inj -> let simple = (eqid = None && deps = [] && occ = None) in if simple && inj then V82.tactic ~nf_evars:false (Ssrelim.perform_injection vc) <*> Tactics.clear (List.map Ssrcommon.hyp_id clear) <*> tclIPATssr ipats else (* macro for "case/v E: x" ---> "case E: x / (v x)" *) let deps, clear, occ = if view <> [] && eqid <> None && deps = [] then [gen], [], None else deps, clear, occ in Ssrelim.ssrelim ~is_case:true deps (`EConstr (clear, occ, vc)) eqid (elim_intro_tac ipats) in if view = [] then conclusion false c clear c else tacVIEW_THEN_GRAB ~simple_types:false view ~conclusion info) in with_dgens dgens (ndefectcasetac view eqid ipats) let ssrscasetoptac = Ssrcommon.tclWITHTOP Ssrelim.ssrscase_or_inj_tac let ssrselimtoptac = Ssrcommon.tclWITHTOP Ssrelim.elimtac (** [move] **************************************************************) let pushmoveeqtac cl c = Goal.enter begin fun g -> let env, sigma = Goal.(env g, sigma g) in let x, t, cl1 = EConstr.destProd sigma cl in let cl2, eqc, sigma = mkEq R2L cl1 c t 1 env sigma in Unsafe.tclEVARS sigma <*> Tactics.apply_type ~typecheck:true (EConstr.mkProd (x, t, cl2)) [c; eqc] end let eqmovetac _ gen = Ssrcommon.pfLIFT (Ssrcommon.pf_interp_gen false gen) >>= fun (cl, c, _) -> pushmoveeqtac cl c ;; let rec eqmoveipats eqpat = function | (IOpSimpl _ | IOpClear _ as ipat) :: ipats -> ipat :: eqmoveipats eqpat ipats | (IOpInaccessibleAll :: _ | []) as ipats -> IOpInaccessible None :: eqpat @ ipats | ipat :: ipats -> ipat :: eqpat @ ipats let ssrsmovetac = Goal.enter begin fun g -> let sigma, concl = Goal.(sigma g, concl g) in match EConstr.kind sigma concl with | Prod _ | LetIn _ -> tclUNIT () | _ -> Tactics.hnf_in_concl end let tclIPAT ip = IpatMachine.main ~first_case_is_dispatch:false ip let ssrmovetac = function | _::_ as view, (_, ({ gens = lastgen :: gens; clr }, ipats)) -> let gentac = V82.tactic ~nf_evars:false (Ssrcommon.genstac (gens, [])) in let conclusion _ t clear ccl = Tactics.apply_type ~typecheck:true ccl [t] <*> Tactics.clear (List.map Ssrcommon.hyp_id clear) in gentac <*> tclLAST_GEN ~to_ind:false lastgen (tacVIEW_THEN_GRAB view ~conclusion) <*> tclIPAT (IOpClear (clr,None) :: IpatMachine.tclCompileIPats ipats) | _::_ as view, (_, ({ gens = []; clr }, ipats)) -> tclIPAT (IOpView (None,view) :: IOpClear (clr,None) :: IpatMachine.tclCompileIPats ipats) | _, (Some pat, (dgens, ipats)) -> let dgentac = with_dgens dgens eqmovetac in dgentac <*> tclIPAT (eqmoveipats (IpatMachine.tclCompileIPats [pat]) (IpatMachine.tclCompileIPats ipats)) | _, (_, ({ gens = (_ :: _ as gens); dgens = []; clr}, ipats)) -> let gentac = V82.tactic ~nf_evars:false (Ssrcommon.genstac (gens, clr)) in gentac <*> tclIPAT (IpatMachine.tclCompileIPats ipats) | _, (_, ({ clr }, ipats)) -> Tacticals.New.tclTHENLIST [ssrsmovetac; Tactics.clear (List.map Ssrcommon.hyp_id clr); tclIPAT (IpatMachine.tclCompileIPats ipats)] (** [abstract: absvar gens] **************************************************) let rec is_Evar_or_CastedMeta sigma x = EConstr.isEvar sigma x || EConstr.isMeta sigma x || (EConstr.isCast sigma x && is_Evar_or_CastedMeta sigma (pi1 (EConstr.destCast sigma x))) let occur_existential_or_casted_meta sigma c = let rec occrec c = match EConstr.kind sigma c with | Evar _ -> raise Not_found | Cast (m,_,_) when EConstr.isMeta sigma m -> raise Not_found | _ -> EConstr.iter sigma occrec c in try occrec c; false with Not_found -> true let tacEXAMINE_ABSTRACT id = Ssrcommon.tacTYPEOF id >>= begin fun tid -> Ssrcommon.tacMK_SSR_CONST "abstract" >>= fun abstract -> Goal.enter_one ~__LOC__ begin fun g -> let sigma, env = Goal.(sigma g, env g) in let err () = Ssrcommon.errorstrm Pp.(strbrk"not a proper abstract constant: "++ Printer.pr_econstr_env env sigma id) in if not (EConstr.isApp sigma tid) then err (); let hd, args_id = EConstr.destApp sigma tid in if not (EConstr.eq_constr_nounivs sigma hd abstract) then err (); if Array.length args_id <> 3 then err (); if not (is_Evar_or_CastedMeta sigma args_id.(2)) then Ssrcommon.errorstrm Pp.(strbrk"abstract constant "++ Printer.pr_econstr_env env sigma id++str" already used"); tclUNIT (tid, args_id) end end let tacFIND_ABSTRACT_PROOF check_lock abstract_n = Ssrcommon.tacMK_SSR_CONST "abstract" >>= fun abstract -> Goal.enter_one ~__LOC__ begin fun g -> let sigma, env = Goal.(sigma g, env g) in let l = Evd.fold_undefined (fun e ei l -> match EConstr.kind sigma ei.Evd.evar_concl with | App(hd, [|ty; n; lock|]) when (not check_lock || (occur_existential_or_casted_meta sigma ty && is_Evar_or_CastedMeta sigma lock)) && EConstr.eq_constr_nounivs sigma hd abstract && EConstr.eq_constr_nounivs sigma n abstract_n -> e :: l | _ -> l) sigma [] in match l with | [e] -> tclUNIT e | _ -> Ssrcommon.errorstrm Pp.(strbrk"abstract constant "++ Printer.pr_econstr_env env sigma abstract_n ++ strbrk" not found in the evar map exactly once. "++ strbrk"Did you tamper with it?") end let ssrabstract dgens = let main _ (_,cid) = Goal.enter begin fun g -> Ssrcommon.tacMK_SSR_CONST "abstract" >>= fun abstract -> Ssrcommon.tacMK_SSR_CONST "abstract_key" >>= fun abstract_key -> Ssrcommon.tacINTERP_CPATTERN cid >>= fun cid -> let id = EConstr.mkVar (Option.get (Ssrmatching.id_of_pattern cid)) in tacEXAMINE_ABSTRACT id >>= fun (idty, args_id) -> let abstract_n = args_id.(1) in tacFIND_ABSTRACT_PROOF true abstract_n >>= fun abstract_proof -> let tacFIND_HOLE = Goal.enter_one ~__LOC__ begin fun g -> let sigma, env, concl = Goal.(sigma g, env g, concl g) in let t = args_id.(0) in match EConstr.kind sigma t with | (Evar _ | Meta _) -> Ssrcommon.tacUNIFY concl t <*> tclUNIT id | Cast(m,_,_) when EConstr.isEvar sigma m || EConstr.isMeta sigma m -> Ssrcommon.tacUNIFY concl t <*> tclUNIT id | _ -> Ssrcommon.errorstrm Pp.(strbrk"abstract constant "++ Printer.pr_econstr_env env sigma abstract_n ++ strbrk" has an unexpected shape. Did you tamper with it?") end in tacFIND_HOLE >>= fun proof -> Ssrcommon.tacUNIFY abstract_key args_id.(2) <*> Ssrcommon.tacTYPEOF idty >>= fun _ -> Unsafe.tclGETGOALS >>= fun goals -> (* Here we jump in the proof tree: we move from the current goal to the evar that inhabits the abstract variable with the current goal *) Unsafe.tclSETGOALS (goals @ [Proofview_monad.with_empty_state abstract_proof]) <*> tclDISPATCH [ Tacticals.New.tclSOLVE [Tactics.apply proof]; Ssrcommon.unfold[abstract;abstract_key] ] end in let interp_gens { gens } ~conclusion = Goal.enter begin fun g -> Ssrcommon.tacSIGMA >>= fun gl0 -> let open Ssrmatching in let ipats = List.map (fun (_,cp) -> match id_of_pattern (interp_cpattern gl0 cp None) with | None -> IPatAnon (One None) | Some id -> IPatId id) (List.tl gens) in conclusion ipats end in interp_gens dgens ~conclusion:(fun ipats -> with_dgens dgens main <*> tclIPATssr ipats) module Internal = struct let pf_find_abstract_proof b gl t = let res = ref None in let _ = V82.of_tactic (tacFIND_ABSTRACT_PROOF b (EConstr.of_constr t) >>= fun x -> res := Some x; tclUNIT ()) gl in match !res with | None -> assert false | Some x -> x let examine_abstract t gl = let res = ref None in let _ = V82.of_tactic (tacEXAMINE_ABSTRACT t >>= fun x -> res := Some x; tclUNIT ()) gl in match !res with | None -> assert false | Some x -> x end (* vim: set filetype=ocaml foldmethod=marker: *) coq-8.11.0/plugins/ssr/ssrparser.mli0000644000175000017500000001117613612664533017313 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Evd.evar_map -> 'a -> 'b -> (Environ.env -> Evd.evar_map -> Notation_gram.tolerability -> 'c) -> 'c val ssrtclarg : Tacexpr.raw_tactic_expr Pcoq.Entry.t val wit_ssrtclarg : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr, Geninterp.Val.t) Genarg.genarg_type val pr_ssrtclarg : Environ.env -> Evd.evar_map -> 'a -> 'b -> (Environ.env -> Evd.evar_map -> Notation_gram.tolerability -> 'c -> 'd) -> 'c -> 'd val add_genarg : string -> (Environ.env -> Evd.evar_map -> 'a -> Pp.t) -> 'a Genarg.uniform_genarg_type (* Parsing witnesses, needed to serialize ssreflect syntax *) open Ssrmatching_plugin open Ssrmatching open Ssrast open Ssrequality type ssrfwdview = ast_closure_term list type ssreqid = ssripat option type ssrarg = ssrfwdview * (ssreqid * (cpattern ssragens * ssripats)) val wit_ssrseqdir : ssrdir Genarg.uniform_genarg_type val wit_ssrseqarg : (Tacexpr.raw_tactic_expr ssrseqarg, Tacexpr.glob_tactic_expr ssrseqarg, Geninterp.Val.t ssrseqarg) Genarg.genarg_type val wit_ssrintrosarg : (Tacexpr.raw_tactic_expr * ssripats, Tacexpr.glob_tactic_expr * ssripats, Geninterp.Val.t * ssripats) Genarg.genarg_type val wit_ssrsufffwd : (Tacexpr.raw_tactic_expr ffwbinders, Tacexpr.glob_tactic_expr ffwbinders, Geninterp.Val.t ffwbinders) Genarg.genarg_type val wit_ssripatrep : ssripat Genarg.uniform_genarg_type val wit_ssrarg : ssrarg Genarg.uniform_genarg_type val wit_ssrrwargs : ssrrwarg list Genarg.uniform_genarg_type val wit_ssrclauses : clauses Genarg.uniform_genarg_type val wit_ssrcasearg : (cpattern ssragens) ssrmovearg Genarg.uniform_genarg_type val wit_ssrmovearg : (cpattern ssragens) ssrmovearg Genarg.uniform_genarg_type val wit_ssrapplyarg : ssrapplyarg Genarg.uniform_genarg_type val wit_ssrhavefwdwbinders : (Tacexpr.raw_tactic_expr fwdbinders, Tacexpr.glob_tactic_expr fwdbinders, Tacinterp.Value.t fwdbinders) Genarg.genarg_type val wit_ssrhintarg : (Tacexpr.raw_tactic_expr ssrhint, Tacexpr.glob_tactic_expr ssrhint, Tacinterp.Value.t ssrhint) Genarg.genarg_type val wit_ssrexactarg : ssrapplyarg Genarg.uniform_genarg_type val wit_ssrcongrarg : ((int * ssrterm) * cpattern ssragens) Genarg.uniform_genarg_type val wit_ssrfwdid : Names.Id.t Genarg.uniform_genarg_type val wit_ssrsetfwd : ((ssrfwdfmt * (cpattern * ast_closure_term option)) * ssrdocc) Genarg.uniform_genarg_type val wit_ssrdoarg : (Tacexpr.raw_tactic_expr ssrdoarg, Tacexpr.glob_tactic_expr ssrdoarg, Tacinterp.Value.t ssrdoarg) Genarg.genarg_type val wit_ssrhint : (Tacexpr.raw_tactic_expr ssrhint, Tacexpr.glob_tactic_expr ssrhint, Tacinterp.Value.t ssrhint) Genarg.genarg_type val wit_ssrhpats : ssrhpats Genarg.uniform_genarg_type val wit_ssrhpats_nobs : ssrhpats Genarg.uniform_genarg_type val wit_ssrhpats_wtransp : ssrhpats_wtransp Genarg.uniform_genarg_type val wit_ssrposefwd : (ssrfwdfmt * ast_closure_term) Genarg.uniform_genarg_type val wit_ssrrpat : ssripat Genarg.uniform_genarg_type val wit_ssrterm : ssrterm Genarg.uniform_genarg_type val wit_ssrunlockarg : (ssrocc * ssrterm) Genarg.uniform_genarg_type val wit_ssrunlockargs : (ssrocc * ssrterm) list Genarg.uniform_genarg_type val wit_ssrwgen : clause Genarg.uniform_genarg_type val wit_ssrwlogfwd : (clause list * (ssrfwdfmt * ast_closure_term)) Genarg.uniform_genarg_type val wit_ssrfixfwd : (Names.Id.t * (ssrfwdfmt * ast_closure_term)) Genarg.uniform_genarg_type val wit_ssrfwd : (ssrfwdfmt * ast_closure_term) Genarg.uniform_genarg_type val wit_ssrfwdfmt : ssrfwdfmt Genarg.uniform_genarg_type val wit_ssrcpat : ssripat Genarg.uniform_genarg_type val wit_ssrdgens : cpattern ssragens Genarg.uniform_genarg_type val wit_ssrdgens_tl : cpattern ssragens Genarg.uniform_genarg_type val wit_ssrdir : ssrdir Genarg.uniform_genarg_type coq-8.11.0/plugins/ssr/ssrvernac.mli0000644000175000017500000000136013612664533017267 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ((ssrfwdfmt * (Ssrmatching_plugin.Ssrmatching.cpattern * ast_closure_term option)) * ssrdocc) -> v82tac val ssrposetac : Id.t * (ssrfwdfmt * ast_closure_term) -> v82tac val havetac : ist -> bool * ((((Ssrast.ssrclear option * Ssrast.ssripat list) * Ssrast.ssripats) * Ssrast.ssripats) * (((Ssrast.ssrfwdkind * 'a) * ast_closure_term) * (bool * Tacinterp.Value.t option list))) -> bool -> bool -> v82tac val basecuttac : string -> EConstr.t -> Goal.goal Evd.sigma -> Evar.t list Evd.sigma val wlogtac : Ltac_plugin.Tacinterp.interp_sign -> ((Ssrast.ssrclear option * Ssrast.ssripats) * 'a) * 'b -> (Ssrast.ssrhyps * ((Ssrast.ssrhyp_or_id * string) * Ssrmatching_plugin.Ssrmatching.cpattern option) option) list * ('c * ast_closure_term) -> Ltac_plugin.Tacinterp.Value.t Ssrast.ssrhint -> bool -> [< `Gen of Names.Id.t option option | `NoGen > `NoGen ] -> Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val sufftac : Ssrast.ist -> (((Ssrast.ssrclear option * Ssrast.ssripats) * Ssrast.ssripat list) * Ssrast.ssripat list) * (('a * ast_closure_term) * (bool * Tacinterp.Value.t option list)) -> Tacmach.tactic (* pad_intro (by default false) indicates whether the intro-pattern "=> i..." must be turned into "=> [i...|i...|i...|]" (n+1 branches, assuming there are n provided tactics in the ssrhint argument "do [...|...|...]"; it is useful when the intro-pattern is "=> *"). Otherwise, "=> i..." is turned into "=> [i...|]". *) val undertac : ?pad_intro:bool -> Ltac_plugin.Tacinterp.interp_sign -> Ssrast.ssripats option -> Ssrequality.ssrrwarg -> Ltac_plugin.Tacinterp.Value.t Ssrast.ssrhint -> unit Proofview.tactic val overtac : unit Proofview.tactic coq-8.11.0/plugins/ssr/plugin_base.dune0000644000175000017500000000033413612664533017723 0ustar treinentreinen(library (name ssreflect_plugin) (public_name coq.plugins.ssreflect) (synopsis "Coq's ssreflect plugin") (modules_without_implementation ssrast) (flags :standard -open Gramlib) (libraries coq.plugins.ssrmatching)) coq-8.11.0/plugins/nsatz/0000755000175000017500000000000013612664533015106 5ustar treinentreinencoq-8.11.0/plugins/nsatz/ideal.ml0000644000175000017500000004631313612664533016525 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int array val make : int array -> t val deg : t -> int val nvar : t -> int val var_mon : int -> int -> t val mult_mon : t -> t -> t val compare_mon : t -> t -> int val div_mon : t -> t -> t val div_mon_test : t -> t -> bool val ppcm_mon : t -> t -> t val const_mon : int -> t end = struct type t = int array type mon = t let repr m = m let make m = m let nvar (m : mon) = Array.length m - 1 let deg (m : mon) = m.(0) let mult_mon (m : mon) (m' : mon) = let d = nvar m in let m'' = Array.make (d+1) 0 in for i=0 to d do m''.(i)<- (m.(i)+m'.(i)); done; m'' let compare_mon (m : mon) (m' : mon) = let d = nvar m in if !lexico then ( (* Comparaison de monomes avec ordre du degre lexicographique = on commence par regarder la 1ere variable*) let res=ref 0 in let i=ref 1 in (* 1 si lexico pur 0 si degre*) while (!res=0) && (!i<=d) do res:= (Int.compare m.(!i) m'.(!i)); i:=!i+1; done; !res) else ( (* degre lexicographique inverse *) match Int.compare m.(0) m'.(0) with | 0 -> (* meme degre total *) let res=ref 0 in let i=ref d in while (!res=0) && (!i>=1) do res:= - (Int.compare m.(!i) m'.(!i)); i:=!i-1; done; !res | x -> x) let div_mon m m' = let d = nvar m in let m'' = Array.make (d+1) 0 in for i=0 to d do m''.(i)<- (m.(i)-m'.(i)); done; m'' (* m' divides m *) let div_mon_test m m' = let d = nvar m in let res=ref true in let i=ref 0 in (*il faut que le degre total soit bien mis sinon, i=ref 1*) while (!res) && (!i<=d) do res:= (m.(!i) >= m'.(!i)); i:=succ !i; done; !res let set_deg m = let d = nvar m in m.(0)<-0; for i=1 to d do m.(0)<- m.(i)+m.(0); done; m (* lcm *) let ppcm_mon m m' = let d = nvar m in let m'' = Array.make (d+1) 0 in for i=1 to d do m''.(i)<- (max m.(i) m'.(i)); done; set_deg m'' (* returns a constant polynom ial with d variables *) let const_mon d = let m = Array.make (d+1) 0 in let m = set_deg m in m let var_mon d i = let m = Array.make (d+1) 0 in m.(i) <- 1; let m = set_deg m in m end (*********************************************************************** Functor *) module Make (P:Polynom.S) = struct type coef = P.t let coef0 = P.of_num (Num.Int 0) let coef1 = P.of_num (Num.Int 1) let string_of_coef c = "["^(P.to_string c)^"]" (*********************************************************************** Monomials array of integers, first is the degree *) open Monomial type mon = Monomial.t type deg = int type poly = (coef * mon) list type polynom = { pol : poly; num : int; } (********************************************************************** Polynomials list of (coefficient, monomial) decreasing order *) let repr p = p let equal = Util.List.for_all2eq (fun (c1,m1) (c2,m2) -> P.equal c1 c2 && m1=m2) let hash p = let c = List.map fst p in let m = List.map snd p in List.fold_left (fun h p -> h * 17 + P.hash p) (Hashtbl.hash m) c module Hashpol = Hashtbl.Make( struct type t = poly let equal = equal let hash = hash end) (* A pretty printer for polynomials, with Maple-like syntax. *) let getvar lv i = try (List.nth lv i) with Failure _ -> (List.fold_left (fun r x -> r^" "^x) "lv= " lv) ^" i="^(string_of_int i) let string_of_pol zeroP hdP tlP coefterm monterm string_of_coef dimmon string_of_exp lvar p = let rec string_of_mon m coefone = let s=ref [] in for i=1 to (dimmon m) do (match (string_of_exp m i) with "0" -> () | "1" -> s:= (!s) @ [(getvar lvar (i-1))] | e -> s:= (!s) @ [((getvar lvar (i-1)) ^ "^" ^ e)]); done; (match !s with [] -> if coefone then "1" else "" | l -> if coefone then (String.concat "*" l) else ( "*" ^ (String.concat "*" l))) and string_of_term t start = let a = coefterm t and m = monterm t in match (string_of_coef a) with "0" -> "" | "1" ->(match start with true -> string_of_mon m true |false -> ( "+ "^ (string_of_mon m true))) | "-1" ->( "-" ^" "^(string_of_mon m true)) | c -> if (String.get c 0)='-' then ( "- "^ (String.sub c 1 ((String.length c)-1))^ (string_of_mon m false)) else (match start with true -> ( c^(string_of_mon m false)) |false -> ( "+ "^ c^(string_of_mon m false))) and stringP p start = if (zeroP p) then (if start then ("0") else "") else ((string_of_term (hdP p) start)^ " "^ (stringP (tlP p) false)) in (stringP p true) let stringP metadata (p : poly) = string_of_pol (fun p -> match p with [] -> true | _ -> false) (fun p -> match p with (t::p) -> t |_ -> failwith "print_pol dans dansideal") (fun p -> match p with (t::p) -> p |_ -> failwith "print_pol dans dansideal") (fun (a,m) -> a) (fun (a,m) -> m) string_of_coef (fun m -> (Array.length (Monomial.repr m))-1) (fun m i -> (string_of_int ((Monomial.repr m).(i)))) metadata.name_var p let nsP2 = 10 let stringPcut metadata (p : poly) = (*Polynomesrec.nsP1:=20;*) let res = if (List.length p)> nsP2 then (stringP metadata [List.hd p])^" + "^(string_of_int (List.length p))^" terms" else stringP metadata p in (*Polynomesrec.nsP1:= max_int;*) res (* Operations *) let zeroP = [] (* returns a constant polynom ial with d variables *) let polconst d c = let m = const_mon d in [(c,m)] let plusP p q = let rec plusP p q accu = match p, q with | [], [] -> List.rev accu | [], _ -> List.rev_append accu q | _, [] -> List.rev_append accu p | t :: p', t' :: q' -> let c = compare_mon (snd t) (snd t') in if c > 0 then plusP p' q (t :: accu) else if c < 0 then plusP p q' (t' :: accu) else let c = P.plusP (fst t) (fst t') in if P.equal c coef0 then plusP p' q' accu else plusP p' q' ((c, (snd t)) :: accu) in plusP p q [] (* multiplication by (a,monomial) *) let mult_t_pol a m p = let map (b, m') = (P.multP a b, mult_mon m m') in CList.map map p let coef_of_int x = P.of_num (Num.Int x) (* variable i *) let gen d i = let m = var_mon d i in [((coef_of_int 1),m)] let oppP p = let rec oppP p = match p with [] -> [] |(b,m')::p -> ((P.oppP b),m')::(oppP p) in oppP p (* multiplication by a coefficient *) let emultP a p = let rec emultP p = match p with [] -> [] |(b,m')::p -> ((P.multP a b),m')::(emultP p) in emultP p let multP p q = let rec aux p accu = match p with [] -> accu |(a,m)::p' -> aux p' (plusP (mult_t_pol a m q) accu) in aux p [] let puisP p n= match p with [] -> [] |_ -> if n = 0 then let d = nvar (snd (List.hd p)) in [coef1, const_mon d] else let rec puisP p n = if n = 1 then p else let q = puisP p (n / 2) in let q = multP q q in if n mod 2 = 0 then q else multP p q in puisP p n (*********************************************************************** Division of polynomials *) type table = { hmon : (mon, poly) Hashtbl.t option; (* coefficients of polynomials when written with initial polynomials *) coefpoldep : ((int * int), poly) Hashtbl.t; mutable nallpol : int; mutable allpol : polynom array; (* list of initial polynomials *) } let pgcdpos a b = P.pgcdP a b let polynom0 = { pol = []; num = 0 } let ppol p = p.pol let lm p = snd (List.hd (ppol p)) let new_allpol table p = table.nallpol <- table.nallpol + 1; if table.nallpol >= Array.length table.allpol then table.allpol <- Array.append table.allpol (Array.make table.nallpol polynom0); let p = { pol = p; num = table.nallpol } in table.allpol.(table.nallpol) <- p; p (* returns a polynomial of l whose head monomial divides m, else [] *) let rec selectdiv m l = match l with [] -> polynom0 |q::r -> let m'= snd (List.hd (ppol q)) in match (div_mon_test m m') with true -> q |false -> selectdiv m r let div_pol p q a b m = plusP (emultP a p) (mult_t_pol b m q) let find_hmon table m = match table.hmon with | None -> raise Not_found | Some hmon -> Hashtbl.find hmon m let add_hmon table m q = match table.hmon with | None -> () | Some hmon -> Hashtbl.add hmon m q let selectdiv table m l = try find_hmon table m with Not_found -> let q = selectdiv m l in let q = ppol q in match q with | [] -> q | _ :: _ -> let () = add_hmon table m q in q let div_coef a b = P.divP a b (* remainder r of the division of p by polynomials of l, returns (c,r) where c is the coefficient for pseudo-division : c p = sum_i q_i p_i + r *) let reduce2 table p l = let l = if nouveaux_pol_en_tete then List.rev l else l in let rec reduce p = match p with [] -> (coef1,[]) |t::p' -> let (a,m)=t in let q = selectdiv table m l in match q with [] -> if reduire_les_queues then let (c,r)=(reduce p') in (c,((P.multP a c,m)::r)) else (coef1,p) |(b,m')::q' -> let c=(pgcdpos a b) in let a'= (div_coef b c) in let b'=(P.oppP (div_coef a c)) in let (e,r)=reduce (div_pol p' q' a' b' (div_mon m m')) in (P.multP a' e,r) in let (c,r) = reduce p in (c,r) (* coef of q in p = sum_i c_i*q_i *) let coefpoldep_find table p q = try (Hashtbl.find table.coefpoldep (p.num,q.num)) with Not_found -> [] let coefpoldep_set table p q c = Hashtbl.add table.coefpoldep (p.num,q.num) c (* keeps trace in coefpoldep divides without pseudodivisions *) let reduce2_trace table p l lcp = let lp = l in let l = if nouveaux_pol_en_tete then List.rev l else l in (* rend (lq,r), ou r = p + sum(lq) *) let rec reduce p = match p with [] -> ([],[]) |t::p' -> let (a,m)=t in let q = selectdiv table m l in match q with [] -> if reduire_les_queues then let (lq,r)=(reduce p') in (lq,((a,m)::r)) else ([],p) |(b,m')::q' -> let b'=(P.oppP (div_coef a b)) in let m''= div_mon m m' in let p1=plusP p' (mult_t_pol b' m'' q') in let (lq,r)=reduce p1 in ((b',m'',q)::lq, r) in let (lq,r) = reduce p in (List.map2 (fun c0 q -> let c = List.fold_left (fun x (a,m,s) -> if equal s (ppol q) then plusP x (mult_t_pol a m (polconst (nvar m) (coef_of_int 1))) else x) c0 lq in c) lcp lp, r) (*********************************************************************** Completion *) let spol0 ps qs= let p = ppol ps in let q = ppol qs in let m = snd (List.hd p) in let m'= snd (List.hd q) in let a = fst (List.hd p) in let b = fst (List.hd q) in let p'= List.tl p in let q'= List.tl q in let c = (pgcdpos a b) in let m''=(ppcm_mon m m') in let m1 = div_mon m'' m in let m2 = div_mon m'' m' in let fsp p' q' = plusP (mult_t_pol (div_coef b c) m1 p') (mult_t_pol (P.oppP (div_coef a c)) m2 q') in let sp = fsp p' q' in let p0 = fsp (polconst (nvar m) (coef_of_int 1)) [] in let q0 = fsp [] (polconst (nvar m) (coef_of_int 1)) in (sp, p0, q0) let etrangers p p'= let m = snd (List.hd p) in let m'= snd (List.hd p') in let d = nvar m in let res=ref true in let i=ref 1 in while (!res) && (!i<=d) do res:= ((Monomial.repr m).(!i) = 0) || ((Monomial.repr m').(!i)=0); i:=!i+1; done; !res let addS x l = l @ [x] (* oblige de mettre en queue sinon le certificat deconne *) (*********************************************************************** critical pairs/s-polynomials *) module CPair = struct type t = (int * int) * Monomial.t let compare ((i1, j1), m1) ((i2, j2), m2) = compare_mon m2 m1 end module Heap : sig type elt = (int * int) * Monomial.t type t val length : t -> int val empty : t val add : elt -> t -> t val pop : t -> (elt * t) option end = struct include Heap.Functional(CPair) let length h = fold (fun _ accu -> accu + 1) h 0 let pop h = try Some (maximum h, remove h) with Heap.EmptyHeap -> None end let ord i j = if i cpair p q r) accu lq let rec cpairs l accu = match l with | [] | [_] -> accu | p :: l -> cpairs l (cpairs1 p l accu) let critere3 table ((i,j),m) lp lcp = List.exists (fun h -> h.num <> i && h.num <> j && (div_mon_test m (lm h)) && (h.num < j || not (m = ppcm_mon (lm (table.allpol.(i))) (lm h))) && (h.num < i || not (m = ppcm_mon (lm (table.allpol.(j))) (lm h)))) lp let infobuch p q = (info (fun () -> Printf.sprintf "[%i,%i]" (List.length p) (Heap.length q))) (* in lp new polynomials are at the end *) type certificate = { coef : coef; power : int; gb_comb : poly list list; last_comb : poly list } type current_problem = { cur_poly : poly; cur_coef : coef; } exception NotInIdealUpdate of current_problem let test_dans_ideal cur_pb table metadata p lp len0 = (* Invariant: [lp] is [List.tl (Array.to_list table.allpol)] *) let (c,r) = reduce2 table cur_pb.cur_poly lp in info (fun () -> "remainder: "^(stringPcut metadata r)); let cur_pb = { cur_coef = P.multP cur_pb.cur_coef c; cur_poly = r; } in match r with | [] -> sinfo "polynomial reduced to 0"; let lcp = List.map (fun q -> []) lp in let c = cur_pb.cur_coef in let (lcq,r) = reduce2_trace table (emultP c p) lp lcp in sinfo "r ok"; info (fun () -> "r: "^(stringP metadata r)); info (fun () -> let fold res cq q = plusP res (multP cq (ppol q)) in let res = List.fold_left2 fold (emultP c p) lcq lp in "verif sum: "^(stringP metadata res) ); info (fun () -> "coefficient: "^(stringP metadata (polconst 1 c))); let coefficient_multiplicateur = c in let liste_des_coefficients_intermediaires = let rec aux accu lp = match lp with | [] -> accu | p :: lp -> let elt = List.map (fun q -> coefpoldep_find table p q) lp in aux (elt :: accu) lp in let lci = aux [] (List.rev lp) in CList.skipn len0 lci in let liste_des_coefficients = List.rev_map (fun cq -> emultP (coef_of_int (-1)) cq) lcq in {coef = coefficient_multiplicateur; power = 1; gb_comb = liste_des_coefficients_intermediaires; last_comb = liste_des_coefficients} | _ -> raise (NotInIdealUpdate cur_pb) let deg_hom p = match p with | [] -> -1 | (a,m)::_ -> Monomial.deg m let pbuchf table metadata cur_pb homogeneous (lp, lpc) p = (* Invariant: [lp] is [List.tl (Array.to_list table.allpol)] *) sinfo "computation of the Groebner basis"; let () = match table.hmon with | None -> () | Some hmon -> Hashtbl.clear hmon in let len0 = List.length lp in let rec pbuchf cur_pb (lp, lpc) = infobuch lp lpc; match Heap.pop lpc with | None -> test_dans_ideal cur_pb table metadata p lp len0 | Some (((i, j), m), lpc2) -> if critere3 table ((i,j),m) lp lpc2 then (sinfo "c"; pbuchf cur_pb (lp, lpc2)) else let (a0, p0, q0) = spol0 table.allpol.(i) table.allpol.(j) in if homogeneous && a0 <>[] && deg_hom a0 > deg_hom cur_pb.cur_poly then (sinfo "h"; pbuchf cur_pb (lp, lpc2)) else (* let sa = a.sugar in*) match reduce2 table a0 lp with _, [] -> sinfo "0";pbuchf cur_pb (lp, lpc2) | ca, _ :: _ -> (* info "pair reduced\n";*) let map q = let r = if q.num == i then p0 else if q.num == j then q0 else [] in emultP ca r in let lcp = List.map map lp in let (lca, a0) = reduce2_trace table (emultP ca a0) lp lcp in (* info "paire re-reduced";*) let a = new_allpol table a0 in List.iter2 (fun c q -> coefpoldep_set table a q c) lca lp; let a0 = a in info (fun () -> "new polynomial: "^(stringPcut metadata (ppol a0))); let nlp = addS a0 lp in try test_dans_ideal cur_pb table metadata p nlp len0 with NotInIdealUpdate cur_pb -> let newlpc = cpairs1 a0 lp lpc2 in pbuchf cur_pb (nlp, newlpc) in pbuchf cur_pb (lp, lpc) let is_homogeneous p = match p with | [] -> true | (a,m)::p1 -> let d = deg m in List.for_all (fun (b,m') -> deg m' =d) p1 (* returns c lp = [pn;...;p1] p lci = [[a(n+1,n);...;a(n+1,1)]; [a(n+2,n+1);...;a(n+2,1)]; ... [a(n+m,n+m-1);...;a(n+m,1)]] lc = [qn+m; ... q1] such that c*p = sum qi*pi where pn+k = a(n+k,n+k-1)*pn+k-1 + ... + a(n+k,1)* p1 *) let in_ideal metadata d lp p = let table = { hmon = None; coefpoldep = Hashtbl.create 51; nallpol = 0; allpol = Array.make 1000 polynom0; } in let homogeneous = List.for_all is_homogeneous (p::lp) in if homogeneous then sinfo "homogeneous polynomials"; info (fun () -> "p: "^(stringPcut metadata p)); info (fun () -> "lp:\n"^(List.fold_left (fun r p -> r^(stringPcut metadata p)^"\n") "" lp)); let lp = List.map (fun c -> new_allpol table c) lp in List.iter (fun p -> coefpoldep_set table p p (polconst d (coef_of_int 1))) lp; let cur_pb = { cur_poly = p; cur_coef = coef1; } in let cert = try pbuchf table metadata cur_pb homogeneous (lp, Heap.empty) p with NotInIdealUpdate cur_pb -> try pbuchf table metadata cur_pb homogeneous (lp, cpairs lp Heap.empty) p with NotInIdealUpdate _ -> raise NotInIdeal in sinfo "computed"; cert end coq-8.11.0/plugins/nsatz/nsatz.ml0000644000175000017500000004125513612664533016606 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 1 let puis = power_big_int_positive_int (* a et b positifs, résultat positif *) let rec pgcd a b = if equal b coef0 then a else if lt a b then pgcd b a else pgcd b (modulo a b) end (* module Ent = struct type t = Entiers.entiers let of_int = Entiers.ent_of_int let of_num x = Entiers.ent_of_string(Num.string_of_num x) let to_num x = Num.num_of_string (Entiers.string_of_ent x) let equal = Entiers.eq_ent let lt = Entiers.lt_ent let le = Entiers.le_ent let abs = Entiers.abs_ent let plus =Entiers.add_ent let mult = Entiers.mult_ent let sub = Entiers.moins_ent let opp = Entiers.opp_ent let div = Entiers.div_ent let modulo = Entiers.mod_ent let coef0 = Entiers.ent0 let coef1 = Entiers.ent1 let to_string = Entiers.string_of_ent let to_int x = Entiers.int_of_ent x let hash x =Entiers.hash_ent x let signe = Entiers.signe_ent let rec puis p n = match n with 0 -> coef1 |_ -> (mult p (puis p (n-1))) (* a et b positifs, résultat positif *) let rec pgcd a b = if equal b coef0 then a else if lt a b then pgcd b a else pgcd b (modulo a b) (* signe du pgcd = signe(a)*signe(b) si non nuls. *) let pgcd2 a b = if equal a coef0 then b else if equal b coef0 then a else let c = pgcd (abs a) (abs b) in if ((lt coef0 a)&&(lt b coef0)) ||((lt coef0 b)&&(lt a coef0)) then opp c else c end *) (* ------------------------------------------------------------------------- *) (* ------------------------------------------------------------------------- *) type vname = string type term = | Zero | Const of Num.num | Var of vname | Opp of term | Add of term * term | Sub of term * term | Mul of term * term | Pow of term * int let const n = if eq_num n num_0 then Zero else Const n let pow(p,i) = if Int.equal i 1 then p else Pow(p,i) let add = function (Zero,q) -> q | (p,Zero) -> p | (p,q) -> Add(p,q) let mul = function (Zero,_) -> Zero | (_,Zero) -> Zero | (p,Const n) when eq_num n num_1 -> p | (Const n,q) when eq_num n num_1 -> q | (p,q) -> Mul(p,q) let gen_constant n = lazy (UnivGen.constr_of_monomorphic_global (Coqlib.lib_ref n)) let tpexpr = gen_constant "plugins.setoid_ring.pexpr" let ttconst = gen_constant "plugins.setoid_ring.const" let ttvar = gen_constant "plugins.setoid_ring.var" let ttadd = gen_constant "plugins.setoid_ring.add" let ttsub = gen_constant "plugins.setoid_ring.sub" let ttmul = gen_constant "plugins.setoid_ring.mul" let ttopp = gen_constant "plugins.setoid_ring.opp" let ttpow = gen_constant "plugins.setoid_ring.pow" let tlist = gen_constant "core.list.type" let lnil = gen_constant "core.list.nil" let lcons = gen_constant "core.list.cons" let tz = gen_constant "num.Z.type" let z0 = gen_constant "num.Z.Z0" let zpos = gen_constant "num.Z.Zpos" let zneg = gen_constant "num.Z.Zneg" let pxI = gen_constant "num.pos.xI" let pxO = gen_constant "num.pos.xO" let pxH = gen_constant "num.pos.xH" let nN0 = gen_constant "num.N.N0" let nNpos = gen_constant "num.N.Npos" let mkt_app name l = mkApp (Lazy.force name, Array.of_list l) let tlp () = mkt_app tlist [mkt_app tpexpr [Lazy.force tz]] let tllp () = mkt_app tlist [tlp()] let rec mkt_pos n = if n =/ num_1 then Lazy.force pxH else if mod_num n num_2 =/ num_0 then mkt_app pxO [mkt_pos (quo_num n num_2)] else mkt_app pxI [mkt_pos (quo_num n num_2)] let mkt_n n = if Num.eq_num n num_0 then Lazy.force nN0 else mkt_app nNpos [mkt_pos n] let mkt_z z = if z =/ num_0 then Lazy.force z0 else if z >/ num_0 then mkt_app zpos [mkt_pos z] else mkt_app zneg [mkt_pos ((Int 0) -/ z)] let rec mkt_term t = match t with | Zero -> mkt_term (Const num_0) | Const r -> let (n,d) = numdom r in mkt_app ttconst [Lazy.force tz; mkt_z n] | Var v -> mkt_app ttvar [Lazy.force tz; mkt_pos (num_of_string v)] | Opp t1 -> mkt_app ttopp [Lazy.force tz; mkt_term t1] | Add (t1,t2) -> mkt_app ttadd [Lazy.force tz; mkt_term t1; mkt_term t2] | Sub (t1,t2) -> mkt_app ttsub [Lazy.force tz; mkt_term t1; mkt_term t2] | Mul (t1,t2) -> mkt_app ttmul [Lazy.force tz; mkt_term t1; mkt_term t2] | Pow (t1,n) -> if Int.equal n 0 then mkt_app ttconst [Lazy.force tz; mkt_z num_1] else mkt_app ttpow [Lazy.force tz; mkt_term t1; mkt_n (num_of_int n)] let rec parse_pos p = match Constr.kind p with | App (a,[|p2|]) -> if Constr.equal a (Lazy.force pxO) then num_2 */ (parse_pos p2) else num_1 +/ (num_2 */ (parse_pos p2)) | _ -> num_1 let parse_z z = match Constr.kind z with | App (a,[|p2|]) -> if Constr.equal a (Lazy.force zpos) then parse_pos p2 else (num_0 -/ (parse_pos p2)) | _ -> num_0 let parse_n z = match Constr.kind z with | App (a,[|p2|]) -> parse_pos p2 | _ -> num_0 let rec parse_term p = match Constr.kind p with | App (a,[|_;p2|]) -> if Constr.equal a (Lazy.force ttvar) then Var (string_of_num (parse_pos p2)) else if Constr.equal a (Lazy.force ttconst) then Const (parse_z p2) else if Constr.equal a (Lazy.force ttopp) then Opp (parse_term p2) else Zero | App (a,[|_;p2;p3|]) -> if Constr.equal a (Lazy.force ttadd) then Add (parse_term p2, parse_term p3) else if Constr.equal a (Lazy.force ttsub) then Sub (parse_term p2, parse_term p3) else if Constr.equal a (Lazy.force ttmul) then Mul (parse_term p2, parse_term p3) else if Constr.equal a (Lazy.force ttpow) then Pow (parse_term p2, int_of_num (parse_n p3)) else Zero | _ -> Zero let rec parse_request lp = match Constr.kind lp with | App (_,[|_|]) -> [] | App (_,[|_;p;lp1|]) -> (parse_term p)::(parse_request lp1) |_-> assert false let set_nvars_term nvars t = let rec aux t nvars = match t with | Zero -> nvars | Const r -> nvars | Var v -> let n = int_of_string v in max nvars n | Opp t1 -> aux t1 nvars | Add (t1,t2) -> aux t2 (aux t1 nvars) | Sub (t1,t2) -> aux t2 (aux t1 nvars) | Mul (t1,t2) -> aux t2 (aux t1 nvars) | Pow (t1,n) -> aux t1 nvars in aux t nvars (*********************************************************************** Coefficients: recursive polynomials *) module Coef = BigInt (*module Coef = Ent*) module Poly = Polynom.Make(Coef) module PIdeal = Ideal.Make(Poly) open PIdeal (* term to sparse polynomial variables <=np are in the coefficients *) let term_pol_sparse nvars np t= let d = nvars in let rec aux t = (* info ("conversion de: "^(string_of_term t)^"\n");*) let res = match t with | Zero -> zeroP | Const r -> if Num.eq_num r num_0 then zeroP else polconst d (Poly.Pint (Coef.of_num r)) | Var v -> let v = int_of_string v in if v <= np then polconst d (Poly.x v) else gen d v | Opp t1 -> oppP (aux t1) | Add (t1,t2) -> plusP (aux t1) (aux t2) | Sub (t1,t2) -> plusP (aux t1) (oppP (aux t2)) | Mul (t1,t2) -> multP (aux t1) (aux t2) | Pow (t1,n) -> puisP (aux t1) n in (* info ("donne: "^(stringP res)^"\n");*) res in let res= aux t in res (* sparse polynomial to term *) let polrec_to_term p = let rec aux p = match p with |Poly.Pint n -> const (Coef.to_num n) |Poly.Prec (v,coefs) -> let fold i c res = add (res, mul (aux c, pow (Var (string_of_int v), i))) in Array.fold_right_i fold coefs Zero in aux p (* approximation of the Horner form used in the tactic ring *) let pol_sparse_to_term n2 p = (* info "pol_sparse_to_term ->\n";*) let p = PIdeal.repr p in let rec aux p = match p with [] -> const (num_of_string "0") | (a,m)::p1 -> let m = Ideal.Monomial.repr m in let n = (Array.length m)-1 in let (i0,e0) = List.fold_left (fun (r,d) (a,m) -> let m = Ideal.Monomial.repr m in let i0= ref 0 in for k=1 to n do if m.(k)>0 then i0:=k done; if Int.equal !i0 0 then (r,d) else if !i0 > r then (!i0, m.(!i0)) else if Int.equal !i0 r && m.(!i0)= e0 then begin let m0 = Array.copy (Ideal.Monomial.repr m) in let () = m0.(i0) <- m0.(i0) - e0 in let m0 = Ideal.Monomial.make m0 in ((a, m0) :: p1, p2) end else (p1, (a, m) :: p2) in let (p1, p2) = List.fold_left fold ([], []) p in let vm = if Int.equal e0 1 then Var (string_of_int (i0)) else pow (Var (string_of_int (i0)),e0) in add (mul(vm, aux (List.rev p1)), aux (List.rev p2)) in (*info "-> pol_sparse_to_term\n";*) aux p (* lq = [cn+m+1 n+m ...cn+m+1 1] lci=[[cn+1 n,...,cn1 1] ... [cn+m n+m-1,...,cn+m 1]] removes intermediate polynomials not useful to compute the last one. *) let remove_zeros lci = let m = List.length lci in let u = Array.make m false in let rec utiles k = (* TODO: Find a more reasonable implementation of this traversal. *) if k >= m || u.(k) then () else let () = u.(k) <- true in let lc = List.nth lci k in let iter i c = if not (PIdeal.equal c zeroP) then utiles (i + k + 1) in List.iteri iter lc in let () = utiles 0 in let filter i l = let f j l = if m <= i + j + 1 then true else u.(i + j + 1) in if u.(i) then Some (List.filteri f l) else None in let lr = CList.map_filter_i filter lci in info (fun () -> Printf.sprintf "useless spolynomials: %i" (m-List.length lr)); info (fun () -> Printf.sprintf "useful spolynomials: %i " (List.length lr)); lr let theoremedeszeros metadata nvars lpol p = let t1 = Unix.gettimeofday() in let m = nvars in let cert = in_ideal metadata m lpol p in info (fun () -> Printf.sprintf "time: @[%10.3f@]s" (Unix.gettimeofday ()-.t1)); cert open Ideal (* Remove zero polynomials and duplicates from the list of polynomials lp Return (clp, lb) where clp is the reduced list and lb is a list of booleans that has the same size than lp and where true indicates an element that has been removed *) let clean_pol lp = let t = Hashpol.create 12 in let find p = try Hashpol.find t p with Not_found -> Hashpol.add t p true; false in let rec aux lp = match lp with | [] -> [], [] | p :: lp1 -> let clp, lb = aux lp1 in if equal p zeroP || find p then clp, true::lb else (p :: clp, false::lb) in aux lp (* Expand the list of polynomials lp putting zeros where the list of booleans lb indicates there is a missing element Warning: the expansion is relative to the end of the list in reversed order lp cannot have less elements than lb *) let expand_pol lb lp = let rec aux lb lp = match lb with | [] -> lp | true :: lb1 -> zeroP :: aux lb1 lp | false :: lb1 -> match lp with [] -> assert false | p :: lp1 -> p :: aux lb1 lp1 in List.rev (aux lb (List.rev lp)) let theoremedeszeros_termes lp = let nvars = List.fold_left set_nvars_term 0 lp in match lp with | Const (Int sugarparam)::Const (Int nparam)::lp -> ((match sugarparam with |0 -> sinfo "computation without sugar"; lexico:=false; |1 -> sinfo "computation with sugar"; lexico:=false; |2 -> sinfo "ordre lexico computation without sugar"; lexico:=true; |3 -> sinfo "ordre lexico computation with sugar"; lexico:=true; |4 -> sinfo "computation without sugar, division by pairs"; lexico:=false; |5 -> sinfo "computation with sugar, division by pairs"; lexico:=false; |6 -> sinfo "ordre lexico computation without sugar, division by pairs"; lexico:=true; |7 -> sinfo "ordre lexico computation with sugar, division by pairs"; lexico:=true; | _ -> user_err Pp.(str "nsatz: bad parameter") ); let lvar = List.init nvars (fun i -> Printf.sprintf "x%i" (i + 1)) in let lvar = ["a";"b";"c";"d";"e";"f";"g";"h";"i";"j";"k";"l";"m";"n";"o";"p";"q";"r";"s";"t";"u";"v";"w";"x";"y";"z"] @ lvar in (* pour macaulay *) let metadata = { name_var = lvar } in let lp = List.map (term_pol_sparse nvars nparam) lp in match lp with | [] -> assert false | p::lp1 -> let lpol = List.rev lp1 in (* preprocessing : we remove zero polynomials and duplicate that are not handled by in_ideal lb is kept in order to fix the certificate in the post-processing *) let lpol, lb = clean_pol lpol in let cert = theoremedeszeros metadata nvars lpol p in sinfo "cert ok"; let lc = cert.last_comb::List.rev cert.gb_comb in match remove_zeros lc with | [] -> assert false | (lq::lci) -> (* post-processing : we apply the correction for the last line *) let lq = expand_pol lb lq in (* lci commence par les nouveaux polynomes *) let m = nvars in let c = pol_sparse_to_term m (polconst m cert.coef) in let r = Pow(Zero,cert.power) in let lci = List.rev lci in (* post-processing we apply the correction for the other lines *) let lci = List.map (expand_pol lb) lci in let lci = List.map (List.map (pol_sparse_to_term m)) lci in let lq = List.map (pol_sparse_to_term m) lq in info (fun () -> Printf.sprintf "number of parameters: %i" nparam); sinfo "term computed"; (c,r,lci,lq) ) |_ -> assert false (* version avec hash-consing du certificat: let nsatz lpol = Hashtbl.clear Dansideal.hmon; Hashtbl.clear Dansideal.coefpoldep; Hashtbl.clear Dansideal.sugartbl; Hashtbl.clear Polynomesrec.hcontentP; init_constants (); let lp= parse_request lpol in let (_lp0,_p,c,r,_lci,_lq as rthz) = theoremedeszeros_termes lp in let certif = certificat_vers_polynome_creux rthz in let certif = hash_certif certif in let certif = certif_term certif in let c = mkt_term c in info "constr computed\n"; (c, certif) *) let nsatz lpol = let lp= parse_request lpol in let (c,r,lci,lq) = theoremedeszeros_termes lp in let res = [c::r::lq]@lci in let res = List.map (fun lx -> List.map mkt_term lx) res in let res = List.fold_right (fun lt r -> let ltterm = List.fold_right (fun t r -> mkt_app lcons [mkt_app tpexpr [Lazy.force tz];t;r]) lt (mkt_app lnil [mkt_app tpexpr [Lazy.force tz]]) in mkt_app lcons [tlp ();ltterm;r]) res (mkt_app lnil [tlp ()]) in sinfo "term computed"; res let return_term t = let a = mkApp (UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.eq.refl",[|tllp ();t|]) in let a = EConstr.of_constr a in generalize [a] let nsatz_compute t = let lpol = try nsatz t with Ideal.NotInIdeal -> user_err Pp.(str "nsatz cannot solve this problem") in return_term lpol coq-8.11.0/plugins/nsatz/Nsatz.v0000644000175000017500000003517713612664533016411 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* x == y. intros x y H; setoid_replace x with ((x - y) + y); simpl; [setoid_rewrite H | idtac]; simpl. cring. cring. Qed. Lemma psos_r1: forall x y, x == y -> x - y == 0. intros x y H; simpl; setoid_rewrite H; simpl; cring. Qed. Lemma nsatzR_diff: forall x y:R, not (x == y) -> not (x - y == 0). intros. intro; apply H. simpl; setoid_replace x with ((x - y) + y). simpl. setoid_rewrite H0. simpl; cring. simpl. simpl; cring. Qed. (* adpatation du code de Benjamin aux setoides *) Export Ring_polynom. Export InitialRing. Definition PolZ := Pol Z. Definition PEZ := PExpr Z. Definition P0Z : PolZ := P0 (C:=Z) 0%Z. Definition PolZadd : PolZ -> PolZ -> PolZ := @Padd Z 0%Z Z.add Zeq_bool. Definition PolZmul : PolZ -> PolZ -> PolZ := @Pmul Z 0%Z 1%Z Z.add Z.mul Zeq_bool. Definition PolZeq := @Peq Z Zeq_bool. Definition norm := @norm_aux Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool. Fixpoint mult_l (la : list PEZ) (lp: list PolZ) : PolZ := match la, lp with | a::la, p::lp => PolZadd (PolZmul (norm a) p) (mult_l la lp) | _, _ => P0Z end. Fixpoint compute_list (lla: list (list PEZ)) (lp:list PolZ) := match lla with | List.nil => lp | la::lla => compute_list lla ((mult_l la lp)::lp) end. Definition check (lpe:list PEZ) (qe:PEZ) (certif: list (list PEZ) * list PEZ) := let (lla, lq) := certif in let lp := List.map norm lpe in PolZeq (norm qe) (mult_l lq (compute_list lla lp)). (* Correction *) Definition PhiR : list R -> PolZ -> R := (Pphi ring0 add mul (InitialRing.gen_phiZ ring0 ring1 add mul opp)). Definition PEevalR : list R -> PEZ -> R := PEeval ring0 ring1 add mul sub opp (gen_phiZ ring0 ring1 add mul opp) N.to_nat pow. Lemma P0Z_correct : forall l, PhiR l P0Z = 0. Proof. trivial. Qed. Lemma Rext: ring_eq_ext add mul opp _==_. Proof. constructor; solve_proper. Qed. Lemma Rset : Setoid_Theory R _==_. apply ring_setoid. Qed. Definition Rtheory:ring_theory ring0 ring1 add mul sub opp _==_. apply mk_rt. apply ring_add_0_l. apply ring_add_comm. apply ring_add_assoc. apply ring_mul_1_l. apply cring_mul_comm. apply ring_mul_assoc. apply ring_distr_l. apply ring_sub_def. apply ring_opp_def. Defined. Lemma PolZadd_correct : forall P' P l, PhiR l (PolZadd P P') == ((PhiR l P) + (PhiR l P')). Proof. unfold PolZadd, PhiR. intros. simpl. refine (Padd_ok Rset Rext (Rth_ARth Rset Rext Rtheory) (gen_phiZ_morph Rset Rext Rtheory) _ _ _). Qed. Lemma PolZmul_correct : forall P P' l, PhiR l (PolZmul P P') == ((PhiR l P) * (PhiR l P')). Proof. unfold PolZmul, PhiR. intros. refine (Pmul_ok Rset Rext (Rth_ARth Rset Rext Rtheory) (gen_phiZ_morph Rset Rext Rtheory) _ _ _). Qed. Lemma R_power_theory : Ring_theory.power_theory ring1 mul _==_ N.to_nat pow. apply Ring_theory.mkpow_th. unfold pow. intros. rewrite Nnat.N2Nat.id. reflexivity. Qed. Lemma norm_correct : forall (l : list R) (pe : PEZ), PEevalR l pe == PhiR l (norm pe). Proof. intros;apply (norm_aux_spec Rset Rext (Rth_ARth Rset Rext Rtheory) (gen_phiZ_morph Rset Rext Rtheory) R_power_theory). Qed. Lemma PolZeq_correct : forall P P' l, PolZeq P P' = true -> PhiR l P == PhiR l P'. Proof. intros;apply (Peq_ok Rset Rext (gen_phiZ_morph Rset Rext Rtheory));trivial. Qed. Fixpoint Cond0 (A:Type) (Interp:A->R) (l:list A) : Prop := match l with | List.nil => True | a::l => Interp a == 0 /\ Cond0 A Interp l end. Lemma mult_l_correct : forall l la lp, Cond0 PolZ (PhiR l) lp -> PhiR l (mult_l la lp) == 0. Proof. induction la;simpl;intros. cring. destruct lp;trivial. simpl. cring. simpl in H;destruct H. rewrite PolZadd_correct. simpl. rewrite PolZmul_correct. simpl. rewrite H. rewrite IHla. cring. trivial. Qed. Lemma compute_list_correct : forall l lla lp, Cond0 PolZ (PhiR l) lp -> Cond0 PolZ (PhiR l) (compute_list lla lp). Proof. induction lla;simpl;intros;trivial. apply IHlla;simpl;split;trivial. apply mult_l_correct;trivial. Qed. Lemma check_correct : forall l lpe qe certif, check lpe qe certif = true -> Cond0 PEZ (PEevalR l) lpe -> PEevalR l qe == 0. Proof. unfold check;intros l lpe qe (lla, lq) H2 H1. apply PolZeq_correct with (l:=l) in H2. rewrite norm_correct, H2. apply mult_l_correct. apply compute_list_correct. clear H2 lq lla qe;induction lpe;simpl;trivial. simpl in H1;destruct H1. rewrite <- norm_correct;auto. Qed. (* fin *) Definition R2:= 1 + 1. Fixpoint IPR p {struct p}: R := match p with xH => ring1 | xO xH => 1+1 | xO p1 => R2*(IPR p1) | xI xH => 1+(1+1) | xI p1 => 1+(R2*(IPR p1)) end. Definition IZR1 z := match z with Z0 => 0 | Zpos p => IPR p | Zneg p => -(IPR p) end. Fixpoint interpret3 t fv {struct t}: R := match t with | (PEadd t1 t2) => let v1 := interpret3 t1 fv in let v2 := interpret3 t2 fv in (v1 + v2) | (PEmul t1 t2) => let v1 := interpret3 t1 fv in let v2 := interpret3 t2 fv in (v1 * v2) | (PEsub t1 t2) => let v1 := interpret3 t1 fv in let v2 := interpret3 t2 fv in (v1 - v2) | (PEopp t1) => let v1 := interpret3 t1 fv in (-v1) | (PEpow t1 t2) => let v1 := interpret3 t1 fv in pow v1 (N.to_nat t2) | (PEc t1) => (IZR1 t1) | PEO => 0 | PEI => 1 | (PEX _ n) => List.nth (pred (Pos.to_nat n)) fv 0 end. End nsatz1. Ltac equality_to_goal H x y:= (* eliminate trivial hypotheses, but it takes time!: let h := fresh "nH" in (assert (h:equality x y); [solve [cring] | clear H; clear h]) || *) try (generalize (@psos_r1 _ _ _ _ _ _ _ _ _ _ _ x y H); clear H) . Ltac equalities_to_goal := lazymatch goal with | H: (_ ?x ?y) |- _ => equality_to_goal H x y | H: (_ _ ?x ?y) |- _ => equality_to_goal H x y | H: (_ _ _ ?x ?y) |- _ => equality_to_goal H x y | H: (_ _ _ _ ?x ?y) |- _ => equality_to_goal H x y (* extension possible :-) *) | H: (?x == ?y) |- _ => equality_to_goal H x y end. (* lp est incluse dans fv. La met en tete. *) Ltac parametres_en_tete fv lp := match fv with | (@nil _) => lp | (@cons _ ?x ?fv1) => let res := AddFvTail x lp in parametres_en_tete fv1 res end. Ltac append1 a l := match l with | (@nil _) => constr:(cons a l) | (cons ?x ?l) => let l' := append1 a l in constr:(cons x l') end. Ltac rev l := match l with |(@nil _) => l | (cons ?x ?l) => let l' := rev l in append1 x l' end. Ltac nsatz_call_n info nparam p rr lp kont := (* idtac "Trying power: " rr;*) let ll := constr:(PEc info :: PEc nparam :: PEpow p rr :: lp) in (* idtac "calcul...";*) nsatz_compute ll; (* idtac "done";*) match goal with | |- (?c::PEpow _ ?r::?lq0)::?lci0 = _ -> _ => intros _; let lci := fresh "lci" in set (lci:=lci0); let lq := fresh "lq" in set (lq:=lq0); kont c rr lq lci end. Ltac nsatz_call radicalmax info nparam p lp kont := let rec try_n n := lazymatch n with | 0%N => fail | _ => (let r := eval compute in (N.sub radicalmax (N.pred n)) in nsatz_call_n info nparam p r lp kont) || let n' := eval compute in (N.pred n) in try_n n' end in try_n radicalmax. Ltac lterm_goal g := match g with ?b1 == ?b2 => constr:(b1::b2::nil) | ?b1 == ?b2 -> ?g => let l := lterm_goal g in constr:(b1::b2::l) end. Ltac reify_goal l le lb:= match le with nil => idtac | ?e::?le1 => match lb with ?b::?lb1 => (* idtac "b="; idtac b;*) let x := fresh "B" in set (x:= b) at 1; change x with (interpret3 e l); clear x; reify_goal l le1 lb1 end end. Ltac get_lpol g := match g with (interpret3 ?p _) == _ => constr:(p::nil) | (interpret3 ?p _) == _ -> ?g => let l := get_lpol g in constr:(p::l) end. Ltac nsatz_generic radicalmax info lparam lvar := let nparam := eval compute in (Z.of_nat (List.length lparam)) in match goal with |- ?g => let lb := lterm_goal g in match (match lvar with |(@nil _) => match lparam with |(@nil _) => let r := eval red in (list_reifyl (lterm:=lb)) in r |_ => match eval red in (list_reifyl (lterm:=lb)) with |(?fv, ?le) => let fv := parametres_en_tete fv lparam in (* we reify a second time, with the good order for variables *) let r := eval red in (list_reifyl (lterm:=lb) (lvar:=fv)) in r end end |_ => let fv := parametres_en_tete lvar lparam in let r := eval red in (list_reifyl (lterm:=lb) (lvar:=fv)) in r end) with |(?fv, ?le) => reify_goal fv le lb ; match goal with |- ?g => let lp := get_lpol g in let lpol := eval compute in (List.rev lp) in intros; let SplitPolyList kont := match lpol with | ?p2::?lp2 => kont p2 lp2 | _ => idtac "polynomial not in the ideal" end in SplitPolyList ltac:(fun p lp => let p21 := fresh "p21" in let lp21 := fresh "lp21" in set (p21:=p) ; set (lp21:=lp); (* idtac "nparam:"; idtac nparam; idtac "p:"; idtac p; idtac "lp:"; idtac lp; *) nsatz_call radicalmax info nparam p lp ltac:(fun c r lq lci => let q := fresh "q" in set (q := PEmul c (PEpow p21 r)); let Hg := fresh "Hg" in assert (Hg:check lp21 q (lci,lq) = true); [ (vm_compute;reflexivity) || idtac "invalid nsatz certificate" | let Hg2 := fresh "Hg" in assert (Hg2: (interpret3 q fv) == 0); [ (*simpl*) idtac; generalize (@check_correct _ _ _ _ _ _ _ _ _ _ _ fv lp21 q (lci,lq) Hg); let cc := fresh "H" in (*simpl*) idtac; intro cc; apply cc; clear cc; (*simpl*) idtac; repeat (split;[assumption|idtac]); exact I | (*simpl in Hg2;*) (*simpl*) idtac; apply Rintegral_domain_pow with (interpret3 c fv) (N.to_nat r); (*simpl*) idtac; try apply integral_domain_one_zero; try apply integral_domain_minus_one_zero; try trivial; try exact integral_domain_one_zero; try exact integral_domain_minus_one_zero || (solve [simpl; unfold R2, equality, eq_notation, addition, add_notation, one, one_notation, multiplication, mul_notation, zero, zero_notation; discrR || omega]) || ((*simpl*) idtac) || idtac "could not prove discrimination result" ] ] ) ) end end end . Ltac nsatz_default:= intros; try apply (@psos_r1b _ _ _ _ _ _ _ _ _ _ _); match goal with |- (@equality ?r _ _ _) => repeat equalities_to_goal; nsatz_generic 6%N 1%Z (@nil r) (@nil r) end. Tactic Notation "nsatz" := nsatz_default. Tactic Notation "nsatz" "with" "radicalmax" ":=" constr(radicalmax) "strategy" ":=" constr(info) "parameters" ":=" constr(lparam) "variables" ":=" constr(lvar):= intros; try apply (@psos_r1b _ _ _ _ _ _ _ _ _ _ _); match goal with |- (@equality ?r _ _ _) => repeat equalities_to_goal; nsatz_generic radicalmax info lparam lvar end. (* Real numbers *) Require Import Reals. Require Import RealField. Lemma Rsth : Setoid_Theory R (@eq R). constructor;red;intros;subst;trivial. Qed. Instance Rops: (@Ring_ops R 0%R 1%R Rplus Rmult Rminus Ropp (@eq R)). Defined. Instance Rri : (Ring (Ro:=Rops)). constructor; try (try apply Rsth; try (unfold respectful, Proper; unfold equality; unfold eq_notation in *; intros; try rewrite H; try rewrite H0; reflexivity)). exact Rplus_0_l. exact Rplus_comm. symmetry. apply Rplus_assoc. exact Rmult_1_l. exact Rmult_1_r. symmetry. apply Rmult_assoc. exact Rmult_plus_distr_r. intros; apply Rmult_plus_distr_l. exact Rplus_opp_r. Defined. Class can_compute_Z (z : Z) := dummy_can_compute_Z : True. Hint Extern 0 (can_compute_Z ?v) => match isZcst v with true => exact I end : typeclass_instances. Instance reify_IZR z lvar {_ : can_compute_Z z} : reify (PEc z) lvar (IZR z). Defined. Lemma R_one_zero: 1%R <> 0%R. discrR. Qed. Instance Rcri: (Cring (Rr:=Rri)). red. exact Rmult_comm. Defined. Instance Rdi : (Integral_domain (Rcr:=Rcri)). constructor. exact Rmult_integral. exact R_one_zero. Defined. (* Rational numbers *) Require Import QArith. Instance Qops: (@Ring_ops Q 0%Q 1%Q Qplus Qmult Qminus Qopp Qeq). Defined. Instance Qri : (Ring (Ro:=Qops)). constructor. try apply Q_Setoid. apply Qplus_comp. apply Qmult_comp. apply Qminus_comp. apply Qopp_comp. exact Qplus_0_l. exact Qplus_comm. apply Qplus_assoc. exact Qmult_1_l. exact Qmult_1_r. apply Qmult_assoc. apply Qmult_plus_distr_l. intros. apply Qmult_plus_distr_r. reflexivity. exact Qplus_opp_r. Defined. Lemma Q_one_zero: not (Qeq 1%Q 0%Q). unfold Qeq. simpl. auto with *. Qed. Instance Qcri: (Cring (Rr:=Qri)). red. exact Qmult_comm. Defined. Instance Qdi : (Integral_domain (Rcr:=Qcri)). constructor. exact Qmult_integral. exact Q_one_zero. Defined. (* Integers *) Lemma Z_one_zero: 1%Z <> 0%Z. omega. Qed. Instance Zcri: (Cring (Rr:=Zr)). red. exact Z.mul_comm. Defined. Instance Zdi : (Integral_domain (Rcr:=Zcri)). constructor. exact Zmult_integral. exact Z_one_zero. Defined. coq-8.11.0/plugins/nsatz/ideal.mli0000644000175000017500000000266513612664533016700 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int array val make : int array -> t end module Make (P : Polynom.S) : sig (* Polynomials *) type deg = int type coef = P.t type poly val repr : poly -> (coef * Monomial.t) list val polconst : int -> coef -> poly val zeroP : poly val gen : int -> int -> poly val equal : poly -> poly -> bool val plusP : poly -> poly -> poly val oppP : poly -> poly val multP : poly -> poly -> poly val puisP : poly -> int -> poly type certificate = { coef : coef; power : int; gb_comb : poly list list; last_comb : poly list } val in_ideal : metadata -> deg -> poly list -> poly -> certificate module Hashpol : Hashtbl.S with type key = poly end exception NotInIdeal val lexico : bool ref coq-8.11.0/plugins/nsatz/nsatz_plugin.mlpack0000644000175000017500000000004213612664533021010 0ustar treinentreinenUtile Polynom Ideal Nsatz G_nsatz coq-8.11.0/plugins/nsatz/utile.ml0000644000175000017500000000043613612664533016565 0ustar treinentreinen(* Printing *) let pr x = if !Flags.debug then (Format.printf "@[%s@]" x; flush(stdout);)else () let prt0 s = () (* print_string s;flush(stdout)*) let sinfo s = if !Flags.debug then Feedback.msg_debug (Pp.str s) let info s = if !Flags.debug then Feedback.msg_debug (Pp.str (s ())) coq-8.11.0/plugins/nsatz/nsatz.mli0000644000175000017500000000133213612664533016747 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit Proofview.tactic coq-8.11.0/plugins/nsatz/polynom.ml0000644000175000017500000004321213612664533017137 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t -> bool val lt : t -> t -> bool val le : t -> t -> bool val abs : t -> t val plus : t -> t -> t val mult : t -> t -> t val sub : t -> t -> t val opp : t -> t val div : t -> t -> t val modulo : t -> t -> t val puis : t -> int -> t val pgcd : t -> t -> t val hash : t -> int val of_num : Num.num -> t val to_string : t -> string end module type S = sig type coef type variable = int type t = Pint of coef | Prec of variable * t array val of_num : Num.num -> t val x : variable -> t val monome : variable -> int -> t val is_constantP : t -> bool val is_zero : t -> bool val max_var_pol : t -> variable val max_var_pol2 : t -> variable val max_var : t array -> variable val equal : t -> t -> bool val norm : t -> t val deg : variable -> t -> int val deg_total : t -> int val copyP : t -> t val coef : variable -> int -> t -> t val plusP : t -> t -> t val content : t -> coef val div_int : t -> coef -> t val vire_contenu : t -> t val vars : t -> variable list val int_of_Pint : t -> coef val multx : int -> variable -> t -> t val multP : t -> t -> t val deriv : variable -> t -> t val oppP : t -> t val moinsP : t -> t -> t val puisP : t -> int -> t val ( @@ ) : t -> t -> t val ( -- ) : t -> t -> t val ( ^^ ) : t -> int -> t val coefDom : variable -> t -> t val coefConst : variable -> t -> t val remP : variable -> t -> t val coef_int_tete : t -> coef val normc : t -> t val coef_constant : t -> coef val univ : bool ref val string_of_var : int -> string val nsP : int ref val to_string : t -> string val printP : t -> unit val print_tpoly : t array -> unit val print_lpoly : t list -> unit val quo_rem_pol : t -> t -> variable -> t * t val div_pol : t -> t -> variable -> t val divP : t -> t -> t val div_pol_rat : t -> t -> bool val pseudo_div : t -> t -> variable -> t * t * int * t val pgcdP : t -> t -> t val pgcd_pol : t -> t -> variable -> t val content_pol : t -> variable -> t val pgcd_coef_pol : t -> t -> variable -> t val pgcd_pol_rec : t -> t -> variable -> t val gcd_sub_res : t -> t -> variable -> t val gcd_sub_res_rec : t -> t -> t -> t -> int -> variable -> t val lazard_power : t -> t -> int -> variable -> t val hash : t -> int module Hashpol : Hashtbl.S with type key=t end (*********************************************************************** 2. Type of polynomials, operations. *) module Make (C:Coef) = struct type coef = C.t let coef_of_int i = C.of_num (Num.Int i) let coef0 = coef_of_int 0 let coef1 = coef_of_int 1 type variable = int type t = Pint of coef (* constant polynomial *) | Prec of variable * (t array) (* coefficients, increasing degree *) (* by default, operations work with normalized polynomials: - variables are positive integers - coefficients of a polynomial in x only use variables < x - no zero coefficient at beginning - no Prec(x,a) where a is constant in x *) (* constant polynomials *) let of_num x = Pint (C.of_num x) let cf0 = of_num (Num.Int 0) let cf1 = of_num (Num.Int 1) (* nth variable *) let x n = Prec (n,[|cf0;cf1|]) (* create v^n *) let monome v n = match n with 0->Pint coef1; |_->let tmp = Array.make (n+1) (Pint coef0) in tmp.(n)<-(Pint coef1); Prec (v, tmp) let is_constantP = function Pint _ -> true | Prec _ -> false let int_of_Pint = function Pint x -> x | _ -> failwith "non" let is_zero p = match p with Pint n -> if C.equal n coef0 then true else false |_-> false let max_var_pol p = match p with Pint _ -> 0 |Prec(x,_) -> x (* p not normalized *) let rec max_var_pol2 p = match p with Pint _ -> 0 |Prec(v,c)-> Array.fold_right (fun q m -> max (max_var_pol2 q) m) c v let max_var l = Array.fold_right (fun p m -> max (max_var_pol2 p) m) l 0 (* equality between polynomials *) let rec equal p q = match (p,q) with (Pint a,Pint b) -> C.equal a b |(Prec(x,p1),Prec(y,q1)) -> (Int.equal x y) && Array.for_all2 equal p1 q1 | (_,_) -> false (* normalize polynomial: remove head zeros, coefficients are normalized if constant, returns the coefficient *) let norm p = match p with Pint _ -> p |Prec (x,a)-> let d = (Array.length a -1) in let n = ref d in while !n>0 && (equal a.(!n) (Pint coef0)) do n:=!n-1; done; if !n<0 then Pint coef0 else if Int.equal !n 0 then a.(0) else if Int.equal !n d then p else (let b=Array.make (!n+1) (Pint coef0) in for i=0 to !n do b.(i)<-a.(i);done; Prec(x,b)) (* degree in v, v >= max var of p *) let deg v p = match p with Prec(x,p1) when Int.equal x v -> Array.length p1 -1 |_ -> 0 (* total degree *) let rec deg_total p = match p with Prec (x,p1) -> let d = ref 0 in Array.iteri (fun i q -> d:= (max !d (i+(deg_total q)))) p1; !d |_ -> 0 let rec copyP p = match p with Pint i -> Pint i |Prec(x,q) -> Prec(x,Array.map copyP q) (* coefficient of degree i in v, v >= max var of p *) let coef v i p = match p with Prec (x,p1) when Int.equal x v -> if i<(Array.length p1) then p1.(i) else Pint coef0 |_ -> if Int.equal i 0 then p else Pint coef0 (* addition *) let rec plusP p q = let res = (match (p,q) with (Pint a,Pint b) -> Pint (C.plus a b) |(Pint a, Prec (y,q1)) -> let q2=Array.map copyP q1 in q2.(0)<- plusP p q1.(0); Prec (y,q2) |(Prec (x,p1),Pint b) -> let p2=Array.map copyP p1 in p2.(0)<- plusP p1.(0) q; Prec (x,p2) |(Prec (x,p1),Prec (y,q1)) -> if xy then (let p2=Array.map copyP p1 in p2.(0)<- plusP p1.(0) q; Prec (x,p2)) else (let n=max (deg x p) (deg x q) in let r=Array.make (n+1) (Pint coef0) in for i=0 to n do r.(i)<- plusP (coef x i p) (coef x i q); done; Prec(x,r))) in norm res (* content, positive integer *) let rec content p = match p with Pint a -> C.abs a | Prec (x ,p1) -> Array.fold_left C.pgcd coef0 (Array.map content p1) let rec div_int p a= match p with Pint b -> Pint (C.div b a) | Prec(x,p1) -> Prec(x,Array.map (fun x -> div_int x a) p1) let vire_contenu p = let c = content p in if C.equal c coef0 then p else div_int p c (* sorted list of variables of a polynomial *) let rec vars=function Pint _->[] | Prec (x,l)->(List.flatten ([x]::(List.map vars (Array.to_list l)))) (* multiply p by v^n, v >= max_var p *) let multx n v p = match p with Prec (x,p1) when Int.equal x v -> let p2= Array.make ((Array.length p1)+n) (Pint coef0) in for i=0 to (Array.length p1)-1 do p2.(i+n)<-p1.(i); done; Prec (x,p2) |_ -> if equal p (Pint coef0) then (Pint coef0) else (let p2=Array.make (n+1) (Pint coef0) in p2.(n)<-p; Prec (v,p2)) (* product *) let rec multP p q = match (p,q) with (Pint a,Pint b) -> Pint (C.mult a b) |(Pint a, Prec (y,q1)) -> if C.equal a coef0 then Pint coef0 else let q2 = Array.map (fun z-> multP p z) q1 in Prec (y,q2) |(Prec (x,p1), Pint b) -> if C.equal b coef0 then Pint coef0 else let p2 = Array.map (fun z-> multP z q) p1 in Prec (x,p2) |(Prec (x,p1), Prec(y,q1)) -> if x multP p z) q1 in Prec (y,q2)) else if x>y then (let p2 = Array.map (fun z-> multP z q) p1 in Prec (x,p2)) else Array.fold_left plusP (Pint coef0) (Array.mapi (fun i z-> (multx i x (multP z q))) p1) (* derive p with variable v, v >= max_var p *) let deriv v p = match p with Pint a -> Pint coef0 | Prec(x,p1) when Int.equal x v -> let d = Array.length p1 -1 in if Int.equal d 1 then p1.(1) else (let p2 = Array.make d (Pint coef0) in for i=0 to d-1 do p2.(i)<- multP (Pint (coef_of_int (i+1))) p1.(i+1); done; Prec (x,p2)) | Prec(x,p1)-> Pint coef0 (* opposite *) let rec oppP p = match p with Pint a -> Pint (C.opp a) |Prec(x,p1) -> Prec(x,Array.map oppP p1) let moinsP p q=plusP p (oppP q) let rec puisP p n = match n with 0 -> cf1 |_ -> (multP p (puisP p (n-1))) (* infix notations *) (*let (++) a b = plusP a b *) let (@@) a b = multP a b let (--) a b = moinsP a b let (^^) a b = puisP a b (* leading coefficient in v, v>= max_var p *) let coefDom v p= coef v (deg v p) p let coefConst v p = coef v 0 p (* tail of a polynomial *) let remP v p = moinsP p (multP (coefDom v p) (puisP (x v) (deg v p))) (* first integer coefficient of p *) let rec coef_int_tete p = let v = max_var_pol p in if v>0 then coef_int_tete (coefDom v p) else (match p with | Pint a -> a |_ -> assert false) (* divide by the content and make the head int coef positive *) let normc p = let p = vire_contenu p in let a = coef_int_tete p in if C.le coef0 a then p else oppP p (* constant coef of normalized polynomial *) let rec coef_constant p = match p with Pint a->a |Prec(_,q)->coef_constant q.(0) (*********************************************************************** 3. Printing polynomials. *) (* if univ = false, we use x,y,z,a,b,c,d... as variables, else x1,x2,... *) let univ=ref true let string_of_var x= if !univ then "u"^(string_of_int x) else if x<=3 then String.make 1 (Char.chr(x+(Char.code 'w'))) else String.make 1 (Char.chr(x-4+(Char.code 'a'))) let nsP = ref 0 let rec string_of_Pcut p = if (!nsP)<=0 then "..." else match p with |Pint a-> nsP:=(!nsP)-1; if C.le coef0 a then C.to_string a else "("^(C.to_string a)^")" |Prec (x,t)-> let v=string_of_var x and s=ref "" and sp=ref "" in let st0 = string_of_Pcut t.(0) in if not (String.equal st0 "0") then s:=st0; let fin = ref false in for i=(Array.length t)-1 downto 1 do if (!nsP)<0 then (sp:="..."; if not (!fin) then s:=(!s)^"+"^(!sp); fin:=true) else ( let si=string_of_Pcut t.(i) in sp:=""; if Int.equal i 1 then ( if not (String.equal si "0") then (nsP:=(!nsP)-1; if String.equal si "1" then sp:=v else (if (String.contains si '+') then sp:="("^si^")*"^v else sp:=si^"*"^v))) else ( if not (String.equal si "0") then (nsP:=(!nsP)-1; if String.equal si "1" then sp:=v^"^"^(string_of_int i) else (if (String.contains si '+') then sp:="("^si^")*"^v^"^"^(string_of_int i) else sp:=si^"*"^v^"^"^(string_of_int i)))); if not (String.is_empty !sp) && not (!fin) then (nsP:=(!nsP)-1; if String.is_empty !s then s:=!sp else s:=(!s)^"+"^(!sp))); done; if String.is_empty !s then (nsP:=(!nsP)-1; (s:="0")); !s let to_string p = nsP:=20; string_of_Pcut p let printP p = Format.printf "@[%s@]" (to_string p) let print_tpoly lp = let s = ref "\n{ " in Array.iter (fun p -> s:=(!s)^(to_string p)^"\n") lp; prt0 ((!s)^"}") let print_lpoly lp = print_tpoly (Array.of_list lp) (*********************************************************************** 4. Exact division of polynomials. *) (* return (s,r) s.t. p = s*q+r *) let rec quo_rem_pol p q x = if Int.equal x 0 then (match (p,q) with |(Pint a, Pint b) -> if C.equal (C.modulo a b) coef0 then (Pint (C.div a b), cf0) else failwith "div_pol1" |_ -> assert false) else let m = deg x q in let b = coefDom x q in let q1 = remP x q in (* q = b*x^m+q1 *) let r = ref p in let s = ref cf0 in let continue =ref true in while (!continue) && (not (equal !r cf0)) do let n = deg x !r in if n false (*********************************************************************** 5. Pseudo-division and gcd with subresultants. *) (* pseudo division : q = c*x^m+q1 returns (r,c,d,s) s.t. c^d*p = s*q + r. *) let pseudo_div p q x = match q with Pint _ -> (cf0, q,1, p) | Prec (v,q1) when not (Int.equal x v) -> (cf0, q,1, p) | Prec (v,q1) -> ( (* pr "pseudo_division: c^d*p = s*q + r";*) let delta = ref 0 in let r = ref p in let c = coefDom x q in let q1 = remP x q in let d' = deg x q in let s = ref cf0 in while (deg x !r)>=(deg x q) do let d = deg x !r in let a = coefDom x !r in let r1=remP x !r in let u = a @@ ((monome x (d-d'))) in r:=(c @@ r1) -- (u @@ q1); s:=plusP (c @@ (!s)) u; delta := (!delta) + 1; done; (* pr ("deg d: "^(string_of_int (!delta))^", deg c: "^(string_of_int (deg_total c))); pr ("deg r:"^(string_of_int (deg_total !r))); *) (!r,c,!delta, !s) ) (* gcd with subresultants *) let rec pgcdP p q = let x = max (max_var_pol p) (max_var_pol q) in pgcd_pol p q x and pgcd_pol p q x = pgcd_pol_rec p q x and content_pol p x = match p with Prec(v,p1) when Int.equal v x -> Array.fold_left (fun a b -> pgcd_pol_rec a b (x-1)) cf0 p1 | _ -> p and pgcd_coef_pol c p x = match p with Prec(v,p1) when Int.equal x v -> Array.fold_left (fun a b -> pgcd_pol_rec a b (x-1)) c p1 |_ -> pgcd_pol_rec c p (x-1) and pgcd_pol_rec p q x = match (p,q) with (Pint a,Pint b) -> Pint (C.pgcd (C.abs a) (C.abs b)) |_ -> if equal p cf0 then q else if equal q cf0 then p else if Int.equal (deg x q) 0 then pgcd_coef_pol q p x else if Int.equal (deg x p) 0 then pgcd_coef_pol p q x else ( let a = content_pol p x in let b = content_pol q x in let c = pgcd_pol_rec a b (x-1) in pr (string_of_int x); let p1 = div_pol p c x in let q1 = div_pol q c x in let r = gcd_sub_res p1 q1 x in let cr = content_pol r x in let res = c @@ (div_pol r cr x) in res ) (* Sub-résultants: ai*Ai = Qi*Ai+1 + bi*Ai+2 deg Ai+2 < deg Ai+1 Ai = ci*X^ni + ... di = ni - ni+1 ai = (- ci+1)^(di + 1) b1 = 1 bi = ci*si^di si i>1 s1 = 1 si+1 = ((ci+1)^di*si)/si^di *) and gcd_sub_res p q x = if equal q cf0 then p else let d = deg x p in let d' = deg x q in if d (C.hash a) | Prec (v,p) -> Array.fold_right (fun q h -> h + hash q) p 0 module Hashpol = Hashtbl.Make( struct type poly = t type t = poly let equal = equal let hash = hash end) end coq-8.11.0/plugins/nsatz/utile.mli0000644000175000017500000000017513612664533016736 0ustar treinentreinen (* Printing *) val pr : string -> unit val prt0 : 'a -> unit val info : (unit -> string) -> unit val sinfo : string -> unit coq-8.11.0/plugins/nsatz/polynom.mli0000644000175000017500000000617413612664533017316 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t -> bool val lt : t -> t -> bool val le : t -> t -> bool val abs : t -> t val plus : t -> t -> t val mult : t -> t -> t val sub : t -> t -> t val opp : t -> t val div : t -> t -> t val modulo : t -> t -> t val puis : t -> int -> t val pgcd : t -> t -> t val hash : t -> int val of_num : Num.num -> t val to_string : t -> string end module type S = sig type coef type variable = int type t = Pint of coef | Prec of variable * t array val of_num : Num.num -> t val x : variable -> t val monome : variable -> int -> t val is_constantP : t -> bool val is_zero : t -> bool val max_var_pol : t -> variable val max_var_pol2 : t -> variable val max_var : t array -> variable val equal : t -> t -> bool val norm : t -> t val deg : variable -> t -> int val deg_total : t -> int val copyP : t -> t val coef : variable -> int -> t -> t val plusP : t -> t -> t val content : t -> coef val div_int : t -> coef -> t val vire_contenu : t -> t val vars : t -> variable list val int_of_Pint : t -> coef val multx : int -> variable -> t -> t val multP : t -> t -> t val deriv : variable -> t -> t val oppP : t -> t val moinsP : t -> t -> t val puisP : t -> int -> t val ( @@ ) : t -> t -> t val ( -- ) : t -> t -> t val ( ^^ ) : t -> int -> t val coefDom : variable -> t -> t val coefConst : variable -> t -> t val remP : variable -> t -> t val coef_int_tete : t -> coef val normc : t -> t val coef_constant : t -> coef val univ : bool ref val string_of_var : int -> string val nsP : int ref val to_string : t -> string val printP : t -> unit val print_tpoly : t array -> unit val print_lpoly : t list -> unit val quo_rem_pol : t -> t -> variable -> t * t val div_pol : t -> t -> variable -> t val divP : t -> t -> t val div_pol_rat : t -> t -> bool val pseudo_div : t -> t -> variable -> t * t * int * t val pgcdP : t -> t -> t val pgcd_pol : t -> t -> variable -> t val content_pol : t -> variable -> t val pgcd_coef_pol : t -> t -> variable -> t val pgcd_pol_rec : t -> t -> variable -> t val gcd_sub_res : t -> t -> variable -> t val gcd_sub_res_rec : t -> t -> t -> t -> int -> variable -> t val lazard_power : t -> t -> int -> variable -> t val hash : t -> int module Hashpol : Hashtbl.S with type key=t end module Make (C:Coef) : S with type coef = C.t coq-8.11.0/plugins/nsatz/g_nsatz.mlg0000644000175000017500000000154213612664533017256 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Nsatz.nsatz_compute (EConstr.Unsafe.to_constr lt) } END coq-8.11.0/plugins/nsatz/plugin_base.dune0000644000175000017500000000021213612664533020246 0ustar treinentreinen(library (name nsatz_plugin) (public_name coq.plugins.nsatz) (synopsis "Coq's nsatz solver plugin") (libraries num coq.plugins.ltac)) coq-8.11.0/plugins/micromega/0000755000175000017500000000000013612664533015712 5ustar treinentreinencoq-8.11.0/plugins/micromega/Refl.v0000644000175000017500000001023613612664533016773 0ustar treinentreinen(* -*- coding: utf-8 -*- *) (************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ' '/\': basic properties *) Fixpoint make_impl (A : Type) (eval : A -> Prop) (l : list A) (goal : Prop) {struct l} : Prop := match l with | nil => goal | cons e l => (eval e) -> (make_impl eval l goal) end. Theorem make_impl_true : forall (A : Type) (eval : A -> Prop) (l : list A), make_impl eval l True. Proof. induction l as [| a l IH]; simpl. trivial. intro; apply IH. Qed. Theorem make_impl_map : forall (A B: Type) (eval : A -> Prop) (eval' : A*B -> Prop) (l : list (A*B)) r (EVAL : forall x, eval' x <-> eval (fst x)), make_impl eval' l r <-> make_impl eval (List.map fst l) r. Proof. induction l as [| a l IH]; simpl. - tauto. - intros. rewrite EVAL. rewrite IH. tauto. auto. Qed. Fixpoint make_conj (A : Type) (eval : A -> Prop) (l : list A) {struct l} : Prop := match l with | nil => True | cons e nil => (eval e) | cons e l2 => ((eval e) /\ (make_conj eval l2)) end. Theorem make_conj_cons : forall (A : Type) (eval : A -> Prop) (a : A) (l : list A), make_conj eval (a :: l) <-> eval a /\ make_conj eval l. Proof. intros; destruct l; simpl; tauto. Qed. Lemma make_conj_impl : forall (A : Type) (eval : A -> Prop) (l : list A) (g : Prop), (make_conj eval l -> g) <-> make_impl eval l g. Proof. induction l. simpl. tauto. simpl. intros. destruct l. simpl. tauto. generalize (IHl g). tauto. Qed. Lemma make_conj_in : forall (A : Type) (eval : A -> Prop) (l : list A), make_conj eval l -> (forall p, In p l -> eval p). Proof. induction l. simpl. tauto. simpl. intros. destruct l. simpl in H0. destruct H0. subst; auto. tauto. destruct H. destruct H0. subst;auto. apply IHl; auto. Qed. Lemma make_conj_app : forall A eval l1 l2, @make_conj A eval (l1 ++ l2) <-> @make_conj A eval l1 /\ @make_conj A eval l2. Proof. induction l1. simpl. tauto. intros. change ((a::l1) ++ l2) with (a :: (l1 ++ l2)). rewrite make_conj_cons. rewrite IHl1. rewrite make_conj_cons. tauto. Qed. Infix "+++" := rev_append (right associativity, at level 60) : list_scope. Lemma make_conj_rapp : forall A eval l1 l2, @make_conj A eval (l1 +++ l2) <-> @make_conj A eval (l1++l2). Proof. induction l1. - simpl. tauto. - intros. simpl rev_append at 1. rewrite IHl1. rewrite make_conj_app. rewrite make_conj_cons. simpl app. rewrite make_conj_cons. rewrite make_conj_app. tauto. Qed. Lemma not_make_conj_cons : forall (A:Type) (t:A) a eval (no_middle_eval : (eval t) \/ ~ (eval t)), ~ make_conj eval (t ::a) <-> ~ (eval t) \/ (~ make_conj eval a). Proof. intros. rewrite make_conj_cons. tauto. Qed. Lemma not_make_conj_app : forall (A:Type) (t:list A) a eval (no_middle_eval : forall d, eval d \/ ~ eval d) , ~ make_conj eval (t ++ a) <-> (~ make_conj eval t) \/ (~ make_conj eval a). Proof. induction t. - simpl. tauto. - intros. simpl ((a::t)++a0). rewrite !not_make_conj_cons by auto. rewrite IHt by auto. tauto. Qed. coq-8.11.0/plugins/micromega/RMicromega.v0000644000175000017500000003054213612664533020132 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Q2R x = Q2R y. Proof. intros. now apply Qeq_eqR, Qeq_bool_eq. Qed. Lemma Qeq_false : forall x y, Qeq_bool x y = false -> Q2R x <> Q2R y. Proof. intros. apply Qeq_bool_neq in H. contradict H. now apply eqR_Qeq. Qed. Lemma Qle_true : forall x y : Q, Qle_bool x y = true -> Q2R x <= Q2R y. Proof. intros. now apply Qle_Rle, Qle_bool_imp_le. Qed. Lemma Q2R_0 : Q2R 0 = 0. Proof. apply Rmult_0_l. Qed. Lemma Q2R_1 : Q2R 1 = 1. Proof. compute. apply Rinv_1. Qed. Lemma Q2R_inv_ext : forall x, Q2R (/ x) = (if Qeq_bool x 0 then 0 else / Q2R x). Proof. intros. case_eq (Qeq_bool x 0). intros. apply Qeq_bool_eq in H. destruct x ; simpl. unfold Qeq in H. simpl in H. rewrite Zmult_1_r in H. rewrite H. apply Rmult_0_l. intros. now apply Q2R_inv, Qeq_bool_neq. Qed. Notation to_nat := N.to_nat. Lemma QSORaddon : @SORaddon R R0 R1 Rplus Rmult Rminus Ropp (@eq R) Rle (* ring elements *) Q 0%Q 1%Q Qplus Qmult Qminus Qopp (* coefficients *) Qeq_bool Qle_bool Q2R nat to_nat pow. Proof. constructor. constructor ; intros ; try reflexivity. apply Q2R_0. apply Q2R_1. apply Q2R_plus. apply Q2R_minus. apply Q2R_mult. apply Q2R_opp. apply Qeq_true ; auto. apply R_power_theory. apply Qeq_false. apply Qle_true. Qed. (* Syntactic ring coefficients. *) Inductive Rcst := | C0 | C1 | CQ (r : Q) | CZ (r : Z) | CPlus (r1 r2 : Rcst) | CMinus (r1 r2 : Rcst) | CMult (r1 r2 : Rcst) | CPow (r1 : Rcst) (z:Z+nat) | CInv (r : Rcst) | COpp (r : Rcst). Definition z_of_exp (z : Z + nat) := match z with | inl z => z | inr n => Z.of_nat n end. Fixpoint Q_of_Rcst (r : Rcst) : Q := match r with | C0 => 0 # 1 | C1 => 1 # 1 | CZ z => z # 1 | CQ q => q | CPlus r1 r2 => Qplus (Q_of_Rcst r1) (Q_of_Rcst r2) | CMinus r1 r2 => Qminus (Q_of_Rcst r1) (Q_of_Rcst r2) | CMult r1 r2 => Qmult (Q_of_Rcst r1) (Q_of_Rcst r2) | CPow r1 z => Qpower (Q_of_Rcst r1) (z_of_exp z) | CInv r => Qinv (Q_of_Rcst r) | COpp r => Qopp (Q_of_Rcst r) end. Definition is_neg (z: Z+nat) := match z with | inl (Zneg _) => true | _ => false end. Lemma is_neg_true : forall z, is_neg z = true -> (z_of_exp z < 0)%Z. Proof. destruct z ; simpl ; try congruence. destruct z ; try congruence. intros. reflexivity. Qed. Lemma is_neg_false : forall z, is_neg z = false -> (z_of_exp z >= 0)%Z. Proof. destruct z ; simpl ; try congruence. destruct z ; try congruence. compute. congruence. compute. congruence. generalize (Zle_0_nat n). auto using Z.le_ge. Qed. Definition CInvR0 (r : Rcst) := Qeq_bool (Q_of_Rcst r) (0 # 1). Definition CPowR0 (z : Z) (r : Rcst) := Z.ltb z Z0 && Qeq_bool (Q_of_Rcst r) (0 # 1). Fixpoint R_of_Rcst (r : Rcst) : R := match r with | C0 => R0 | C1 => R1 | CZ z => IZR z | CQ q => Q2R q | CPlus r1 r2 => (R_of_Rcst r1) + (R_of_Rcst r2) | CMinus r1 r2 => (R_of_Rcst r1) - (R_of_Rcst r2) | CMult r1 r2 => (R_of_Rcst r1) * (R_of_Rcst r2) | CPow r1 z => match z with | inl z => if CPowR0 z r1 then R0 else powerRZ (R_of_Rcst r1) z | inr n => pow (R_of_Rcst r1) n end | CInv r => if CInvR0 r then R0 else Rinv (R_of_Rcst r) | COpp r => - (R_of_Rcst r) end. Add Morphism Q2R with signature Qeq ==> @eq R as Q2R_m. exact Qeq_eqR. Qed. Lemma Q2R_pow_pos : forall q p, Q2R (pow_pos Qmult q p) = pow_pos Rmult (Q2R q) p. Proof. induction p ; simpl;auto; rewrite <- IHp; repeat rewrite Q2R_mult; reflexivity. Qed. Lemma Q2R_pow_N : forall q n, Q2R (pow_N 1%Q Qmult q n) = pow_N 1 Rmult (Q2R q) n. Proof. destruct n ; simpl. - apply Q2R_1. - apply Q2R_pow_pos. Qed. Lemma Qmult_integral : forall q r, q * r == 0 -> q == 0 \/ r == 0. Proof. intros. destruct (Qeq_dec q 0)%Q. - left ; apply q0. - apply Qmult_integral_l in H ; tauto. Qed. Lemma Qpower_positive_eq_zero : forall q p, Qpower_positive q p == 0 -> q == 0. Proof. unfold Qpower_positive. induction p ; simpl; intros; repeat match goal with | H : _ * _ == 0 |- _ => apply Qmult_integral in H; destruct H end; tauto. Qed. Lemma Qpower_positive_zero : forall p, Qpower_positive 0 p == 0%Q. Proof. induction p ; simpl; try rewrite IHp ; reflexivity. Qed. Lemma Q2RpowerRZ : forall q z (DEF : not (q == 0)%Q \/ (z >= Z0)%Z), Q2R (q ^ z) = powerRZ (Q2R q) z. Proof. intros. destruct Qpower_theory. destruct R_power_theory. unfold Qpower, powerRZ. destruct z. - apply Q2R_1. - change (Qpower_positive q p) with (Qpower q (Zpos p)). rewrite <- N2Z.inj_pos. rewrite <- positive_N_nat. rewrite rpow_pow_N. rewrite rpow_pow_N0. apply Q2R_pow_N. - rewrite Q2R_inv. unfold Qpower_positive. rewrite <- positive_N_nat. rewrite rpow_pow_N0. unfold pow_N. rewrite Q2R_pow_pos. auto. intro. apply Qpower_positive_eq_zero in H. destruct DEF ; auto with arith. Qed. Lemma Qpower0 : forall z, (z <> 0)%Z -> (0 ^ z == 0)%Q. Proof. unfold Qpower. destruct z;intros. - congruence. - apply Qpower_positive_zero. - rewrite Qpower_positive_zero. reflexivity. Qed. Lemma Q_of_RcstR : forall c, Q2R (Q_of_Rcst c) = R_of_Rcst c. Proof. induction c ; simpl ; try (rewrite <- IHc1 ; rewrite <- IHc2). - apply Q2R_0. - apply Q2R_1. - reflexivity. - unfold Q2R. simpl. rewrite Rinv_1. reflexivity. - apply Q2R_plus. - apply Q2R_minus. - apply Q2R_mult. - destruct z. destruct (CPowR0 z c) eqn:C; unfold CPowR0 in C. + rewrite andb_true_iff in C. destruct C as (C1 & C2). rewrite Z.ltb_lt in C1. apply Qeq_bool_eq in C2. rewrite C2. simpl. rewrite Qpower0. apply Q2R_0. intro ; subst ; slia C1 C1. + rewrite Q2RpowerRZ. rewrite IHc. reflexivity. rewrite andb_false_iff in C. destruct C. simpl. apply Z.ltb_ge in H. right ; normZ. slia H H0. left ; apply Qeq_bool_neq; auto. + simpl. rewrite <- IHc. destruct Qpower_theory. rewrite <- nat_N_Z. rewrite rpow_pow_N. destruct R_power_theory. rewrite <- (Nnat.Nat2N.id n) at 2. rewrite rpow_pow_N0. apply Q2R_pow_N. - rewrite <- IHc. unfold CInvR0. apply Q2R_inv_ext. - rewrite <- IHc. apply Q2R_opp. Qed. Require Import EnvRing. Definition INZ (n:N) : R := match n with | N0 => IZR 0%Z | Npos p => IZR (Zpos p) end. Definition Reval_expr := eval_pexpr Rplus Rmult Rminus Ropp R_of_Rcst N.to_nat pow. Definition Reval_op2 (o:Op2) : R -> R -> Prop := match o with | OpEq => @eq R | OpNEq => fun x y => ~ x = y | OpLe => Rle | OpGe => Rge | OpLt => Rlt | OpGt => Rgt end. Definition Reval_formula (e: PolEnv R) (ff : Formula Rcst) := let (lhs,o,rhs) := ff in Reval_op2 o (Reval_expr e lhs) (Reval_expr e rhs). Definition Reval_formula' := eval_sformula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt N.to_nat pow R_of_Rcst. Definition QReval_formula := eval_formula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt Q2R N.to_nat pow . Lemma Reval_formula_compat : forall env f, Reval_formula env f <-> Reval_formula' env f. Proof. intros. unfold Reval_formula. destruct f. unfold Reval_formula'. unfold Reval_expr. split ; destruct Fop ; simpl ; auto. apply Rge_le. apply Rle_ge. Qed. Definition Qeval_nformula := eval_nformula 0 Rplus Rmult (@eq R) Rle Rlt Q2R. Lemma Reval_nformula_dec : forall env d, (Qeval_nformula env d) \/ ~ (Qeval_nformula env d). Proof. exact (fun env d =>eval_nformula_dec Rsor Q2R env d). Qed. Definition RWitness := Psatz Q. Definition RWeakChecker := check_normalised_formulas 0%Q 1%Q Qplus Qmult Qeq_bool Qle_bool. Require Import List. Lemma RWeakChecker_sound : forall (l : list (NFormula Q)) (cm : RWitness), RWeakChecker l cm = true -> forall env, make_impl (Qeval_nformula env) l False. Proof. intros l cm H. intro. unfold Qeval_nformula. apply (checker_nf_sound Rsor QSORaddon l cm). unfold RWeakChecker in H. exact H. Qed. Require Import Coq.micromega.Tauto. Definition Rnormalise := @cnf_normalise Q 0%Q 1%Q Qplus Qmult Qminus Qopp Qeq_bool Qle_bool. Definition Rnegate := @cnf_negate Q 0%Q 1%Q Qplus Qmult Qminus Qopp Qeq_bool Qle_bool. Definition runsat := check_inconsistent 0%Q Qeq_bool Qle_bool. Definition rdeduce := nformula_plus_nformula 0%Q Qplus Qeq_bool. Definition RTautoChecker (f : BFormula (Formula Rcst)) (w: list RWitness) : bool := @tauto_checker (Formula Q) (NFormula Q) unit runsat rdeduce (Rnormalise unit) (Rnegate unit) RWitness (fun cl => RWeakChecker (List.map fst cl)) (map_bformula (map_Formula Q_of_Rcst) f) w. Lemma RTautoChecker_sound : forall f w, RTautoChecker f w = true -> forall env, eval_bf (Reval_formula env) f. Proof. intros f w. unfold RTautoChecker. intros TC env. apply tauto_checker_sound with (eval:=QReval_formula) (eval':= Qeval_nformula) (env := env) in TC. - change (eval_f (fun x : Prop => x) (QReval_formula env)) with (eval_bf (QReval_formula env)) in TC. rewrite eval_bf_map in TC. unfold eval_bf in TC. rewrite eval_f_morph with (ev':= Reval_formula env) in TC ; auto. intro. unfold QReval_formula. rewrite <- eval_formulaSC with (phiS := R_of_Rcst). rewrite Reval_formula_compat. tauto. intro. rewrite Q_of_RcstR. reflexivity. - apply Reval_nformula_dec. - destruct t. apply (check_inconsistent_sound Rsor QSORaddon) ; auto. - unfold rdeduce. intros. revert H. eapply (nformula_plus_nformula_correct Rsor QSORaddon); eauto. - now apply (cnf_normalise_correct Rsor QSORaddon). - intros. now eapply (cnf_negate_correct Rsor QSORaddon); eauto. - intros t w0. unfold eval_tt. intros. rewrite make_impl_map with (eval := Qeval_nformula env0). eapply RWeakChecker_sound; eauto. tauto. Qed. (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/QMicromega.v0000644000175000017500000001621513612664533020132 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* x) (fun x => x) (pow_N 1 Qmult). Proof. constructor. constructor ; intros ; try reflexivity. apply Qeq_bool_eq; auto. constructor. reflexivity. intros x y. apply Qeq_bool_neq ; auto. apply Qle_bool_imp_le. Qed. (*Definition Zeval_expr := eval_pexpr 0 Z.add Z.mul Z.sub Z.opp (fun x => x) (fun x => Z.of_N x) (Z.pow).*) Require Import EnvRing. Fixpoint Qeval_expr (env: PolEnv Q) (e: PExpr Q) : Q := match e with | PEc c => c | PEX j => env j | PEadd pe1 pe2 => (Qeval_expr env pe1) + (Qeval_expr env pe2) | PEsub pe1 pe2 => (Qeval_expr env pe1) - (Qeval_expr env pe2) | PEmul pe1 pe2 => (Qeval_expr env pe1) * (Qeval_expr env pe2) | PEopp pe1 => - (Qeval_expr env pe1) | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z.of_N n) end. Lemma Qeval_expr_simpl : forall env e, Qeval_expr env e = match e with | PEc c => c | PEX j => env j | PEadd pe1 pe2 => (Qeval_expr env pe1) + (Qeval_expr env pe2) | PEsub pe1 pe2 => (Qeval_expr env pe1) - (Qeval_expr env pe2) | PEmul pe1 pe2 => (Qeval_expr env pe1) * (Qeval_expr env pe2) | PEopp pe1 => - (Qeval_expr env pe1) | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z.of_N n) end. Proof. destruct e ; reflexivity. Qed. Definition Qeval_expr' := eval_pexpr Qplus Qmult Qminus Qopp (fun x => x) (fun x => x) (pow_N 1 Qmult). Lemma QNpower : forall r n, r ^ Z.of_N n = pow_N 1 Qmult r n. Proof. destruct n ; reflexivity. Qed. Lemma Qeval_expr_compat : forall env e, Qeval_expr env e = Qeval_expr' env e. Proof. induction e ; simpl ; subst ; try congruence. reflexivity. rewrite IHe. apply QNpower. Qed. Definition Qeval_op2 (o : Op2) : Q -> Q -> Prop := match o with | OpEq => Qeq | OpNEq => fun x y => ~ x == y | OpLe => Qle | OpGe => fun x y => Qle y x | OpLt => Qlt | OpGt => fun x y => Qlt y x end. Definition Qeval_formula (e:PolEnv Q) (ff : Formula Q) := let (lhs,o,rhs) := ff in Qeval_op2 o (Qeval_expr e lhs) (Qeval_expr e rhs). Definition Qeval_formula' := eval_formula Qplus Qmult Qminus Qopp Qeq Qle Qlt (fun x => x) (fun x => x) (pow_N 1 Qmult). Lemma Qeval_formula_compat : forall env f, Qeval_formula env f <-> Qeval_formula' env f. Proof. intros. unfold Qeval_formula. destruct f. repeat rewrite Qeval_expr_compat. unfold Qeval_formula'. unfold Qeval_expr'. split ; destruct Fop ; simpl; auto. Qed. Definition Qeval_nformula := eval_nformula 0 Qplus Qmult Qeq Qle Qlt (fun x => x) . Definition Qeval_op1 (o : Op1) : Q -> Prop := match o with | Equal => fun x : Q => x == 0 | NonEqual => fun x : Q => ~ x == 0 | Strict => fun x : Q => 0 < x | NonStrict => fun x : Q => 0 <= x end. Lemma Qeval_nformula_dec : forall env d, (Qeval_nformula env d) \/ ~ (Qeval_nformula env d). Proof. exact (fun env d =>eval_nformula_dec Qsor (fun x => x) env d). Qed. Definition QWitness := Psatz Q. Definition QWeakChecker := check_normalised_formulas 0 1 Qplus Qmult Qeq_bool Qle_bool. Require Import List. Lemma QWeakChecker_sound : forall (l : list (NFormula Q)) (cm : QWitness), QWeakChecker l cm = true -> forall env, make_impl (Qeval_nformula env) l False. Proof. intros l cm H. intro. unfold Qeval_nformula. apply (checker_nf_sound Qsor QSORaddon l cm). unfold QWeakChecker in H. exact H. Qed. Require Import Coq.micromega.Tauto. Definition Qnormalise := @cnf_normalise Q 0 1 Qplus Qmult Qminus Qopp Qeq_bool Qle_bool. Definition Qnegate := @cnf_negate Q 0 1 Qplus Qmult Qminus Qopp Qeq_bool Qle_bool. Definition qunsat := check_inconsistent 0 Qeq_bool Qle_bool. Definition qdeduce := nformula_plus_nformula 0 Qplus Qeq_bool. Definition normQ := norm 0 1 Qplus Qmult Qminus Qopp Qeq_bool. Declare Equivalent Keys normQ RingMicromega.norm. Definition cnfQ (Annot TX AF: Type) (f: TFormula (Formula Q) Annot TX AF) := rxcnf qunsat qdeduce (Qnormalise Annot) (Qnegate Annot) true f. Definition QTautoChecker (f : BFormula (Formula Q)) (w: list QWitness) : bool := @tauto_checker (Formula Q) (NFormula Q) unit qunsat qdeduce (Qnormalise unit) (Qnegate unit) QWitness (fun cl => QWeakChecker (List.map fst cl)) f w. Lemma QTautoChecker_sound : forall f w, QTautoChecker f w = true -> forall env, eval_bf (Qeval_formula env) f. Proof. intros f w. unfold QTautoChecker. apply tauto_checker_sound with (eval:= Qeval_formula) (eval':= Qeval_nformula). - apply Qeval_nformula_dec. - intros until env. unfold eval_nformula. unfold RingMicromega.eval_nformula. destruct t. apply (check_inconsistent_sound Qsor QSORaddon) ; auto. - unfold qdeduce. intros. revert H. apply (nformula_plus_nformula_correct Qsor QSORaddon);auto. - intros. rewrite Qeval_formula_compat. unfold Qeval_formula'. now eapply (cnf_normalise_correct Qsor QSORaddon);eauto. - intros. rewrite Qeval_formula_compat. unfold Qeval_formula'. now eapply (cnf_negate_correct Qsor QSORaddon);eauto. - intros t w0. unfold eval_tt. intros. rewrite make_impl_map with (eval := Qeval_nformula env). eapply QWeakChecker_sound; eauto. tauto. Qed. (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/Fourier.v0000644000175000017500000000020113612664533017505 0ustar treinentreinenRequire Import Lra. Require Export Fourier_util. #[deprecated(since = "8.9.0", note = "Use lra instead.")] Ltac fourier := lra. coq-8.11.0/plugins/micromega/micromega.mli0000644000175000017500000004155313612664533020370 0ustar treinentreinentype __ = Obj.t type unit0 = Tt val negb : bool -> bool type nat = O | S of nat type ('a, 'b) sum = Inl of 'a | Inr of 'b val fst : 'a1 * 'a2 -> 'a1 val snd : 'a1 * 'a2 -> 'a2 val app : 'a1 list -> 'a1 list -> 'a1 list type comparison = Eq | Lt | Gt val compOpp : comparison -> comparison val add : nat -> nat -> nat val nth : nat -> 'a1 list -> 'a1 -> 'a1 val rev_append : 'a1 list -> 'a1 list -> 'a1 list val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list val fold_left : ('a1 -> 'a2 -> 'a1) -> 'a2 list -> 'a1 -> 'a1 val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1 type positive = XI of positive | XO of positive | XH type n = N0 | Npos of positive type z = Z0 | Zpos of positive | Zneg of positive module Pos : sig type mask = IsNul | IsPos of positive | IsNeg end module Coq_Pos : sig val succ : positive -> positive val add : positive -> positive -> positive val add_carry : positive -> positive -> positive val pred_double : positive -> positive type mask = Pos.mask = IsNul | IsPos of positive | IsNeg val succ_double_mask : mask -> mask val double_mask : mask -> mask val double_pred_mask : positive -> mask val sub_mask : positive -> positive -> mask val sub_mask_carry : positive -> positive -> mask val sub : positive -> positive -> positive val mul : positive -> positive -> positive val iter : ('a1 -> 'a1) -> 'a1 -> positive -> 'a1 val size_nat : positive -> nat val compare_cont : comparison -> positive -> positive -> comparison val compare : positive -> positive -> comparison val max : positive -> positive -> positive val leb : positive -> positive -> bool val gcdn : nat -> positive -> positive -> positive val gcd : positive -> positive -> positive val of_succ_nat : nat -> positive end module N : sig val of_nat : nat -> n end val pow_pos : ('a1 -> 'a1 -> 'a1) -> 'a1 -> positive -> 'a1 module Z : sig val double : z -> z val succ_double : z -> z val pred_double : z -> z val pos_sub : positive -> positive -> z val add : z -> z -> z val opp : z -> z val sub : z -> z -> z val mul : z -> z -> z val pow_pos : z -> positive -> z val pow : z -> z -> z val compare : z -> z -> comparison val leb : z -> z -> bool val ltb : z -> z -> bool val gtb : z -> z -> bool val max : z -> z -> z val abs : z -> z val to_N : z -> n val of_nat : nat -> z val of_N : n -> z val pos_div_eucl : positive -> z -> z * z val div_eucl : z -> z -> z * z val div : z -> z -> z val gcd : z -> z -> z end val zeq_bool : z -> z -> bool type 'c pol = | Pc of 'c | Pinj of positive * 'c pol | PX of 'c pol * positive * 'c pol val p0 : 'a1 -> 'a1 pol val p1 : 'a1 -> 'a1 pol val peq : ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> bool val mkPinj : positive -> 'a1 pol -> 'a1 pol val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol val mkPX : 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol val mkXi : 'a1 -> 'a1 -> positive -> 'a1 pol val mkX : 'a1 -> 'a1 -> 'a1 pol val popp : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol val paddC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol val psubC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol val paddI : ('a1 -> 'a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol val psubI : ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol val paddX : 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol val psubX : 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol val padd : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol val psub : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol val pmulC_aux : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 pol val pmulC : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 pol val pmulI : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol val pmul : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol val psquare : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol type 'c pExpr = | PEc of 'c | PEX of positive | PEadd of 'c pExpr * 'c pExpr | PEsub of 'c pExpr * 'c pExpr | PEmul of 'c pExpr * 'c pExpr | PEopp of 'c pExpr | PEpow of 'c pExpr * n val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol val ppow_pos : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 pol val ppow_N : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol val norm_aux : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol type ('tA, 'tX, 'aA, 'aF) gFormula = | TT | FF | X of 'tX | A of 'tA * 'aA | Cj of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula | D of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula | N of ('tA, 'tX, 'aA, 'aF) gFormula | I of ('tA, 'tX, 'aA, 'aF) gFormula * 'aF option * ('tA, 'tX, 'aA, 'aF) gFormula val mapX : ('a2 -> 'a2) -> ('a1, 'a2, 'a3, 'a4) gFormula -> ('a1, 'a2, 'a3, 'a4) gFormula val foldA : ('a5 -> 'a3 -> 'a5) -> ('a1, 'a2, 'a3, 'a4) gFormula -> 'a5 -> 'a5 val cons_id : 'a1 option -> 'a1 list -> 'a1 list val ids_of_formula : ('a1, 'a2, 'a3, 'a4) gFormula -> 'a4 list val collect_annot : ('a1, 'a2, 'a3, 'a4) gFormula -> 'a3 list type 'a bFormula = ('a, __, unit0, unit0) gFormula val map_bformula : ('a1 -> 'a2) -> ('a1, 'a3, 'a4, 'a5) gFormula -> ('a2, 'a3, 'a4, 'a5) gFormula type ('x, 'annot) clause = ('x * 'annot) list type ('x, 'annot) cnf = ('x, 'annot) clause list val cnf_tt : ('a1, 'a2) cnf val cnf_ff : ('a1, 'a2) cnf val add_term : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 * 'a2 -> ('a1, 'a2) clause -> ('a1, 'a2) clause option val or_clause : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) clause -> ('a1, 'a2) clause option val xor_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf val or_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf val or_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf val and_cnf : ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf type ('term, 'annot, 'tX, 'aF) tFormula = ('term, 'tX, 'annot, 'aF) gFormula val is_cnf_tt : ('a1, 'a2) cnf -> bool val is_cnf_ff : ('a1, 'a2) cnf -> bool val and_cnf_opt : ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf val or_cnf_opt : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf val xcnf : ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> bool -> ('a1, 'a3, 'a4, 'a5) tFormula -> ('a2, 'a3) cnf val radd_term : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 * 'a2 -> ('a1, 'a2) clause -> (('a1, 'a2) clause, 'a2 list) sum val ror_clause : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause -> (('a1, 'a2) clause, 'a2 list) sum val xror_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list * 'a2 list val ror_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list * 'a2 list val ror_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list -> ('a1, 'a2) cnf * 'a2 list val ror_cnf_opt : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf * 'a2 list val ratom : ('a1, 'a2) cnf -> 'a2 -> ('a1, 'a2) cnf * 'a2 list val rxcnf : ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> bool -> ('a1, 'a3, 'a4, 'a5) tFormula -> ('a2, 'a3) cnf * 'a3 list type ('term, 'annot, 'tX) to_constrT = { mkTT : 'tX ; mkFF : 'tX ; mkA : 'term -> 'annot -> 'tX ; mkCj : 'tX -> 'tX -> 'tX ; mkD : 'tX -> 'tX -> 'tX ; mkI : 'tX -> 'tX -> 'tX ; mkN : 'tX -> 'tX } val aformula : ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> 'a3 val is_X : ('a1, 'a2, 'a3, 'a4) tFormula -> 'a3 option val abs_and : ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ( ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula) -> ('a1, 'a3, 'a2, 'a4) gFormula val abs_or : ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ( ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula) -> ('a1, 'a3, 'a2, 'a4) gFormula val mk_arrow : 'a4 option -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula val abst_form : ('a1, 'a2, 'a3) to_constrT -> ('a2 -> bool) -> bool -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a3, 'a2, 'a4) gFormula val cnf_checker : (('a1 * 'a2) list -> 'a3 -> bool) -> ('a1, 'a2) cnf -> 'a3 list -> bool val tauto_checker : ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> (('a2 * 'a3) list -> 'a4 -> bool) -> ('a1, __, 'a3, unit0) gFormula -> 'a4 list -> bool val cneqb : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool val cltb : ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool type 'c polC = 'c pol type op1 = Equal | NonEqual | Strict | NonStrict type 'c nFormula = 'c polC * op1 val opMult : op1 -> op1 -> op1 option val opAdd : op1 -> op1 -> op1 option type 'c psatz = | PsatzIn of nat | PsatzSquare of 'c polC | PsatzMulC of 'c polC * 'c psatz | PsatzMulE of 'c psatz * 'c psatz | PsatzAdd of 'c psatz * 'c psatz | PsatzC of 'c | PsatzZ val map_option : ('a1 -> 'a2 option) -> 'a1 option -> 'a2 option val map_option2 : ('a1 -> 'a2 -> 'a3 option) -> 'a1 option -> 'a2 option -> 'a3 option val pexpr_times_nformula : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option val nformula_times_nformula : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option val nformula_plus_nformula : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option val eval_Psatz : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1 nFormula option val check_inconsistent : 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> bool val check_normalised_formulas : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool type op2 = OpEq | OpNEq | OpLe | OpGe | OpLt | OpGt type 't formula = {flhs : 't pExpr; fop : op2; frhs : 't pExpr} val norm : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol val psub0 : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol val padd0 : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol val popp0 : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol val normalise : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula val xnormalise : ('a1 -> 'a1) -> 'a1 nFormula -> 'a1 nFormula list val xnegate : ('a1 -> 'a1) -> 'a1 nFormula -> 'a1 nFormula list val cnf_of_list : 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a2 -> ('a1 nFormula, 'a2) cnf val cnf_normalise : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula, 'a2) cnf val cnf_negate : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula, 'a2) cnf val xdenorm : positive -> 'a1 pol -> 'a1 pExpr val denorm : 'a1 pol -> 'a1 pExpr val map_PExpr : ('a2 -> 'a1) -> 'a2 pExpr -> 'a1 pExpr val map_Formula : ('a2 -> 'a1) -> 'a2 formula -> 'a1 formula val simpl_cone : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz -> 'a1 psatz type q = {qnum : z; qden : positive} val qeq_bool : q -> q -> bool val qle_bool : q -> q -> bool val qplus : q -> q -> q val qmult : q -> q -> q val qopp : q -> q val qminus : q -> q -> q val qinv : q -> q val qpower_positive : q -> positive -> q val qpower : q -> z -> q type 'a t = Empty | Elt of 'a | Branch of 'a t * 'a * 'a t val find : 'a1 -> 'a1 t -> positive -> 'a1 val singleton : 'a1 -> positive -> 'a1 -> 'a1 t val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t val zeval_const : z pExpr -> z option type zWitness = z psatz val zWeakChecker : z nFormula list -> z psatz -> bool val psub1 : z pol -> z pol -> z pol val padd1 : z pol -> z pol -> z pol val normZ : z pExpr -> z pol val zunsat : z nFormula -> bool val zdeduce : z nFormula -> z nFormula -> z nFormula option val xnnormalise : z formula -> z nFormula val xnormalise0 : z nFormula -> z nFormula list val cnf_of_list0 : 'a1 -> z nFormula list -> (z nFormula * 'a1) list list val normalise0 : z formula -> 'a1 -> (z nFormula, 'a1) cnf val xnegate0 : z nFormula -> z nFormula list val negate : z formula -> 'a1 -> (z nFormula, 'a1) cnf val cnfZ : (z formula, 'a1, 'a2, 'a3) tFormula -> (z nFormula, 'a1) cnf * 'a1 list val ceiling : z -> z -> z type zArithProof = | DoneProof | RatProof of zWitness * zArithProof | CutProof of zWitness * zArithProof | EnumProof of zWitness * zWitness * zArithProof list | ExProof of positive * zArithProof val zgcdM : z -> z -> z val zgcd_pol : z polC -> z * z val zdiv_pol : z polC -> z -> z polC val makeCuttingPlane : z polC -> z polC * z val genCuttingPlane : z nFormula -> ((z polC * z) * op1) option val nformula_of_cutting_plane : (z polC * z) * op1 -> z nFormula val is_pol_Z0 : z polC -> bool val eval_Psatz0 : z nFormula list -> zWitness -> z nFormula option val valid_cut_sign : op1 -> bool val bound_var : positive -> z formula val mk_eq_pos : positive -> positive -> positive -> z formula val max_var : positive -> z pol -> positive val max_var_nformulae : z nFormula list -> positive val zChecker : z nFormula list -> zArithProof -> bool val zTautoChecker : z formula bFormula -> zArithProof list -> bool type qWitness = q psatz val qWeakChecker : q nFormula list -> q psatz -> bool val qnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf val qnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf val qunsat : q nFormula -> bool val qdeduce : q nFormula -> q nFormula -> q nFormula option val normQ : q pExpr -> q pol val cnfQ : (q formula, 'a1, 'a2, 'a3) tFormula -> (q nFormula, 'a1) cnf * 'a1 list val qTautoChecker : q formula bFormula -> qWitness list -> bool type rcst = | C0 | C1 | CQ of q | CZ of z | CPlus of rcst * rcst | CMinus of rcst * rcst | CMult of rcst * rcst | CPow of rcst * (z, nat) sum | CInv of rcst | COpp of rcst val z_of_exp : (z, nat) sum -> z val q_of_Rcst : rcst -> q type rWitness = q psatz val rWeakChecker : q nFormula list -> q psatz -> bool val rnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf val rnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf val runsat : q nFormula -> bool val rdeduce : q nFormula -> q nFormula -> q nFormula option val rTautoChecker : rcst formula bFormula -> rWitness list -> bool coq-8.11.0/plugins/micromega/EnvRing.v0000644000175000017500000007554713612664533017473 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* PExpr | PEX : positive -> PExpr | PEadd : PExpr -> PExpr -> PExpr | PEsub : PExpr -> PExpr -> PExpr | PEmul : PExpr -> PExpr -> PExpr | PEopp : PExpr -> PExpr | PEpow : PExpr -> N -> PExpr. Arguments PExpr : clear implicits. (* Definition of multivariable polynomials with coefficients in C : Type [Pol] represents [X1 ... Xn]. The representation is Horner's where a [n] variable polynomial (C[X1..Xn]) is seen as a polynomial on [X1] which coefficients are polynomials with [n-1] variables (C[X2..Xn]). There are several optimisations to make the repr compacter: - [Pc c] is the constant polynomial of value c == c*X1^0*..*Xn^0 - [Pinj j Q] is a polynomial constant w.r.t the [j] first variables. variable indices are shifted of j in Q. == X1^0 *..* Xj^0 * Q{X1 <- Xj+1;..; Xn-j <- Xn} - [PX P i Q] is an optimised Horner form of P*X^i + Q with P not the null polynomial == P * X1^i + Q{X1 <- X2; ..; Xn-1 <- Xn} In addition: - polynomials of the form (PX (PX P i (Pc 0)) j Q) are forbidden since they can be represented by the simpler form (PX P (i+j) Q) - (Pinj i (Pinj j P)) is (Pinj (i+j) P) - (Pinj i (Pc c)) is (Pc c) *) #[universes(template)] Inductive Pol {C} : Type := | Pc : C -> Pol | Pinj : positive -> Pol -> Pol | PX : Pol -> positive -> Pol -> Pol. Arguments Pol : clear implicits. Section MakeRingPol. (* Ring elements *) Variable R:Type. Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R->R). Variable req : R -> R -> Prop. (* Ring properties *) Variable Rsth : Equivalence req. Variable Reqe : ring_eq_ext radd rmul ropp req. Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req. (* Coefficients *) Variable C: Type. Variable (cO cI: C) (cadd cmul csub : C->C->C) (copp : C->C). Variable ceqb : C->C->bool. Variable phi : C -> R. Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi. (* Power coefficients *) Variable Cpow : Type. Variable Cp_phi : N -> Cpow. Variable rpow : R -> Cpow -> R. Variable pow_th : power_theory rI rmul req Cp_phi rpow. (* R notations *) Notation "0" := rO. Notation "1" := rI. Infix "+" := radd. Infix "*" := rmul. Infix "-" := rsub. Notation "- x" := (ropp x). Infix "==" := req. Infix "^" := (pow_pos rmul). (* C notations *) Infix "+!" := cadd. Infix "*!" := cmul. Infix "-! " := csub. Notation "-! x" := (copp x). Infix "?=!" := ceqb. Notation "[ x ]" := (phi x). (* Useful tactics *) Add Morphism radd with signature (req ==> req ==> req) as radd_ext. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext. Proof. exact (Ropp_ext Reqe). Qed. Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed. Ltac rsimpl := gen_srewrite Rsth Reqe ARth. Ltac add_push := gen_add_push radd Rsth Reqe ARth. Ltac mul_push := gen_mul_push rmul Rsth Reqe ARth. Ltac add_permut_rec t := match t with | ?x + ?y => add_permut_rec y || add_permut_rec x | _ => add_push t; apply (Radd_ext Reqe); [|reflexivity] end. Ltac add_permut := repeat (reflexivity || match goal with |- ?t == _ => add_permut_rec t end). Ltac mul_permut_rec t := match t with | ?x * ?y => mul_permut_rec y || mul_permut_rec x | _ => mul_push t; apply (Rmul_ext Reqe); [|reflexivity] end. Ltac mul_permut := repeat (reflexivity || match goal with |- ?t == _ => mul_permut_rec t end). Notation PExpr := (PExpr C). Notation Pol := (Pol C). Implicit Types pe : PExpr. Implicit Types P : Pol. Definition P0 := Pc cO. Definition P1 := Pc cI. Fixpoint Peq (P P' : Pol) {struct P'} : bool := match P, P' with | Pc c, Pc c' => c ?=! c' | Pinj j Q, Pinj j' Q' => match j ?= j' with | Eq => Peq Q Q' | _ => false end | PX P i Q, PX P' i' Q' => match i ?= i' with | Eq => if Peq P P' then Peq Q Q' else false | _ => false end | _, _ => false end. Infix "?==" := Peq. Definition mkPinj j P := match P with | Pc _ => P | Pinj j' Q => Pinj (j + j') Q | _ => Pinj j P end. Definition mkPinj_pred j P := match j with | xH => P | xO j => Pinj (Pos.pred_double j) P | xI j => Pinj (xO j) P end. Definition mkPX P i Q := match P with | Pc c => if c ?=! cO then mkPinj xH Q else PX P i Q | Pinj _ _ => PX P i Q | PX P' i' Q' => if Q' ?== P0 then PX P' (i' + i) Q else PX P i Q end. Definition mkXi i := PX P1 i P0. Definition mkX := mkXi 1. (** Opposite of addition *) Fixpoint Popp (P:Pol) : Pol := match P with | Pc c => Pc (-! c) | Pinj j Q => Pinj j (Popp Q) | PX P i Q => PX (Popp P) i (Popp Q) end. Notation "-- P" := (Popp P). (** Addition et subtraction *) Fixpoint PaddC (P:Pol) (c:C) : Pol := match P with | Pc c1 => Pc (c1 +! c) | Pinj j Q => Pinj j (PaddC Q c) | PX P i Q => PX P i (PaddC Q c) end. Fixpoint PsubC (P:Pol) (c:C) : Pol := match P with | Pc c1 => Pc (c1 -! c) | Pinj j Q => Pinj j (PsubC Q c) | PX P i Q => PX P i (PsubC Q c) end. Section PopI. Variable Pop : Pol -> Pol -> Pol. Variable Q : Pol. Fixpoint PaddI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PaddC Q c) | Pinj j' Q' => match Z.pos_sub j' j with | Zpos k => mkPinj j (Pop (Pinj k Q') Q) | Z0 => mkPinj j (Pop Q' Q) | Zneg k => mkPinj j' (PaddI k Q') end | PX P i Q' => match j with | xH => PX P i (Pop Q' Q) | xO j => PX P i (PaddI (Pos.pred_double j) Q') | xI j => PX P i (PaddI (xO j) Q') end end. Fixpoint PsubI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PaddC (--Q) c) | Pinj j' Q' => match Z.pos_sub j' j with | Zpos k => mkPinj j (Pop (Pinj k Q') Q) | Z0 => mkPinj j (Pop Q' Q) | Zneg k => mkPinj j' (PsubI k Q') end | PX P i Q' => match j with | xH => PX P i (Pop Q' Q) | xO j => PX P i (PsubI (Pos.pred_double j) Q') | xI j => PX P i (PsubI (xO j) Q') end end. Variable P' : Pol. Fixpoint PaddX (i':positive) (P:Pol) : Pol := match P with | Pc c => PX P' i' P | Pinj j Q' => match j with | xH => PX P' i' Q' | xO j => PX P' i' (Pinj (Pos.pred_double j) Q') | xI j => PX P' i' (Pinj (xO j) Q') end | PX P i Q' => match Z.pos_sub i i' with | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' | Z0 => mkPX (Pop P P') i Q' | Zneg k => mkPX (PaddX k P) i Q' end end. Fixpoint PsubX (i':positive) (P:Pol) : Pol := match P with | Pc c => PX (--P') i' P | Pinj j Q' => match j with | xH => PX (--P') i' Q' | xO j => PX (--P') i' (Pinj (Pos.pred_double j) Q') | xI j => PX (--P') i' (Pinj (xO j) Q') end | PX P i Q' => match Z.pos_sub i i' with | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' | Z0 => mkPX (Pop P P') i Q' | Zneg k => mkPX (PsubX k P) i Q' end end. End PopI. Fixpoint Padd (P P': Pol) {struct P'} : Pol := match P' with | Pc c' => PaddC P c' | Pinj j' Q' => PaddI Padd Q' j' P | PX P' i' Q' => match P with | Pc c => PX P' i' (PaddC Q' c) | Pinj j Q => match j with | xH => PX P' i' (Padd Q Q') | xO j => PX P' i' (Padd (Pinj (Pos.pred_double j) Q) Q') | xI j => PX P' i' (Padd (Pinj (xO j) Q) Q') end | PX P i Q => match Z.pos_sub i i' with | Zpos k => mkPX (Padd (PX P k P0) P') i' (Padd Q Q') | Z0 => mkPX (Padd P P') i (Padd Q Q') | Zneg k => mkPX (PaddX Padd P' k P) i (Padd Q Q') end end end. Infix "++" := Padd. Fixpoint Psub (P P': Pol) {struct P'} : Pol := match P' with | Pc c' => PsubC P c' | Pinj j' Q' => PsubI Psub Q' j' P | PX P' i' Q' => match P with | Pc c => PX (--P') i' (*(--(PsubC Q' c))*) (PaddC (--Q') c) | Pinj j Q => match j with | xH => PX (--P') i' (Psub Q Q') | xO j => PX (--P') i' (Psub (Pinj (Pos.pred_double j) Q) Q') | xI j => PX (--P') i' (Psub (Pinj (xO j) Q) Q') end | PX P i Q => match Z.pos_sub i i' with | Zpos k => mkPX (Psub (PX P k P0) P') i' (Psub Q Q') | Z0 => mkPX (Psub P P') i (Psub Q Q') | Zneg k => mkPX (PsubX Psub P' k P) i (Psub Q Q') end end end. Infix "--" := Psub. (** Multiplication *) Fixpoint PmulC_aux (P:Pol) (c:C) : Pol := match P with | Pc c' => Pc (c' *! c) | Pinj j Q => mkPinj j (PmulC_aux Q c) | PX P i Q => mkPX (PmulC_aux P c) i (PmulC_aux Q c) end. Definition PmulC P c := if c ?=! cO then P0 else if c ?=! cI then P else PmulC_aux P c. Section PmulI. Variable Pmul : Pol -> Pol -> Pol. Variable Q : Pol. Fixpoint PmulI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PmulC Q c) | Pinj j' Q' => match Z.pos_sub j' j with | Zpos k => mkPinj j (Pmul (Pinj k Q') Q) | Z0 => mkPinj j (Pmul Q' Q) | Zneg k => mkPinj j' (PmulI k Q') end | PX P' i' Q' => match j with | xH => mkPX (PmulI xH P') i' (Pmul Q' Q) | xO j' => mkPX (PmulI j P') i' (PmulI (Pos.pred_double j') Q') | xI j' => mkPX (PmulI j P') i' (PmulI (xO j') Q') end end. End PmulI. Fixpoint Pmul (P P'' : Pol) {struct P''} : Pol := match P'' with | Pc c => PmulC P c | Pinj j' Q' => PmulI Pmul Q' j' P | PX P' i' Q' => match P with | Pc c => PmulC P'' c | Pinj j Q => let QQ' := match j with | xH => Pmul Q Q' | xO j => Pmul (Pinj (Pos.pred_double j) Q) Q' | xI j => Pmul (Pinj (xO j) Q) Q' end in mkPX (Pmul P P') i' QQ' | PX P i Q=> let QQ' := Pmul Q Q' in let PQ' := PmulI Pmul Q' xH P in let QP' := Pmul (mkPinj xH Q) P' in let PP' := Pmul P P' in (mkPX (mkPX PP' i P0 ++ QP') i' P0) ++ mkPX PQ' i QQ' end end. Infix "**" := Pmul. Fixpoint Psquare (P:Pol) : Pol := match P with | Pc c => Pc (c *! c) | Pinj j Q => Pinj j (Psquare Q) | PX P i Q => let twoPQ := Pmul P (mkPinj xH (PmulC Q (cI +! cI))) in let Q2 := Psquare Q in let P2 := Psquare P in mkPX (mkPX P2 i P0 ++ twoPQ) i Q2 end. (** Monomial **) (** A monomial is X1^k1...Xi^ki. Its representation is a simplified version of the polynomial representation: - [mon0] correspond to the polynom [P1]. - [(zmon j M)] corresponds to [(Pinj j ...)], i.e. skip j variable indices. - [(vmon i M)] is X^i*M with X the current variable, its corresponds to (PX P1 i ...)] *) Inductive Mon: Set := | mon0: Mon | zmon: positive -> Mon -> Mon | vmon: positive -> Mon -> Mon. Definition mkZmon j M := match M with mon0 => mon0 | _ => zmon j M end. Definition zmon_pred j M := match j with xH => M | _ => mkZmon (Pos.pred j) M end. Definition mkVmon i M := match M with | mon0 => vmon i mon0 | zmon j m => vmon i (zmon_pred j m) | vmon i' m => vmon (i+i') m end. Fixpoint MFactor (P: Pol) (M: Mon) : Pol * Pol := match P, M with _, mon0 => (Pc cO, P) | Pc _, _ => (P, Pc cO) | Pinj j1 P1, zmon j2 M1 => match (j1 ?= j2) with Eq => let (R,S) := MFactor P1 M1 in (mkPinj j1 R, mkPinj j1 S) | Lt => let (R,S) := MFactor P1 (zmon (j2 - j1) M1) in (mkPinj j1 R, mkPinj j1 S) | Gt => (P, Pc cO) end | Pinj _ _, vmon _ _ => (P, Pc cO) | PX P1 i Q1, zmon j M1 => let M2 := zmon_pred j M1 in let (R1, S1) := MFactor P1 M in let (R2, S2) := MFactor Q1 M2 in (mkPX R1 i R2, mkPX S1 i S2) | PX P1 i Q1, vmon j M1 => match (i ?= j) with Eq => let (R1,S1) := MFactor P1 (mkZmon xH M1) in (mkPX R1 i Q1, S1) | Lt => let (R1,S1) := MFactor P1 (vmon (j - i) M1) in (mkPX R1 i Q1, S1) | Gt => let (R1,S1) := MFactor P1 (mkZmon xH M1) in (mkPX R1 i Q1, mkPX S1 (i-j) (Pc cO)) end end. Definition POneSubst (P1: Pol) (M1: Mon) (P2: Pol): option Pol := let (Q1,R1) := MFactor P1 M1 in match R1 with (Pc c) => if c ?=! cO then None else Some (Padd Q1 (Pmul P2 R1)) | _ => Some (Padd Q1 (Pmul P2 R1)) end. Fixpoint PNSubst1 (P1: Pol) (M1: Mon) (P2: Pol) (n: nat) : Pol := match POneSubst P1 M1 P2 with Some P3 => match n with S n1 => PNSubst1 P3 M1 P2 n1 | _ => P3 end | _ => P1 end. Definition PNSubst (P1: Pol) (M1: Mon) (P2: Pol) (n: nat): option Pol := match POneSubst P1 M1 P2 with Some P3 => match n with S n1 => Some (PNSubst1 P3 M1 P2 n1) | _ => None end | _ => None end. Fixpoint PSubstL1 (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) : Pol := match LM1 with cons (M1,P2) LM2 => PSubstL1 (PNSubst1 P1 M1 P2 n) LM2 n | _ => P1 end. Fixpoint PSubstL (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) : option Pol := match LM1 with cons (M1,P2) LM2 => match PNSubst P1 M1 P2 n with Some P3 => Some (PSubstL1 P3 LM2 n) | None => PSubstL P1 LM2 n end | _ => None end. Fixpoint PNSubstL (P1: Pol) (LM1: list (Mon * Pol)) (m n: nat) : Pol := match PSubstL P1 LM1 n with Some P3 => match m with S m1 => PNSubstL P3 LM1 m1 n | _ => P3 end | _ => P1 end. (** Evaluation of a polynomial towards R *) Fixpoint Pphi(l:Env R) (P:Pol) : R := match P with | Pc c => [c] | Pinj j Q => Pphi (jump j l) Q | PX P i Q => Pphi l P * (hd l) ^ i + Pphi (tail l) Q end. Reserved Notation "P @ l " (at level 10, no associativity). Notation "P @ l " := (Pphi l P). (** Evaluation of a monomial towards R *) Fixpoint Mphi(l:Env R) (M: Mon) : R := match M with | mon0 => rI | zmon j M1 => Mphi (jump j l) M1 | vmon i M1 => Mphi (tail l) M1 * (hd l) ^ i end. Notation "M @@ l" := (Mphi l M) (at level 10, no associativity). (** Proofs *) Ltac destr_pos_sub := match goal with |- context [Z.pos_sub ?x ?y] => generalize (Z.pos_sub_discr x y); destruct (Z.pos_sub x y) end. Lemma Peq_ok P P' : (P ?== P') = true -> forall l, P@l == P'@ l. Proof. revert P';induction P;destruct P';simpl; intros H l; try easy. - now apply (morph_eq CRmorph). - destruct (Pos.compare_spec p p0); [ subst | easy | easy ]. now rewrite IHP. - specialize (IHP1 P'1); specialize (IHP2 P'2). destruct (Pos.compare_spec p p0); [ subst | easy | easy ]. destruct (P2 ?== P'1); [|easy]. rewrite H in *. now rewrite IHP1, IHP2. Qed. Lemma Peq_spec P P' : BoolSpec (forall l, P@l == P'@l) True (P ?== P'). Proof. generalize (Peq_ok P P'). destruct (P ?== P'); auto. Qed. Lemma Pphi0 l : P0@l == 0. Proof. simpl;apply (morph0 CRmorph). Qed. Lemma Pphi1 l : P1@l == 1. Proof. simpl;apply (morph1 CRmorph). Qed. Lemma env_morph p e1 e2 : (forall x, e1 x = e2 x) -> p @ e1 = p @ e2. Proof. revert e1 e2. induction p ; simpl. - reflexivity. - intros e1 e2 EQ. apply IHp. intros. apply EQ. - intros e1 e2 EQ. f_equal; [f_equal|]. + now apply IHp1. + f_equal. apply EQ. + apply IHp2. intros; apply EQ. Qed. Lemma Pjump_add P i j l : P @ (jump (i + j) l) = P @ (jump j (jump i l)). Proof. apply env_morph. intros. rewrite <- jump_add. f_equal. apply Pos.add_comm. Qed. Lemma Pjump_xO_tail P p l : P @ (jump (xO p) (tail l)) = P @ (jump (xI p) l). Proof. apply env_morph. intros. now jump_simpl. Qed. Lemma Pjump_pred_double P p l : P @ (jump (Pos.pred_double p) (tail l)) = P @ (jump (xO p) l). Proof. apply env_morph. intros. rewrite jump_pred_double. now jump_simpl. Qed. Lemma mkPinj_ok j l P : (mkPinj j P)@l == P@(jump j l). Proof. destruct P;simpl;rsimpl. now rewrite Pjump_add. Qed. Lemma pow_pos_add x i j : x^(j + i) == x^i * x^j. Proof. rewrite Pos.add_comm. apply (pow_pos_add Rsth (Rmul_ext Reqe) (ARmul_assoc ARth)). Qed. Lemma ceqb_spec c c' : BoolSpec ([c] == [c']) True (c ?=! c'). Proof. generalize (morph_eq CRmorph c c'). destruct (c ?=! c'); auto. Qed. Lemma mkPX_ok l P i Q : (mkPX P i Q)@l == P@l * (hd l)^i + Q@(tail l). Proof. unfold mkPX. destruct P. - case ceqb_spec; intros H; simpl; try reflexivity. rewrite H, (morph0 CRmorph), mkPinj_ok; rsimpl. - reflexivity. - case Peq_spec; intros H; simpl; try reflexivity. rewrite H, Pphi0, Pos.add_comm, pow_pos_add; rsimpl. Qed. Hint Rewrite Pphi0 Pphi1 mkPinj_ok mkPX_ok (morph0 CRmorph) (morph1 CRmorph) (morph0 CRmorph) (morph_add CRmorph) (morph_mul CRmorph) (morph_sub CRmorph) (morph_opp CRmorph) : Esimpl. (* Quicker than autorewrite with Esimpl :-) *) Ltac Esimpl := try rewrite_db Esimpl; rsimpl; simpl. Lemma PaddC_ok c P l : (PaddC P c)@l == P@l + [c]. Proof. revert l;induction P;simpl;intros;Esimpl;trivial. rewrite IHP2;rsimpl. Qed. Lemma PsubC_ok c P l : (PsubC P c)@l == P@l - [c]. Proof. revert l;induction P;simpl;intros. - Esimpl. - rewrite IHP;rsimpl. - rewrite IHP2;rsimpl. Qed. Lemma PmulC_aux_ok c P l : (PmulC_aux P c)@l == P@l * [c]. Proof. revert l;induction P;simpl;intros;Esimpl;trivial. rewrite IHP1, IHP2;rsimpl. add_permut. mul_permut. Qed. Lemma PmulC_ok c P l : (PmulC P c)@l == P@l * [c]. Proof. unfold PmulC. case ceqb_spec; intros H. - rewrite H; Esimpl. - case ceqb_spec; intros H'. + rewrite H'; Esimpl. + apply PmulC_aux_ok. Qed. Lemma Popp_ok P l : (--P)@l == - P@l. Proof. revert l;induction P;simpl;intros. - Esimpl. - apply IHP. - rewrite IHP1, IHP2;rsimpl. Qed. Hint Rewrite PaddC_ok PsubC_ok PmulC_ok Popp_ok : Esimpl. Lemma PaddX_ok P' P k l : (forall P l, (P++P')@l == P@l + P'@l) -> (PaddX Padd P' k P) @ l == P@l + P'@l * (hd l)^k. Proof. intros IHP'. revert k l. induction P;simpl;intros. - add_permut. - destruct p; simpl; rewrite ?Pjump_xO_tail, ?Pjump_pred_double; add_permut. - destr_pos_sub; intros ->;Esimpl. + rewrite IHP';rsimpl. add_permut. + rewrite IHP', pow_pos_add;simpl;Esimpl. add_permut. + rewrite IHP1, pow_pos_add;rsimpl. add_permut. Qed. Lemma Padd_ok P' P l : (P ++ P')@l == P@l + P'@l. Proof. revert P l; induction P';simpl;intros;Esimpl. - revert p l; induction P;simpl;intros. + Esimpl; add_permut. + destr_pos_sub; intros ->;Esimpl. * now rewrite IHP'. * rewrite IHP';Esimpl. now rewrite Pjump_add. * rewrite IHP. now rewrite Pjump_add. + destruct p0;simpl. * rewrite IHP2;simpl. rsimpl. rewrite Pjump_xO_tail. Esimpl. * rewrite IHP2;simpl. rewrite Pjump_pred_double. rsimpl. * rewrite IHP'. rsimpl. - destruct P;simpl. + Esimpl. add_permut. + destruct p0;simpl;Esimpl; rewrite IHP'2; simpl. * rewrite Pjump_xO_tail. rsimpl. add_permut. * rewrite Pjump_pred_double. rsimpl. add_permut. * rsimpl. unfold tail. add_permut. + destr_pos_sub; intros ->; Esimpl. * rewrite IHP'1, IHP'2;rsimpl. add_permut. * rewrite IHP'1, IHP'2;simpl;Esimpl. rewrite pow_pos_add;rsimpl. add_permut. * rewrite PaddX_ok by trivial; rsimpl. rewrite IHP'2, pow_pos_add; rsimpl. add_permut. Qed. Lemma PsubX_ok P' P k l : (forall P l, (P--P')@l == P@l - P'@l) -> (PsubX Psub P' k P) @ l == P@l - P'@l * (hd l)^k. Proof. intros IHP'. revert k l. induction P;simpl;intros. - rewrite Popp_ok;rsimpl; add_permut. - destruct p; simpl; rewrite Popp_ok;rsimpl; rewrite ?Pjump_xO_tail, ?Pjump_pred_double; add_permut. - destr_pos_sub; intros ->; Esimpl. + rewrite IHP';rsimpl. add_permut. + rewrite IHP', pow_pos_add;simpl;Esimpl. add_permut. + rewrite IHP1, pow_pos_add;rsimpl. add_permut. Qed. Lemma Psub_ok P' P l : (P -- P')@l == P@l - P'@l. Proof. revert P l; induction P';simpl;intros;Esimpl. - revert p l; induction P;simpl;intros. + Esimpl; add_permut. + destr_pos_sub; intros ->;Esimpl. * rewrite IHP';rsimpl. * rewrite IHP';Esimpl. now rewrite Pjump_add. * rewrite IHP. now rewrite Pjump_add. + destruct p0;simpl. * rewrite IHP2;simpl. rsimpl. rewrite Pjump_xO_tail. Esimpl. * rewrite IHP2;simpl. rewrite Pjump_pred_double. rsimpl. * rewrite IHP'. rsimpl. - destruct P;simpl. + Esimpl; add_permut. + destruct p0;simpl;Esimpl; rewrite IHP'2; simpl. * rewrite Pjump_xO_tail. rsimpl. add_permut. * rewrite Pjump_pred_double. rsimpl. add_permut. * rsimpl. unfold tail. add_permut. + destr_pos_sub; intros ->; Esimpl. * rewrite IHP'1, IHP'2;rsimpl. add_permut. * rewrite IHP'1, IHP'2;simpl;Esimpl. rewrite pow_pos_add;rsimpl. add_permut. * rewrite PsubX_ok by trivial;rsimpl. rewrite IHP'2, pow_pos_add;rsimpl. add_permut. Qed. Lemma PmulI_ok P' : (forall P l, (Pmul P P') @ l == P @ l * P' @ l) -> forall P p l, (PmulI Pmul P' p P) @ l == P @ l * P' @ (jump p l). Proof. intros IHP'. induction P;simpl;intros. - Esimpl; mul_permut. - destr_pos_sub; intros ->;Esimpl. + now rewrite IHP'. + now rewrite IHP', Pjump_add. + now rewrite IHP, Pjump_add. - destruct p0;Esimpl; rewrite ?IHP1, ?IHP2; rsimpl. + rewrite Pjump_xO_tail. f_equiv. mul_permut. + rewrite Pjump_pred_double. f_equiv. mul_permut. + rewrite IHP'. f_equiv. mul_permut. Qed. Lemma Pmul_ok P P' l : (P**P')@l == P@l * P'@l. Proof. revert P l;induction P';simpl;intros. - apply PmulC_ok. - apply PmulI_ok;trivial. - destruct P. + rewrite (ARmul_comm ARth). Esimpl. + Esimpl. rewrite IHP'1;Esimpl. f_equiv. destruct p0;rewrite IHP'2;Esimpl. * now rewrite Pjump_xO_tail. * rewrite Pjump_pred_double; Esimpl. + rewrite Padd_ok, !mkPX_ok, Padd_ok, !mkPX_ok, !IHP'1, !IHP'2, PmulI_ok; trivial. simpl. Esimpl. unfold tail. add_permut; f_equiv; mul_permut. Qed. Lemma Psquare_ok P l : (Psquare P)@l == P@l * P@l. Proof. revert l;induction P;simpl;intros;Esimpl. - apply IHP. - rewrite Padd_ok, Pmul_ok;Esimpl. rewrite IHP1, IHP2. mul_push ((hd l)^p). now mul_push (P2@l). Qed. Lemma Mphi_morph M e1 e2 : (forall x, e1 x = e2 x) -> M @@ e1 = M @@ e2. Proof. revert e1 e2; induction M; simpl; intros e1 e2 EQ; trivial. - apply IHM. intros; apply EQ. - f_equal. * apply IHM. intros; apply EQ. * f_equal. apply EQ. Qed. Lemma Mjump_xO_tail M p l : M @@ (jump (xO p) (tail l)) = M @@ (jump (xI p) l). Proof. apply Mphi_morph. intros. now jump_simpl. Qed. Lemma Mjump_pred_double M p l : M @@ (jump (Pos.pred_double p) (tail l)) = M @@ (jump (xO p) l). Proof. apply Mphi_morph. intros. rewrite jump_pred_double. now jump_simpl. Qed. Lemma Mjump_add M i j l : M @@ (jump (i + j) l) = M @@ (jump j (jump i l)). Proof. apply Mphi_morph. intros. now rewrite <- jump_add, Pos.add_comm. Qed. Lemma mkZmon_ok M j l : (mkZmon j M) @@ l == (zmon j M) @@ l. Proof. destruct M; simpl; rsimpl. Qed. Lemma zmon_pred_ok M j l : (zmon_pred j M) @@ (tail l) == (zmon j M) @@ l. Proof. destruct j; simpl; rewrite ?mkZmon_ok; simpl; rsimpl. - now rewrite Mjump_xO_tail. - rewrite Mjump_pred_double; rsimpl. Qed. Lemma mkVmon_ok M i l : (mkVmon i M)@@l == M@@l * (hd l)^i. Proof. destruct M;simpl;intros;rsimpl. - rewrite zmon_pred_ok;simpl;rsimpl. - rewrite pow_pos_add;rsimpl. Qed. Ltac destr_mfactor R S := match goal with | H : context [MFactor ?P _] |- context [MFactor ?P ?M] => specialize (H M); destruct MFactor as (R,S) end. Lemma Mphi_ok P M l : let (Q,R) := MFactor P M in P@l == Q@l + M@@l * R@l. Proof. revert M l; induction P; destruct M; intros l; simpl; auto; Esimpl. - case Pos.compare_spec; intros He; simpl. * destr_mfactor R1 S1. now rewrite IHP, He, !mkPinj_ok. * destr_mfactor R1 S1. rewrite IHP; simpl. now rewrite !mkPinj_ok, <- Mjump_add, Pos.add_comm, Pos.sub_add. * Esimpl. - destr_mfactor R1 S1. destr_mfactor R2 S2. rewrite IHP1, IHP2, !mkPX_ok, zmon_pred_ok; simpl; rsimpl. add_permut. - case Pos.compare_spec; intros He; simpl; destr_mfactor R1 S1; rewrite ?He, IHP1, mkPX_ok, ?mkZmon_ok; simpl; rsimpl; unfold tail; add_permut; mul_permut. * rewrite <- pow_pos_add, Pos.add_comm, Pos.sub_add by trivial; rsimpl. * rewrite mkPX_ok. simpl. Esimpl. mul_permut. rewrite <- pow_pos_add, Pos.sub_add by trivial; rsimpl. Qed. Lemma POneSubst_ok P1 M1 P2 P3 l : POneSubst P1 M1 P2 = Some P3 -> M1@@l == P2@l -> P1@l == P3@l. Proof. unfold POneSubst. assert (H := Mphi_ok P1). destr_mfactor R1 S1. rewrite H; clear H. intros EQ EQ'. replace P3 with (R1 ++ P2 ** S1). - rewrite EQ', Padd_ok, Pmul_ok; rsimpl. - revert EQ. destruct S1; try now injection 1. case ceqb_spec; now inversion 2. Qed. Lemma PNSubst1_ok n P1 M1 P2 l : M1@@l == P2@l -> P1@l == (PNSubst1 P1 M1 P2 n)@l. Proof. revert P1. induction n; simpl; intros P1; generalize (POneSubst_ok P1 M1 P2); destruct POneSubst; intros; rewrite <- ?IHn; auto; reflexivity. Qed. Lemma PNSubst_ok n P1 M1 P2 l P3 : PNSubst P1 M1 P2 n = Some P3 -> M1@@l == P2@l -> P1@l == P3@l. Proof. unfold PNSubst. assert (H := POneSubst_ok P1 M1 P2); destruct POneSubst; try discriminate. destruct n; inversion_clear 1. intros. rewrite <- PNSubst1_ok; auto. Qed. Fixpoint MPcond (LM1: list (Mon * Pol)) (l: Env R) : Prop := match LM1 with | cons (M1,P2) LM2 => (M1@@l == P2@l) /\ MPcond LM2 l | _ => True end. Lemma PSubstL1_ok n LM1 P1 l : MPcond LM1 l -> P1@l == (PSubstL1 P1 LM1 n)@l. Proof. revert P1; induction LM1 as [|(M2,P2) LM2 IH]; simpl; intros. - reflexivity. - rewrite <- IH by intuition. now apply PNSubst1_ok. Qed. Lemma PSubstL_ok n LM1 P1 P2 l : PSubstL P1 LM1 n = Some P2 -> MPcond LM1 l -> P1@l == P2@l. Proof. revert P1. induction LM1 as [|(M2,P2') LM2 IH]; simpl; intros. - discriminate. - assert (H':=PNSubst_ok n P3 M2 P2'). destruct PNSubst. * injection H as [= <-]. rewrite <- PSubstL1_ok; intuition. * now apply IH. Qed. Lemma PNSubstL_ok m n LM1 P1 l : MPcond LM1 l -> P1@l == (PNSubstL P1 LM1 m n)@l. Proof. revert LM1 P1. induction m; simpl; intros; assert (H' := PSubstL_ok n LM1 P2); destruct PSubstL; auto; try reflexivity. rewrite <- IHm; auto. Qed. (** evaluation of polynomial expressions towards R *) Definition mk_X j := mkPinj_pred j mkX. (** evaluation of polynomial expressions towards R *) Fixpoint PEeval (l:Env R) (pe:PExpr) : R := match pe with | PEc c => phi c | PEX j => nth j l | PEadd pe1 pe2 => (PEeval l pe1) + (PEeval l pe2) | PEsub pe1 pe2 => (PEeval l pe1) - (PEeval l pe2) | PEmul pe1 pe2 => (PEeval l pe1) * (PEeval l pe2) | PEopp pe1 => - (PEeval l pe1) | PEpow pe1 n => rpow (PEeval l pe1) (Cp_phi n) end. (** Correctness proofs *) Lemma mkX_ok p l : nth p l == (mk_X p) @ l. Proof. destruct p;simpl;intros;Esimpl;trivial. rewrite nth_spec ; auto. unfold hd. now rewrite <- nth_pred_double, nth_jump. Qed. Hint Rewrite Padd_ok Psub_ok : Esimpl. Section POWER. Variable subst_l : Pol -> Pol. Fixpoint Ppow_pos (res P:Pol) (p:positive) : Pol := match p with | xH => subst_l (res ** P) | xO p => Ppow_pos (Ppow_pos res P p) P p | xI p => subst_l ((Ppow_pos (Ppow_pos res P p) P p) ** P) end. Definition Ppow_N P n := match n with | N0 => P1 | Npos p => Ppow_pos P1 P p end. Lemma Ppow_pos_ok l : (forall P, subst_l P@l == P@l) -> forall res P p, (Ppow_pos res P p)@l == res@l * (pow_pos Pmul P p)@l. Proof. intros subst_l_ok res P p. revert res. induction p;simpl;intros; rewrite ?subst_l_ok, ?Pmul_ok, ?IHp; mul_permut. Qed. Lemma Ppow_N_ok l : (forall P, subst_l P@l == P@l) -> forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l. Proof. destruct n;simpl. - reflexivity. - rewrite Ppow_pos_ok by trivial. Esimpl. Qed. End POWER. (** Normalization and rewriting *) Section NORM_SUBST_REC. Variable n : nat. Variable lmp:list (Mon*Pol). Let subst_l P := PNSubstL P lmp n n. Let Pmul_subst P1 P2 := subst_l (Pmul P1 P2). Let Ppow_subst := Ppow_N subst_l. Fixpoint norm_aux (pe:PExpr) : Pol := match pe with | PEc c => Pc c | PEX j => mk_X j | PEadd (PEopp pe1) pe2 => Psub (norm_aux pe2) (norm_aux pe1) | PEadd pe1 (PEopp pe2) => Psub (norm_aux pe1) (norm_aux pe2) | PEadd pe1 pe2 => Padd (norm_aux pe1) (norm_aux pe2) | PEsub pe1 pe2 => Psub (norm_aux pe1) (norm_aux pe2) | PEmul pe1 pe2 => Pmul (norm_aux pe1) (norm_aux pe2) | PEopp pe1 => Popp (norm_aux pe1) | PEpow pe1 n => Ppow_N (fun p => p) (norm_aux pe1) n end. Definition norm_subst pe := subst_l (norm_aux pe). (** Internally, [norm_aux] is expanded in a large number of cases. To speed-up proofs, we use an alternative definition. *) Definition get_PEopp pe := match pe with | PEopp pe' => Some pe' | _ => None end. Lemma norm_aux_PEadd pe1 pe2 : norm_aux (PEadd pe1 pe2) = match get_PEopp pe1, get_PEopp pe2 with | Some pe1', _ => (norm_aux pe2) -- (norm_aux pe1') | None, Some pe2' => (norm_aux pe1) -- (norm_aux pe2') | None, None => (norm_aux pe1) ++ (norm_aux pe2) end. Proof. simpl (norm_aux (PEadd _ _)). destruct pe1; [ | | | | | reflexivity | ]; destruct pe2; simpl get_PEopp; reflexivity. Qed. Lemma norm_aux_PEopp pe : match get_PEopp pe with | Some pe' => norm_aux pe = -- (norm_aux pe') | None => True end. Proof. now destruct pe. Qed. Lemma norm_aux_spec l pe : PEeval l pe == (norm_aux pe)@l. Proof. intros. induction pe. - reflexivity. - apply mkX_ok. - simpl PEeval. rewrite IHpe1, IHpe2. assert (H1 := norm_aux_PEopp pe1). assert (H2 := norm_aux_PEopp pe2). rewrite norm_aux_PEadd. do 2 destruct get_PEopp; rewrite ?H1, ?H2; Esimpl; add_permut. - simpl. rewrite IHpe1, IHpe2. Esimpl. - simpl. rewrite IHpe1, IHpe2. now rewrite Pmul_ok. - simpl. rewrite IHpe. Esimpl. - simpl. rewrite Ppow_N_ok by reflexivity. rewrite (rpow_pow_N pow_th). destruct n0; simpl; Esimpl. induction p;simpl; now rewrite ?IHp, ?IHpe, ?Pms_ok, ?Pmul_ok. Qed. End NORM_SUBST_REC. End MakeRingPol. coq-8.11.0/plugins/micromega/sos_lib.mli0000644000175000017500000000600013612664533020043 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'b) -> ('c -> 'a) -> 'c -> 'b val num_1 : Num.num val pow10 : int -> Num.num val pow2 : int -> Num.num val implode : string list -> string val explode : string -> string list val funpow : int -> ('a -> 'a) -> 'a -> 'a val tryfind : ('a -> 'b) -> 'a list -> 'b type ('a, 'b) func = | Empty | Leaf of int * ('a * 'b) list | Branch of int * int * ('a, 'b) func * ('a, 'b) func val undefined : ('a, 'b) func val is_undefined : ('a, 'b) func -> bool val ( |-> ) : 'a -> 'b -> ('a, 'b) func -> ('a, 'b) func val ( |=> ) : 'a -> 'b -> ('a, 'b) func val choose : ('a, 'b) func -> 'a * 'b val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> ('b, 'a) func -> ('b, 'a) func -> ('b, 'a) func val ( -- ) : int -> int -> int list val tryapplyd : ('a, 'b) func -> 'a -> 'b -> 'b val apply : ('a, 'b) func -> 'a -> 'b val foldl : ('a -> 'b -> 'c -> 'a) -> 'a -> ('b, 'c) func -> 'a val foldr : ('a -> 'b -> 'c -> 'c) -> ('a, 'b) func -> 'c -> 'c val mapf : ('a -> 'b) -> ('c, 'a) func -> ('c, 'b) func val undefine : 'a -> ('a, 'b) func -> ('a, 'b) func val dom : ('a, 'b) func -> 'a list val graph : ('a, 'b) func -> ('a * 'b) list val union : 'a list -> 'a list -> 'a list val subtract : 'a list -> 'a list -> 'a list val sort : ('a -> 'a -> bool) -> 'a list -> 'a list val setify : 'a list -> 'a list val increasing : ('a -> 'b) -> 'a -> 'a -> bool val allpairs : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list val gcd_num : Num.num -> Num.num -> Num.num val lcm_num : Num.num -> Num.num -> Num.num val numerator : Num.num -> Num.num val denominator : Num.num -> Num.num val end_itlist : ('a -> 'a -> 'a) -> 'a list -> 'a val ( >> ) : ('a -> 'b * 'c) -> ('b -> 'd) -> 'a -> 'd * 'c val ( ++ ) : ('a -> 'b * 'c) -> ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e val a : 'a -> 'a list -> 'a * 'a list val many : ('a -> 'b * 'a) -> 'a -> 'b list * 'a val some : ('a -> bool) -> 'a list -> 'a * 'a list val possibly : ('a -> 'b * 'a) -> 'a -> 'b list * 'a val isspace : string -> bool val parser_or : ('a -> 'b) -> ('a -> 'b) -> 'a -> 'b val isnum : string -> bool val atleast : int -> ('a -> 'b * 'a) -> 'a -> 'b list * 'a val listof : ('a -> 'b * 'c) -> ('c -> 'd * 'a) -> string -> 'a -> 'b list * 'c val temp_path : string val string_of_file : string -> string val file_of_string : string -> string -> unit val deepen_until : int -> (int -> 'a) -> int -> 'a exception TooDeep coq-8.11.0/plugins/micromega/itv.mli0000644000175000017500000000172713612664533017226 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* interval -> unit val inter : interval -> interval -> interval option val range : interval -> num option val smaller_itv : interval -> interval -> bool val in_bound : interval -> num -> bool val norm_itv : interval -> interval option coq-8.11.0/plugins/micromega/zify_plugin.mlpack0000644000175000017500000000001413612664533021435 0ustar treinentreinenZify G_zify coq-8.11.0/plugins/micromega/ZCoeff.v0000644000175000017500000001267113612664533017264 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* R -> R. Variable ropp : R -> R. Variables req rle rlt : R -> R -> Prop. Variable sor : SOR rO rI rplus rtimes rminus ropp req rle rlt. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (rplus x y). Notation "x * y " := (rtimes x y). Notation "x - y " := (rminus x y). Notation "- x" := (ropp x). Notation "x == y" := (req x y). Notation "x ~= y" := (~ req x y). Notation "x <= y" := (rle x y). Notation "x < y" := (rlt x y). Lemma req_refl : forall x, req x x. Proof. destruct (SORsetoid sor) as (Equivalence_Reflexive,_,_). apply Equivalence_Reflexive. Qed. Lemma req_sym : forall x y, req x y -> req y x. Proof. destruct (SORsetoid sor) as (_,Equivalence_Symmetric,_). apply Equivalence_Symmetric. Qed. Lemma req_trans : forall x y z, req x y -> req y z -> req x z. Proof. destruct (SORsetoid sor) as (_,_,Equivalence_Transitive). apply Equivalence_Transitive. Qed. Add Relation R req reflexivity proved by (@Equivalence_Reflexive _ _ (SORsetoid sor)) symmetry proved by (@Equivalence_Symmetric _ _ (SORsetoid sor)) transitivity proved by (@Equivalence_Transitive _ _ (SORsetoid sor)) as sor_setoid. Add Morphism rplus with signature req ==> req ==> req as rplus_morph. Proof. exact (SORplus_wd sor). Qed. Add Morphism rtimes with signature req ==> req ==> req as rtimes_morph. Proof. exact (SORtimes_wd sor). Qed. Add Morphism ropp with signature req ==> req as ropp_morph. Proof. exact (SORopp_wd sor). Qed. Add Morphism rle with signature req ==> req ==> iff as rle_morph. Proof. exact (SORle_wd sor). Qed. Add Morphism rlt with signature req ==> req ==> iff as rlt_morph. Proof. exact (SORlt_wd sor). Qed. Add Morphism rminus with signature req ==> req ==> req as rminus_morph. Proof. exact (rminus_morph sor). Qed. Ltac le_less := rewrite (Rle_lt_eq sor); left; try assumption. Ltac le_equal := rewrite (Rle_lt_eq sor); right; try reflexivity; try assumption. Definition gen_order_phi_Z : Z -> R := gen_phiZ 0 1 rplus rtimes ropp. Declare Equivalent Keys gen_order_phi_Z gen_phiZ. Notation phi_pos := (gen_phiPOS 1 rplus rtimes). Notation phi_pos1 := (gen_phiPOS1 1 rplus rtimes). Notation "[ x ]" := (gen_order_phi_Z x). Lemma ring_ops_wd : ring_eq_ext rplus rtimes ropp req. Proof. constructor. exact rplus_morph. exact rtimes_morph. exact ropp_morph. Qed. Lemma Zring_morph : ring_morph 0 1 rplus rtimes rminus ropp req 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool gen_order_phi_Z. Proof. exact (gen_phiZ_morph (SORsetoid sor) ring_ops_wd (SORrt sor)). Qed. Lemma phi_pos1_pos : forall x : positive, 0 < phi_pos1 x. Proof. induction x as [x IH | x IH |]; simpl; try apply (Rplus_pos_pos sor); try apply (Rtimes_pos_pos sor); try apply (Rplus_pos_pos sor); try apply (Rlt_0_1 sor); assumption. Qed. Lemma phi_pos1_succ : forall x : positive, phi_pos1 (Pos.succ x) == 1 + phi_pos1 x. Proof. exact (ARgen_phiPOS_Psucc (SORsetoid sor) ring_ops_wd (Rth_ARth (SORsetoid sor) ring_ops_wd (SORrt sor))). Qed. Lemma clt_pos_morph : forall x y : positive, (x < y)%positive -> phi_pos1 x < phi_pos1 y. Proof. intros x y H. pattern y; apply Pos.lt_ind with x. rewrite phi_pos1_succ; apply (Rlt_succ_r sor). clear y H; intros y _ H. rewrite phi_pos1_succ. now apply (Rlt_lt_succ sor). assumption. Qed. Lemma clt_morph : forall x y : Z, (x < y)%Z -> [x] < [y]. Proof. intros x y H. do 2 rewrite (same_genZ (SORsetoid sor) ring_ops_wd (SORrt sor)); destruct x; destruct y; simpl in *; try discriminate. apply phi_pos1_pos. now apply clt_pos_morph. apply <- (Ropp_neg_pos sor); apply phi_pos1_pos. apply (Rlt_trans sor) with 0. apply <- (Ropp_neg_pos sor); apply phi_pos1_pos. apply phi_pos1_pos. apply -> (Ropp_lt_mono sor); apply clt_pos_morph. red. now rewrite Pos.compare_antisym. Qed. Lemma Zcleb_morph : forall x y : Z, Z.leb x y = true -> [x] <= [y]. Proof. unfold Z.leb; intros x y H. case_eq (x ?= y)%Z; intro H1; rewrite H1 in H. le_equal. apply (morph_eq Zring_morph). unfold Zeq_bool; now rewrite H1. le_less. now apply clt_morph. discriminate. Qed. Lemma Zcneqb_morph : forall x y : Z, Zeq_bool x y = false -> [x] ~= [y]. Proof. intros x y H. unfold Zeq_bool in H. case_eq (Z.compare x y); intro H1; rewrite H1 in *; (discriminate || clear H). apply (Rlt_neq sor). now apply clt_morph. fold (x > y)%Z in H1. rewrite Z.gt_lt_iff in H1. apply (Rneq_symm sor). apply (Rlt_neq sor). now apply clt_morph. Qed. End InitialMorphism. coq-8.11.0/plugins/micromega/itv.ml0000644000175000017500000000472213612664533017053 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* \]-oo,+oo\[ - None , Some v -> \]-oo,v\] - Some v, None -> \[v,+oo\[ - Some v, Some v' -> \[v,v'\] Intervals needs to be explicitly normalised. *) let pp o (n1, n2) = ( match n1 with | None -> output_string o "]-oo" | Some n -> Printf.fprintf o "[%s" (string_of_num n) ); output_string o ","; match n2 with | None -> output_string o "+oo[" | Some n -> Printf.fprintf o "%s]" (string_of_num n) (** if then interval [itv] is empty, [norm_itv itv] returns [None] otherwise, it returns [Some itv] *) let norm_itv itv = match itv with | Some a, Some b -> if a <=/ b then Some itv else None | _ -> Some itv (** [inter i1 i2 = None] if the intersection of intervals is empty [inter i1 i2 = Some i] if [i] is the intersection of the intervals [i1] and [i2] *) let inter i1 i2 = let l1, r1 = i1 and l2, r2 = i2 in let inter f o1 o2 = match (o1, o2) with | None, None -> None | Some _, None -> o1 | None, Some _ -> o2 | Some n1, Some n2 -> Some (f n1 n2) in norm_itv (inter max_num l1 l2, inter min_num r1 r2) let range = function | None, _ | _, None -> None | Some i, Some j -> Some (floor_num j -/ ceiling_num i +/ Int 1) let smaller_itv i1 i2 = match (range i1, range i2) with | None, _ -> false | _, None -> true | Some i, Some j -> i <=/ j (** [in_bound bnd v] checks whether [v] is within the bounds [bnd] *) let in_bound bnd v = let l, r = bnd in match (l, r) with | None, None -> true | None, Some a -> v <=/ a | Some a, None -> a <=/ v | Some a, Some b -> a <=/ v && v <=/ b coq-8.11.0/plugins/micromega/ZifyPow.v0000644000175000017500000000034413612664533017511 0ustar treinentreinenRequire Import Arith Max Min BinInt BinNat Znat Nnat. Require Import ZifyClasses. Require Export ZifyInst. Instance Op_Z_pow_pos : BinOp Z.pow_pos := { TBOp := Z.pow ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_pow_pos. coq-8.11.0/plugins/micromega/VarMap.v0000644000175000017500000000524013612664533017270 0ustar treinentreinen(* -*- coding: utf-8 -*- *) (************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t | Branch : t -> A -> t -> t . Arguments t : clear implicits. Section MakeVarMap. Variable A : Type. Variable default : A. Notation t := (t A). Fixpoint find (vm : t) (p:positive) {struct vm} : A := match vm with | Empty => default | Elt i => i | Branch l e r => match p with | xH => e | xO p => find l p | xI p => find r p end end. Fixpoint singleton (x:positive) (v : A) : t := match x with | xH => Elt v | xO p => Branch (singleton p v) default Empty | xI p => Branch Empty default (singleton p v) end. Fixpoint vm_add (x: positive) (v : A) (m : t) {struct m} : t := match m with | Empty => singleton x v | Elt vl => match x with | xH => Elt v | xO p => Branch (singleton p v) vl Empty | xI p => Branch Empty vl (singleton p v) end | Branch l o r => match x with | xH => Branch l v r | xI p => Branch l o (vm_add p v r) | xO p => Branch (vm_add p v l) o r end end. End MakeVarMap. coq-8.11.0/plugins/micromega/sos.mli0000644000175000017500000000243213612664533017222 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool val poly_neg : poly -> poly val poly_mul : poly -> poly -> poly val poly_pow : poly -> int -> poly val poly_const : Num.num -> poly val poly_of_term : term -> poly val term_of_poly : poly -> term val term_of_sos : positivstellensatz * (Num.num * poly) list -> positivstellensatz val string_of_poly : poly -> string val real_positivnullstellensatz_general : bool -> int -> poly list -> (poly * positivstellensatz) list -> poly -> poly list * (positivstellensatz * (Num.num * poly) list) list val sumofsquares : poly -> Num.num * (Num.num * poly) list coq-8.11.0/plugins/micromega/zify.ml0000644000175000017500000007630613612664533017241 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Z) q ? *) let get_type_of env evd e = Tacred.cbv_beta env evd (Retyping.get_type_of env evd e) (** [unsafe_to_constr c] returns a [Constr.t] without considering an evar_map. This is useful for calling Constr.hash *) let unsafe_to_constr = EConstr.Unsafe.to_constr let pr_constr env evd e = Printer.pr_econstr_env env evd e let rec find_option pred l = match l with | [] -> raise Not_found | e :: l -> ( match pred e with Some r -> r | None -> find_option pred l ) (** [HConstr] is a map indexed by EConstr.t. It should only be used using closed terms. *) module HConstr = struct module M = Map.Make (struct type t = EConstr.t let compare c c' = Constr.compare (unsafe_to_constr c) (unsafe_to_constr c') end) type 'a t = 'a list M.t let lfind h m = try M.find h m with Not_found -> [] let add h e m = let l = lfind h m in M.add h (e :: l) m let empty = M.empty let find h m = match lfind h m with e :: _ -> e | [] -> raise Not_found let find_all = lfind let fold f m acc = M.fold (fun k l acc -> List.fold_left (fun acc e -> f k e acc) acc l) m acc end (** [get_projections_from_constant (evd,c) ] returns an array of constr [| a1,.. an|] such that [c] is defined as Definition c := mk a1 .. an with mk a constructor. ai is therefore either a type parameter or a projection. *) let get_projections_from_constant (evd, i) = match EConstr.kind evd (Reductionops.clos_whd_flags CClosure.all (Global.env ()) evd i) with | App (c, a) -> Some a | _ -> raise (CErrors.user_err Pp.( str "The hnf of term " ++ pr_constr (Global.env ()) evd i ++ str " should be an application i.e. (c a1 ... an)")) (** An instance of type, say T, is registered into a hashtable, say TableT. *) type 'a decl = { decl : EConstr.t ; (* Registered type instance *) deriv : 'a (* Projections of insterest *) } module EInjT = struct type t = { isid : bool ; (* S = T -> inj = fun x -> x*) source : EConstr.t ; (* S *) target : EConstr.t ; (* T *) (* projections *) inj : EConstr.t ; (* S -> T *) pred : EConstr.t ; (* T -> Prop *) cstr : EConstr.t option (* forall x, pred (inj x) *) } end module EBinOpT = struct type t = { (* Op : source1 -> source2 -> source3 *) source1 : EConstr.t ; source2 : EConstr.t ; source3 : EConstr.t ; target : EConstr.t ; inj1 : EConstr.t ; (* InjTyp source1 target *) inj2 : EConstr.t ; (* InjTyp source2 target *) inj3 : EConstr.t ; (* InjTyp source3 target *) tbop : EConstr.t (* TBOpInj *) } end module ECstOpT = struct type t = {source : EConstr.t; target : EConstr.t; inj : EConstr.t} end module EUnOpT = struct type t = { source1 : EConstr.t ; source2 : EConstr.t ; target : EConstr.t ; inj1_t : EConstr.t ; inj2_t : EConstr.t ; unop : EConstr.t } end module EBinRelT = struct type t = {source : EConstr.t; target : EConstr.t; inj : EConstr.t; brel : EConstr.t} end module EPropBinOpT = struct type t = EConstr.t end module EPropUnOpT = struct type t = EConstr.t end module ESatT = struct type t = {parg1 : EConstr.t; parg2 : EConstr.t; satOK : EConstr.t} end (* Different type of declarations *) type decl_kind = | PropOp of EPropBinOpT.t decl | PropUnOp of EPropUnOpT.t decl | InjTyp of EInjT.t decl | BinRel of EBinRelT.t decl | BinOp of EBinOpT.t decl | UnOp of EUnOpT.t decl | CstOp of ECstOpT.t decl | Saturate of ESatT.t decl let get_decl = function | PropOp d -> d.decl | PropUnOp d -> d.decl | InjTyp d -> d.decl | BinRel d -> d.decl | BinOp d -> d.decl | UnOp d -> d.decl | CstOp d -> d.decl | Saturate d -> d.decl type term_kind = Application of EConstr.constr | OtherTerm of EConstr.constr module type Elt = sig type elt val name : string (** name *) val table : (term_kind * decl_kind) HConstr.t ref val cast : elt decl -> decl_kind val dest : decl_kind -> elt decl option val get_key : int (** [get_key] is the type-index used as key for the instance *) val mk_elt : Evd.evar_map -> EConstr.t -> EConstr.t array -> elt (** [mk_elt evd i [a0,..,an] returns the element of the table built from the type-instance i and the arguments (type indexes and projections) of the type-class constructor. *) (* val arity : int*) end let table = Summary.ref ~name:"zify_table" HConstr.empty let saturate = Summary.ref ~name:"zify_saturate" HConstr.empty let table_cache = ref HConstr.empty let saturate_cache = ref HConstr.empty (** Each type-class gives rise to a different table. They only differ on how projections are extracted. *) module EInj = struct open EInjT type elt = EInjT.t let name = "EInj" let table = table let cast x = InjTyp x let dest = function InjTyp x -> Some x | _ -> None let mk_elt evd i (a : EConstr.t array) = let isid = EConstr.eq_constr evd a.(0) a.(1) in { isid ; source = a.(0) ; target = a.(1) ; inj = a.(2) ; pred = a.(3) ; cstr = (if isid then None else Some a.(4)) } let get_key = 0 end module EBinOp = struct type elt = EBinOpT.t open EBinOpT let name = "BinOp" let table = table let mk_elt evd i a = { source1 = a.(0) ; source2 = a.(1) ; source3 = a.(2) ; target = a.(3) ; inj1 = a.(5) ; inj2 = a.(6) ; inj3 = a.(7) ; tbop = a.(9) } let get_key = 4 let cast x = BinOp x let dest = function BinOp x -> Some x | _ -> None end module ECstOp = struct type elt = ECstOpT.t open ECstOpT let name = "CstOp" let table = table let cast x = CstOp x let dest = function CstOp x -> Some x | _ -> None let mk_elt evd i a = {source = a.(0); target = a.(1); inj = a.(3)} let get_key = 2 end module EUnOp = struct type elt = EUnOpT.t open EUnOpT let name = "UnOp" let table = table let cast x = UnOp x let dest = function UnOp x -> Some x | _ -> None let mk_elt evd i a = { source1 = a.(0) ; source2 = a.(1) ; target = a.(2) ; inj1_t = a.(4) ; inj2_t = a.(5) ; unop = a.(6) } let get_key = 3 end module EBinRel = struct type elt = EBinRelT.t open EBinRelT let name = "BinRel" let table = table let cast x = BinRel x let dest = function BinRel x -> Some x | _ -> None let mk_elt evd i a = {source = a.(0); target = a.(1); inj = a.(3); brel = a.(4)} let get_key = 2 end module EPropOp = struct type elt = EConstr.t let name = "PropBinOp" let table = table let cast x = PropOp x let dest = function PropOp x -> Some x | _ -> None let mk_elt evd i a = i let get_key = 0 end module EPropUnOp = struct type elt = EConstr.t let name = "PropUnOp" let table = table let cast x = PropUnOp x let dest = function PropUnOp x -> Some x | _ -> None let mk_elt evd i a = i let get_key = 0 end let constr_of_term_kind = function Application c -> c | OtherTerm c -> c let fold_declared_const f evd acc = HConstr.fold (fun _ (_, e) acc -> f (fst (EConstr.destConst evd (get_decl e))) acc) !table_cache acc module type S = sig val register : Constrexpr.constr_expr -> unit val print : unit -> unit end module MakeTable (E : Elt) = struct (** Given a term [c] and its arguments ai, we construct a HConstr.t table that is indexed by ai for i = E.get_key. The elements of the table are built using E.mk_elt c [|a0,..,an|] *) let make_elt (evd, i) = match get_projections_from_constant (evd, i) with | None -> let env = Global.env () in let t = string_of_ppcmds (pr_constr env evd i) in failwith ("Cannot register term " ^ t) | Some a -> E.mk_elt evd i a let register_hint evd t elt = match EConstr.kind evd t with | App (c, _) -> E.table := HConstr.add c (Application t, E.cast elt) !E.table | _ -> E.table := HConstr.add t (OtherTerm t, E.cast elt) !E.table let register_constr env evd c = let c = EConstr.of_constr c in let t = get_type_of env evd c in match EConstr.kind evd t with | App (intyp, args) -> let styp = args.(E.get_key) in let elt = {decl = c; deriv = make_elt (evd, c)} in register_hint evd styp elt | _ -> let env = Global.env () in raise (CErrors.user_err Pp.( str ": Cannot register term " ++ pr_constr env evd c ++ str ". It has type " ++ pr_constr env evd t ++ str " which should be of the form [F X1 .. Xn]")) let register_obj : Constr.constr -> Libobject.obj = let cache_constr (_, c) = let env = Global.env () in let evd = Evd.from_env env in register_constr env evd c in let subst_constr (subst, c) = Mod_subst.subst_mps subst c in Libobject.declare_object @@ Libobject.superglobal_object_nodischarge ("register-zify-" ^ E.name) ~cache:cache_constr ~subst:(Some subst_constr) (** [register c] is called from the VERNACULAR ADD [name] constr(t). The term [c] is interpreted and registered as a [superglobal_object_nodischarge]. TODO: pre-compute [get_type_of] - [cache_constr] is using another environment. *) let register c = let env = Global.env () in let evd = Evd.from_env env in let evd, c = Constrintern.interp_open_constr env evd c in let _ = Lib.add_anonymous_leaf (register_obj (EConstr.to_constr evd c)) in () let pp_keys () = let env = Global.env () in let evd = Evd.from_env env in HConstr.fold (fun _ (k, d) acc -> match E.dest d with | None -> acc | Some _ -> Pp.(pr_constr env evd (constr_of_term_kind k) ++ str " " ++ acc)) !E.table (Pp.str "") let print () = Feedback.msg_info (pp_keys ()) end module InjTable = MakeTable (EInj) module ESat = struct type elt = ESatT.t open ESatT let name = "Saturate" let table = saturate let cast x = Saturate x let dest = function Saturate x -> Some x | _ -> None let mk_elt evd i a = {parg1 = a.(2); parg2 = a.(3); satOK = a.(5)} let get_key = 1 end module BinOp = MakeTable (EBinOp) module UnOp = MakeTable (EUnOp) module CstOp = MakeTable (ECstOp) module BinRel = MakeTable (EBinRel) module PropOp = MakeTable (EPropOp) module PropUnOp = MakeTable (EPropUnOp) module Saturate = MakeTable (ESat) let init_cache () = table_cache := !table; saturate_cache := !saturate (** The module [Spec] is used to register the instances of [BinOpSpec], [UnOpSpec]. They are not indexed and stored in a list. *) module Spec = struct let table = Summary.ref ~name:"zify_Spec" [] let register_obj : Constr.constr -> Libobject.obj = let cache_constr (_, c) = table := EConstr.of_constr c :: !table in let subst_constr (subst, c) = Mod_subst.subst_mps subst c in Libobject.declare_object @@ Libobject.superglobal_object_nodischarge "register-zify-Spec" ~cache:cache_constr ~subst:(Some subst_constr) let register c = let env = Global.env () in let evd = Evd.from_env env in let _, c = Constrintern.interp_open_constr env evd c in let _ = Lib.add_anonymous_leaf (register_obj (EConstr.to_constr evd c)) in () let get () = !table let print () = let env = Global.env () in let evd = Evd.from_env env in let constr_of_spec c = let t = get_type_of env evd c in match EConstr.kind evd t with | App (intyp, args) -> pr_constr env evd args.(2) | _ -> Pp.str "" in let l = List.fold_left (fun acc c -> Pp.(constr_of_spec c ++ str " " ++ acc)) (Pp.str "") !table in Feedback.msg_notice l end let unfold_decl evd = let f cst acc = cst :: acc in fold_declared_const f evd [] open EInjT (** Get constr of lemma and projections in ZifyClasses. *) let zify str = EConstr.of_constr (UnivGen.constr_of_monomorphic_global (Coqlib.lib_ref ("ZifyClasses." ^ str))) let locate_const str = let rf = "ZifyClasses." ^ str in match Coqlib.lib_ref rf with | GlobRef.ConstRef c -> c | _ -> CErrors.anomaly Pp.(str rf ++ str " should be a constant") (* The following [constr] are necessary for constructing the proof terms *) let mkapp2 = lazy (zify "mkapp2") let mkapp = lazy (zify "mkapp") let mkapp0 = lazy (zify "mkapp0") let mkdp = lazy (zify "mkinjterm") let eq_refl = lazy (zify "eq_refl") let mkrel = lazy (zify "mkrel") let mkprop_op = lazy (zify "mkprop_op") let mkuprop_op = lazy (zify "mkuprop_op") let mkdpP = lazy (zify "mkinjprop") let iff_refl = lazy (zify "iff_refl") let q = lazy (zify "target_prop") let ieq = lazy (zify "injprop_ok") let iff = lazy (zify "iff") (* A super-set of the previous are needed to unfold the generated proof terms. *) let to_unfold = lazy (List.rev_map locate_const [ "source_prop" ; "target_prop" ; "uop_iff" ; "op_iff" ; "mkuprop_op" ; "TUOp" ; "inj_ok" ; "TRInj" ; "inj" ; "source" ; "injprop_ok" ; "TR" ; "TBOp" ; "TCst" ; "target" ; "mkrel" ; "mkapp2" ; "mkapp" ; "mkapp0" ; "mkprop_op" ]) (** Module [CstrTable] records terms [x] injected into [inj x] together with the corresponding type constraint. The terms are stored by side-effect during the traversal of the goal. It must therefore be cleared before calling the main tactic. *) module CstrTable = struct module HConstr = Hashtbl.Make (struct type t = EConstr.t let hash c = Constr.hash (unsafe_to_constr c) let equal c c' = Constr.equal (unsafe_to_constr c) (unsafe_to_constr c') end) let table : EConstr.t HConstr.t = HConstr.create 10 let register evd t (i : EConstr.t) = HConstr.add table t i let get () = let l = HConstr.fold (fun k i acc -> (k, i) :: acc) table [] in HConstr.clear table; l (** [gen_cstr table] asserts (cstr k) for each element of the table (k,cstr). NB: the constraint is only asserted if it does not already exist in the context. *) let gen_cstr table = Proofview.Goal.enter (fun gl -> let evd = Tacmach.New.project gl in (* Build the table of existing hypotheses *) let has_hyp = let hyps_table = HConstr.create 20 in List.iter (fun (_, (t : EConstr.types)) -> HConstr.add hyps_table t ()) (Tacmach.New.pf_hyps_types gl); fun c -> HConstr.mem hyps_table c in (* Add the constraint (cstr k) if it is not already present *) let gen k cstr = Proofview.Goal.enter (fun gl -> let env = Tacmach.New.pf_env gl in let term = EConstr.mkApp (cstr, [|k|]) in let types = get_type_of env evd term in if has_hyp types then Tacticals.New.tclIDTAC else let n = Tactics.fresh_id_in_env Id.Set.empty (Names.Id.of_string "cstr") env in Tactics.pose_proof (Names.Name n) term) in List.fold_left (fun acc (k, i) -> Tacticals.New.tclTHEN (gen k i) acc) Tacticals.New.tclIDTAC table) end let mkvar red evd inj v = ( if not red then match inj.cstr with None -> () | Some ctr -> CstrTable.register evd v ctr ); let iv = EConstr.mkApp (inj.inj, [|v|]) in let iv = if red then Tacred.compute (Global.env ()) evd iv else iv in EConstr.mkApp ( force mkdp , [| inj.source ; inj.target ; inj.inj ; v ; iv ; EConstr.mkApp (force eq_refl, [|inj.target; iv|]) |] ) type texpr = | Var of EInj.elt * EConstr.t (** Var is a term that cannot be injected further *) | Constant of EInj.elt * EConstr.t (** Constant is a term that is solely built from constructors *) | Injterm of EConstr.t (** Injected is an injected term represented by a term of type [injterm] *) let is_constant = function Constant _ -> true | _ -> false let constr_of_texpr = function | Constant (i, e) | Var (i, e) -> if i.isid then Some e else None | _ -> None let inj_term_of_texpr evd = function | Injterm e -> e | Var (inj, e) -> mkvar false evd inj e | Constant (inj, e) -> mkvar true evd inj e let mkapp2_id evd i (* InjTyp S3 T *) inj (* deriv i *) t (* S1 -> S2 -> S3 *) b (* Binop S1 S2 S3 t ... *) dbop (* deriv b *) e1 e2 = let default () = let e1' = inj_term_of_texpr evd e1 in let e2' = inj_term_of_texpr evd e2 in EBinOpT.( Injterm (EConstr.mkApp ( force mkapp2 , [| dbop.source1 ; dbop.source2 ; dbop.source3 ; dbop.target ; t ; dbop.inj1 ; dbop.inj2 ; dbop.inj3 ; b ; e1' ; e2' |] ))) in if not inj.isid then default () else match (e1, e2) with | Constant (_, e1), Constant (_, e2) |Var (_, e1), Var (_, e2) |Constant (_, e1), Var (_, e2) |Var (_, e1), Constant (_, e2) -> Var (inj, EConstr.mkApp (t, [|e1; e2|])) | _, _ -> default () let mkapp_id evd i inj (unop, u) f e1 = EUnOpT.( if EConstr.eq_constr evd u.unop f then (* Injection does nothing *) match e1 with | Constant (_, e) | Var (_, e) -> Var (inj, EConstr.mkApp (f, [|e|])) | Injterm e1 -> Injterm (EConstr.mkApp ( force mkapp , [| u.source1 ; u.source2 ; u.target ; f ; u.inj1_t ; u.inj2_t ; unop ; e1 |] )) else let e1 = inj_term_of_texpr evd e1 in Injterm (EConstr.mkApp ( force mkapp , [|u.source1; u.source2; u.target; f; u.inj1_t; u.inj2_t; unop; e1|] ))) type typed_constr = {constr : EConstr.t; typ : EConstr.t} let get_injection env evd t = match snd (HConstr.find t !table_cache) with | InjTyp i -> i | _ -> raise Not_found (* [arrow] is the term (fun (x:Prop) (y : Prop) => x -> y) *) let arrow = let name x = Context.make_annot (Name.mk_name (Names.Id.of_string x)) Sorts.Relevant in EConstr.mkLambda ( name "x" , EConstr.mkProp , EConstr.mkLambda ( name "y" , EConstr.mkProp , EConstr.mkProd ( Context.make_annot Names.Anonymous Sorts.Relevant , EConstr.mkRel 2 , EConstr.mkRel 2 ) ) ) let is_prop env sigma term = let sort = Retyping.get_sort_of env sigma term in Sorts.is_prop sort (** [get_application env evd e] expresses [e] as an application (c a) where c is the head symbol and [a] is the array of arguments. The function also transforms (x -> y) as (arrow x y) *) let get_operator env evd e = let is_arrow a p1 p2 = is_prop env evd p1 && is_prop (EConstr.push_rel (Context.Rel.Declaration.LocalAssum (a, p1)) env) evd p2 && (a.Context.binder_name = Names.Anonymous || EConstr.Vars.noccurn evd 1 p2) in match EConstr.kind evd e with | Prod (a, p1, p2) when is_arrow a p1 p2 -> (arrow, [|p1; p2|]) | App (c, a) -> (c, a) | _ -> (e, [||]) let is_convertible env evd k t = Reductionops.check_conv env evd k t (** [match_operator env evd hd arg (t,d)] - hd is head operator of t - If t = OtherTerm _, then t = hd - If t = Application _, then we extract the relevant number of arguments from arg and check for convertibility *) let match_operator env evd hd args (t, d) = let decomp t i = let n = Array.length args in let t' = EConstr.mkApp (hd, Array.sub args 0 (n - i)) in if is_convertible env evd t' t then Some (d, t) else None in match t with | OtherTerm t -> Some (d, t) | Application t -> ( match d with | CstOp _ -> decomp t 0 | UnOp _ -> decomp t 1 | BinOp _ -> decomp t 2 | BinRel _ -> decomp t 2 | PropOp _ -> decomp t 2 | PropUnOp _ -> decomp t 1 | _ -> None ) let rec trans_expr env evd e = (* Get the injection *) let {decl = i; deriv = inj} = get_injection env evd e.typ in let e = e.constr in if EConstr.isConstruct evd e then Constant (inj, e) (* Evaluate later *) else let c, a = get_operator env evd e in try let k, t = find_option (match_operator env evd c a) (HConstr.find_all c !table_cache) in let n = Array.length a in match k with | CstOp {decl = c'} -> Injterm (EConstr.mkApp (force mkapp0, [|inj.source; inj.target; e; i; c'|])) | UnOp {decl = unop; deriv = u} -> let a' = trans_expr env evd {constr = a.(n - 1); typ = u.EUnOpT.source1} in if is_constant a' && EConstr.isConstruct evd t then Constant (inj, e) else mkapp_id evd i inj (unop, u) t a' | BinOp {decl = binop; deriv = b} -> let a0 = trans_expr env evd {constr = a.(n - 2); typ = b.EBinOpT.source1} in let a1 = trans_expr env evd {constr = a.(n - 1); typ = b.EBinOpT.source2} in if is_constant a0 && is_constant a1 && EConstr.isConstruct evd t then Constant (inj, e) else mkapp2_id evd i inj t binop b a0 a1 | d -> Var (inj, e) with Not_found -> Var (inj, e) let trans_expr env evd e = try trans_expr env evd e with Not_found -> raise (CErrors.user_err ( Pp.str "Missing injection for type " ++ Printer.pr_leconstr_env env evd e.typ )) type tprop = | TProp of EConstr.t (** Transformed proposition *) | IProp of EConstr.t (** Identical proposition *) let mk_iprop e = EConstr.mkApp (force mkdpP, [|e; e; EConstr.mkApp (force iff_refl, [|e|])|]) let inj_prop_of_tprop = function TProp p -> p | IProp e -> mk_iprop e let rec trans_prop env evd e = let c, a = get_operator env evd e in try let k, t = find_option (match_operator env evd c a) (HConstr.find_all c !table_cache) in let n = Array.length a in match k with | PropOp {decl = rop} -> ( try let t1 = trans_prop env evd a.(n - 2) in let t2 = trans_prop env evd a.(n - 1) in match (t1, t2) with | IProp _, IProp _ -> IProp e | _, _ -> let t1 = inj_prop_of_tprop t1 in let t2 = inj_prop_of_tprop t2 in TProp (EConstr.mkApp (force mkprop_op, [|t; rop; t1; t2|])) with Not_found -> IProp e ) | BinRel {decl = br; deriv = rop} -> ( try let a1 = trans_expr env evd {constr = a.(n - 2); typ = rop.EBinRelT.source} in let a2 = trans_expr env evd {constr = a.(n - 1); typ = rop.EBinRelT.source} in if EConstr.eq_constr evd t rop.EBinRelT.brel then match (constr_of_texpr a1, constr_of_texpr a2) with | Some e1, Some e2 -> IProp (EConstr.mkApp (t, [|e1; e2|])) | _, _ -> let a1 = inj_term_of_texpr evd a1 in let a2 = inj_term_of_texpr evd a2 in TProp (EConstr.mkApp ( force mkrel , [| rop.EBinRelT.source ; rop.EBinRelT.target ; t ; rop.EBinRelT.inj ; br ; a1 ; a2 |] )) else let a1 = inj_term_of_texpr evd a1 in let a2 = inj_term_of_texpr evd a2 in TProp (EConstr.mkApp ( force mkrel , [| rop.EBinRelT.source ; rop.EBinRelT.target ; t ; rop.EBinRelT.inj ; br ; a1 ; a2 |] )) with Not_found -> IProp e ) | PropUnOp {decl = rop} -> ( try let t1 = trans_prop env evd a.(n - 1) in match t1 with | IProp _ -> IProp e | _ -> let t1 = inj_prop_of_tprop t1 in TProp (EConstr.mkApp (force mkuprop_op, [|t; rop; t1|])) with Not_found -> IProp e ) | _ -> IProp e with Not_found -> IProp e let unfold n env evd c = let cbv l = CClosure.RedFlags.( Tacred.cbv_norm_flags (mkflags (fBETA :: fMATCH :: fFIX :: fCOFIX :: fZETA :: List.rev_map fCONST l))) in let unfold_decl = unfold_decl evd in (* Unfold the let binding *) let c = match n with | None -> c | Some n -> Tacred.unfoldn [(Locus.AllOccurrences, Names.EvalVarRef n)] env evd c in (* Reduce the term *) let c = cbv (List.rev_append (force to_unfold) unfold_decl) env evd c in c let trans_check_prop env evd t = if is_prop env evd t then (*let t = Tacred.unfoldn [Locus.AllOccurrences, Names.EvalConstRef coq_not] env evd t in*) match trans_prop env evd t with IProp e -> None | TProp e -> Some e else None let trans_hyps env evd l = List.fold_left (fun acc (h, p) -> match trans_check_prop env evd p with | None -> acc | Some p' -> (h, p, p') :: acc) [] (List.rev l) (* Only used if a direct rewrite fails *) let trans_hyp h t = Tactics.( Tacticals.New.( Proofview.Goal.enter (fun gl -> let env = Tacmach.New.pf_env gl in let n = fresh_id_in_env Id.Set.empty (Names.Id.of_string "__zify") env in let h' = fresh_id_in_env Id.Set.empty h env in tclTHENLIST [ letin_tac None (Names.Name n) t None Locus.{onhyps = None; concl_occs = NoOccurrences} ; assert_by (Name.Name h') (EConstr.mkApp (force q, [|EConstr.mkVar n|])) (tclTHEN (Equality.rewriteRL (EConstr.mkApp (force ieq, [|EConstr.mkVar n|]))) (exact_check (EConstr.mkVar h))) ; reduct_in_hyp ~check:true ~reorder:false (unfold (Some n)) (h', Locus.InHyp) ; clear [n] ; (* [clear H] may fail if [h] has dependencies *) tclTRY (clear [h]) ]))) let is_progress_rewrite evd t rew = match EConstr.kind evd rew with | App (c, [|lhs; rhs|]) -> if EConstr.eq_constr evd (force iff) c then (* This is a successful rewriting *) not (EConstr.eq_constr evd lhs rhs) else CErrors.anomaly Pp.( str "is_progress_rewrite: not a rewrite" ++ pr_constr (Global.env ()) evd rew) | _ -> failwith "is_progress_rewrite: not even an application" let trans_hyp h t0 t = Tacticals.New.( Proofview.Goal.enter (fun gl -> let env = Tacmach.New.pf_env gl in let evd = Tacmach.New.project gl in let t' = unfold None env evd (EConstr.mkApp (force ieq, [|t|])) in if is_progress_rewrite evd t0 (get_type_of env evd t') then tclFIRST [ Equality.general_rewrite_in true Locus.AllOccurrences true false h t' false ; trans_hyp h t ] else tclIDTAC)) let trans_concl t = Tacticals.New.( Proofview.Goal.enter (fun gl -> let concl = Tacmach.New.pf_concl gl in let env = Tacmach.New.pf_env gl in let evd = Tacmach.New.project gl in let t' = unfold None env evd (EConstr.mkApp (force ieq, [|t|])) in if is_progress_rewrite evd concl (get_type_of env evd t') then Equality.general_rewrite true Locus.AllOccurrences true false t' else tclIDTAC)) let tclTHENOpt e tac tac' = match e with None -> tac' | Some e' -> Tacticals.New.tclTHEN (tac e') tac' let zify_tac = Proofview.Goal.enter (fun gl -> Coqlib.check_required_library ["Coq"; "micromega"; "ZifyClasses"]; Coqlib.check_required_library ["Coq"; "micromega"; "ZifyInst"]; init_cache (); let evd = Tacmach.New.project gl in let env = Tacmach.New.pf_env gl in let concl = trans_check_prop env evd (Tacmach.New.pf_concl gl) in let hyps = trans_hyps env evd (Tacmach.New.pf_hyps_types gl) in let l = CstrTable.get () in tclTHENOpt concl trans_concl (Tacticals.New.tclTHEN (Tacticals.New.tclTHENLIST (List.rev_map (fun (h, p, t) -> trans_hyp h p t) hyps)) (CstrTable.gen_cstr l))) let iter_specs tac = Tacticals.New.tclTHENLIST (List.fold_left (fun acc d -> tac d :: acc) [] (Spec.get ())) let iter_specs (tac : Ltac_plugin.Tacinterp.Value.t) = iter_specs (fun c -> Ltac_plugin.Tacinterp.Value.apply tac [Ltac_plugin.Tacinterp.Value.of_constr c]) let find_hyp evd t l = try Some (fst (List.find (fun (h, t') -> EConstr.eq_constr evd t t') l)) with Not_found -> None let sat_constr c d = Proofview.Goal.enter (fun gl -> let evd = Tacmach.New.project gl in let env = Tacmach.New.pf_env gl in let hyps = Tacmach.New.pf_hyps_types gl in match EConstr.kind evd c with | App (c, args) -> if Array.length args = 2 then let h1 = Tacred.cbv_beta env evd (EConstr.mkApp (d.ESatT.parg1, [|args.(0)|])) in let h2 = Tacred.cbv_beta env evd (EConstr.mkApp (d.ESatT.parg2, [|args.(1)|])) in match (find_hyp evd h1 hyps, find_hyp evd h2 hyps) with | Some h1, Some h2 -> let n = Tactics.fresh_id_in_env Id.Set.empty (Names.Id.of_string "__sat") env in let trm = EConstr.mkApp ( d.ESatT.satOK , [|args.(0); args.(1); EConstr.mkVar h1; EConstr.mkVar h2|] ) in Tactics.pose_proof (Names.Name n) trm | _, _ -> Tacticals.New.tclIDTAC else Tacticals.New.tclIDTAC | _ -> Tacticals.New.tclIDTAC) let get_all_sat env evd c = List.fold_left (fun acc e -> match e with _, Saturate s -> s :: acc | _ -> acc) [] (HConstr.find_all c !saturate_cache) let saturate = Proofview.Goal.enter (fun gl -> init_cache (); let table = CstrTable.HConstr.create 20 in let concl = Tacmach.New.pf_concl gl in let hyps = Tacmach.New.pf_hyps_types gl in let evd = Tacmach.New.project gl in let env = Tacmach.New.pf_env gl in let rec sat t = match EConstr.kind evd t with | App (c, args) -> sat c; Array.iter sat args; if Array.length args = 2 then let ds = get_all_sat env evd c in if ds = [] then () else List.iter (fun x -> CstrTable.HConstr.add table t x.deriv) ds else () | Prod (a, t1, t2) when a.Context.binder_name = Names.Anonymous -> sat t1; sat t2 | _ -> () in (* Collect all the potential saturation lemma *) sat concl; List.iter (fun (_, t) -> sat t) hyps; Tacticals.New.tclTHENLIST (CstrTable.HConstr.fold (fun c d acc -> sat_constr c d :: acc) table [])) coq-8.11.0/plugins/micromega/DeclConstant.v0000644000175000017500000000540713612664533020470 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* T2) (A : T1) : DeclaredConstant F -> GT A -> GT (F A). Defined. Instance GT_APP2 {T1 T2 T3: Type} (F : T1 -> T2 -> T3) {A1 : T1} {A2 : T2} {DC:DeclaredConstant F} : GT A1 -> GT A2 -> GT (F A1 A2). Defined. Require Import QArith_base. Instance DO : DeclaredConstant O := {}. Instance DS : DeclaredConstant S := {}. Instance DxH: DeclaredConstant xH := {}. Instance DxI: DeclaredConstant xI := {}. Instance DxO: DeclaredConstant xO := {}. Instance DZO: DeclaredConstant Z0 := {}. Instance DZpos: DeclaredConstant Zpos := {}. Instance DZneg: DeclaredConstant Zneg := {}. Instance DZpow_pos : DeclaredConstant Z.pow_pos := {}. Instance DZpow : DeclaredConstant Z.pow := {}. Instance DQ : DeclaredConstant Qmake := {}. coq-8.11.0/plugins/micromega/Lqa.v0000644000175000017500000000420713612664533016621 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ((sos_Q rchecker) || (psatz_Q d rchecker)) | _ => fail "Unsupported domain" end in tac. Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n. Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1). (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/persistent_cache.mli0000644000175000017500000000326113612664533021742 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a t (** [open_in f] rebuilds a table from the records stored in file [f]. As marshaling is not type-safe, it might segfault. *) val find : 'a t -> key -> 'a (** find has the specification of Hashtable.find *) val add : 'a t -> key -> 'a -> unit (** [add tbl key elem] adds the binding [key] [elem] to the table [tbl]. (and writes the binding to the file associated with [tbl].) If [key] is already bound, raises KeyAlreadyBound *) val memo : string -> (key -> 'a) -> key -> 'a (** [memo cache f] returns a memo function for [f] using file [cache] as persistent table. Note that the cache will only be loaded when the function is used for the first time *) val memo_cond : string -> (key -> bool) -> (key -> 'a) -> key -> 'a (** [memo cache cond f] only use the cache if [cond k] holds for the key [k]. *) end module PHashtable (Key : HashedType) : PHashtable with type key = Key.t coq-8.11.0/plugins/micromega/micromega_plugin.mlpack0000644000175000017500000000016713612664533022430 0ustar treinentreinenMicromega Mutils Itv Vect Sos_types Polynomial Mfourier Simplex Certificate Persistent_cache Coq_micromega G_micromega coq-8.11.0/plugins/micromega/OrderedRing.v0000644000175000017500000003367613612664533020324 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* R -> R) (ropp : R -> R). Variable req rle rlt : R -> R -> Prop. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (rplus x y). Notation "x * y " := (rtimes x y). Notation "x - y " := (rminus x y). Notation "- x" := (ropp x). Notation "x == y" := (req x y). Notation "x ~= y" := (~ req x y). Notation "x <= y" := (rle x y). Notation "x < y" := (rlt x y). Record SOR : Type := mk_SOR_theory { SORsetoid : Setoid_Theory R req; SORplus_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 + y1 == x2 + y2; SORtimes_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2; SORopp_wd : forall x1 x2, x1 == x2 -> -x1 == -x2; SORle_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> (x1 <= y1 <-> x2 <= y2); SORlt_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> (x1 < y1 <-> x2 < y2); SORrt : ring_theory rO rI rplus rtimes rminus ropp req; SORle_refl : forall n : R, n <= n; SORle_antisymm : forall n m : R, n <= m -> m <= n -> n == m; SORle_trans : forall n m p : R, n <= m -> m <= p -> n <= p; SORlt_le_neq : forall n m : R, n < m <-> n <= m /\ n ~= m; SORlt_trichotomy : forall n m : R, n < m \/ n == m \/ m < n; SORplus_le_mono_l : forall n m p : R, n <= m -> p + n <= p + m; SORtimes_pos_pos : forall n m : R, 0 < n -> 0 < m -> 0 < n * m; SORneq_0_1 : 0 ~= 1 }. (* We cannot use Relation_Definitions.order.ord_antisym and Relations_1.Antisymmetric because they refer to Leibniz equality *) End DEFINITIONS. Section STRICT_ORDERED_RING. Variable R : Type. Variable (rO rI : R) (rplus rtimes rminus: R -> R -> R) (ropp : R -> R). Variable req rle rlt : R -> R -> Prop. Variable sor : SOR rO rI rplus rtimes rminus ropp req rle rlt. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (rplus x y). Notation "x * y " := (rtimes x y). Notation "x - y " := (rminus x y). Notation "- x" := (ropp x). Notation "x == y" := (req x y). Notation "x ~= y" := (~ req x y). Notation "x <= y" := (rle x y). Notation "x < y" := (rlt x y). Add Relation R req reflexivity proved by (@Equivalence_Reflexive _ _ (SORsetoid sor)) symmetry proved by (@Equivalence_Symmetric _ _ (SORsetoid sor)) transitivity proved by (@Equivalence_Transitive _ _ (SORsetoid sor)) as sor_setoid. Add Morphism rplus with signature req ==> req ==> req as rplus_morph. Proof. exact (SORplus_wd sor). Qed. Add Morphism rtimes with signature req ==> req ==> req as rtimes_morph. Proof. exact (SORtimes_wd sor). Qed. Add Morphism ropp with signature req ==> req as ropp_morph. Proof. exact (SORopp_wd sor). Qed. Add Morphism rle with signature req ==> req ==> iff as rle_morph. Proof. exact (SORle_wd sor). Qed. Add Morphism rlt with signature req ==> req ==> iff as rlt_morph. Proof. exact (SORlt_wd sor). Qed. Add Ring SOR : (SORrt sor). Add Morphism rminus with signature req ==> req ==> req as rminus_morph. Proof. intros x1 x2 H1 y1 y2 H2. rewrite ((Rsub_def (SORrt sor)) x1 y1). rewrite ((Rsub_def (SORrt sor)) x2 y2). rewrite H1; now rewrite H2. Qed. Theorem Rneq_symm : forall n m : R, n ~= m -> m ~= n. Proof. intros n m H1 H2; rewrite H2 in H1; now apply H1. Qed. (* Properties of plus, minus and opp *) Theorem Rplus_0_l : forall n : R, 0 + n == n. Proof. intro; ring. Qed. Theorem Rplus_0_r : forall n : R, n + 0 == n. Proof. intro; ring. Qed. Theorem Rtimes_0_r : forall n : R, n * 0 == 0. Proof. intro; ring. Qed. Theorem Rplus_comm : forall n m : R, n + m == m + n. Proof. intros; ring. Qed. Theorem Rtimes_0_l : forall n : R, 0 * n == 0. Proof. intro; ring. Qed. Theorem Rtimes_comm : forall n m : R, n * m == m * n. Proof. intros; ring. Qed. Theorem Rminus_eq_0 : forall n m : R, n - m == 0 <-> n == m. Proof. intros n m. split; intro H. setoid_replace n with ((n - m) + m) by ring. rewrite H. now rewrite Rplus_0_l. rewrite H; ring. Qed. Theorem Rplus_cancel_l : forall n m p : R, p + n == p + m <-> n == m. Proof. intros n m p; split; intro H. setoid_replace n with (- p + (p + n)) by ring. setoid_replace m with (- p + (p + m)) by ring. now rewrite H. now rewrite H. Qed. (* Relations *) Theorem Rle_refl : forall n : R, n <= n. Proof (SORle_refl sor). Theorem Rle_antisymm : forall n m : R, n <= m -> m <= n -> n == m. Proof (SORle_antisymm sor). Theorem Rle_trans : forall n m p : R, n <= m -> m <= p -> n <= p. Proof (SORle_trans sor). Theorem Rlt_trichotomy : forall n m : R, n < m \/ n == m \/ m < n. Proof (SORlt_trichotomy sor). Theorem Rlt_le_neq : forall n m : R, n < m <-> n <= m /\ n ~= m. Proof (SORlt_le_neq sor). Theorem Rneq_0_1 : 0 ~= 1. Proof (SORneq_0_1 sor). Theorem Req_em : forall n m : R, n == m \/ n ~= m. Proof. intros n m. destruct (Rlt_trichotomy n m) as [H | [H | H]]; try rewrite Rlt_le_neq in H. right; now destruct H. now left. right; apply Rneq_symm; now destruct H. Qed. Theorem Req_dne : forall n m : R, ~ ~ n == m <-> n == m. Proof. intros n m; destruct (Req_em n m) as [H | H]. split; auto. split. intro H1; false_hyp H H1. auto. Qed. Theorem Rle_lt_eq : forall n m : R, n <= m <-> n < m \/ n == m. Proof. intros n m; rewrite Rlt_le_neq. split; [intro H | intros [[H1 H2] | H]]. destruct (Req_em n m) as [H1 | H1]. now right. left; now split. assumption. rewrite H; apply Rle_refl. Qed. Ltac le_less := rewrite Rle_lt_eq; left; try assumption. Ltac le_equal := rewrite Rle_lt_eq; right; try reflexivity; try assumption. Ltac le_elim H := rewrite Rle_lt_eq in H; destruct H as [H | H]. Theorem Rlt_trans : forall n m p : R, n < m -> m < p -> n < p. Proof. intros n m p; repeat rewrite Rlt_le_neq; intros [H1 H2] [H3 H4]; split. now apply Rle_trans with m. intro H. rewrite H in H1. pose proof (Rle_antisymm H3 H1). now apply H4. Qed. Theorem Rle_lt_trans : forall n m p : R, n <= m -> m < p -> n < p. Proof. intros n m p H1 H2; le_elim H1. now apply Rlt_trans with (m := m). now rewrite H1. Qed. Theorem Rlt_le_trans : forall n m p : R, n < m -> m <= p -> n < p. Proof. intros n m p H1 H2; le_elim H2. now apply Rlt_trans with (m := m). now rewrite <- H2. Qed. Theorem Rle_gt_cases : forall n m : R, n <= m \/ m < n. Proof. intros n m; destruct (Rlt_trichotomy n m) as [H | [H | H]]. left; now le_less. left; now le_equal. now right. Qed. Theorem Rlt_neq : forall n m : R, n < m -> n ~= m. Proof. intros n m; rewrite Rlt_le_neq; now intros [_ H]. Qed. Theorem Rle_ngt : forall n m : R, n <= m <-> ~ m < n. Proof. intros n m; split. intros H H1; assert (H2 : n < n) by now apply Rle_lt_trans with m. now apply (Rlt_neq H2). intro H. destruct (Rle_gt_cases n m) as [H1 | H1]. assumption. false_hyp H1 H. Qed. Theorem Rlt_nge : forall n m : R, n < m <-> ~ m <= n. Proof. intros n m; split. intros H H1; assert (H2 : n < n) by now apply Rlt_le_trans with m. now apply (Rlt_neq H2). intro H. destruct (Rle_gt_cases m n) as [H1 | H1]. false_hyp H1 H. assumption. Qed. (* Plus, minus and order *) Theorem Rplus_le_mono_l : forall n m p : R, n <= m <-> p + n <= p + m. Proof. intros n m p; split. apply (SORplus_le_mono_l sor). intro H. apply ((SORplus_le_mono_l sor) (p + n) (p + m) (- p)) in H. setoid_replace (- p + (p + n)) with n in H by ring. setoid_replace (- p + (p + m)) with m in H by ring. assumption. Qed. Theorem Rplus_le_mono_r : forall n m p : R, n <= m <-> n + p <= m + p. Proof. intros n m p; rewrite (Rplus_comm n p); rewrite (Rplus_comm m p). apply Rplus_le_mono_l. Qed. Theorem Rplus_lt_mono_l : forall n m p : R, n < m <-> p + n < p + m. Proof. intros n m p; do 2 rewrite Rlt_le_neq. rewrite Rplus_cancel_l. now rewrite <- Rplus_le_mono_l. Qed. Theorem Rplus_lt_mono_r : forall n m p : R, n < m <-> n + p < m + p. Proof. intros n m p. rewrite (Rplus_comm n p); rewrite (Rplus_comm m p); apply Rplus_lt_mono_l. Qed. Theorem Rplus_lt_mono : forall n m p q : R, n < m -> p < q -> n + p < m + q. Proof. intros n m p q H1 H2. apply Rlt_trans with (m + p); [now apply -> Rplus_lt_mono_r | now apply -> Rplus_lt_mono_l]. Qed. Theorem Rplus_le_mono : forall n m p q : R, n <= m -> p <= q -> n + p <= m + q. Proof. intros n m p q H1 H2. apply Rle_trans with (m + p); [now apply -> Rplus_le_mono_r | now apply -> Rplus_le_mono_l]. Qed. Theorem Rplus_lt_le_mono : forall n m p q : R, n < m -> p <= q -> n + p < m + q. Proof. intros n m p q H1 H2. apply Rlt_le_trans with (m + p); [now apply -> Rplus_lt_mono_r | now apply -> Rplus_le_mono_l]. Qed. Theorem Rplus_le_lt_mono : forall n m p q : R, n <= m -> p < q -> n + p < m + q. Proof. intros n m p q H1 H2. apply Rle_lt_trans with (m + p); [now apply -> Rplus_le_mono_r | now apply -> Rplus_lt_mono_l]. Qed. Theorem Rplus_pos_pos : forall n m : R, 0 < n -> 0 < m -> 0 < n + m. Proof. intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_lt_mono. Qed. Theorem Rplus_pos_nonneg : forall n m : R, 0 < n -> 0 <= m -> 0 < n + m. Proof. intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_lt_le_mono. Qed. Theorem Rplus_nonneg_pos : forall n m : R, 0 <= n -> 0 < m -> 0 < n + m. Proof. intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_le_lt_mono. Qed. Theorem Rplus_nonneg_nonneg : forall n m : R, 0 <= n -> 0 <= m -> 0 <= n + m. Proof. intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_le_mono. Qed. Theorem Rle_le_minus : forall n m : R, n <= m <-> 0 <= m - n. Proof. intros n m. rewrite (@Rplus_le_mono_r n m (- n)). setoid_replace (n + - n) with 0 by ring. now setoid_replace (m + - n) with (m - n) by ring. Qed. Theorem Rlt_lt_minus : forall n m : R, n < m <-> 0 < m - n. Proof. intros n m. rewrite (@Rplus_lt_mono_r n m (- n)). setoid_replace (n + - n) with 0 by ring. now setoid_replace (m + - n) with (m - n) by ring. Qed. Theorem Ropp_lt_mono : forall n m : R, n < m <-> - m < - n. Proof. intros n m. split; intro H. apply -> (@Rplus_lt_mono_l n m (- n - m)) in H. setoid_replace (- n - m + n) with (- m) in H by ring. now setoid_replace (- n - m + m) with (- n) in H by ring. apply -> (@Rplus_lt_mono_l (- m) (- n) (n + m)) in H. setoid_replace (n + m + - m) with n in H by ring. now setoid_replace (n + m + - n) with m in H by ring. Qed. Theorem Ropp_pos_neg : forall n : R, 0 < - n <-> n < 0. Proof. intro n; rewrite (Ropp_lt_mono n 0). now setoid_replace (- 0) with 0 by ring. Qed. (* Times and order *) Theorem Rtimes_pos_pos : forall n m : R, 0 < n -> 0 < m -> 0 < n * m. Proof (SORtimes_pos_pos sor). Theorem Rtimes_nonneg_nonneg : forall n m : R, 0 <= n -> 0 <= m -> 0 <= n * m. Proof. intros n m H1 H2. le_elim H1. le_elim H2. le_less; now apply Rtimes_pos_pos. rewrite <- H2; rewrite Rtimes_0_r; le_equal. rewrite <- H1; rewrite Rtimes_0_l; le_equal. Qed. Theorem Rtimes_pos_neg : forall n m : R, 0 < n -> m < 0 -> n * m < 0. Proof. intros n m H1 H2. apply -> Ropp_pos_neg. setoid_replace (- (n * m)) with (n * (- m)) by ring. apply Rtimes_pos_pos. assumption. now apply <- Ropp_pos_neg. Qed. Theorem Rtimes_neg_neg : forall n m : R, n < 0 -> m < 0 -> 0 < n * m. Proof. intros n m H1 H2. setoid_replace (n * m) with ((- n) * (- m)) by ring. apply Rtimes_pos_pos; now apply <- Ropp_pos_neg. Qed. Theorem Rtimes_square_nonneg : forall n : R, 0 <= n * n. Proof. intro n; destruct (Rlt_trichotomy 0 n) as [H | [H | H]]. le_less; now apply Rtimes_pos_pos. rewrite <- H, Rtimes_0_l; le_equal. le_less; now apply Rtimes_neg_neg. Qed. Theorem Rtimes_neq_0 : forall n m : R, n ~= 0 /\ m ~= 0 -> n * m ~= 0. Proof. intros n m [H1 H2]. destruct (Rlt_trichotomy n 0) as [H3 | [H3 | H3]]; destruct (Rlt_trichotomy m 0) as [H4 | [H4 | H4]]; try (false_hyp H3 H1); try (false_hyp H4 H2). apply Rneq_symm. apply Rlt_neq. now apply Rtimes_neg_neg. apply Rlt_neq. rewrite Rtimes_comm. now apply Rtimes_pos_neg. apply Rlt_neq. now apply Rtimes_pos_neg. apply Rneq_symm. apply Rlt_neq. now apply Rtimes_pos_pos. Qed. (* The following theorems are used to build a morphism from Z to R and prove its properties in ZCoeff.v. They are not used in RingMicromega.v. *) (* Surprisingly, multilication is needed to prove the following theorem *) Theorem Ropp_neg_pos : forall n : R, - n < 0 <-> 0 < n. Proof. intro n; setoid_replace n with (- - n) by ring. rewrite Ropp_pos_neg. now setoid_replace (- - n) with n by ring. Qed. Theorem Rlt_0_1 : 0 < 1. Proof. apply <- Rlt_le_neq. split. setoid_replace 1 with (1 * 1) by ring. apply Rtimes_square_nonneg. apply Rneq_0_1. Qed. Theorem Rlt_succ_r : forall n : R, n < 1 + n. Proof. intro n. rewrite <- (Rplus_0_l n); setoid_replace (1 + (0 + n)) with (1 + n) by ring. apply -> Rplus_lt_mono_r. apply Rlt_0_1. Qed. Theorem Rlt_lt_succ : forall n m : R, n < m -> n < 1 + m. Proof. intros n m H; apply Rlt_trans with m. assumption. apply Rlt_succ_r. Qed. (*Theorem Rtimes_lt_mono_pos_l : forall n m p : R, 0 < p -> n < m -> p * n < p * m. Proof. intros n m p H1 H2. apply <- Rlt_lt_minus. setoid_replace (p * m - p * n) with (p * (m - n)) by ring. apply Rtimes_pos_pos. assumption. now apply -> Rlt_lt_minus. Qed.*) End STRICT_ORDERED_RING. coq-8.11.0/plugins/micromega/simplex.mli0000644000175000017500000000167613612664533020110 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* cstr list -> (Num.num option * Num.num option) option val find_point : cstr list -> Vect.t option val find_unsat_certificate : cstr list -> Vect.t option val integer_solver : (cstr * ProofFormat.prf_rule) list -> ProofFormat.proof option coq-8.11.0/plugins/micromega/Fourier_util.v0000644000175000017500000000127313612664533020554 0ustar treinentreinenRequire Export Rbase. Require Import Lra. Local Open Scope R_scope. Lemma Rlt_mult_inv_pos : forall x y:R, 0 < x -> 0 < y -> 0 < x * / y. intros x y H H0; try assumption. replace 0 with (x * 0). apply Rmult_lt_compat_l; auto with real. ring. Qed. Lemma Rlt_zero_pos_plus1 : forall x:R, 0 < x -> 0 < 1 + x. intros x H; try assumption. rewrite Rplus_comm. apply Rle_lt_0_plus_1. red; auto with real. Qed. Lemma Rle_zero_pos_plus1 : forall x:R, 0 <= x -> 0 <= 1 + x. intros; lra. Qed. Lemma Rle_mult_inv_pos : forall x y:R, 0 <= x -> 0 < y -> 0 <= x * / y. intros x y H H0; try assumption. case H; intros. red; left. apply Rlt_mult_inv_pos; auto with real. rewrite <- H1. red; right; ring. Qed. coq-8.11.0/plugins/micromega/ZifyBool.v0000644000175000017500000001702013612664533017636 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 0 <= x <= 1) ; cstr := Z_of_bool_bound}. Add InjTyp Inj_bool_Z. (** Boolean operators *) Instance Op_andb : BinOp andb := { TBOp := Z.min ; TBOpInj := ltac: (destruct n,m; reflexivity)}. Add BinOp Op_andb. Instance Op_orb : BinOp orb := { TBOp := Z.max ; TBOpInj := ltac:(destruct n,m; reflexivity)}. Add BinOp Op_orb. Instance Op_implb : BinOp implb := { TBOp := fun x y => Z.max (1 - x) y; TBOpInj := ltac:(destruct n,m; reflexivity) }. Add BinOp Op_implb. Instance Op_xorb : BinOp xorb := { TBOp := fun x y => Z.max (x - y) (y - x); TBOpInj := ltac:(destruct n,m; reflexivity) }. Add BinOp Op_xorb. Instance Op_negb : UnOp negb := { TUOp := fun x => 1 - x ; TUOpInj := ltac:(destruct x; reflexivity)}. Add UnOp Op_negb. Instance Op_eq_bool : BinRel (@eq bool) := {TR := @eq Z ; TRInj := ltac:(destruct n,m; simpl ; intuition congruence) }. Add BinRel Op_eq_bool. Instance Op_true : CstOp true := { TCst := 1 ; TCstInj := eq_refl }. Add CstOp Op_true. Instance Op_false : CstOp false := { TCst := 0 ; TCstInj := eq_refl }. Add CstOp Op_false. (** Comparisons are encoded using the predicates [isZero] and [isLeZero].*) Definition isZero (z : Z) := Z_of_bool (Z.eqb z 0). Definition isLeZero (x : Z) := Z_of_bool (Z.leb x 0). (* Some intermediate lemma *) Lemma Z_eqb_isZero : forall n m, Z_of_bool (n =? m) = isZero (n - m). Proof. intros ; unfold isZero. destruct ( n =? m) eqn:EQ. - simpl. rewrite Z.eqb_eq in EQ. rewrite EQ. rewrite Z.sub_diag. reflexivity. - destruct (n - m =? 0) eqn:EQ'. rewrite Z.eqb_neq in EQ. rewrite Z.eqb_eq in EQ'. apply Zminus_eq in EQ'. congruence. reflexivity. Qed. Lemma Z_leb_sub : forall x y, x <=? y = ((x - y) <=? 0). Proof. intros. destruct (x <=?y) eqn:B1 ; destruct (x - y <=?0) eqn:B2 ; auto. - rewrite Z.leb_le in B1. rewrite Z.leb_nle in B2. rewrite Z.le_sub_0 in B2. tauto. - rewrite Z.leb_nle in B1. rewrite Z.leb_le in B2. rewrite Z.le_sub_0 in B2. tauto. Qed. Lemma Z_ltb_leb : forall x y, x isZero (Z.sub x y) ; TBOpInj := Z_eqb_isZero}. Instance Op_Zleb : BinOp Z.leb := { TBOp := fun x y => isLeZero (x-y) ; TBOpInj := ltac: (intros;unfold isLeZero; rewrite Z_leb_sub; auto) }. Add BinOp Op_Zleb. Instance Op_Zgeb : BinOp Z.geb := { TBOp := fun x y => isLeZero (y-x) ; TBOpInj := ltac:( intros; unfold isLeZero; rewrite Z.geb_leb; rewrite Z_leb_sub; auto) }. Add BinOp Op_Zgeb. Instance Op_Zltb : BinOp Z.ltb := { TBOp := fun x y => isLeZero (x+1-y) ; TBOpInj := ltac:( intros; unfold isLeZero; rewrite Z_ltb_leb; rewrite <- Z_leb_sub; reflexivity) }. Instance Op_Zgtb : BinOp Z.gtb := { TBOp := fun x y => isLeZero (y-x+1) ; TBOpInj := ltac:( intros; unfold isLeZero; rewrite Z.gtb_ltb; rewrite Z_ltb_leb; rewrite Z_leb_sub; rewrite Z.add_sub_swap; reflexivity) }. Add BinOp Op_Zgtb. (** Comparison over nat *) Lemma Z_of_nat_eqb_iff : forall n m, (n =? m)%nat = (Z.of_nat n =? Z.of_nat m). Proof. intros. rewrite Nat.eqb_compare. rewrite Z.eqb_compare. rewrite Nat2Z.inj_compare. reflexivity. Qed. Lemma Z_of_nat_leb_iff : forall n m, (n <=? m)%nat = (Z.of_nat n <=? Z.of_nat m). Proof. intros. rewrite Nat.leb_compare. rewrite Z.leb_compare. rewrite Nat2Z.inj_compare. reflexivity. Qed. Lemma Z_of_nat_ltb_iff : forall n m, (n isZero (Z.sub x y) ; TBOpInj := ltac:( intros; simpl; rewrite <- Z_eqb_isZero; f_equal; apply Z_of_nat_eqb_iff) }. Add BinOp Op_nat_eqb. Instance Op_nat_leb : BinOp Nat.leb := { TBOp := fun x y => isLeZero (x-y) ; TBOpInj := ltac:( intros; rewrite Z_of_nat_leb_iff; unfold isLeZero; rewrite Z_leb_sub; auto) }. Add BinOp Op_nat_leb. Instance Op_nat_ltb : BinOp Nat.ltb := { TBOp := fun x y => isLeZero (x+1-y) ; TBOpInj := ltac:( intros; rewrite Z_of_nat_ltb_iff; unfold isLeZero; rewrite Z_ltb_leb; rewrite <- Z_leb_sub; reflexivity) }. Add BinOp Op_nat_ltb. (** Injected boolean operators *) Lemma Z_eqb_ZSpec_ok : forall x, 0 <= isZero x <= 1 /\ (x = 0 <-> isZero x = 1). Proof. intros. unfold isZero. destruct (x =? 0) eqn:EQ. - apply Z.eqb_eq in EQ. simpl. intuition try congruence; compute ; congruence. - apply Z.eqb_neq in EQ. simpl. intuition try congruence; compute ; congruence. Qed. Instance Z_eqb_ZSpec : UnOpSpec isZero := {| UPred := fun n r => 0 <= r <= 1 /\ (n = 0 <-> isZero n = 1) ; USpec := Z_eqb_ZSpec_ok |}. Add Spec Z_eqb_ZSpec. Lemma leZeroSpec_ok : forall x, x <= 0 /\ isLeZero x = 1 \/ x > 0 /\ isLeZero x = 0. Proof. intros. unfold isLeZero. destruct (x <=? 0) eqn:EQ. - apply Z.leb_le in EQ. simpl. intuition congruence. - simpl. apply Z.leb_nle in EQ. apply Zorder.Znot_le_gt in EQ. tauto. Qed. Instance leZeroSpec : UnOpSpec isLeZero := {| UPred := fun n r => (n<=0 /\ r = 1) \/ (n > 0 /\ r = 0) ; USpec := leZeroSpec_ok|}. Add Spec leZeroSpec. coq-8.11.0/plugins/micromega/polynomial.mli0000644000175000017500000002604213612664533020604 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int -> 'a -> 'a) -> t -> 'a -> 'a (** [fold f m acc] folds over the variables with multiplicities *) val degree : t -> int (** [degree m] is the sum of the degrees of each variable *) val const : t (** [const] @return the empty monomial i.e. without any variable *) val is_const : t -> bool val var : var -> t (** [var x] @return the monomial x^1 *) val prod : t -> t -> t (** [prod n m] @return the monomial n*m *) val sqrt : t -> t option (** [sqrt m] @return [Some r] iff r^2 = m *) val is_var : t -> bool (** [is_var m] @return [true] iff m = x^1 for some variable x *) val get_var : t -> var option (** [get_var m] @return [x] iff m = x^1 for variable x *) val div : t -> t -> t * int (** [div m1 m2] @return a pair [mr,n] such that mr * (m2)^n = m1 where n is maximum *) val compare : t -> t -> int (** [compare m1 m2] provides a total order over monomials*) val variables : t -> ISet.t (** [variables m] @return the set of variables with (strictly) positive multiplicities *) end module MonMap : sig include Map.S with type key = Monomial.t val union : (Monomial.t -> 'a -> 'a -> 'a option) -> 'a t -> 'a t -> 'a t end module Poly : sig (** Representation of polonomial with rational coefficient. a1.m1 + ... + c where - ai are rational constants (num type) - mi are monomials - c is a rational constant *) type t val constant : Num.num -> t (** [constant c] @return the constant polynomial c *) val variable : var -> t (** [variable x] @return the polynomial 1.x^1 *) val addition : t -> t -> t (** [addition p1 p2] @return the polynomial p1+p2 *) val product : t -> t -> t (** [product p1 p2] @return the polynomial p1*p2 *) val uminus : t -> t (** [uminus p] @return the polynomial -p i.e product by -1 *) val get : Monomial.t -> t -> Num.num (** [get mi p] @return the coefficient ai of the monomial mi. *) val fold : (Monomial.t -> Num.num -> 'a -> 'a) -> t -> 'a -> 'a (** [fold f p a] folds f over the monomials of p with non-zero coefficient *) val add : Monomial.t -> Num.num -> t -> t (** [add m n p] @return the polynomial n*m + p *) end type cstr = {coeffs : Vect.t; op : op; cst : Num.num} (* Representation of linear constraints *) and op = Eq | Ge | Gt val eval_op : op -> Num.num -> Num.num -> bool (*val opMult : op -> op -> op*) val opAdd : op -> op -> op val is_strict : cstr -> bool (** [is_strict c] @return whether the constraint is strict i.e. c.op = Gt *) exception Strict module LinPoly : sig (** Linear(ised) polynomials represented as a [Vect.t] i.e a sorted association list. The constant is the coefficient of the variable 0 Each linear polynomial can be interpreted as a multi-variate polynomial. There is a bijection mapping between a linear variable and a monomial (see module [MonT]) *) type t = Vect.t (** Each variable of a linear polynomial is mapped to a monomial. This is done using the monomial tables of the module MonT. *) module MonT : sig val clear : unit -> unit (** [clear ()] clears the mapping. *) val reserve : int -> unit (** [reserve i] reserves the integer i *) val get_fresh : unit -> int (** [get_fresh ()] return the first fresh variable *) val retrieve : int -> Monomial.t (** [retrieve x] @return the monomial corresponding to the variable [x] *) val register : Monomial.t -> int (** [register m] @return the variable index for the monomial m *) end val linpol_of_pol : Poly.t -> t (** [linpol_of_pol p] linearise the polynomial p *) val var : var -> t (** [var x] @return 1.y where y is the variable index of the monomial x^1. *) val coq_poly_of_linpol : (Num.num -> 'a) -> t -> 'a Mc.pExpr (** [coq_poly_of_linpol c p] @param p is a multi-variate polynomial. @param c maps a rational to a Coq polynomial coefficient. @return the coq expression corresponding to polynomial [p].*) val of_monomial : Monomial.t -> t (** [of_monomial m] @returns 1.x where x is the variable (index) for monomial m *) val of_vect : Vect.t -> t (** [of_vect v] @returns a1.x1 + ... + an.xn This is not the identity because xi is the variable index of xi^1 *) val variables : t -> ISet.t (** [variables p] @return the set of variables of the polynomial p interpreted as a multi-variate polynomial *) val is_variable : t -> var option (** [is_variable p] @return Some x if p = a.x for a >= 0 *) val is_linear : t -> bool (** [is_linear p] @return whether the multi-variate polynomial is linear. *) val is_linear_for : var -> t -> bool (** [is_linear_for x p] @return true if the polynomial is linear in x i.e can be written c*x+r where c is a constant and r is independent from x *) val constant : Num.num -> t (** [constant c] @return the constant polynomial c *) (** [search_linear pred p] @return a variable x such p = a.x + b such that p is linear in x i.e x does not occur in b and a is a constant such that [pred a] *) val search_linear : (Num.num -> bool) -> t -> var option val search_all_linear : (Num.num -> bool) -> t -> var list (** [search_all_linear pred p] @return all the variables x such p = a.x + b such that p is linear in x i.e x does not occur in b and a is a constant such that [pred a] *) val product : t -> t -> t (** [product p q] @return the product of the polynomial [p*q] *) val factorise : var -> t -> t * t (** [factorise x p] @return [a,b] such that [p = a.x + b] and [x] does not occur in [b] *) val collect_square : t -> Monomial.t MonMap.t (** [collect_square p] @return a mapping m such that m[s] = s^2 for every s^2 that is a monomial of [p] *) val monomials : t -> ISet.t (** [monomials p] @return the set of monomials. *) val degree : t -> int (** [degree p] @return return the maximum degree *) val pp_var : out_channel -> var -> unit (** [pp_var o v] pretty-prints a monomial indexed by v. *) val pp : out_channel -> t -> unit (** [pp o p] pretty-prints a polynomial. *) val pp_goal : string -> out_channel -> (t * op) list -> unit (** [pp_goal typ o l] pretty-prints the list of constraints as a Coq goal. *) end module ProofFormat : sig (** Proof format used by the proof-generating procedures. It is fairly close to Coq format but a bit more liberal. It is used for proofs over Z, Q, R. However, certain constructions e.g. [CutPrf] are only relevant for Z. *) type prf_rule = | Annot of string * prf_rule | Hyp of int | Def of int | Cst of Num.num | Zero | Square of Vect.t | MulC of Vect.t * prf_rule | Gcd of Big_int.big_int * prf_rule | MulPrf of prf_rule * prf_rule | AddPrf of prf_rule * prf_rule | CutPrf of prf_rule type proof = | Done | Step of int * prf_rule * proof | Enum of int * prf_rule * Vect.t * prf_rule * proof list | ExProof of int * int * int * var * var * var * proof (* x = z - t, z >= 0, t >= 0 *) val pr_size : prf_rule -> Num.num val pr_rule_max_id : prf_rule -> int val proof_max_id : proof -> int val normalise_proof : int -> proof -> int * proof val output_prf_rule : out_channel -> prf_rule -> unit val output_proof : out_channel -> proof -> unit val add_proof : prf_rule -> prf_rule -> prf_rule val mul_cst_proof : Num.num -> prf_rule -> prf_rule val mul_proof : prf_rule -> prf_rule -> prf_rule val compile_proof : int list -> proof -> Micromega.zArithProof val cmpl_prf_rule : ('a Micromega.pExpr -> 'a Micromega.pol) -> (Num.num -> 'a) -> int list -> prf_rule -> 'a Micromega.psatz val proof_of_farkas : prf_rule IMap.t -> Vect.t -> prf_rule val eval_prf_rule : (int -> LinPoly.t * op) -> prf_rule -> LinPoly.t * op val eval_proof : (LinPoly.t * op) IMap.t -> proof -> bool end val output_cstr : out_channel -> cstr -> unit val opMult : op -> op -> op (** [module WithProof] constructs polynomials packed with the proof that their sign is correct. *) module WithProof : sig type t = (LinPoly.t * op) * ProofFormat.prf_rule exception InvalidProof (** [InvalidProof] is raised if the operation is invalid. *) val annot : string -> t -> t val of_cstr : cstr * ProofFormat.prf_rule -> t val output : out_channel -> t -> unit (** [out_channel chan c] pretty-prints the constraint [c] over the channel [chan] *) val output_sys : out_channel -> t list -> unit val zero : t (** [zero] represents the tautology (0=0) *) val const : Num.num -> t (** [const n] represents the tautology (n>=0) *) val product : t -> t -> t (** [product p q] @return the polynomial p*q with its sign and proof *) val addition : t -> t -> t (** [addition p q] @return the polynomial p+q with its sign and proof *) val mult : LinPoly.t -> t -> t (** [mult p q] @return the polynomial p*q with its sign and proof. @raise InvalidProof if p is not a constant and p is not an equality *) val cutting_plane : t -> t option (** [cutting_plane p] does integer reasoning and adjust the constant to be integral *) val linear_pivot : t list -> t -> Vect.var -> t -> t option (** [linear_pivot sys p x q] @return the polynomial [q] where [x] is eliminated using the polynomial [p] The pivoting operation is only defined if - p is linear in x i.e p = a.x+b and x neither occurs in a and b - The pivoting also requires some sign conditions for [a] *) (** [subst sys] performs the equivalent of the 'subst' tactic of Coq. For every p=0 \in sys such that p is linear in x with coefficient +/- 1 i.e. p = 0 <-> x = e and x \notin e. Replace x by e in sys NB: performing this transformation may hinders the non-linear prover to find a proof. [elim_simple_linear_equality] is much more careful. *) val subst : t list -> t list val subst1 : t list -> t list (** [subst1 sys] performs a single substitution *) val saturate_subst : bool -> t list -> t list val is_substitution : bool -> t -> var option end coq-8.11.0/plugins/micromega/g_micromega.mlg0000644000175000017500000000621313612664533020666 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Tactics.red_in_concl } END TACTIC EXTEND PsatzZ | [ "psatz_Z" int_or_var(i) tactic(t) ] -> { (Coq_micromega.psatz_Z i (Tacinterp.tactic_of_value ist t)) } | [ "psatz_Z" tactic(t)] -> { (Coq_micromega.psatz_Z (-1)) (Tacinterp.tactic_of_value ist t) } END TACTIC EXTEND Lia | [ "xlia" tactic(t) ] -> { (Coq_micromega.xlia (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND Nia | [ "xnlia" tactic(t) ] -> { (Coq_micromega.xnlia (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND NRA | [ "xnra" tactic(t) ] -> { (Coq_micromega.nra (Tacinterp.tactic_of_value ist t))} END TACTIC EXTEND NQA | [ "xnqa" tactic(t) ] -> { (Coq_micromega.nqa (Tacinterp.tactic_of_value ist t))} END TACTIC EXTEND Sos_Z | [ "sos_Z" tactic(t) ] -> { (Coq_micromega.sos_Z (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND Sos_Q | [ "sos_Q" tactic(t) ] -> { (Coq_micromega.sos_Q (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND Sos_R | [ "sos_R" tactic(t) ] -> { (Coq_micromega.sos_R (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND LRA_Q | [ "lra_Q" tactic(t) ] -> { (Coq_micromega.lra_Q (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND LRA_R | [ "lra_R" tactic(t) ] -> { (Coq_micromega.lra_R (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND PsatzR | [ "psatz_R" int_or_var(i) tactic(t) ] -> { (Coq_micromega.psatz_R i (Tacinterp.tactic_of_value ist t)) } | [ "psatz_R" tactic(t) ] -> { (Coq_micromega.psatz_R (-1) (Tacinterp.tactic_of_value ist t)) } END TACTIC EXTEND PsatzQ | [ "psatz_Q" int_or_var(i) tactic(t) ] -> { (Coq_micromega.psatz_Q i (Tacinterp.tactic_of_value ist t)) } | [ "psatz_Q" tactic(t) ] -> { (Coq_micromega.psatz_Q (-1) (Tacinterp.tactic_of_value ist t)) } END coq-8.11.0/plugins/micromega/ZMicromega.v0000644000175000017500000013711313612664533020144 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* destruct (andb_prop _ _ id); clear id | [ id : (_ || _)%bool = true |- _ ] => destruct (orb_prop _ _ id); clear id end. Ltac inv H := inversion H ; try subst ; clear H. Lemma eq_le_iff : forall x, 0 = x <-> (0 <= x /\ x <= 0). Proof. intros. split ; intros. - subst. compute. intuition congruence. - destruct H. apply Z.le_antisymm; auto. Qed. Lemma lt_le_iff : forall x, 0 < x <-> 0 <= x - 1. Proof. split ; intros. - apply Zlt_succ_le. ring_simplify. auto. - apply Zle_lt_succ in H. ring_simplify in H. auto. Qed. Lemma le_0_iff : forall x y, x <= y <-> 0 <= y - x. Proof. split ; intros. - apply Zle_minus_le_0; auto. - apply Zle_0_minus_le; auto. Qed. Lemma le_neg : forall x, ((0 <= x) -> False) <-> 0 < -x. Proof. intro. rewrite lt_le_iff. split ; intros. - apply Znot_le_gt in H. apply Zgt_le_succ in H. rewrite le_0_iff in H. ring_simplify in H; auto. - assert (C := (Z.add_le_mono _ _ _ _ H H0)). ring_simplify in C. compute in C. apply C ; reflexivity. Qed. Lemma eq_cnf : forall x, (0 <= x - 1 -> False) /\ (0 <= -1 - x -> False) <-> x = 0. Proof. intros. rewrite Z.eq_sym_iff. rewrite eq_le_iff. rewrite (le_0_iff x 0). rewrite !le_neg. rewrite !lt_le_iff. replace (- (x - 1) -1) with (-x) by ring. replace (- (-1 - x) -1) with x by ring. split ; intros (H1 & H2); auto. Qed. Require Import EnvRing. Lemma Zsor : SOR 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt. Proof. constructor ; intros ; subst; try reflexivity. apply Zsth. apply Zth. auto using Z.le_antisymm. eauto using Z.le_trans. apply Z.le_neq. destruct (Z.lt_trichotomy n m) ; intuition. apply Z.add_le_mono_l; assumption. apply Z.mul_pos_pos ; auto. discriminate. Qed. Lemma ZSORaddon : SORaddon 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le (* ring elements *) 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (* coefficients *) Zeq_bool Z.leb (fun x => x) (fun x => x) (pow_N 1 Z.mul). Proof. constructor. constructor ; intros ; try reflexivity. apply Zeq_bool_eq ; auto. constructor. reflexivity. intros x y. apply Zeq_bool_neq ; auto. apply Zle_bool_imp_le. Qed. Fixpoint Zeval_expr (env : PolEnv Z) (e: PExpr Z) : Z := match e with | PEc c => c | PEX x => env x | PEadd e1 e2 => Zeval_expr env e1 + Zeval_expr env e2 | PEmul e1 e2 => Zeval_expr env e1 * Zeval_expr env e2 | PEpow e1 n => Z.pow (Zeval_expr env e1) (Z.of_N n) | PEsub e1 e2 => (Zeval_expr env e1) - (Zeval_expr env e2) | PEopp e => Z.opp (Zeval_expr env e) end. Definition eval_expr := eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x => x) (fun x => x) (pow_N 1 Z.mul). Fixpoint Zeval_const (e: PExpr Z) : option Z := match e with | PEc c => Some c | PEX x => None | PEadd e1 e2 => map_option2 (fun x y => Some (x + y)) (Zeval_const e1) (Zeval_const e2) | PEmul e1 e2 => map_option2 (fun x y => Some (x * y)) (Zeval_const e1) (Zeval_const e2) | PEpow e1 n => map_option (fun x => Some (Z.pow x (Z.of_N n))) (Zeval_const e1) | PEsub e1 e2 => map_option2 (fun x y => Some (x - y)) (Zeval_const e1) (Zeval_const e2) | PEopp e => map_option (fun x => Some (Z.opp x)) (Zeval_const e) end. Lemma ZNpower : forall r n, r ^ Z.of_N n = pow_N 1 Z.mul r n. Proof. destruct n. reflexivity. simpl. unfold Z.pow_pos. replace (pow_pos Z.mul r p) with (1 * (pow_pos Z.mul r p)) by ring. generalize 1. induction p; simpl ; intros ; repeat rewrite IHp ; ring. Qed. Lemma Zeval_expr_compat : forall env e, Zeval_expr env e = eval_expr env e. Proof. induction e ; simpl ; try congruence. reflexivity. rewrite ZNpower. congruence. Qed. Definition Zeval_op2 (o : Op2) : Z -> Z -> Prop := match o with | OpEq => @eq Z | OpNEq => fun x y => ~ x = y | OpLe => Z.le | OpGe => Z.ge | OpLt => Z.lt | OpGt => Z.gt end. Definition Zeval_formula (env : PolEnv Z) (f : Formula Z):= let (lhs, op, rhs) := f in (Zeval_op2 op) (Zeval_expr env lhs) (Zeval_expr env rhs). Definition Zeval_formula' := eval_formula Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt (fun x => x) (fun x => x) (pow_N 1 Z.mul). Lemma Zeval_formula_compat' : forall env f, Zeval_formula env f <-> Zeval_formula' env f. Proof. intros. unfold Zeval_formula. destruct f. repeat rewrite Zeval_expr_compat. unfold Zeval_formula' ; simpl. unfold eval_expr. generalize (eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x) (fun x : N => x) (pow_N 1 Z.mul) env Flhs). generalize ((eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x) (fun x : N => x) (pow_N 1 Z.mul) env Frhs)). destruct Fop ; simpl; intros; intuition auto using Z.le_ge, Z.ge_le, Z.lt_gt, Z.gt_lt. Qed. Definition eval_nformula := eval_nformula 0 Z.add Z.mul (@eq Z) Z.le Z.lt (fun x => x) . Definition Zeval_op1 (o : Op1) : Z -> Prop := match o with | Equal => fun x : Z => x = 0 | NonEqual => fun x : Z => x <> 0 | Strict => fun x : Z => 0 < x | NonStrict => fun x : Z => 0 <= x end. Lemma Zeval_nformula_dec : forall env d, (eval_nformula env d) \/ ~ (eval_nformula env d). Proof. intros. apply (eval_nformula_dec Zsor). Qed. Definition ZWitness := Psatz Z. Definition ZWeakChecker := check_normalised_formulas 0 1 Z.add Z.mul Zeq_bool Z.leb. Lemma ZWeakChecker_sound : forall (l : list (NFormula Z)) (cm : ZWitness), ZWeakChecker l cm = true -> forall env, make_impl (eval_nformula env) l False. Proof. intros l cm H. intro. unfold eval_nformula. apply (checker_nf_sound Zsor ZSORaddon l cm). unfold ZWeakChecker in H. exact H. Qed. Definition psub := psub Z0 Z.add Z.sub Z.opp Zeq_bool. Declare Equivalent Keys psub RingMicromega.psub. Definition padd := padd Z0 Z.add Zeq_bool. Declare Equivalent Keys padd RingMicromega.padd. Definition pmul := pmul 0 1 Z.add Z.mul Zeq_bool. Definition normZ := norm 0 1 Z.add Z.mul Z.sub Z.opp Zeq_bool. Declare Equivalent Keys normZ RingMicromega.norm. Definition eval_pol := eval_pol Z.add Z.mul (fun x => x). Declare Equivalent Keys eval_pol RingMicromega.eval_pol. Lemma eval_pol_sub : forall env lhs rhs, eval_pol env (psub lhs rhs) = eval_pol env lhs - eval_pol env rhs. Proof. intros. apply (eval_pol_sub Zsor ZSORaddon). Qed. Lemma eval_pol_add : forall env lhs rhs, eval_pol env (padd lhs rhs) = eval_pol env lhs + eval_pol env rhs. Proof. intros. apply (eval_pol_add Zsor ZSORaddon). Qed. Lemma eval_pol_mul : forall env lhs rhs, eval_pol env (pmul lhs rhs) = eval_pol env lhs * eval_pol env rhs. Proof. intros. apply (eval_pol_mul Zsor ZSORaddon). Qed. Lemma eval_pol_norm : forall env e, eval_expr env e = eval_pol env (normZ e) . Proof. intros. apply (eval_pol_norm Zsor ZSORaddon). Qed. Definition Zunsat := check_inconsistent 0 Zeq_bool Z.leb. Definition Zdeduce := nformula_plus_nformula 0 Z.add Zeq_bool. Lemma Zunsat_sound : forall f, Zunsat f = true -> forall env, eval_nformula env f -> False. Proof. unfold Zunsat. intros. destruct f. eapply check_inconsistent_sound with (1 := Zsor) (2 := ZSORaddon) in H; eauto. Qed. Definition xnnormalise (t : Formula Z) : NFormula Z := let (lhs,o,rhs) := t in let lhs := normZ lhs in let rhs := normZ rhs in match o with | OpEq => (psub rhs lhs, Equal) | OpNEq => (psub rhs lhs, NonEqual) | OpGt => (psub lhs rhs, Strict) | OpLt => (psub rhs lhs, Strict) | OpGe => (psub lhs rhs, NonStrict) | OpLe => (psub rhs lhs, NonStrict) end. Lemma xnnormalise_correct : forall env f, eval_nformula env (xnnormalise f) <-> Zeval_formula env f. Proof. intros. rewrite Zeval_formula_compat'. unfold xnnormalise. destruct f as [lhs o rhs]. destruct o eqn:O ; cbn ; rewrite ?eval_pol_sub; rewrite <- !eval_pol_norm ; simpl in *; unfold eval_expr; generalize ( eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x) (fun x : N => x) (pow_N 1 Z.mul) env lhs); generalize (eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x) (fun x : N => x) (pow_N 1 Z.mul) env rhs); intros. - split ; intros. + assert (z0 + (z - z0) = z0 + 0) by congruence. rewrite Z.add_0_r in H0. rewrite <- H0. ring. + subst. ring. - split ; repeat intro. subst. apply H. ring. apply H. assert (z0 + (z - z0) = z0 + 0) by congruence. rewrite Z.add_0_r in H1. rewrite <- H1. ring. - split ; intros. + apply Zle_0_minus_le; auto. + apply Zle_minus_le_0; auto. - split ; intros. + apply Zle_0_minus_le; auto. + apply Zle_minus_le_0; auto. - split ; intros. + apply Zlt_0_minus_lt; auto. + apply Zlt_left_lt in H. apply H. - split ; intros. + apply Zlt_0_minus_lt ; auto. + apply Zlt_left_lt in H. apply H. Qed. Definition xnormalise (f: NFormula Z) : list (NFormula Z) := let (e,o) := f in match o with | Equal => (psub e (Pc 1),NonStrict) :: (psub (Pc (-1)) e, NonStrict) :: nil | NonStrict => ((psub (Pc (-1)) e,NonStrict)::nil) | Strict => ((psub (Pc 0)) e, NonStrict)::nil | NonEqual => (e, Equal)::nil end. Lemma eval_pol_Pc : forall env z, eval_pol env (Pc z) = z. Proof. reflexivity. Qed. Ltac iff_ring := match goal with | |- ?F 0 ?X <-> ?F 0 ?Y => replace X with Y by ring ; tauto end. Lemma xnormalise_correct : forall env f, (make_conj (fun x => eval_nformula env x -> False) (xnormalise f)) <-> eval_nformula env f. Proof. intros. destruct f as [e o]; destruct o eqn:Op; cbn - [psub]; repeat rewrite eval_pol_sub; fold eval_pol; repeat rewrite eval_pol_Pc; generalize (eval_pol env e) as x; intro. - apply eq_cnf. - unfold not. tauto. - rewrite le_neg. iff_ring. - rewrite le_neg. rewrite lt_le_iff. iff_ring. Qed. Require Import Coq.micromega.Tauto BinNums. Definition cnf_of_list {T: Type} (tg : T) (l : list (NFormula Z)) := List.fold_right (fun x acc => if Zunsat x then acc else ((x,tg)::nil)::acc) (cnf_tt _ _) l. Lemma cnf_of_list_correct : forall {T : Type} (tg:T) (f : list (NFormula Z)) env, eval_cnf eval_nformula env (cnf_of_list tg f) <-> make_conj (fun x : NFormula Z => eval_nformula env x -> False) f. Proof. unfold cnf_of_list. intros. set (F := (fun (x : NFormula Z) (acc : list (list (NFormula Z * T))) => if Zunsat x then acc else ((x, tg) :: nil) :: acc)). set (E := ((fun x : NFormula Z => eval_nformula env x -> False))). induction f. - compute. tauto. - rewrite make_conj_cons. simpl. unfold F at 1. destruct (Zunsat a) eqn:EQ. + rewrite IHf. unfold E at 1. specialize (Zunsat_sound _ EQ env). tauto. + rewrite <- eval_cnf_cons_iff. rewrite IHf. simpl. unfold E at 2. unfold eval_tt. simpl. tauto. Qed. Definition normalise {T : Type} (t:Formula Z) (tg:T) : cnf (NFormula Z) T := let f := xnnormalise t in if Zunsat f then cnf_ff _ _ else cnf_of_list tg (xnormalise f). Lemma normalise_correct : forall (T: Type) env t (tg:T), eval_cnf eval_nformula env (normalise t tg) <-> Zeval_formula env t. Proof. intros. rewrite <- xnnormalise_correct. unfold normalise. generalize (xnnormalise t) as f;intro. destruct (Zunsat f) eqn:U. - assert (US := Zunsat_sound _ U env). rewrite eval_cnf_ff. tauto. - rewrite cnf_of_list_correct. apply xnormalise_correct. Qed. Definition xnegate (f:NFormula Z) : list (NFormula Z) := let (e,o) := f in match o with | Equal => (e,Equal) :: nil | NonEqual => (psub e (Pc 1),NonStrict) :: (psub (Pc (-1)) e, NonStrict) :: nil | NonStrict => (e,NonStrict)::nil | Strict => (psub e (Pc 1),NonStrict)::nil end. Definition negate {T : Type} (t:Formula Z) (tg:T) : cnf (NFormula Z) T := let f := xnnormalise t in if Zunsat f then cnf_tt _ _ else cnf_of_list tg (xnegate f). Lemma xnegate_correct : forall env f, (make_conj (fun x => eval_nformula env x -> False) (xnegate f)) <-> ~ eval_nformula env f. Proof. intros. destruct f as [e o]; destruct o eqn:Op; cbn - [psub]; repeat rewrite eval_pol_sub; fold eval_pol; repeat rewrite eval_pol_Pc; generalize (eval_pol env e) as x; intro. - tauto. - rewrite eq_cnf. destruct (Z.eq_decidable x 0);tauto. - rewrite lt_le_iff. tauto. - tauto. Qed. Lemma negate_correct : forall T env t (tg:T), eval_cnf eval_nformula env (negate t tg) <-> ~ Zeval_formula env t. Proof. intros. rewrite <- xnnormalise_correct. unfold negate. generalize (xnnormalise t) as f;intro. destruct (Zunsat f) eqn:U. - assert (US := Zunsat_sound _ U env). rewrite eval_cnf_tt. tauto. - rewrite cnf_of_list_correct. apply xnegate_correct. Qed. Definition cnfZ (Annot: Type) (TX : Type) (AF : Type) (f : TFormula (Formula Z) Annot TX AF) := rxcnf Zunsat Zdeduce normalise negate true f. Definition ZweakTautoChecker (w: list ZWitness) (f : BFormula (Formula Z)) : bool := @tauto_checker (Formula Z) (NFormula Z) unit Zunsat Zdeduce normalise negate ZWitness (fun cl => ZWeakChecker (List.map fst cl)) f w. (* To get a complete checker, the proof format has to be enriched *) Require Import Zdiv. Local Open Scope Z_scope. Definition ceiling (a b:Z) : Z := let (q,r) := Z.div_eucl a b in match r with | Z0 => q | _ => q + 1 end. Require Import Znumtheory. Lemma Zdivide_ceiling : forall a b, (b | a) -> ceiling a b = Z.div a b. Proof. unfold ceiling. intros. apply Zdivide_mod in H. case_eq (Z.div_eucl a b). intros. change z with (fst (z,z0)). rewrite <- H0. change (fst (Z.div_eucl a b)) with (Z.div a b). change z0 with (snd (z,z0)). rewrite <- H0. change (snd (Z.div_eucl a b)) with (Z.modulo a b). rewrite H. reflexivity. Qed. Lemma narrow_interval_lower_bound a b x : a > 0 -> a * x >= b -> x >= ceiling b a. Proof. rewrite !Z.ge_le_iff. unfold ceiling. intros Ha H. generalize (Z_div_mod b a Ha). destruct (Z.div_eucl b a) as (q,r). intros (->,(H1,H2)). destruct r as [|r|r]. - rewrite Z.add_0_r in H. apply Z.mul_le_mono_pos_l in H; auto with zarith. - assert (0 < Z.pos r) by easy. rewrite Z.add_1_r, Z.le_succ_l. apply Z.mul_lt_mono_pos_l with a. auto using Z.gt_lt. eapply Z.lt_le_trans. 2: eassumption. now apply Z.lt_add_pos_r. - now elim H1. Qed. (** NB: narrow_interval_upper_bound is Zdiv.Zdiv_le_lower_bound *) Require Import QArith. Inductive ZArithProof := | DoneProof | RatProof : ZWitness -> ZArithProof -> ZArithProof | CutProof : ZWitness -> ZArithProof -> ZArithProof | EnumProof : ZWitness -> ZWitness -> list ZArithProof -> ZArithProof | ExProof : positive -> ZArithProof -> ZArithProof (*ExProof x : exists z t, x = z - t /\ z >= 0 /\ t >= 0 *) . (*| SplitProof : PolC Z -> ZArithProof -> ZArithProof -> ZArithProof.*) (* n/d <= x -> d*x - n >= 0 *) (* In order to compute the 'cut', we need to express a polynomial P as a * Q + b. - b is the constant - a is the gcd of the other coefficient. *) Require Import Znumtheory. Definition isZ0 (x:Z) := match x with | Z0 => true | _ => false end. Lemma isZ0_0 : forall x, isZ0 x = true <-> x = 0. Proof. destruct x ; simpl ; intuition congruence. Qed. Lemma isZ0_n0 : forall x, isZ0 x = false <-> x <> 0. Proof. destruct x ; simpl ; intuition congruence. Qed. Definition ZgcdM (x y : Z) := Z.max (Z.gcd x y) 1. Fixpoint Zgcd_pol (p : PolC Z) : (Z * Z) := match p with | Pc c => (0,c) | Pinj _ p => Zgcd_pol p | PX p _ q => let (g1,c1) := Zgcd_pol p in let (g2,c2) := Zgcd_pol q in (ZgcdM (ZgcdM g1 c1) g2 , c2) end. (*Eval compute in (Zgcd_pol ((PX (Pc (-2)) 1 (Pc 4)))).*) Fixpoint Zdiv_pol (p:PolC Z) (x:Z) : PolC Z := match p with | Pc c => Pc (Z.div c x) | Pinj j p => Pinj j (Zdiv_pol p x) | PX p j q => PX (Zdiv_pol p x) j (Zdiv_pol q x) end. Inductive Zdivide_pol (x:Z): PolC Z -> Prop := | Zdiv_Pc : forall c, (x | c) -> Zdivide_pol x (Pc c) | Zdiv_Pinj : forall p, Zdivide_pol x p -> forall j, Zdivide_pol x (Pinj j p) | Zdiv_PX : forall p q, Zdivide_pol x p -> Zdivide_pol x q -> forall j, Zdivide_pol x (PX p j q). Lemma Zdiv_pol_correct : forall a p, 0 < a -> Zdivide_pol a p -> forall env, eval_pol env p = a * eval_pol env (Zdiv_pol p a). Proof. intros until 2. induction H0. (* Pc *) simpl. intros. apply Zdivide_Zdiv_eq ; auto. (* Pinj *) simpl. intros. apply IHZdivide_pol. (* PX *) simpl. intros. rewrite IHZdivide_pol1. rewrite IHZdivide_pol2. ring. Qed. Lemma Zgcd_pol_ge : forall p, fst (Zgcd_pol p) >= 0. Proof. induction p. 1-2: easy. simpl. case_eq (Zgcd_pol p1). case_eq (Zgcd_pol p3). intros. simpl. unfold ZgcdM. apply Z.le_ge; transitivity 1. easy. apply Z.le_max_r. Qed. Lemma Zdivide_pol_Zdivide : forall p x y, Zdivide_pol x p -> (y | x) -> Zdivide_pol y p. Proof. intros. induction H. constructor. apply Z.divide_trans with (1:= H0) ; assumption. constructor. auto. constructor ; auto. Qed. Lemma Zdivide_pol_one : forall p, Zdivide_pol 1 p. Proof. induction p ; constructor ; auto. exists c. ring. Qed. Lemma Zgcd_minus : forall a b c, (a | c - b ) -> (Z.gcd a b | c). Proof. intros a b c (q,Hq). destruct (Zgcd_is_gcd a b) as [(a',Ha) (b',Hb) _]. set (g:=Z.gcd a b) in *; clearbody g. exists (q * a' + b'). symmetry in Hq. rewrite <- Z.add_move_r in Hq. rewrite <- Hq, Hb, Ha. ring. Qed. Lemma Zdivide_pol_sub : forall p a b, 0 < Z.gcd a b -> Zdivide_pol a (PsubC Z.sub p b) -> Zdivide_pol (Z.gcd a b) p. Proof. induction p. simpl. intros. inversion H0. constructor. apply Zgcd_minus ; auto. intros. constructor. simpl in H0. inversion H0 ; subst; clear H0. apply IHp ; auto. simpl. intros. inv H0. constructor. apply Zdivide_pol_Zdivide with (1:= H3). destruct (Zgcd_is_gcd a b) ; assumption. apply IHp2 ; assumption. Qed. Lemma Zdivide_pol_sub_0 : forall p a, Zdivide_pol a (PsubC Z.sub p 0) -> Zdivide_pol a p. Proof. induction p. simpl. intros. inversion H. constructor. rewrite Z.sub_0_r in *. assumption. intros. constructor. simpl in H. inversion H ; subst; clear H. apply IHp ; auto. simpl. intros. inv H. constructor. auto. apply IHp2 ; assumption. Qed. Lemma Zgcd_pol_div : forall p g c, Zgcd_pol p = (g, c) -> Zdivide_pol g (PsubC Z.sub p c). Proof. induction p ; simpl. (* Pc *) intros. inv H. constructor. exists 0. now ring. (* Pinj *) intros. constructor. apply IHp ; auto. (* PX *) intros g c. case_eq (Zgcd_pol p1) ; case_eq (Zgcd_pol p3) ; intros. inv H1. unfold ZgcdM at 1. destruct (Zmax_spec (Z.gcd (ZgcdM z1 z2) z) 1) as [HH1 | HH1]; destruct HH1 as [HH1 HH1'] ; rewrite HH1'. constructor. apply Zdivide_pol_Zdivide with (x:= ZgcdM z1 z2). unfold ZgcdM. destruct (Zmax_spec (Z.gcd z1 z2) 1) as [HH2 | HH2]. destruct HH2. rewrite H2. apply Zdivide_pol_sub ; auto. apply Z.lt_le_trans with 1. reflexivity. now apply Z.ge_le. destruct HH2. rewrite H2. apply Zdivide_pol_one. unfold ZgcdM in HH1. unfold ZgcdM. destruct (Zmax_spec (Z.gcd z1 z2) 1) as [HH2 | HH2]. destruct HH2. rewrite H2 in *. destruct (Zgcd_is_gcd (Z.gcd z1 z2) z); auto. destruct HH2. rewrite H2. destruct (Zgcd_is_gcd 1 z); auto. apply Zdivide_pol_Zdivide with (x:= z). apply (IHp2 _ _ H); auto. destruct (Zgcd_is_gcd (ZgcdM z1 z2) z); auto. constructor. apply Zdivide_pol_one. apply Zdivide_pol_one. Qed. Lemma Zgcd_pol_correct_lt : forall p env g c, Zgcd_pol p = (g,c) -> 0 < g -> eval_pol env p = g * (eval_pol env (Zdiv_pol (PsubC Z.sub p c) g)) + c. Proof. intros. rewrite <- Zdiv_pol_correct ; auto. rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). unfold eval_pol. ring. (**) apply Zgcd_pol_div ; auto. Qed. Definition makeCuttingPlane (p : PolC Z) : PolC Z * Z := let (g,c) := Zgcd_pol p in if Z.gtb g Z0 then (Zdiv_pol (PsubC Z.sub p c) g , Z.opp (ceiling (Z.opp c) g)) else (p,Z0). Definition genCuttingPlane (f : NFormula Z) : option (PolC Z * Z * Op1) := let (e,op) := f in match op with | Equal => let (g,c) := Zgcd_pol e in if andb (Z.gtb g Z0) (andb (negb (Zeq_bool c Z0)) (negb (Zeq_bool (Z.gcd g c) g))) then None (* inconsistent *) else (* Could be optimised Zgcd_pol is recomputed *) let (p,c) := makeCuttingPlane e in Some (p,c,Equal) | NonEqual => Some (e,Z0,op) | Strict => let (p,c) := makeCuttingPlane (PsubC Z.sub e 1) in Some (p,c,NonStrict) | NonStrict => let (p,c) := makeCuttingPlane e in Some (p,c,NonStrict) end. Definition nformula_of_cutting_plane (t : PolC Z * Z * Op1) : NFormula Z := let (e_z, o) := t in let (e,z) := e_z in (padd e (Pc z) , o). Definition is_pol_Z0 (p : PolC Z) : bool := match p with | Pc Z0 => true | _ => false end. Lemma is_pol_Z0_eval_pol : forall p, is_pol_Z0 p = true -> forall env, eval_pol env p = 0. Proof. unfold is_pol_Z0. destruct p ; try discriminate. destruct z ; try discriminate. reflexivity. Qed. Definition eval_Psatz : list (NFormula Z) -> ZWitness -> option (NFormula Z) := eval_Psatz 0 1 Z.add Z.mul Zeq_bool Z.leb. Definition valid_cut_sign (op:Op1) := match op with | Equal => true | NonStrict => true | _ => false end. Definition bound_var (v : positive) : Formula Z := Build_Formula (PEX v) OpGe (PEc 0). Definition mk_eq_pos (x : positive) (y:positive) (t : positive) : Formula Z := Build_Formula (PEX x) OpEq (PEsub (PEX y) (PEX t)). Fixpoint vars (jmp : positive) (p : Pol Z) : list positive := match p with | Pc c => nil | Pinj j p => vars (Pos.add j jmp) p | PX p j q => jmp::(vars jmp p)++vars (Pos.succ jmp) q end. Fixpoint max_var (jmp : positive) (p : Pol Z) : positive := match p with | Pc _ => jmp | Pinj j p => max_var (Pos.add j jmp) p | PX p j q => Pos.max (max_var jmp p) (max_var (Pos.succ jmp) q) end. Lemma pos_le_add : forall y x, (x <= y + x)%positive. Proof. intros. assert ((Z.pos x) <= Z.pos (x + y))%Z. rewrite <- (Z.add_0_r (Zpos x)). rewrite <- Pos2Z.add_pos_pos. apply Z.add_le_mono_l. compute. congruence. rewrite Pos.add_comm in H. apply H. Qed. Lemma max_var_le : forall p v, (v <= max_var v p)%positive. Proof. induction p; simpl. - intros. apply Pos.le_refl. - intros. specialize (IHp (p+v)%positive). eapply Pos.le_trans ; eauto. assert (xH + v <= p + v)%positive. { apply Pos.add_le_mono. apply Pos.le_1_l. apply Pos.le_refl. } eapply Pos.le_trans ; eauto. apply pos_le_add. - intros. apply Pos.max_case_strong;intros ; auto. specialize (IHp2 (Pos.succ v)%positive). eapply Pos.le_trans ; eauto. Qed. Lemma max_var_correct : forall p j v, In v (vars j p) -> Pos.le v (max_var j p). Proof. induction p; simpl. - tauto. - auto. - intros. rewrite in_app_iff in H. destruct H as [H |[ H | H]]. + subst. apply Pos.max_case_strong;intros ; auto. apply max_var_le. eapply Pos.le_trans ; eauto. apply max_var_le. + apply Pos.max_case_strong;intros ; auto. eapply Pos.le_trans ; eauto. + apply Pos.max_case_strong;intros ; auto. eapply Pos.le_trans ; eauto. Qed. Definition max_var_nformulae (l : list (NFormula Z)) := List.fold_left (fun acc f => Pos.max acc (max_var xH (fst f))) l xH. Section MaxVar. Definition F (acc : positive) (f : Pol Z * Op1) := Pos.max acc (max_var 1 (fst f)). Lemma max_var_nformulae_mono_aux : forall l v acc, (v <= acc -> v <= fold_left F l acc)%positive. Proof. induction l ; simpl ; [easy|]. intros. apply IHl. unfold F. apply Pos.max_case_strong;intros ; auto. eapply Pos.le_trans ; eauto. Qed. Lemma max_var_nformulae_mono_aux' : forall l acc acc', (acc <= acc' -> fold_left F l acc <= fold_left F l acc')%positive. Proof. induction l ; simpl ; [easy|]. intros. apply IHl. unfold F. apply Pos.max_le_compat_r; auto. Qed. Lemma max_var_nformulae_correct_aux : forall l p o v, In (p,o) l -> In v (vars xH p) -> Pos.le v (fold_left F l 1)%positive. Proof. intros. generalize 1%positive as acc. revert p o v H H0. induction l. - simpl. tauto. - simpl. intros. destruct H ; subst. + unfold F at 2. simpl. apply max_var_correct in H0. apply max_var_nformulae_mono_aux. apply Pos.max_case_strong;intros ; auto. eapply Pos.le_trans ; eauto. + eapply IHl ; eauto. Qed. End MaxVar. Lemma max_var_nformalae_correct : forall l p o v, In (p,o) l -> In v (vars xH p) -> Pos.le v (max_var_nformulae l)%positive. Proof. intros l p o v. apply max_var_nformulae_correct_aux. Qed. Fixpoint max_var_psatz (w : Psatz Z) : positive := match w with | PsatzIn _ n => xH | PsatzSquare p => max_var xH (Psquare 0 1 Z.add Z.mul Zeq_bool p) | PsatzMulC p w => Pos.max (max_var xH p) (max_var_psatz w) | PsatzMulE w1 w2 => Pos.max (max_var_psatz w1) (max_var_psatz w2) | PsatzAdd w1 w2 => Pos.max (max_var_psatz w1) (max_var_psatz w2) | _ => xH end. Fixpoint max_var_prf (w : ZArithProof) : positive := match w with | DoneProof => xH | RatProof w pf | CutProof w pf => Pos.max (max_var_psatz w) (max_var_prf pf) | EnumProof w1 w2 l => List.fold_left (fun acc prf => Pos.max acc (max_var_prf prf)) l (Pos.max (max_var_psatz w1) (max_var_psatz w2)) | ExProof _ pf => max_var_prf pf end. Fixpoint ZChecker (l:list (NFormula Z)) (pf : ZArithProof) {struct pf} : bool := match pf with | DoneProof => false | RatProof w pf => match eval_Psatz l w with | None => false | Some f => if Zunsat f then true else ZChecker (f::l) pf end | CutProof w pf => match eval_Psatz l w with | None => false | Some f => match genCuttingPlane f with | None => true | Some cp => ZChecker (nformula_of_cutting_plane cp::l) pf end end | ExProof x prf => let fr := max_var_nformulae l in if Pos.leb x fr then let z := Pos.succ fr in let t := Pos.succ z in let nfx := xnnormalise (mk_eq_pos x z t) in let posz := xnnormalise (bound_var z) in let post := xnnormalise (bound_var t) in ZChecker (nfx::posz::post::l) prf else false | EnumProof w1 w2 pf => match eval_Psatz l w1 , eval_Psatz l w2 with | Some f1 , Some f2 => match genCuttingPlane f1 , genCuttingPlane f2 with |Some (e1,z1,op1) , Some (e2,z2,op2) => if (valid_cut_sign op1 && valid_cut_sign op2 && is_pol_Z0 (padd e1 e2)) then (fix label (pfs:list ZArithProof) := fun lb ub => match pfs with | nil => if Z.gtb lb ub then true else false | pf::rsr => andb (ZChecker ((psub e1 (Pc lb), Equal) :: l) pf) (label rsr (Z.add lb 1%Z) ub) end) pf (Z.opp z1) z2 else false | _ , _ => true end | _ , _ => false end end. Fixpoint bdepth (pf : ZArithProof) : nat := match pf with | DoneProof => O | RatProof _ p => S (bdepth p) | CutProof _ p => S (bdepth p) | EnumProof _ _ l => S (List.fold_right (fun pf x => Nat.max (bdepth pf) x) O l) | ExProof _ p => S (bdepth p) end. Require Import Wf_nat. Lemma in_bdepth : forall l a b y, In y l -> ltof ZArithProof bdepth y (EnumProof a b l). Proof. induction l. (* nil *) simpl. tauto. (* cons *) simpl. intros. destruct H. subst. unfold ltof. simpl. generalize ( (fold_right (fun (pf : ZArithProof) (x : nat) => Nat.max (bdepth pf) x) 0%nat l)). intros. generalize (bdepth y) ; intros. rewrite Nat.lt_succ_r. apply Nat.le_max_l. generalize (IHl a0 b y H). unfold ltof. simpl. generalize ( (fold_right (fun (pf : ZArithProof) (x : nat) => Nat.max (bdepth pf) x) 0%nat l)). intros. eapply lt_le_trans. eassumption. rewrite <- Nat.succ_le_mono. apply Nat.le_max_r. Qed. Lemma eval_Psatz_sound : forall env w l f', make_conj (eval_nformula env) l -> eval_Psatz l w = Some f' -> eval_nformula env f'. Proof. intros. apply (eval_Psatz_Sound Zsor ZSORaddon) with (l:=l) (e:= w) ; auto. apply make_conj_in ; auto. Qed. Lemma makeCuttingPlane_ns_sound : forall env e e' c, eval_nformula env (e, NonStrict) -> makeCuttingPlane e = (e',c) -> eval_nformula env (nformula_of_cutting_plane (e', c, NonStrict)). Proof. unfold nformula_of_cutting_plane. unfold eval_nformula. unfold RingMicromega.eval_nformula. unfold eval_op1. intros. rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). simpl. (**) unfold makeCuttingPlane in H0. revert H0. case_eq (Zgcd_pol e) ; intros g c0. generalize (Zgt_cases g 0) ; destruct (Z.gtb g 0). intros. inv H2. change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in *. apply Zgcd_pol_correct_lt with (env:=env) in H1. 2: auto using Z.gt_lt. apply Z.le_add_le_sub_l, Z.ge_le; rewrite Z.add_0_r. apply (narrow_interval_lower_bound g (- c0) (eval_pol env (Zdiv_pol (PsubC Z.sub e c0) g)) H0). apply Z.le_ge. rewrite <- Z.sub_0_l. apply Z.le_sub_le_add_r. rewrite <- H1. assumption. (* g <= 0 *) intros. inv H2. auto with zarith. Qed. Lemma cutting_plane_sound : forall env f p, eval_nformula env f -> genCuttingPlane f = Some p -> eval_nformula env (nformula_of_cutting_plane p). Proof. unfold genCuttingPlane. destruct f as [e op]. destruct op. (* Equal *) destruct p as [[e' z] op]. case_eq (Zgcd_pol e) ; intros g c. case_eq (Z.gtb g 0 && (negb (Zeq_bool c 0) && negb (Zeq_bool (Z.gcd g c) g))) ; [discriminate|]. case_eq (makeCuttingPlane e). intros. inv H3. unfold makeCuttingPlane in H. rewrite H1 in H. revert H. change (eval_pol env e = 0) in H2. case_eq (Z.gtb g 0). intros. rewrite <- Zgt_is_gt_bool in H. rewrite Zgcd_pol_correct_lt with (1:= H1) in H2. 2: auto using Z.gt_lt. unfold nformula_of_cutting_plane. change (eval_pol env (padd e' (Pc z)) = 0). inv H3. rewrite eval_pol_add. set (x:=eval_pol env (Zdiv_pol (PsubC Z.sub e c) g)) in *; clearbody x. simpl. rewrite andb_false_iff in H0. destruct H0. rewrite Zgt_is_gt_bool in H ; congruence. rewrite andb_false_iff in H0. destruct H0. rewrite negb_false_iff in H0. apply Zeq_bool_eq in H0. subst. simpl. rewrite Z.add_0_r, Z.mul_eq_0 in H2. intuition subst; easy. rewrite negb_false_iff in H0. apply Zeq_bool_eq in H0. assert (HH := Zgcd_is_gcd g c). rewrite H0 in HH. inv HH. apply Zdivide_opp_r in H4. rewrite Zdivide_ceiling ; auto. apply Z.sub_move_0_r. apply Z.div_unique_exact. now intros ->. now rewrite Z.add_move_0_r in H2. intros. unfold nformula_of_cutting_plane. inv H3. change (eval_pol env (padd e' (Pc 0)) = 0). rewrite eval_pol_add. simpl. now rewrite Z.add_0_r. (* NonEqual *) intros. inv H0. unfold eval_nformula in *. unfold RingMicromega.eval_nformula in *. unfold nformula_of_cutting_plane. unfold eval_op1 in *. rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). simpl. now rewrite Z.add_0_r. (* Strict *) destruct p as [[e' z] op]. case_eq (makeCuttingPlane (PsubC Z.sub e 1)). intros. inv H1. apply makeCuttingPlane_ns_sound with (env:=env) (2:= H). simpl in *. rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). now apply Z.lt_le_pred. (* NonStrict *) destruct p as [[e' z] op]. case_eq (makeCuttingPlane e). intros. inv H1. apply makeCuttingPlane_ns_sound with (env:=env) (2:= H). assumption. Qed. Lemma genCuttingPlaneNone : forall env f, genCuttingPlane f = None -> eval_nformula env f -> False. Proof. unfold genCuttingPlane. destruct f. destruct o. case_eq (Zgcd_pol p) ; intros g c. case_eq (Z.gtb g 0 && (negb (Zeq_bool c 0) && negb (Zeq_bool (Z.gcd g c) g))). intros. flatten_bool. rewrite negb_true_iff in H5. apply Zeq_bool_neq in H5. rewrite <- Zgt_is_gt_bool in H3. rewrite negb_true_iff in H. apply Zeq_bool_neq in H. change (eval_pol env p = 0) in H2. rewrite Zgcd_pol_correct_lt with (1:= H0) in H2. 2: auto using Z.gt_lt. set (x:=eval_pol env (Zdiv_pol (PsubC Z.sub p c) g)) in *; clearbody x. contradict H5. apply Zis_gcd_gcd. apply Z.lt_le_incl, Z.gt_lt; assumption. constructor; auto with zarith. exists (-x). rewrite Z.mul_opp_l, Z.mul_comm. now apply Z.add_move_0_l. (**) destruct (makeCuttingPlane p); discriminate. discriminate. destruct (makeCuttingPlane (PsubC Z.sub p 1)) ; discriminate. destruct (makeCuttingPlane p) ; discriminate. Qed. Lemma eval_nformula_mk_eq_pos : forall env x z t, env x = env z - env t -> eval_nformula env (xnnormalise (mk_eq_pos x z t)). Proof. intros. rewrite xnnormalise_correct. simpl. auto. Qed. Lemma eval_nformula_bound_var : forall env x, env x >= 0 -> eval_nformula env (xnnormalise (bound_var x)). Proof. intros. rewrite xnnormalise_correct. simpl. auto. Qed. Definition agree_env (fr : positive) (env env' : positive -> Z) : Prop := forall x, Pos.le x fr -> env x = env' x. Lemma agree_env_subset : forall v1 v2 env env', agree_env v1 env env' -> Pos.le v2 v1 -> agree_env v2 env env'. Proof. unfold agree_env. intros. apply H. eapply Pos.le_trans ; eauto. Qed. Lemma agree_env_jump : forall fr j env env', agree_env (fr + j) env env' -> agree_env fr (Env.jump j env) (Env.jump j env'). Proof. intros. unfold agree_env ; intro. intros. unfold Env.jump. apply H. apply Pos.add_le_mono_r; auto. Qed. Lemma agree_env_tail : forall fr env env', agree_env (Pos.succ fr) env env' -> agree_env fr (Env.tail env) (Env.tail env'). Proof. intros. unfold Env.tail. apply agree_env_jump. rewrite <- Pos.add_1_r in H. apply H. Qed. Lemma max_var_acc : forall p i j, (max_var (i + j) p = max_var i p + j)%positive. Proof. induction p; simpl. - reflexivity. - intros. rewrite ! IHp. rewrite Pos.add_assoc. reflexivity. - intros. rewrite !Pplus_one_succ_l. rewrite ! IHp1. rewrite ! IHp2. rewrite ! Pos.add_assoc. rewrite <- Pos.add_max_distr_r. reflexivity. Qed. Lemma agree_env_eval_nformula : forall env env' e (AGREE : agree_env (max_var xH (fst e)) env env'), eval_nformula env e <-> eval_nformula env' e. Proof. destruct e. simpl; intros. assert ((RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x) env p) = (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x) env' p)). { revert env env' AGREE. generalize xH. induction p ; simpl. - reflexivity. - intros. apply IHp with (p := p1%positive). apply agree_env_jump. eapply agree_env_subset; eauto. rewrite (Pos.add_comm p). rewrite max_var_acc. apply Pos.le_refl. - intros. f_equal. f_equal. { apply IHp1 with (p:= p). eapply agree_env_subset; eauto. apply Pos.le_max_l. } f_equal. { unfold Env.hd. unfold Env.nth. apply AGREE. apply Pos.le_1_l. } { apply IHp2 with (p := p). apply agree_env_tail. eapply agree_env_subset; eauto. rewrite !Pplus_one_succ_r. rewrite max_var_acc. apply Pos.le_max_r. } } rewrite H. tauto. Qed. Lemma agree_env_eval_nformulae : forall env env' l (AGREE : agree_env (max_var_nformulae l) env env'), make_conj (eval_nformula env) l <-> make_conj (eval_nformula env') l. Proof. induction l. - simpl. tauto. - intros. rewrite ! make_conj_cons. assert (eval_nformula env a <-> eval_nformula env' a). { apply agree_env_eval_nformula. eapply agree_env_subset ; eauto. unfold max_var_nformulae. simpl. rewrite Pos.max_1_l. apply max_var_nformulae_mono_aux. apply Pos.le_refl. } rewrite H. apply and_iff_compat_l. apply IHl. eapply agree_env_subset ; eauto. unfold max_var_nformulae. simpl. apply max_var_nformulae_mono_aux'. apply Pos.le_1_l. Qed. Lemma eq_true_iff_eq : forall b1 b2 : bool, (b1 = true <-> b2 = true) <-> b1 = b2. Proof. destruct b1,b2 ; intuition congruence. Qed. Ltac pos_tac := repeat match goal with | |- false = _ => symmetry | |- Pos.eqb ?X ?Y = false => rewrite Pos.eqb_neq ; intro | H : @eq positive ?X ?Y |- _ => apply Zpos_eq in H | H : context[Z.pos (Pos.succ ?X)] |- _ => rewrite (Pos2Z.inj_succ X) in H | H : Pos.leb ?X ?Y = true |- _ => rewrite Pos.leb_le in H ; apply (Pos2Z.pos_le_pos X Y) in H end. Lemma ZChecker_sound : forall w l, ZChecker l w = true -> forall env, make_impl (eval_nformula env) l False. Proof. induction w using (well_founded_ind (well_founded_ltof _ bdepth)). destruct w as [ | w pf | w pf | w1 w2 pf | x pf]. - (* DoneProof *) simpl. discriminate. - (* RatProof *) simpl. intros l. case_eq (eval_Psatz l w) ; [| discriminate]. intros f Hf. case_eq (Zunsat f). intros. apply (checker_nf_sound Zsor ZSORaddon l w). unfold check_normalised_formulas. unfold eval_Psatz in Hf. rewrite Hf. unfold Zunsat in H0. assumption. intros. assert (make_impl (eval_nformula env) (f::l) False). apply H with (2:= H1). unfold ltof. simpl. auto with arith. destruct f. rewrite <- make_conj_impl in H2. rewrite make_conj_cons in H2. rewrite <- make_conj_impl. intro. apply H2. split ; auto. apply eval_Psatz_sound with (2:= Hf) ; assumption. - (* CutProof *) simpl. intros l. case_eq (eval_Psatz l w) ; [ | discriminate]. intros f' Hlc. case_eq (genCuttingPlane f'). intros. assert (make_impl (eval_nformula env) (nformula_of_cutting_plane p::l) False). eapply (H pf) ; auto. unfold ltof. simpl. auto with arith. rewrite <- make_conj_impl in H2. rewrite make_conj_cons in H2. rewrite <- make_conj_impl. intro. apply H2. split ; auto. apply eval_Psatz_sound with (env:=env) in Hlc. apply cutting_plane_sound with (1:= Hlc) (2:= H0). auto. (* genCuttingPlane = None *) intros. rewrite <- make_conj_impl. intros. apply eval_Psatz_sound with (2:= Hlc) in H2. apply genCuttingPlaneNone with (2:= H2) ; auto. - (* EnumProof *) intros l. simpl. case_eq (eval_Psatz l w1) ; [ | discriminate]. case_eq (eval_Psatz l w2) ; [ | discriminate]. intros f1 Hf1 f2 Hf2. case_eq (genCuttingPlane f2). destruct p as [ [p1 z1] op1]. case_eq (genCuttingPlane f1). destruct p as [ [p2 z2] op2]. case_eq (valid_cut_sign op1 && valid_cut_sign op2 && is_pol_Z0 (padd p1 p2)). intros Hcond. flatten_bool. rename H1 into HZ0. rename H2 into Hop1. rename H3 into Hop2. intros HCutL HCutR Hfix env. (* get the bounds of the enum *) rewrite <- make_conj_impl. intro. assert (-z1 <= eval_pol env p1 <= z2). split. apply eval_Psatz_sound with (env:=env) in Hf2 ; auto. apply cutting_plane_sound with (1:= Hf2) in HCutR. unfold nformula_of_cutting_plane in HCutR. unfold eval_nformula in HCutR. unfold RingMicromega.eval_nformula in HCutR. change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in HCutR. unfold eval_op1 in HCutR. destruct op1 ; simpl in Hop1 ; try discriminate; rewrite eval_pol_add in HCutR; simpl in HCutR. rewrite Z.add_move_0_l in HCutR; rewrite HCutR, Z.opp_involutive; reflexivity. now apply Z.le_sub_le_add_r in HCutR. (**) apply is_pol_Z0_eval_pol with (env := env) in HZ0. rewrite eval_pol_add, Z.add_move_r, Z.sub_0_l in HZ0. rewrite HZ0. apply eval_Psatz_sound with (env:=env) in Hf1 ; auto. apply cutting_plane_sound with (1:= Hf1) in HCutL. unfold nformula_of_cutting_plane in HCutL. unfold eval_nformula in HCutL. unfold RingMicromega.eval_nformula in HCutL. change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in HCutL. unfold eval_op1 in HCutL. rewrite eval_pol_add in HCutL. simpl in HCutL. destruct op2 ; simpl in Hop2 ; try discriminate. rewrite Z.add_move_r, Z.sub_0_l in HCutL. now rewrite HCutL, Z.opp_involutive. now rewrite <- Z.le_sub_le_add_l in HCutL. revert Hfix. match goal with | |- context[?F pf (-z1) z2 = true] => set (FF := F) end. intros. assert (HH :forall x, -z1 <= x <= z2 -> exists pr, (In pr pf /\ ZChecker ((PsubC Z.sub p1 x,Equal) :: l) pr = true)%Z). clear HZ0 Hop1 Hop2 HCutL HCutR H0 H1. revert Hfix. generalize (-z1). clear z1. intro z1. revert z1 z2. induction pf;simpl ;intros. revert Hfix. now case (Z.gtb_spec); [ | easy ]; intros LT; elim (Zlt_not_le _ _ LT); transitivity x. flatten_bool. destruct (Z_le_lt_eq_dec _ _ (proj1 H0)) as [ LT | -> ]. 2: exists a; auto. rewrite <- Z.le_succ_l in LT. assert (LE: (Z.succ z1 <= x <= z2)%Z) by intuition. elim IHpf with (2:=H2) (3:= LE). intros. exists x0 ; split;tauto. intros until 1. apply H ; auto. unfold ltof in *. simpl in *. PreOmega.zify. intuition subst. assumption. eapply Z.lt_le_trans. eassumption. apply Z.add_le_mono_r. assumption. (*/asser *) destruct (HH _ H1) as [pr [Hin Hcheker]]. assert (make_impl (eval_nformula env) ((PsubC Z.sub p1 (eval_pol env p1),Equal) :: l) False). eapply (H pr) ;auto. apply in_bdepth ; auto. rewrite <- make_conj_impl in H2. apply H2. rewrite make_conj_cons. split ;auto. unfold eval_nformula. unfold RingMicromega.eval_nformula. simpl. rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). unfold eval_pol. ring. discriminate. (* No cutting plane *) intros. rewrite <- make_conj_impl. intros. apply eval_Psatz_sound with (2:= Hf1) in H3. apply genCuttingPlaneNone with (2:= H3) ; auto. (* No Cutting plane (bis) *) intros. rewrite <- make_conj_impl. intros. apply eval_Psatz_sound with (2:= Hf2) in H2. apply genCuttingPlaneNone with (2:= H2) ; auto. - intros l. unfold ZChecker. fold ZChecker. set (fr := (max_var_nformulae l)%positive). set (z1 := (Pos.succ fr)) in *. set (t1 := (Pos.succ z1)) in *. destruct (x <=? fr)%positive eqn:LE ; [|congruence]. intros. set (env':= fun v => if Pos.eqb v z1 then if Z.leb (env x) 0 then 0 else env x else if Pos.eqb v t1 then if Z.leb (env x) 0 then -(env x) else 0 else env v). apply H with (env:=env') in H0. + rewrite <- make_conj_impl in *. intro. rewrite !make_conj_cons in H0. apply H0 ; repeat split. * apply eval_nformula_mk_eq_pos. unfold env'. rewrite! Pos.eqb_refl. replace (x=?z1)%positive with false. replace (x=?t1)%positive with false. replace (t1=?z1)%positive with false. destruct (env x <=? 0); ring. { unfold t1. pos_tac; normZ. lia (Hyp H2). } { unfold t1, z1. pos_tac; normZ. lia (Add (Hyp LE) (Hyp H3)). } { unfold z1. pos_tac; normZ. lia (Add (Hyp LE) (Hyp H3)). } * apply eval_nformula_bound_var. unfold env'. rewrite! Pos.eqb_refl. destruct (env x <=? 0) eqn:EQ. compute. congruence. rewrite Z.leb_gt in EQ. normZ. lia (Add (Hyp EQ) (Hyp H2)). * apply eval_nformula_bound_var. unfold env'. rewrite! Pos.eqb_refl. replace (t1 =? z1)%positive with false. destruct (env x <=? 0) eqn:EQ. rewrite Z.leb_le in EQ. normZ. lia (Add (Hyp EQ) (Hyp H2)). compute; congruence. unfold t1. clear. pos_tac; normZ. lia (Hyp H). * rewrite agree_env_eval_nformulae with (env':= env') in H1;auto. unfold agree_env; intros. unfold env'. replace (x0 =? z1)%positive with false. replace (x0 =? t1)%positive with false. reflexivity. { unfold t1, z1. unfold fr in *. apply Pos2Z.pos_le_pos in H2. pos_tac; normZ. lia (Add (Hyp H2) (Hyp H4)). } { unfold z1, fr in *. apply Pos2Z.pos_le_pos in H2. pos_tac; normZ. lia (Add (Hyp H2) (Hyp H4)). } + unfold ltof. simpl. apply Nat.lt_succ_diag_r. Qed. Definition ZTautoChecker (f : BFormula (Formula Z)) (w: list ZArithProof): bool := @tauto_checker (Formula Z) (NFormula Z) unit Zunsat Zdeduce normalise negate ZArithProof (fun cl => ZChecker (List.map fst cl)) f w. Lemma ZTautoChecker_sound : forall f w, ZTautoChecker f w = true -> forall env, eval_bf (Zeval_formula env) f. Proof. intros f w. unfold ZTautoChecker. apply tauto_checker_sound with (eval' := eval_nformula). - apply Zeval_nformula_dec. - intros until env. unfold eval_nformula. unfold RingMicromega.eval_nformula. destruct t. apply (check_inconsistent_sound Zsor ZSORaddon) ; auto. - unfold Zdeduce. intros. revert H. apply (nformula_plus_nformula_correct Zsor ZSORaddon); auto. - intros. rewrite normalise_correct in H. auto. - intros. rewrite negate_correct in H ; auto. - intros t w0. unfold eval_tt. intros. rewrite make_impl_map with (eval := eval_nformula env). eapply ZChecker_sound; eauto. tauto. Qed. Fixpoint xhyps_of_pt (base:nat) (acc : list nat) (pt:ZArithProof) : list nat := match pt with | DoneProof => acc | RatProof c pt => xhyps_of_pt (S base ) (xhyps_of_psatz base acc c) pt | CutProof c pt => xhyps_of_pt (S base ) (xhyps_of_psatz base acc c) pt | EnumProof c1 c2 l => let acc := xhyps_of_psatz base (xhyps_of_psatz base acc c2) c1 in List.fold_left (xhyps_of_pt (S base)) l acc | ExProof _ pt => xhyps_of_pt (S (S (S base ))) acc pt end. Definition hyps_of_pt (pt : ZArithProof) : list nat := xhyps_of_pt 0 nil pt. Open Scope Z_scope. (** To ease bindings from ml code **) Definition make_impl := Refl.make_impl. Definition make_conj := Refl.make_conj. Require VarMap. (*Definition varmap_type := VarMap.t Z. *) Definition env := PolEnv Z. Definition node := @VarMap.Branch Z. Definition empty := @VarMap.Empty Z. Definition leaf := @VarMap.Elt Z. Definition coneMember := ZWitness. Definition eval := eval_formula. Definition prod_pos_nat := prod positive nat. Notation n_of_Z := Z.to_N (only parsing). (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/mfourier.mli0000644000175000017500000000216113612664533020245 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (Vect.t, proof) Util.union val optimise : Vect.t -> Polynomial.cstr list -> Itv.interval option end val pp_proof : out_channel -> proof -> unit module Proof : sig val mk_proof : Polynomial.cstr list -> proof -> (Vect.t * Polynomial.cstr) list val add_op : Polynomial.op -> Polynomial.op -> Polynomial.op end exception TimeOut coq-8.11.0/plugins/micromega/Ztac.v0000644000175000017500000000702713612664533017010 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* x <= y /\ y <= x. Proof. intros; split; apply Z.eq_le_incl; auto. Qed. Lemma elim_concl_eq : forall x y, (x < y \/ y < x -> False) -> x = y. Proof. intros. destruct (Z_lt_le_dec x y). exfalso. apply H ; auto. destruct (Zle_lt_or_eq y x);auto. exfalso. apply H ; auto. Qed. Lemma elim_concl_le : forall x y, (y < x -> False) -> x <= y. Proof. intros. destruct (Z_lt_le_dec y x). exfalso ; auto. auto. Qed. Lemma elim_concl_lt : forall x y, (y <= x -> False) -> x < y. Proof. intros. destruct (Z_lt_le_dec x y). auto. exfalso ; auto. Qed. Lemma Zlt_le_add_1 : forall n m : Z, n < m -> n + 1 <= m. Proof. exact (Zlt_le_succ). Qed. Ltac normZ := repeat match goal with | H : _ < _ |- _ => apply Zlt_le_add_1 in H | H : ?Y <= _ |- _ => lazymatch Y with | 0 => fail | _ => apply Zle_minus_le_0 in H end | H : _ >= _ |- _ => apply Z.ge_le in H | H : _ > _ |- _ => apply Z.gt_lt in H | H : _ = _ |- _ => apply eq_incl in H ; destruct H | |- @eq Z _ _ => apply elim_concl_eq ; let H := fresh "HZ" in intros [H|H] | |- _ <= _ => apply elim_concl_le ; intros | |- _ < _ => apply elim_concl_lt ; intros | |- _ >= _ => apply Z.le_ge end. Inductive proof := | Hyp (e : Z) (prf : 0 <= e) | Add (p1 p2: proof) | Mul (p1 p2: proof) | Cst (c : Z) . Lemma add_le : forall e1 e2, 0 <= e1 -> 0 <= e2 -> 0 <= e1+e2. Proof. intros. change 0 with (0+ 0). apply Z.add_le_mono; auto. Qed. Lemma mul_le : forall e1 e2, 0 <= e1 -> 0 <= e2 -> 0 <= e1*e2. Proof. intros. change 0 with (0* e2). apply Zmult_le_compat_r; auto. Qed. Fixpoint eval_proof (p : proof) : { e : Z | 0 <= e} := match p with | Hyp e prf => exist _ e prf | Add p1 p2 => let (e1,p1) := eval_proof p1 in let (e2,p2) := eval_proof p2 in exist _ _ (add_le _ _ p1 p2) | Mul p1 p2 => let (e1,p1) := eval_proof p1 in let (e2,p2) := eval_proof p2 in exist _ _ (mul_le _ _ p1 p2) | Cst c => match Z_le_dec 0 c with | left prf => exist _ _ prf | _ => exist _ _ Z.le_0_1 end end. Ltac lia_step p := let H := fresh in let prf := (eval cbn - [Z.le Z.mul Z.opp Z.sub Z.add] in (eval_proof p)) in match prf with | @exist _ _ _ ?P => pose proof P as H end ; ring_simplify in H. Ltac lia_contr := match goal with | H : 0 <= - (Zpos _) |- _ => rewrite <- Z.leb_le in H; compute in H ; discriminate | H : 0 <= (Zneg _) |- _ => rewrite <- Z.leb_le in H; compute in H ; discriminate end. Ltac lia p := lia_step p ; lia_contr. Ltac slia H1 H2 := normZ ; lia (Add (Hyp _ H1) (Hyp _ H2)). Arguments Hyp {_} prf. coq-8.11.0/plugins/micromega/RingMicromega.v0000644000175000017500000010231613612664533020627 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* R -> R. Variable ropp : R -> R. Variables req rle rlt : R -> R -> Prop. Variable sor : SOR rO rI rplus rtimes rminus ropp req rle rlt. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (rplus x y). Notation "x * y " := (rtimes x y). Notation "x - y " := (rminus x y). Notation "- x" := (ropp x). Notation "x == y" := (req x y). Notation "x ~= y" := (~ req x y). Notation "x <= y" := (rle x y). Notation "x < y" := (rlt x y). (* Assume we have a type of coefficients C and a morphism from C to R *) Variable C : Type. Variables cO cI : C. Variables cplus ctimes cminus: C -> C -> C. Variable copp : C -> C. Variables ceqb cleb : C -> C -> bool. Variable phi : C -> R. (* Power coefficients *) Variable E : Type. (* the type of exponents *) Variable pow_phi : N -> E. Variable rpow : R -> E -> R. Notation "[ x ]" := (phi x). Notation "x [=] y" := (ceqb x y). Notation "x [<=] y" := (cleb x y). (* Let's collect all hypotheses in addition to the ordered ring axioms into one structure *) Record SORaddon := mk_SOR_addon { SORrm : ring_morph 0 1 rplus rtimes rminus ropp req cO cI cplus ctimes cminus copp ceqb phi; SORpower : power_theory rI rtimes req pow_phi rpow; SORcneqb_morph : forall x y : C, x [=] y = false -> [x] ~= [y]; SORcleb_morph : forall x y : C, x [<=] y = true -> [x] <= [y] }. Variable addon : SORaddon. Add Relation R req reflexivity proved by (@Equivalence_Reflexive _ _ (SORsetoid sor)) symmetry proved by (@Equivalence_Symmetric _ _ (SORsetoid sor)) transitivity proved by (@Equivalence_Transitive _ _ (SORsetoid sor)) as micomega_sor_setoid. Add Morphism rplus with signature req ==> req ==> req as rplus_morph. Proof. exact (SORplus_wd sor). Qed. Add Morphism rtimes with signature req ==> req ==> req as rtimes_morph. Proof. exact (SORtimes_wd sor). Qed. Add Morphism ropp with signature req ==> req as ropp_morph. Proof. exact (SORopp_wd sor). Qed. Add Morphism rle with signature req ==> req ==> iff as rle_morph. Proof. exact (SORle_wd sor). Qed. Add Morphism rlt with signature req ==> req ==> iff as rlt_morph. Proof. exact (SORlt_wd sor). Qed. Add Morphism rminus with signature req ==> req ==> req as rminus_morph. Proof. exact (rminus_morph sor). (* We already proved that minus is a morphism in OrderedRing.v *) Qed. Definition cneqb (x y : C) := negb (ceqb x y). Definition cltb (x y : C) := (cleb x y) && (cneqb x y). Notation "x [~=] y" := (cneqb x y). Notation "x [<] y" := (cltb x y). Ltac le_less := rewrite (Rle_lt_eq sor); left; try assumption. Ltac le_equal := rewrite (Rle_lt_eq sor); right; try reflexivity; try assumption. Ltac le_elim H := rewrite (Rle_lt_eq sor) in H; destruct H as [H | H]. Lemma cleb_sound : forall x y : C, x [<=] y = true -> [x] <= [y]. Proof. exact (SORcleb_morph addon). Qed. Lemma cneqb_sound : forall x y : C, x [~=] y = true -> [x] ~= [y]. Proof. intros x y H1. apply (SORcneqb_morph addon). unfold cneqb, negb in H1. destruct (ceqb x y); now try discriminate. Qed. Lemma cltb_sound : forall x y : C, x [<] y = true -> [x] < [y]. Proof. intros x y H. unfold cltb in H. apply andb_prop in H. destruct H as [H1 H2]. apply cleb_sound in H1. apply cneqb_sound in H2. apply <- (Rlt_le_neq sor). now split. Qed. (* Begin Micromega *) Definition PolC := Pol C. (* polynomials in generalized Horner form, defined in Ring_polynom or EnvRing *) Definition PolEnv := Env R. (* For interpreting PolC *) Definition eval_pol : PolEnv -> PolC -> R := Pphi rplus rtimes phi. Inductive Op1 : Set := (* relations with 0 *) | Equal (* == 0 *) | NonEqual (* ~= 0 *) | Strict (* > 0 *) | NonStrict (* >= 0 *). Definition NFormula := (PolC * Op1)%type. (* normalized formula *) Definition eval_op1 (o : Op1) : R -> Prop := match o with | Equal => fun x => x == 0 | NonEqual => fun x : R => x ~= 0 | Strict => fun x : R => 0 < x | NonStrict => fun x : R => 0 <= x end. Definition eval_nformula (env : PolEnv) (f : NFormula) : Prop := let (p, op) := f in eval_op1 op (eval_pol env p). (** Rule of "signs" for addition and multiplication. An arbitrary result is coded buy None. *) Definition OpMult (o o' : Op1) : option Op1 := match o with | Equal => Some Equal | NonStrict => match o' with | Equal => Some Equal | NonEqual => None | Strict => Some NonStrict | NonStrict => Some NonStrict end | Strict => match o' with | NonEqual => None | _ => Some o' end | NonEqual => match o' with | Equal => Some Equal | NonEqual => Some NonEqual | _ => None end end. Definition OpAdd (o o': Op1) : option Op1 := match o with | Equal => Some o' | NonStrict => match o' with | Strict => Some Strict | NonEqual => None | _ => Some NonStrict end | Strict => match o' with | NonEqual => None | _ => Some Strict end | NonEqual => match o' with | Equal => Some NonEqual | _ => None end end. Lemma OpMult_sound : forall (o o' om: Op1) (x y : R), eval_op1 o x -> eval_op1 o' y -> OpMult o o' = Some om -> eval_op1 om (x * y). Proof. unfold eval_op1; destruct o; simpl; intros o' om x y H1 H2 H3. (* x == 0 *) inversion H3. rewrite H1. now rewrite (Rtimes_0_l sor). (* x ~= 0 *) destruct o' ; inversion H3. (* y == 0 *) rewrite H2. now rewrite (Rtimes_0_r sor). (* y ~= 0 *) apply (Rtimes_neq_0 sor) ; auto. (* 0 < x *) destruct o' ; inversion H3. (* y == 0 *) rewrite H2; now rewrite (Rtimes_0_r sor). (* 0 < y *) now apply (Rtimes_pos_pos sor). (* 0 <= y *) apply (Rtimes_nonneg_nonneg sor); [le_less | assumption]. (* 0 <= x *) destruct o' ; inversion H3. (* y == 0 *) rewrite H2; now rewrite (Rtimes_0_r sor). (* 0 < y *) apply (Rtimes_nonneg_nonneg sor); [assumption | le_less ]. (* 0 <= y *) now apply (Rtimes_nonneg_nonneg sor). Qed. Lemma OpAdd_sound : forall (o o' oa : Op1) (e e' : R), eval_op1 o e -> eval_op1 o' e' -> OpAdd o o' = Some oa -> eval_op1 oa (e + e'). Proof. unfold eval_op1; destruct o; simpl; intros o' oa e e' H1 H2 Hoa. (* e == 0 *) inversion Hoa. rewrite <- H0. destruct o' ; rewrite H1 ; now rewrite (Rplus_0_l sor). (* e ~= 0 *) destruct o'. (* e' == 0 *) inversion Hoa. rewrite H2. now rewrite (Rplus_0_r sor). (* e' ~= 0 *) discriminate. (* 0 < e' *) discriminate. (* 0 <= e' *) discriminate. (* 0 < e *) destruct o'. (* e' == 0 *) inversion Hoa. rewrite H2. now rewrite (Rplus_0_r sor). (* e' ~= 0 *) discriminate. (* 0 < e' *) inversion Hoa. now apply (Rplus_pos_pos sor). (* 0 <= e' *) inversion Hoa. now apply (Rplus_pos_nonneg sor). (* 0 <= e *) destruct o'. (* e' == 0 *) inversion Hoa. now rewrite H2, (Rplus_0_r sor). (* e' ~= 0 *) discriminate. (* 0 < e' *) inversion Hoa. now apply (Rplus_nonneg_pos sor). (* 0 <= e' *) inversion Hoa. now apply (Rplus_nonneg_nonneg sor). Qed. Inductive Psatz : Type := | PsatzIn : nat -> Psatz | PsatzSquare : PolC -> Psatz | PsatzMulC : PolC -> Psatz -> Psatz | PsatzMulE : Psatz -> Psatz -> Psatz | PsatzAdd : Psatz -> Psatz -> Psatz | PsatzC : C -> Psatz | PsatzZ : Psatz. (** Given a list [l] of NFormula and an extended polynomial expression [e], if [eval_Psatz l e] succeeds (= Some f) then [f] is a logic consequence of the conjunction of the formulae in l. Moreover, the polynomial expression is obtained by replacing the (PsatzIn n) by the nth polynomial expression in [l] and the sign is computed by the "rule of sign" *) (* Might be defined elsewhere *) Definition map_option (A B:Type) (f : A -> option B) (o : option A) : option B := match o with | None => None | Some x => f x end. Arguments map_option [A B] f o. Definition map_option2 (A B C : Type) (f : A -> B -> option C) (o: option A) (o': option B) : option C := match o , o' with | None , _ => None | _ , None => None | Some x , Some x' => f x x' end. Arguments map_option2 [A B C] f o o'. Definition Rops_wd := mk_reqe (*rplus rtimes ropp req*) (SORplus_wd sor) (SORtimes_wd sor) (SORopp_wd sor). Definition pexpr_times_nformula (e: PolC) (f : NFormula) : option NFormula := let (ef,o) := f in match o with | Equal => Some (Pmul cO cI cplus ctimes ceqb e ef , Equal) | _ => None end. Definition nformula_times_nformula (f1 f2 : NFormula) : option NFormula := let (e1,o1) := f1 in let (e2,o2) := f2 in map_option (fun x => (Some (Pmul cO cI cplus ctimes ceqb e1 e2,x))) (OpMult o1 o2). Definition nformula_plus_nformula (f1 f2 : NFormula) : option NFormula := let (e1,o1) := f1 in let (e2,o2) := f2 in map_option (fun x => (Some (Padd cO cplus ceqb e1 e2,x))) (OpAdd o1 o2). Fixpoint eval_Psatz (l : list NFormula) (e : Psatz) {struct e} : option NFormula := match e with | PsatzIn n => Some (nth n l (Pc cO, Equal)) | PsatzSquare e => Some (Psquare cO cI cplus ctimes ceqb e , NonStrict) | PsatzMulC re e => map_option (pexpr_times_nformula re) (eval_Psatz l e) | PsatzMulE f1 f2 => map_option2 nformula_times_nformula (eval_Psatz l f1) (eval_Psatz l f2) | PsatzAdd f1 f2 => map_option2 nformula_plus_nformula (eval_Psatz l f1) (eval_Psatz l f2) | PsatzC c => if cltb cO c then Some (Pc c, Strict) else None (* This could be 0, or <> 0 -- but these cases are useless *) | PsatzZ => Some (Pc cO, Equal) (* Just to make life easier *) end. Lemma pexpr_times_nformula_correct : forall (env: PolEnv) (e: PolC) (f f' : NFormula), eval_nformula env f -> pexpr_times_nformula e f = Some f' -> eval_nformula env f'. Proof. unfold pexpr_times_nformula. destruct f. intros. destruct o ; inversion H0 ; try discriminate. simpl in *. unfold eval_pol in *. rewrite (Pmul_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)). rewrite H. apply (Rtimes_0_r sor). Qed. Lemma nformula_times_nformula_correct : forall (env:PolEnv) (f1 f2 f : NFormula), eval_nformula env f1 -> eval_nformula env f2 -> nformula_times_nformula f1 f2 = Some f -> eval_nformula env f. Proof. unfold nformula_times_nformula. destruct f1 ; destruct f2. case_eq (OpMult o o0) ; simpl ; try discriminate. intros. inversion H2 ; simpl. unfold eval_pol. destruct o1; simpl; rewrite (Pmul_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)); apply OpMult_sound with (3:= H);assumption. Qed. Lemma nformula_plus_nformula_correct : forall (env:PolEnv) (f1 f2 f : NFormula), eval_nformula env f1 -> eval_nformula env f2 -> nformula_plus_nformula f1 f2 = Some f -> eval_nformula env f. Proof. unfold nformula_plus_nformula. destruct f1 ; destruct f2. case_eq (OpAdd o o0) ; simpl ; try discriminate. intros. inversion H2 ; simpl. unfold eval_pol. destruct o1; simpl; rewrite (Padd_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)); apply OpAdd_sound with (3:= H);assumption. Qed. Lemma eval_Psatz_Sound : forall (l : list NFormula) (env : PolEnv), (forall (f : NFormula), In f l -> eval_nformula env f) -> forall (e : Psatz) (f : NFormula), eval_Psatz l e = Some f -> eval_nformula env f. Proof. induction e. (* PsatzIn *) simpl ; intros. destruct (nth_in_or_default n l (Pc cO, Equal)) as [Hin|Heq]. (* index is in bounds *) apply H. congruence. (* index is out-of-bounds *) inversion H0. rewrite Heq. simpl. now apply (morph0 (SORrm addon)). (* PsatzSquare *) simpl. intros. inversion H0. simpl. unfold eval_pol. rewrite (Psquare_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)); now apply (Rtimes_square_nonneg sor). (* PsatzMulC *) simpl. intro. case_eq (eval_Psatz l e) ; simpl ; intros. apply IHe in H0. apply pexpr_times_nformula_correct with (1:=H0) (2:= H1). discriminate. (* PsatzMulC *) simpl ; intro. case_eq (eval_Psatz l e1) ; simpl ; try discriminate. case_eq (eval_Psatz l e2) ; simpl ; try discriminate. intros. apply IHe1 in H1. apply IHe2 in H0. apply (nformula_times_nformula_correct env n0 n) ; assumption. (* PsatzAdd *) simpl ; intro. case_eq (eval_Psatz l e1) ; simpl ; try discriminate. case_eq (eval_Psatz l e2) ; simpl ; try discriminate. intros. apply IHe1 in H1. apply IHe2 in H0. apply (nformula_plus_nformula_correct env n0 n) ; assumption. (* PsatzC *) simpl. intro. case_eq (cO [<] c). intros. inversion H1. simpl. rewrite <- (morph0 (SORrm addon)). now apply cltb_sound. discriminate. (* PsatzZ *) simpl. intros. inversion H0. simpl. apply (morph0 (SORrm addon)). Qed. Fixpoint ge_bool (n m : nat) : bool := match n with | O => match m with | O => true | S _ => false end | S n => match m with | O => true | S m => ge_bool n m end end. Lemma ge_bool_cases : forall n m, (if ge_bool n m then n >= m else n < m)%nat. Proof. induction n; destruct m ; simpl; auto with arith. specialize (IHn m). destruct (ge_bool); auto with arith. Qed. Fixpoint xhyps_of_psatz (base:nat) (acc : list nat) (prf : Psatz) : list nat := match prf with | PsatzC _ | PsatzZ | PsatzSquare _ => acc | PsatzMulC _ prf => xhyps_of_psatz base acc prf | PsatzAdd e1 e2 | PsatzMulE e1 e2 => xhyps_of_psatz base (xhyps_of_psatz base acc e2) e1 | PsatzIn n => if ge_bool n base then (n::acc) else acc end. Fixpoint nhyps_of_psatz (prf : Psatz) : list nat := match prf with | PsatzC _ | PsatzZ | PsatzSquare _ => nil | PsatzMulC _ prf => nhyps_of_psatz prf | PsatzAdd e1 e2 | PsatzMulE e1 e2 => nhyps_of_psatz e1 ++ nhyps_of_psatz e2 | PsatzIn n => n :: nil end. Fixpoint extract_hyps (l: list NFormula) (ln : list nat) : list NFormula := match ln with | nil => nil | n::ln => nth n l (Pc cO, Equal) :: extract_hyps l ln end. Lemma extract_hyps_app : forall l ln1 ln2, extract_hyps l (ln1 ++ ln2) = (extract_hyps l ln1) ++ (extract_hyps l ln2). Proof. induction ln1. reflexivity. simpl. intros. rewrite IHln1. reflexivity. Qed. Ltac inv H := inversion H ; try subst ; clear H. Lemma nhyps_of_psatz_correct : forall (env : PolEnv) (e:Psatz) (l : list NFormula) (f: NFormula), eval_Psatz l e = Some f -> ((forall f', In f' (extract_hyps l (nhyps_of_psatz e)) -> eval_nformula env f') -> eval_nformula env f). Proof. induction e ; intros. (*PsatzIn*) simpl in *. apply H0. intuition congruence. (* PsatzSquare *) simpl in *. inv H. simpl. unfold eval_pol. rewrite (Psquare_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)); now apply (Rtimes_square_nonneg sor). (* PsatzMulC *) simpl in *. case_eq (eval_Psatz l e). intros. rewrite H1 in H. simpl in H. apply pexpr_times_nformula_correct with (2:= H). apply IHe with (1:= H1); auto. intros. rewrite H1 in H. simpl in H ; discriminate. (* PsatzMulE *) simpl in *. revert H. case_eq (eval_Psatz l e1). case_eq (eval_Psatz l e2) ; simpl ; intros. apply nformula_times_nformula_correct with (3:= H2). apply IHe1 with (1:= H1) ; auto. intros. apply H0. rewrite extract_hyps_app. apply in_or_app. tauto. apply IHe2 with (1:= H) ; auto. intros. apply H0. rewrite extract_hyps_app. apply in_or_app. tauto. discriminate. simpl. discriminate. (* PsatzAdd *) simpl in *. revert H. case_eq (eval_Psatz l e1). case_eq (eval_Psatz l e2) ; simpl ; intros. apply nformula_plus_nformula_correct with (3:= H2). apply IHe1 with (1:= H1) ; auto. intros. apply H0. rewrite extract_hyps_app. apply in_or_app. tauto. apply IHe2 with (1:= H) ; auto. intros. apply H0. rewrite extract_hyps_app. apply in_or_app. tauto. discriminate. simpl. discriminate. (* PsatzC *) simpl in H. case_eq (cO [<] c). intros. rewrite H1 in H. inv H. unfold eval_nformula. simpl. rewrite <- (morph0 (SORrm addon)). now apply cltb_sound. intros. rewrite H1 in H. discriminate. (* PsatzZ *) simpl in *. inv H. unfold eval_nformula. simpl. apply (morph0 (SORrm addon)). Qed. (* roughly speaking, normalise_pexpr_correct is a proof of forall env p, eval_pexpr env p == eval_pol env (normalise_pexpr p) *) (*****) Definition paddC := PaddC cplus. Definition psubC := PsubC cminus. Definition PsubC_ok : forall c P env, eval_pol env (psubC P c) == eval_pol env P - [c] := let Rops_wd := mk_reqe (*rplus rtimes ropp req*) (SORplus_wd sor) (SORtimes_wd sor) (SORopp_wd sor) in PsubC_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon). Definition PaddC_ok : forall c P env, eval_pol env (paddC P c) == eval_pol env P + [c] := let Rops_wd := mk_reqe (*rplus rtimes ropp req*) (SORplus_wd sor) (SORtimes_wd sor) (SORopp_wd sor) in PaddC_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon). (* Check that a formula f is inconsistent by normalizing and comparing the resulting constant with 0 *) Definition check_inconsistent (f : NFormula) : bool := let (e, op) := f in match e with | Pc c => match op with | Equal => cneqb c cO | NonStrict => c [<] cO | Strict => c [<=] cO | NonEqual => c [=] cO end | _ => false (* not a constant *) end. Lemma check_inconsistent_sound : forall (p : PolC) (op : Op1), check_inconsistent (p, op) = true -> forall env, ~ eval_op1 op (eval_pol env p). Proof. intros p op H1 env. unfold check_inconsistent in H1. destruct op; simpl ; (*****) destruct p ; simpl; try discriminate H1; try rewrite <- (morph0 (SORrm addon)); trivial. now apply cneqb_sound. apply (morph_eq (SORrm addon)) in H1. congruence. apply cleb_sound in H1. now apply -> (Rle_ngt sor). apply cltb_sound in H1. now apply -> (Rlt_nge sor). Qed. Definition check_normalised_formulas : list NFormula -> Psatz -> bool := fun l cm => match eval_Psatz l cm with | None => false | Some f => check_inconsistent f end. Lemma checker_nf_sound : forall (l : list NFormula) (cm : Psatz), check_normalised_formulas l cm = true -> forall env : PolEnv, make_impl (eval_nformula env) l False. Proof. intros l cm H env. unfold check_normalised_formulas in H. revert H. case_eq (eval_Psatz l cm) ; [|discriminate]. intros nf. intros. rewrite <- make_conj_impl. intro. assert (H1' := make_conj_in _ _ H1). assert (Hnf := @eval_Psatz_Sound _ _ H1' _ _ H). destruct nf. apply (@check_inconsistent_sound _ _ H0 env Hnf). Qed. (** Normalisation of formulae **) Inductive Op2 : Set := (* binary relations *) | OpEq | OpNEq | OpLe | OpGe | OpLt | OpGt. Definition eval_op2 (o : Op2) : R -> R -> Prop := match o with | OpEq => req | OpNEq => fun x y : R => x ~= y | OpLe => rle | OpGe => fun x y : R => y <= x | OpLt => fun x y : R => x < y | OpGt => fun x y : R => y < x end. Definition eval_pexpr : PolEnv -> PExpr C -> R := PEeval rplus rtimes rminus ropp phi pow_phi rpow. #[universes(template)] Record Formula (T:Type) : Type := { Flhs : PExpr T; Fop : Op2; Frhs : PExpr T }. Definition eval_formula (env : PolEnv) (f : Formula C) : Prop := let (lhs, op, rhs) := f in (eval_op2 op) (eval_pexpr env lhs) (eval_pexpr env rhs). (* We normalize Formulas by moving terms to one side *) Definition norm := norm_aux cO cI cplus ctimes cminus copp ceqb. Definition psub := Psub cO cplus cminus copp ceqb. Definition padd := Padd cO cplus ceqb. Definition pmul := Pmul cO cI cplus ctimes ceqb. Definition popp := Popp copp. Definition normalise (f : Formula C) : NFormula := let (lhs, op, rhs) := f in let lhs := norm lhs in let rhs := norm rhs in match op with | OpEq => (psub lhs rhs, Equal) | OpNEq => (psub lhs rhs, NonEqual) | OpLe => (psub rhs lhs, NonStrict) | OpGe => (psub lhs rhs, NonStrict) | OpGt => (psub lhs rhs, Strict) | OpLt => (psub rhs lhs, Strict) end. Definition negate (f : Formula C) : NFormula := let (lhs, op, rhs) := f in let lhs := norm lhs in let rhs := norm rhs in match op with | OpEq => (psub rhs lhs, NonEqual) | OpNEq => (psub rhs lhs, Equal) | OpLe => (psub lhs rhs, Strict) (* e <= e' == ~ e > e' *) | OpGe => (psub rhs lhs, Strict) | OpGt => (psub rhs lhs, NonStrict) | OpLt => (psub lhs rhs, NonStrict) end. Lemma eval_pol_sub : forall env lhs rhs, eval_pol env (psub lhs rhs) == eval_pol env lhs - eval_pol env rhs. Proof. intros. apply (Psub_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)). Qed. Lemma eval_pol_add : forall env lhs rhs, eval_pol env (padd lhs rhs) == eval_pol env lhs + eval_pol env rhs. Proof. intros. apply (Padd_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)). Qed. Lemma eval_pol_mul : forall env lhs rhs, eval_pol env (pmul lhs rhs) == eval_pol env lhs * eval_pol env rhs. Proof. intros. apply (Pmul_ok sor.(SORsetoid) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd sor.(SORrt)) addon.(SORrm)). Qed. Lemma eval_pol_opp : forall env e, eval_pol env (popp e) == - eval_pol env e. Proof. intros. apply (Popp_ok (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon)). Qed. Lemma eval_pol_norm : forall env lhs, eval_pexpr env lhs == eval_pol env (norm lhs). Proof. intros. apply (norm_aux_spec (SORsetoid sor) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd (SORrt sor)) (SORrm addon) (SORpower addon) ). Qed. Theorem normalise_sound : forall (env : PolEnv) (f : Formula C), eval_formula env f <-> eval_nformula env (normalise f). Proof. intros env f; destruct f as [lhs op rhs]; simpl in *. destruct op; simpl in *; rewrite eval_pol_sub ; rewrite <- eval_pol_norm ; rewrite <- eval_pol_norm. - symmetry. now apply (Rminus_eq_0 sor). - rewrite (Rminus_eq_0 sor). tauto. - now apply (Rle_le_minus sor). - now apply (Rle_le_minus sor). - now apply (Rlt_lt_minus sor). - now apply (Rlt_lt_minus sor). Qed. Theorem negate_correct : forall (env : PolEnv) (f : Formula C), eval_formula env f <-> ~ (eval_nformula env (negate f)). Proof. intros env f; destruct f as [lhs op rhs]; simpl. destruct op; simpl in *; rewrite eval_pol_sub ; rewrite <- eval_pol_norm ; rewrite <- eval_pol_norm. - symmetry. rewrite (Rminus_eq_0 sor). split; intro H; [symmetry; now apply -> (Req_dne sor) | symmetry in H; now apply <- (Req_dne sor)]. - rewrite (Rminus_eq_0 sor). split; intro; now apply (Rneq_symm sor). - rewrite <- (Rlt_lt_minus sor). now rewrite <- (Rle_ngt sor). - rewrite <- (Rlt_lt_minus sor). now rewrite <- (Rle_ngt sor). - rewrite <- (Rle_le_minus sor). now rewrite <- (Rlt_nge sor). - rewrite <- (Rle_le_minus sor). now rewrite <- (Rlt_nge sor). Qed. (** Another normalisation - this is used for cnf conversion **) Definition xnormalise (f:NFormula) : list (NFormula) := let (e,o) := f in match o with | Equal => (e , Strict) :: (popp e, Strict) :: nil | NonEqual => (e , Equal) :: nil | Strict => (popp e, NonStrict) :: nil | NonStrict => (popp e, Strict) :: nil end. Definition xnegate (t:NFormula) : list (NFormula) := let (e,o) := t in match o with | Equal => (e,Equal) :: nil | NonEqual => (e,Strict)::(popp e,Strict)::nil | Strict => (e,Strict) :: nil | NonStrict => (e,NonStrict) :: nil end. Import Coq.micromega.Tauto. Definition cnf_of_list {T : Type} (l:list NFormula) (tg : T) : cnf NFormula T := List.fold_right (fun x acc => if check_inconsistent x then acc else ((x,tg)::nil)::acc) (cnf_tt _ _) l. Add Ring SORRing : (SORrt sor). Lemma cnf_of_list_correct : forall (T : Type) env l tg, eval_cnf (Annot:=T) eval_nformula env (cnf_of_list l tg) <-> make_conj (fun x : NFormula => eval_nformula env x -> False) l. Proof. unfold cnf_of_list. intros T env l tg. set (F := (fun (x : NFormula) (acc : list (list (NFormula * T))) => if check_inconsistent x then acc else ((x, tg) :: nil) :: acc)). set (G := ((fun x : NFormula => eval_nformula env x -> False))). induction l. - compute. tauto. - rewrite make_conj_cons. simpl. unfold F at 1. destruct (check_inconsistent a) eqn:EQ. + rewrite IHl. unfold G. destruct a. specialize (check_inconsistent_sound _ _ EQ env). simpl. tauto. + rewrite <- eval_cnf_cons_iff. simpl. unfold eval_tt. simpl. rewrite IHl. unfold G at 2. tauto. Qed. Definition cnf_normalise {T: Type} (t: Formula C) (tg: T) : cnf NFormula T := let f := normalise t in if check_inconsistent f then cnf_ff _ _ else cnf_of_list (xnormalise f) tg. Definition cnf_negate {T: Type} (t: Formula C) (tg: T) : cnf NFormula T := let f := normalise t in if check_inconsistent f then cnf_tt _ _ else cnf_of_list (xnegate f) tg. Lemma eq0_cnf : forall x, (0 < x -> False) /\ (0 < - x -> False) <-> x == 0. Proof. split ; intros. + apply (SORle_antisymm sor). * now rewrite (Rle_ngt sor). * rewrite (Rle_ngt sor). rewrite (Rlt_lt_minus sor). setoid_replace (0 - x) with (-x) by ring. tauto. + split; intro. * rewrite (SORlt_le_neq sor) in H0. apply (proj2 H0). now rewrite H. * rewrite (SORlt_le_neq sor) in H0. apply (proj2 H0). rewrite H. ring. Qed. Lemma xnormalise_correct : forall env f, (make_conj (fun x => eval_nformula env x -> False) (xnormalise f)) <-> eval_nformula env f. Proof. intros env f. destruct f as [e o]; destruct o eqn:Op; cbn - [psub]; repeat rewrite eval_pol_sub; fold eval_pol; repeat rewrite eval_pol_Pc; repeat rewrite eval_pol_opp; generalize (eval_pol env e) as x; intro. - apply eq0_cnf. - unfold not. tauto. - symmetry. rewrite (Rlt_nge sor). rewrite (Rle_le_minus sor). setoid_replace (0 - x) with (-x) by ring. tauto. - rewrite (Rle_ngt sor). symmetry. rewrite (Rlt_lt_minus sor). setoid_replace (0 - x) with (-x) by ring. tauto. Qed. Lemma xnegate_correct : forall env f, (make_conj (fun x => eval_nformula env x -> False) (xnegate f)) <-> ~ eval_nformula env f. Proof. intros env f. destruct f as [e o]; destruct o eqn:Op; cbn - [psub]; repeat rewrite eval_pol_sub; fold eval_pol; repeat rewrite eval_pol_Pc; repeat rewrite eval_pol_opp; generalize (eval_pol env e) as x; intro. - tauto. - rewrite eq0_cnf. rewrite (Req_dne sor). tauto. - tauto. - tauto. Qed. Lemma cnf_normalise_correct : forall (T : Type) env t tg, eval_cnf (Annot:=T) eval_nformula env (cnf_normalise t tg) <-> eval_formula env t. Proof. intros T env t tg. unfold cnf_normalise. rewrite normalise_sound. generalize (normalise t) as f;intro. destruct (check_inconsistent f) eqn:U. - destruct f as [e op]. assert (US := check_inconsistent_sound _ _ U env). rewrite eval_cnf_ff. tauto. - intros. rewrite cnf_of_list_correct. now apply xnormalise_correct. Qed. Lemma cnf_negate_correct : forall (T : Type) env t (tg:T), eval_cnf eval_nformula env (cnf_negate t tg) <-> ~ eval_formula env t. Proof. intros T env t tg. rewrite normalise_sound. unfold cnf_negate. generalize (normalise t) as f;intro. destruct (check_inconsistent f) eqn:U. - destruct f as [e o]. assert (US := check_inconsistent_sound _ _ U env). rewrite eval_cnf_tt. tauto. - rewrite cnf_of_list_correct. apply xnegate_correct. Qed. Lemma eval_nformula_dec : forall env d, (eval_nformula env d) \/ ~ (eval_nformula env d). Proof. intros. destruct d ; simpl. generalize (eval_pol env p); intros. destruct o ; simpl. apply (Req_em sor r 0). destruct (Req_em sor r 0) ; tauto. rewrite <- (Rle_ngt sor r 0). generalize (Rle_gt_cases sor r 0). tauto. rewrite <- (Rlt_nge sor r 0). generalize (Rle_gt_cases sor 0 r). tauto. Qed. (** Reverse transformation *) Fixpoint xdenorm (jmp : positive) (p: Pol C) : PExpr C := match p with | Pc c => PEc c | Pinj j p => xdenorm (Pos.add j jmp ) p | PX p j q => PEadd (PEmul (xdenorm jmp p) (PEpow (PEX jmp) (Npos j))) (xdenorm (Pos.succ jmp) q) end. Lemma xdenorm_correct : forall p i env, eval_pol (jump i env) p == eval_pexpr env (xdenorm (Pos.succ i) p). Proof. unfold eval_pol. induction p. simpl. reflexivity. (* Pinj *) simpl. intros. rewrite Pos.add_succ_r. rewrite <- IHp. symmetry. rewrite Pos.add_comm. rewrite Pjump_add. reflexivity. (* PX *) simpl. intros. rewrite <- IHp1, <- IHp2. unfold Env.tail , Env.hd. rewrite <- Pjump_add. rewrite Pos.add_1_r. unfold Env.nth. unfold jump at 2. rewrite <- Pos.add_1_l. rewrite (rpow_pow_N (SORpower addon)). unfold pow_N. ring. Qed. Definition denorm := xdenorm xH. Lemma denorm_correct : forall p env, eval_pol env p == eval_pexpr env (denorm p). Proof. unfold denorm. induction p. reflexivity. simpl. rewrite Pos.add_1_r. apply xdenorm_correct. simpl. intros. rewrite IHp1. unfold Env.tail. rewrite xdenorm_correct. change (Pos.succ xH) with 2%positive. rewrite (rpow_pow_N (SORpower addon)). simpl. reflexivity. Qed. (** Sometimes it is convenient to make a distinction between "syntactic" coefficients and "real" coefficients that are used to actually compute *) Variable S : Type. Variable C_of_S : S -> C. Variable phiS : S -> R. Variable phi_C_of_S : forall c, phiS c = phi (C_of_S c). Fixpoint map_PExpr (e : PExpr S) : PExpr C := match e with | PEc c => PEc (C_of_S c) | PEX p => PEX p | PEadd e1 e2 => PEadd (map_PExpr e1) (map_PExpr e2) | PEsub e1 e2 => PEsub (map_PExpr e1) (map_PExpr e2) | PEmul e1 e2 => PEmul (map_PExpr e1) (map_PExpr e2) | PEopp e => PEopp (map_PExpr e) | PEpow e n => PEpow (map_PExpr e) n end. Definition map_Formula (f : Formula S) : Formula C := let (l,o,r) := f in Build_Formula (map_PExpr l) o (map_PExpr r). Definition eval_sexpr : PolEnv -> PExpr S -> R := PEeval rplus rtimes rminus ropp phiS pow_phi rpow. Definition eval_sformula (env : PolEnv) (f : Formula S) : Prop := let (lhs, op, rhs) := f in (eval_op2 op) (eval_sexpr env lhs) (eval_sexpr env rhs). Lemma eval_pexprSC : forall env s, eval_sexpr env s = eval_pexpr env (map_PExpr s). Proof. unfold eval_pexpr, eval_sexpr. induction s ; simpl ; try (rewrite IHs1 ; rewrite IHs2) ; try reflexivity. apply phi_C_of_S. rewrite IHs. reflexivity. rewrite IHs. reflexivity. Qed. (** equality might be (too) strong *) Lemma eval_formulaSC : forall env f, eval_sformula env f = eval_formula env (map_Formula f). Proof. destruct f. simpl. repeat rewrite eval_pexprSC. reflexivity. Qed. (** Some syntactic simplifications of expressions *) Definition simpl_cone (e:Psatz) : Psatz := match e with | PsatzSquare t => match t with | Pc c => if ceqb cO c then PsatzZ else PsatzC (ctimes c c) | _ => PsatzSquare t end | PsatzMulE t1 t2 => match t1 , t2 with | PsatzZ , x => PsatzZ | x , PsatzZ => PsatzZ | PsatzC c , PsatzC c' => PsatzC (ctimes c c') | PsatzC p1 , PsatzMulE (PsatzC p2) x => PsatzMulE (PsatzC (ctimes p1 p2)) x | PsatzC p1 , PsatzMulE x (PsatzC p2) => PsatzMulE (PsatzC (ctimes p1 p2)) x | PsatzMulE (PsatzC p2) x , PsatzC p1 => PsatzMulE (PsatzC (ctimes p1 p2)) x | PsatzMulE x (PsatzC p2) , PsatzC p1 => PsatzMulE (PsatzC (ctimes p1 p2)) x | PsatzC x , PsatzAdd y z => PsatzAdd (PsatzMulE (PsatzC x) y) (PsatzMulE (PsatzC x) z) | PsatzC c , _ => if ceqb cI c then t2 else PsatzMulE t1 t2 | _ , PsatzC c => if ceqb cI c then t1 else PsatzMulE t1 t2 | _ , _ => e end | PsatzAdd t1 t2 => match t1 , t2 with | PsatzZ , x => x | x , PsatzZ => x | x , y => PsatzAdd x y end | _ => e end. End Micromega. (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/sos_types.ml0000644000175000017500000000531313612664533020276 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* output_string o "0" | Const n -> output_string o (string_of_num n) | Var n -> Printf.fprintf o "v%s" n | Opp t -> Printf.fprintf o "- (%a)" output_term t | Add (t1, t2) -> Printf.fprintf o "(%a)+(%a)" output_term t1 output_term t2 | Sub (t1, t2) -> Printf.fprintf o "(%a)-(%a)" output_term t1 output_term t2 | Mul (t1, t2) -> Printf.fprintf o "(%a)*(%a)" output_term t1 output_term t2 | Pow (t1, i) -> Printf.fprintf o "(%a)^(%i)" output_term t1 i (* ------------------------------------------------------------------------- *) (* Data structure for Positivstellensatz refutations. *) (* ------------------------------------------------------------------------- *) type positivstellensatz = | Axiom_eq of int | Axiom_le of int | Axiom_lt of int | Rational_eq of num | Rational_le of num | Rational_lt of num | Square of term | Monoid of int list | Eqmul of term * positivstellensatz | Sum of positivstellensatz * positivstellensatz | Product of positivstellensatz * positivstellensatz let rec output_psatz o = function | Axiom_eq i -> Printf.fprintf o "Aeq(%i)" i | Axiom_le i -> Printf.fprintf o "Ale(%i)" i | Axiom_lt i -> Printf.fprintf o "Alt(%i)" i | Rational_eq n -> Printf.fprintf o "eq(%s)" (string_of_num n) | Rational_le n -> Printf.fprintf o "le(%s)" (string_of_num n) | Rational_lt n -> Printf.fprintf o "lt(%s)" (string_of_num n) | Square t -> Printf.fprintf o "(%a)^2" output_term t | Monoid l -> Printf.fprintf o "monoid" | Eqmul (t, ps) -> Printf.fprintf o "%a * %a" output_term t output_psatz ps | Sum (t1, t2) -> Printf.fprintf o "%a + %a" output_psatz t1 output_psatz t2 | Product (t1, t2) -> Printf.fprintf o "%a * %a" output_psatz t1 output_psatz t2 coq-8.11.0/plugins/micromega/csdpcert.ml0000644000175000017500000001506513612664533020062 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Const (C2Ml.q_to_num z) | PEX v -> Var ("x" ^ string_of_int (C2Ml.index v)) | PEmul (p1, p2) -> let p1 = expr_to_term p1 in let p2 = expr_to_term p2 in let res = Mul (p1, p2) in res | PEadd (p1, p2) -> Add (expr_to_term p1, expr_to_term p2) | PEsub (p1, p2) -> Sub (expr_to_term p1, expr_to_term p2) | PEpow (p, n) -> Pow (expr_to_term p, C2Ml.n n) | PEopp p -> Opp (expr_to_term p) end open M let partition_expr l = let rec f i = function | [] -> ([], [], []) | (e, k) :: l -> ( let eq, ge, neq = f (i + 1) l in match k with | Mc.Equal -> ((e, i) :: eq, ge, neq) | Mc.NonStrict -> (eq, (e, Axiom_le i) :: ge, neq) | Mc.Strict -> (* e > 0 == e >= 0 /\ e <> 0 *) (eq, (e, Axiom_lt i) :: ge, (e, Axiom_lt i) :: neq) | Mc.NonEqual -> (eq, ge, (e, Axiom_eq i) :: neq) ) (* Not quite sure -- Coq interface has changed *) in f 0 l let rec sets_of_list l = match l with | [] -> [[]] | e :: l -> let s = sets_of_list l in s @ List.map (fun s0 -> e :: s0) s (* The exploration is probably not complete - for simple cases, it works... *) let real_nonlinear_prover d l = let l = List.map (fun (e, op) -> (Mc.denorm e, op)) l in try let eq, ge, neq = partition_expr l in let rec elim_const = function | [] -> [] | (x, y) :: l -> let p = poly_of_term (expr_to_term x) in if poly_isconst p then elim_const l else (p, y) :: elim_const l in let eq = elim_const eq in let peq = List.map fst eq in let pge = List.map (fun (e, psatz) -> (poly_of_term (expr_to_term e), psatz)) ge in let monoids = List.map (fun m -> ( List.fold_right (fun (p, kd) y -> let p = poly_of_term (expr_to_term p) in match kd with | Axiom_lt i -> poly_mul p y | Axiom_eq i -> poly_mul (poly_pow p 2) y | _ -> failwith "monoids") m (poly_const (Int 1)) , List.map snd m )) (sets_of_list neq) in let cert_ideal, cert_cone, monoid = deepen_until d (fun d -> tryfind (fun m -> let ci, cc = real_positivnullstellensatz_general false d peq pge (poly_neg (fst m)) in (ci, cc, snd m)) monoids) 0 in let proofs_ideal = List.map2 (fun q i -> Eqmul (term_of_poly q, Axiom_eq i)) cert_ideal (List.map snd eq) in let proofs_cone = List.map term_of_sos cert_cone in let proof_ne = let neq, lt = List.partition (function Axiom_eq _ -> true | _ -> false) monoid in let sq = match List.map (function Axiom_eq i -> i | _ -> failwith "error") neq with | [] -> Rational_lt (Int 1) | l -> Monoid l in List.fold_right (fun x y -> Product (x, y)) lt sq in let proof = end_itlist (fun s t -> Sum (s, t)) ((proof_ne :: proofs_ideal) @ proofs_cone) in S (Some proof) with | Sos_lib.TooDeep -> S None | any -> F (Printexc.to_string any) (* This is somewhat buggy, over Z, strict inequality vanish... *) let pure_sos l = let l = List.map (fun (e, o) -> (Mc.denorm e, o)) l in (* If there is no strict inequality, I should nonetheless be able to try something - over Z > is equivalent to -1 >= *) try let l = List.combine l (CList.interval 0 (List.length l - 1)) in let lt, i = try List.find (fun (x, _) -> snd x = Mc.Strict) l with Not_found -> List.hd l in let plt = poly_neg (poly_of_term (expr_to_term (fst lt))) in let n, polys = sumofsquares plt in (* n * (ci * pi^2) *) let pos = Product ( Rational_lt n , List.fold_right (fun (c, p) rst -> Sum (Product (Rational_lt c, Square (term_of_poly p)), rst)) polys (Rational_lt (Int 0)) ) in let proof = Sum (Axiom_lt i, pos) in (* let s,proof' = scale_certificate proof in let cert = snd (cert_of_pos proof') in *) S (Some proof) with (* | Sos.CsdpNotFound -> F "Sos.CsdpNotFound" *) | any -> (* May be that could be refined *) S None let run_prover prover pb = match prover with | "real_nonlinear_prover", Some d -> real_nonlinear_prover d pb | "pure_sos", None -> pure_sos pb | prover, _ -> Printf.printf "unknown prover: %s\n" prover; exit 1 let main () = try let (prover, poly) = (input_value stdin : provername * micromega_polys) in let cert = run_prover prover poly in (* Printf.fprintf chan "%a -> %a" print_list_term poly output_csdp_certificate cert ; close_out chan ; *) output_value stdout (cert : csdp_certificate); flush stdout; Marshal.to_channel chan (cert : csdp_certificate) []; flush chan; exit 0 with any -> Printf.fprintf chan "error %s" (Printexc.to_string any); exit 1 ;; let _ = main () in () (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/coq_micromega.ml0000644000175000017500000023717513612664533021070 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Some !vref) ; optwrite = (fun x -> vref := match x with None -> max_depth | Some v -> v) } in let lia_enum_opt = { optdepr = false ; optname = "Lia Enum" ; optkey = ["Lia"; "Enum"] ; optread = (fun () -> !lia_enum) ; optwrite = (fun x -> lia_enum := x) } in let solver_opt = { optdepr = false ; optname = "Use the Simplex instead of Fourier elimination" ; optkey = ["Simplex"] ; optread = (fun () -> !Certificate.use_simplex) ; optwrite = (fun x -> Certificate.use_simplex := x) } in let dump_file_opt = { optdepr = false ; optname = "Generate Coq goals in file from calls to 'lia' 'nia'" ; optkey = ["Dump"; "Arith"] ; optread = (fun () -> !Certificate.dump_file) ; optwrite = (fun x -> Certificate.dump_file := x) } in let lia_cache_opt = { optdepr = false ; optname = "cache of lia (.lia.cache)" ; optkey = ["Lia"; "Cache"] ; optread = (fun () -> !use_lia_cache) ; optwrite = (fun x -> use_lia_cache := x) } in let nia_cache_opt = { optdepr = false ; optname = "cache of nia (.nia.cache)" ; optkey = ["Nia"; "Cache"] ; optread = (fun () -> !use_nia_cache) ; optwrite = (fun x -> use_nia_cache := x) } in let nra_cache_opt = { optdepr = false ; optname = "cache of nra (.nra.cache)" ; optkey = ["Nra"; "Cache"] ; optread = (fun () -> !use_nra_cache) ; optwrite = (fun x -> use_nra_cache := x) } in let () = declare_bool_option solver_opt in let () = declare_bool_option lia_cache_opt in let () = declare_bool_option nia_cache_opt in let () = declare_bool_option nra_cache_opt in let () = declare_stringopt_option dump_file_opt in let () = declare_int_option (int_opt ["Lra"; "Depth"] lra_proof_depth) in let () = declare_int_option (int_opt ["Lia"; "Depth"] lia_proof_depth) in let () = declare_bool_option lia_enum_opt in () (** * Initialize a tag type to the Tag module declaration (see Mutils). *) type tag = Tag.t module Mc = Micromega (** * An atom is of the form: * pExpr1 \{<,>,=,<>,<=,>=\} pExpr2 * where pExpr1, pExpr2 are polynomial expressions (see Micromega). pExprs are * parametrized by 'cst, which is used as the type of constants. *) type 'cst atom = 'cst Mc.formula type 'cst formula = ('cst atom, EConstr.constr, tag * EConstr.constr, Names.Id.t) Mc.gFormula type 'cst clause = ('cst Mc.nFormula, tag * EConstr.constr) Mc.clause type 'cst cnf = ('cst Mc.nFormula, tag * EConstr.constr) Mc.cnf let rec pp_formula o (f : 'cst formula) = Mc.( match f with | TT -> output_string o "tt" | FF -> output_string o "ff" | X c -> output_string o "X " | A (_, (t, _)) -> Printf.fprintf o "A(%a)" Tag.pp t | Cj (f1, f2) -> Printf.fprintf o "C(%a,%a)" pp_formula f1 pp_formula f2 | D (f1, f2) -> Printf.fprintf o "D(%a,%a)" pp_formula f1 pp_formula f2 | I (f1, n, f2) -> Printf.fprintf o "I(%a,%s,%a)" pp_formula f1 (match n with Some id -> Names.Id.to_string id | None -> "") pp_formula f2 | N f -> Printf.fprintf o "N(%a)" pp_formula f) (** * Given a set of integers s=\{i0,...,iN\} and a list m, return the list of * elements of m that are at position i0,...,iN. *) let selecti s m = let rec xselecti i m = match m with | [] -> [] | e :: m -> if ISet.mem i s then e :: xselecti (i + 1) m else xselecti (i + 1) m in xselecti 0 m (** * MODULE: Mapping of the Coq data-strustures into Caml and Caml extracted * code. This includes initializing Caml variables based on Coq terms, parsing * various Coq expressions into Caml, and dumping Caml expressions into Coq. * * Opened here and in csdpcert.ml. *) (** * MODULE END: M *) module M = struct (** * Location of the Coq libraries. *) let logic_dir = ["Coq"; "Logic"; "Decidable"] let mic_modules = [ ["Coq"; "Lists"; "List"] ; ["Coq"; "micromega"; "ZMicromega"] ; ["Coq"; "micromega"; "Tauto"] ; ["Coq"; "micromega"; "DeclConstant"] ; ["Coq"; "micromega"; "RingMicromega"] ; ["Coq"; "micromega"; "EnvRing"] ; ["Coq"; "micromega"; "ZMicromega"] ; ["Coq"; "micromega"; "RMicromega"] ; ["Coq"; "micromega"; "Tauto"] ; ["Coq"; "micromega"; "RingMicromega"] ; ["Coq"; "micromega"; "EnvRing"] ; ["Coq"; "QArith"; "QArith_base"] ; ["Coq"; "Reals"; "Rdefinitions"] ; ["Coq"; "Reals"; "Rpow_def"] ; ["LRing_normalise"] ] [@@@ocaml.warning "-3"] let coq_modules = Coqlib.( init_modules @ [logic_dir] @ arith_modules @ zarith_base_modules @ mic_modules) let bin_module = [["Coq"; "Numbers"; "BinNums"]] let r_modules = [ ["Coq"; "Reals"; "Rdefinitions"] ; ["Coq"; "Reals"; "Rpow_def"] ; ["Coq"; "Reals"; "Raxioms"] ; ["Coq"; "QArith"; "Qreals"] ] let z_modules = [["Coq"; "ZArith"; "BinInt"]] (** * Initialization : a large amount of Caml symbols are derived from * ZMicromega.v *) let gen_constant_in_modules s m n = EConstr.of_constr ( UnivGen.constr_of_monomorphic_global @@ Coqlib.gen_reference_in_modules s m n ) let init_constant = gen_constant_in_modules "ZMicromega" Coqlib.init_modules [@@@ocaml.warning "+3"] let constant = gen_constant_in_modules "ZMicromega" coq_modules let bin_constant = gen_constant_in_modules "ZMicromega" bin_module let r_constant = gen_constant_in_modules "ZMicromega" r_modules let z_constant = gen_constant_in_modules "ZMicromega" z_modules let m_constant = gen_constant_in_modules "ZMicromega" mic_modules let coq_and = lazy (init_constant "and") let coq_or = lazy (init_constant "or") let coq_not = lazy (init_constant "not") let coq_iff = lazy (init_constant "iff") let coq_True = lazy (init_constant "True") let coq_False = lazy (init_constant "False") let coq_cons = lazy (constant "cons") let coq_nil = lazy (constant "nil") let coq_list = lazy (constant "list") let coq_O = lazy (init_constant "O") let coq_S = lazy (init_constant "S") let coq_nat = lazy (init_constant "nat") let coq_unit = lazy (init_constant "unit") (* let coq_option = lazy (init_constant "option")*) let coq_None = lazy (init_constant "None") let coq_tt = lazy (init_constant "tt") let coq_Inl = lazy (init_constant "inl") let coq_Inr = lazy (init_constant "inr") let coq_N0 = lazy (bin_constant "N0") let coq_Npos = lazy (bin_constant "Npos") let coq_xH = lazy (bin_constant "xH") let coq_xO = lazy (bin_constant "xO") let coq_xI = lazy (bin_constant "xI") let coq_Z = lazy (bin_constant "Z") let coq_ZERO = lazy (bin_constant "Z0") let coq_POS = lazy (bin_constant "Zpos") let coq_NEG = lazy (bin_constant "Zneg") let coq_Q = lazy (constant "Q") let coq_R = lazy (constant "R") let coq_Qmake = lazy (constant "Qmake") let coq_Rcst = lazy (constant "Rcst") let coq_C0 = lazy (m_constant "C0") let coq_C1 = lazy (m_constant "C1") let coq_CQ = lazy (m_constant "CQ") let coq_CZ = lazy (m_constant "CZ") let coq_CPlus = lazy (m_constant "CPlus") let coq_CMinus = lazy (m_constant "CMinus") let coq_CMult = lazy (m_constant "CMult") let coq_CPow = lazy (m_constant "CPow") let coq_CInv = lazy (m_constant "CInv") let coq_COpp = lazy (m_constant "COpp") let coq_R0 = lazy (constant "R0") let coq_R1 = lazy (constant "R1") let coq_proofTerm = lazy (constant "ZArithProof") let coq_doneProof = lazy (constant "DoneProof") let coq_ratProof = lazy (constant "RatProof") let coq_cutProof = lazy (constant "CutProof") let coq_enumProof = lazy (constant "EnumProof") let coq_ExProof = lazy (constant "ExProof") let coq_Zgt = lazy (z_constant "Z.gt") let coq_Zge = lazy (z_constant "Z.ge") let coq_Zle = lazy (z_constant "Z.le") let coq_Zlt = lazy (z_constant "Z.lt") let coq_Eq = lazy (init_constant "eq") let coq_Zplus = lazy (z_constant "Z.add") let coq_Zminus = lazy (z_constant "Z.sub") let coq_Zopp = lazy (z_constant "Z.opp") let coq_Zmult = lazy (z_constant "Z.mul") let coq_Zpower = lazy (z_constant "Z.pow") let coq_Qle = lazy (constant "Qle") let coq_Qlt = lazy (constant "Qlt") let coq_Qeq = lazy (constant "Qeq") let coq_Qplus = lazy (constant "Qplus") let coq_Qminus = lazy (constant "Qminus") let coq_Qopp = lazy (constant "Qopp") let coq_Qmult = lazy (constant "Qmult") let coq_Qpower = lazy (constant "Qpower") let coq_Rgt = lazy (r_constant "Rgt") let coq_Rge = lazy (r_constant "Rge") let coq_Rle = lazy (r_constant "Rle") let coq_Rlt = lazy (r_constant "Rlt") let coq_Rplus = lazy (r_constant "Rplus") let coq_Rminus = lazy (r_constant "Rminus") let coq_Ropp = lazy (r_constant "Ropp") let coq_Rmult = lazy (r_constant "Rmult") let coq_Rinv = lazy (r_constant "Rinv") let coq_Rpower = lazy (r_constant "pow") let coq_powerZR = lazy (r_constant "powerRZ") let coq_IZR = lazy (r_constant "IZR") let coq_IQR = lazy (r_constant "Q2R") let coq_PEX = lazy (constant "PEX") let coq_PEc = lazy (constant "PEc") let coq_PEadd = lazy (constant "PEadd") let coq_PEopp = lazy (constant "PEopp") let coq_PEmul = lazy (constant "PEmul") let coq_PEsub = lazy (constant "PEsub") let coq_PEpow = lazy (constant "PEpow") let coq_PX = lazy (constant "PX") let coq_Pc = lazy (constant "Pc") let coq_Pinj = lazy (constant "Pinj") let coq_OpEq = lazy (constant "OpEq") let coq_OpNEq = lazy (constant "OpNEq") let coq_OpLe = lazy (constant "OpLe") let coq_OpLt = lazy (constant "OpLt") let coq_OpGe = lazy (constant "OpGe") let coq_OpGt = lazy (constant "OpGt") let coq_PsatzIn = lazy (constant "PsatzIn") let coq_PsatzSquare = lazy (constant "PsatzSquare") let coq_PsatzMulE = lazy (constant "PsatzMulE") let coq_PsatzMultC = lazy (constant "PsatzMulC") let coq_PsatzAdd = lazy (constant "PsatzAdd") let coq_PsatzC = lazy (constant "PsatzC") let coq_PsatzZ = lazy (constant "PsatzZ") (* let coq_GT = lazy (m_constant "GT")*) let coq_DeclaredConstant = lazy (m_constant "DeclaredConstant") let coq_TT = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "TT") let coq_FF = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "FF") let coq_And = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "Cj") let coq_Or = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "D") let coq_Neg = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "N") let coq_Atom = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "A") let coq_X = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "X") let coq_Impl = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "I") let coq_Formula = lazy (gen_constant_in_modules "ZMicromega" [["Coq"; "micromega"; "Tauto"]; ["Tauto"]] "BFormula") (** * Initialization : a few Caml symbols are derived from other libraries; * QMicromega, ZArithRing, RingMicromega. *) let coq_QWitness = lazy (gen_constant_in_modules "QMicromega" [["Coq"; "micromega"; "QMicromega"]] "QWitness") let coq_Build = lazy (gen_constant_in_modules "RingMicromega" [["Coq"; "micromega"; "RingMicromega"]; ["RingMicromega"]] "Build_Formula") let coq_Cstr = lazy (gen_constant_in_modules "RingMicromega" [["Coq"; "micromega"; "RingMicromega"]; ["RingMicromega"]] "Formula") (** * Parsing and dumping : transformation functions between Caml and Coq * data-structures. * * dump_* functions go from Micromega to Coq terms * parse_* functions go from Coq to Micromega terms * pp_* functions pretty-print Coq terms. *) exception ParseError (* A simple but useful getter function *) let get_left_construct sigma term = match EConstr.kind sigma term with | Construct ((_, i), _) -> (i, [||]) | App (l, rst) -> ( match EConstr.kind sigma l with | Construct ((_, i), _) -> (i, rst) | _ -> raise ParseError ) | _ -> raise ParseError (* Access the Micromega module *) (* parse/dump/print from numbers up to expressions and formulas *) let rec parse_nat sigma term = let i, c = get_left_construct sigma term in match i with | 1 -> Mc.O | 2 -> Mc.S (parse_nat sigma c.(0)) | i -> raise ParseError let pp_nat o n = Printf.fprintf o "%i" (CoqToCaml.nat n) let rec dump_nat x = match x with | Mc.O -> Lazy.force coq_O | Mc.S p -> EConstr.mkApp (Lazy.force coq_S, [|dump_nat p|]) let rec parse_positive sigma term = let i, c = get_left_construct sigma term in match i with | 1 -> Mc.XI (parse_positive sigma c.(0)) | 2 -> Mc.XO (parse_positive sigma c.(0)) | 3 -> Mc.XH | i -> raise ParseError let rec dump_positive x = match x with | Mc.XH -> Lazy.force coq_xH | Mc.XO p -> EConstr.mkApp (Lazy.force coq_xO, [|dump_positive p|]) | Mc.XI p -> EConstr.mkApp (Lazy.force coq_xI, [|dump_positive p|]) let pp_positive o x = Printf.fprintf o "%i" (CoqToCaml.positive x) let dump_n x = match x with | Mc.N0 -> Lazy.force coq_N0 | Mc.Npos p -> EConstr.mkApp (Lazy.force coq_Npos, [|dump_positive p|]) (** [is_ground_term env sigma term] holds if the term [term] is an instance of the typeclass [DeclConstant.GT term] i.e. built from user-defined constants and functions. NB: This mechanism can be used to customise the reification process to decide what to consider as a constant (see [parse_constant]) *) let is_declared_term env evd t = match EConstr.kind evd t with | Const _ | Construct _ -> ( (* Restrict typeclass resolution to trivial cases *) let typ = Retyping.get_type_of env evd t in try ignore (Typeclasses.resolve_one_typeclass env evd (EConstr.mkApp (Lazy.force coq_DeclaredConstant, [|typ; t|]))); true with Not_found -> false ) | _ -> false let rec is_ground_term env evd term = match EConstr.kind evd term with | App (c, args) -> is_declared_term env evd c && Array.for_all (is_ground_term env evd) args | Const _ | Construct _ -> is_declared_term env evd term | _ -> false let parse_z sigma term = let i, c = get_left_construct sigma term in match i with | 1 -> Mc.Z0 | 2 -> Mc.Zpos (parse_positive sigma c.(0)) | 3 -> Mc.Zneg (parse_positive sigma c.(0)) | i -> raise ParseError let dump_z x = match x with | Mc.Z0 -> Lazy.force coq_ZERO | Mc.Zpos p -> EConstr.mkApp (Lazy.force coq_POS, [|dump_positive p|]) | Mc.Zneg p -> EConstr.mkApp (Lazy.force coq_NEG, [|dump_positive p|]) let pp_z o x = Printf.fprintf o "%s" (Big_int.string_of_big_int (CoqToCaml.z_big_int x)) let dump_q q = EConstr.mkApp ( Lazy.force coq_Qmake , [|dump_z q.Micromega.qnum; dump_positive q.Micromega.qden|] ) let parse_q sigma term = match EConstr.kind sigma term with | App (c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then { Mc.qnum = parse_z sigma args.(0) ; Mc.qden = parse_positive sigma args.(1) } else raise ParseError | _ -> raise ParseError let rec pp_Rcst o cst = match cst with | Mc.C0 -> output_string o "C0" | Mc.C1 -> output_string o "C1" | Mc.CQ q -> output_string o "CQ _" | Mc.CZ z -> pp_z o z | Mc.CPlus (x, y) -> Printf.fprintf o "(%a + %a)" pp_Rcst x pp_Rcst y | Mc.CMinus (x, y) -> Printf.fprintf o "(%a - %a)" pp_Rcst x pp_Rcst y | Mc.CMult (x, y) -> Printf.fprintf o "(%a * %a)" pp_Rcst x pp_Rcst y | Mc.CPow (x, y) -> Printf.fprintf o "(%a ^ _)" pp_Rcst x | Mc.CInv t -> Printf.fprintf o "(/ %a)" pp_Rcst t | Mc.COpp t -> Printf.fprintf o "(- %a)" pp_Rcst t let rec dump_Rcst cst = match cst with | Mc.C0 -> Lazy.force coq_C0 | Mc.C1 -> Lazy.force coq_C1 | Mc.CQ q -> EConstr.mkApp (Lazy.force coq_CQ, [|dump_q q|]) | Mc.CZ z -> EConstr.mkApp (Lazy.force coq_CZ, [|dump_z z|]) | Mc.CPlus (x, y) -> EConstr.mkApp (Lazy.force coq_CPlus, [|dump_Rcst x; dump_Rcst y|]) | Mc.CMinus (x, y) -> EConstr.mkApp (Lazy.force coq_CMinus, [|dump_Rcst x; dump_Rcst y|]) | Mc.CMult (x, y) -> EConstr.mkApp (Lazy.force coq_CMult, [|dump_Rcst x; dump_Rcst y|]) | Mc.CPow (x, y) -> EConstr.mkApp ( Lazy.force coq_CPow , [| dump_Rcst x ; ( match y with | Mc.Inl z -> EConstr.mkApp ( Lazy.force coq_Inl , [|Lazy.force coq_Z; Lazy.force coq_nat; dump_z z|] ) | Mc.Inr n -> EConstr.mkApp ( Lazy.force coq_Inr , [|Lazy.force coq_Z; Lazy.force coq_nat; dump_nat n|] ) ) |] ) | Mc.CInv t -> EConstr.mkApp (Lazy.force coq_CInv, [|dump_Rcst t|]) | Mc.COpp t -> EConstr.mkApp (Lazy.force coq_COpp, [|dump_Rcst t|]) let rec dump_list typ dump_elt l = match l with | [] -> EConstr.mkApp (Lazy.force coq_nil, [|typ|]) | e :: l -> EConstr.mkApp (Lazy.force coq_cons, [|typ; dump_elt e; dump_list typ dump_elt l|]) let pp_list op cl elt o l = let rec _pp o l = match l with | [] -> () | [e] -> Printf.fprintf o "%a" elt e | e :: l -> Printf.fprintf o "%a ,%a" elt e _pp l in Printf.fprintf o "%s%a%s" op _pp l cl let dump_var = dump_positive let dump_expr typ dump_z e = let rec dump_expr e = match e with | Mc.PEX n -> EConstr.mkApp (Lazy.force coq_PEX, [|typ; dump_var n|]) | Mc.PEc z -> EConstr.mkApp (Lazy.force coq_PEc, [|typ; dump_z z|]) | Mc.PEadd (e1, e2) -> EConstr.mkApp (Lazy.force coq_PEadd, [|typ; dump_expr e1; dump_expr e2|]) | Mc.PEsub (e1, e2) -> EConstr.mkApp (Lazy.force coq_PEsub, [|typ; dump_expr e1; dump_expr e2|]) | Mc.PEopp e -> EConstr.mkApp (Lazy.force coq_PEopp, [|typ; dump_expr e|]) | Mc.PEmul (e1, e2) -> EConstr.mkApp (Lazy.force coq_PEmul, [|typ; dump_expr e1; dump_expr e2|]) | Mc.PEpow (e, n) -> EConstr.mkApp (Lazy.force coq_PEpow, [|typ; dump_expr e; dump_n n|]) in dump_expr e let dump_pol typ dump_c e = let rec dump_pol e = match e with | Mc.Pc n -> EConstr.mkApp (Lazy.force coq_Pc, [|typ; dump_c n|]) | Mc.Pinj (p, pol) -> EConstr.mkApp (Lazy.force coq_Pinj, [|typ; dump_positive p; dump_pol pol|]) | Mc.PX (pol1, p, pol2) -> EConstr.mkApp ( Lazy.force coq_PX , [|typ; dump_pol pol1; dump_positive p; dump_pol pol2|] ) in dump_pol e let pp_pol pp_c o e = let rec pp_pol o e = match e with | Mc.Pc n -> Printf.fprintf o "Pc %a" pp_c n | Mc.Pinj (p, pol) -> Printf.fprintf o "Pinj(%a,%a)" pp_positive p pp_pol pol | Mc.PX (pol1, p, pol2) -> Printf.fprintf o "PX(%a,%a,%a)" pp_pol pol1 pp_positive p pp_pol pol2 in pp_pol o e (* let pp_clause pp_c o (f: 'cst clause) = List.iter (fun ((p,_),(t,_)) -> Printf.fprintf o "(%a @%a)" (pp_pol pp_c) p Tag.pp t) f *) let pp_clause_tag o (f : 'cst clause) = List.iter (fun ((p, _), (t, _)) -> Printf.fprintf o "(_ @%a)" Tag.pp t) f (* let pp_cnf pp_c o (f:'cst cnf) = List.iter (fun l -> Printf.fprintf o "[%a]" (pp_clause pp_c) l) f *) let pp_cnf_tag o (f : 'cst cnf) = List.iter (fun l -> Printf.fprintf o "[%a]" pp_clause_tag l) f let dump_psatz typ dump_z e = let z = Lazy.force typ in let rec dump_cone e = match e with | Mc.PsatzIn n -> EConstr.mkApp (Lazy.force coq_PsatzIn, [|z; dump_nat n|]) | Mc.PsatzMulC (e, c) -> EConstr.mkApp (Lazy.force coq_PsatzMultC, [|z; dump_pol z dump_z e; dump_cone c|]) | Mc.PsatzSquare e -> EConstr.mkApp (Lazy.force coq_PsatzSquare, [|z; dump_pol z dump_z e|]) | Mc.PsatzAdd (e1, e2) -> EConstr.mkApp (Lazy.force coq_PsatzAdd, [|z; dump_cone e1; dump_cone e2|]) | Mc.PsatzMulE (e1, e2) -> EConstr.mkApp (Lazy.force coq_PsatzMulE, [|z; dump_cone e1; dump_cone e2|]) | Mc.PsatzC p -> EConstr.mkApp (Lazy.force coq_PsatzC, [|z; dump_z p|]) | Mc.PsatzZ -> EConstr.mkApp (Lazy.force coq_PsatzZ, [|z|]) in dump_cone e let pp_psatz pp_z o e = let rec pp_cone o e = match e with | Mc.PsatzIn n -> Printf.fprintf o "(In %a)%%nat" pp_nat n | Mc.PsatzMulC (e, c) -> Printf.fprintf o "( %a [*] %a)" (pp_pol pp_z) e pp_cone c | Mc.PsatzSquare e -> Printf.fprintf o "(%a^2)" (pp_pol pp_z) e | Mc.PsatzAdd (e1, e2) -> Printf.fprintf o "(%a [+] %a)" pp_cone e1 pp_cone e2 | Mc.PsatzMulE (e1, e2) -> Printf.fprintf o "(%a [*] %a)" pp_cone e1 pp_cone e2 | Mc.PsatzC p -> Printf.fprintf o "(%a)%%positive" pp_z p | Mc.PsatzZ -> Printf.fprintf o "0" in pp_cone o e let dump_op = function | Mc.OpEq -> Lazy.force coq_OpEq | Mc.OpNEq -> Lazy.force coq_OpNEq | Mc.OpLe -> Lazy.force coq_OpLe | Mc.OpGe -> Lazy.force coq_OpGe | Mc.OpGt -> Lazy.force coq_OpGt | Mc.OpLt -> Lazy.force coq_OpLt let dump_cstr typ dump_constant {Mc.flhs = e1; Mc.fop = o; Mc.frhs = e2} = EConstr.mkApp ( Lazy.force coq_Build , [| typ ; dump_expr typ dump_constant e1 ; dump_op o ; dump_expr typ dump_constant e2 |] ) let assoc_const sigma x l = try snd (List.find (fun (x', y) -> EConstr.eq_constr sigma x (Lazy.force x')) l) with Not_found -> raise ParseError let zop_table = [ (coq_Zgt, Mc.OpGt) ; (coq_Zge, Mc.OpGe) ; (coq_Zlt, Mc.OpLt) ; (coq_Zle, Mc.OpLe) ] let rop_table = [ (coq_Rgt, Mc.OpGt) ; (coq_Rge, Mc.OpGe) ; (coq_Rlt, Mc.OpLt) ; (coq_Rle, Mc.OpLe) ] let qop_table = [(coq_Qlt, Mc.OpLt); (coq_Qle, Mc.OpLe); (coq_Qeq, Mc.OpEq)] type gl = {env : Environ.env; sigma : Evd.evar_map} let is_convertible gl t1 t2 = Reductionops.is_conv gl.env gl.sigma t1 t2 let parse_zop gl (op, args) = let sigma = gl.sigma in match args with | [|a1; a2|] -> (assoc_const sigma op zop_table, a1, a2) | [|ty; a1; a2|] -> if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl ty (Lazy.force coq_Z) then (Mc.OpEq, args.(1), args.(2)) else raise ParseError | _ -> raise ParseError let parse_rop gl (op, args) = let sigma = gl.sigma in match args with | [|a1; a2|] -> (assoc_const sigma op rop_table, a1, a2) | [|ty; a1; a2|] -> if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl ty (Lazy.force coq_R) then (Mc.OpEq, a1, a2) else raise ParseError | _ -> raise ParseError let parse_qop gl (op, args) = if Array.length args = 2 then (assoc_const gl.sigma op qop_table, args.(0), args.(1)) else raise ParseError type 'a op = | Binop of ('a Mc.pExpr -> 'a Mc.pExpr -> 'a Mc.pExpr) | Opp | Power | Ukn of string let assoc_ops sigma x l = try snd (List.find (fun (x', y) -> EConstr.eq_constr sigma x (Lazy.force x')) l) with Not_found -> Ukn "Oups" (** * MODULE: Env is for environment. *) module Env = struct type t = { vars : EConstr.t list ; (* The list represents a mapping from EConstr.t to indexes. *) gl : gl (* The evar_map may be updated due to unification of universes *) } let empty gl = {vars = []; gl} (** [eq_constr gl x y] returns an updated [gl] if x and y can be unified *) let eq_constr gl x y = let evd = gl.sigma in match EConstr.eq_constr_universes_proj gl.env evd x y with | Some csts -> ( let csts = UnivProblem.to_constraints ~force_weak:false (Evd.universes evd) csts in match Evd.add_constraints evd csts with | evd -> Some {gl with sigma = evd} | exception Univ.UniverseInconsistency _ -> None ) | None -> None let compute_rank_add env v = let rec _add gl vars n v = match vars with | [] -> (gl, [v], n) | e :: l -> ( match eq_constr gl e v with | Some gl' -> (gl', vars, n) | None -> let gl, l', n = _add gl l (n + 1) v in (gl, e :: l', n) ) in let gl', vars', n = _add env.gl env.vars 1 v in ({vars = vars'; gl = gl'}, CamlToCoq.positive n) let get_rank env v = let gl = env.gl in let rec _get_rank env n = match env with | [] -> raise (Invalid_argument "get_rank") | e :: l -> ( match eq_constr gl e v with Some _ -> n | None -> _get_rank l (n + 1) ) in _get_rank env.vars 1 let elements env = env.vars (* let string_of_env gl env = let rec string_of_env i env acc = match env with | [] -> acc | e::env -> string_of_env (i+1) env (IMap.add i (Pp.string_of_ppcmds (Printer.pr_econstr_env gl.env gl.sigma e)) acc) in string_of_env 1 env IMap.empty *) let pp gl env = let ppl = List.mapi (fun i e -> Pp.str "x" ++ Pp.int (i + 1) ++ Pp.str ":" ++ Printer.pr_econstr_env gl.env gl.sigma e) env in List.fold_right (fun e p -> e ++ Pp.str " ; " ++ p) ppl (Pp.str "\n") end (* MODULE END: Env *) (** * This is the big generic function for expression parsers. *) let parse_expr gl parse_constant parse_exp ops_spec env term = if debug then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.pr_leconstr_env gl.env gl.sigma term); let parse_variable env term = let env, n = Env.compute_rank_add env term in (Mc.PEX n, env) in let rec parse_expr env term = let combine env op (t1, t2) = let expr1, env = parse_expr env t1 in let expr2, env = parse_expr env t2 in (op expr1 expr2, env) in try (Mc.PEc (parse_constant gl term), env) with ParseError -> ( match EConstr.kind gl.sigma term with | App (t, args) -> ( match EConstr.kind gl.sigma t with | Const c -> ( match assoc_ops gl.sigma t ops_spec with | Binop f -> combine env f (args.(0), args.(1)) | Opp -> let expr, env = parse_expr env args.(0) in (Mc.PEopp expr, env) | Power -> ( try let expr, env = parse_expr env args.(0) in let power = parse_exp expr args.(1) in (power, env) with ParseError -> (* if the exponent is a variable *) let env, n = Env.compute_rank_add env term in (Mc.PEX n, env) ) | Ukn s -> if debug then ( Printf.printf "unknown op: %s\n" s; flush stdout ); let env, n = Env.compute_rank_add env term in (Mc.PEX n, env) ) | _ -> parse_variable env term ) | _ -> parse_variable env term ) in parse_expr env term let zop_spec = [ (coq_Zplus, Binop (fun x y -> Mc.PEadd (x, y))) ; (coq_Zminus, Binop (fun x y -> Mc.PEsub (x, y))) ; (coq_Zmult, Binop (fun x y -> Mc.PEmul (x, y))) ; (coq_Zopp, Opp) ; (coq_Zpower, Power) ] let qop_spec = [ (coq_Qplus, Binop (fun x y -> Mc.PEadd (x, y))) ; (coq_Qminus, Binop (fun x y -> Mc.PEsub (x, y))) ; (coq_Qmult, Binop (fun x y -> Mc.PEmul (x, y))) ; (coq_Qopp, Opp) ; (coq_Qpower, Power) ] let rop_spec = [ (coq_Rplus, Binop (fun x y -> Mc.PEadd (x, y))) ; (coq_Rminus, Binop (fun x y -> Mc.PEsub (x, y))) ; (coq_Rmult, Binop (fun x y -> Mc.PEmul (x, y))) ; (coq_Ropp, Opp) ; (coq_Rpower, Power) ] let parse_constant parse gl t = parse gl.sigma t (** [parse_more_constant parse gl t] returns the reification of term [t]. If [t] is a ground term, then it is first reduced to normal form before using a 'syntactic' parser *) let parse_more_constant parse gl t = try parse gl t with ParseError -> if debug then Feedback.msg_debug Pp.(str "try harder"); if is_ground_term gl.env gl.sigma t then parse gl (Redexpr.cbv_vm gl.env gl.sigma t) else raise ParseError let zconstant = parse_constant parse_z let qconstant = parse_constant parse_q let nconstant = parse_constant parse_nat (** [parse_more_zexpr parse_constant gl] improves the parsing of exponent which can be arithmetic expressions (without variables). [parse_constant_expr] returns a constant if the argument is an expression without variables. *) let rec parse_zexpr gl = parse_expr gl zconstant (fun expr (x : EConstr.t) -> let z = parse_zconstant gl x in match z with | Mc.Zneg _ -> Mc.PEc Mc.Z0 | _ -> Mc.PEpow (expr, Mc.Z.to_N z)) zop_spec and parse_zconstant gl e = let e, _ = parse_zexpr gl (Env.empty gl) e in match Mc.zeval_const e with None -> raise ParseError | Some z -> z (* NB: R is a different story. Because it is axiomatised, reducing would not be effective. Therefore, there is a specific parser for constant over R *) let rconst_assoc = [ (coq_Rplus, fun x y -> Mc.CPlus (x, y)) ; (coq_Rminus, fun x y -> Mc.CMinus (x, y)) ; (coq_Rmult, fun x y -> Mc.CMult (x, y)) (* coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;*) ] let rconstant gl term = let sigma = gl.sigma in let rec rconstant term = match EConstr.kind sigma term with | Const x -> if EConstr.eq_constr sigma term (Lazy.force coq_R0) then Mc.C0 else if EConstr.eq_constr sigma term (Lazy.force coq_R1) then Mc.C1 else raise ParseError | App (op, args) -> ( try (* the evaluation order is important in the following *) let f = assoc_const sigma op rconst_assoc in let a = rconstant args.(0) in let b = rconstant args.(1) in f a b with ParseError -> ( match op with | op when EConstr.eq_constr sigma op (Lazy.force coq_Rinv) -> let arg = rconstant args.(0) in if Mc.qeq_bool (Mc.q_of_Rcst arg) {Mc.qnum = Mc.Z0; Mc.qden = Mc.XH} then raise ParseError (* This is a division by zero -- no semantics *) else Mc.CInv arg | op when EConstr.eq_constr sigma op (Lazy.force coq_Rpower) -> Mc.CPow ( rconstant args.(0) , Mc.Inr (parse_more_constant nconstant gl args.(1)) ) | op when EConstr.eq_constr sigma op (Lazy.force coq_IQR) -> Mc.CQ (qconstant gl args.(0)) | op when EConstr.eq_constr sigma op (Lazy.force coq_IZR) -> Mc.CZ (parse_more_constant zconstant gl args.(0)) | _ -> raise ParseError ) ) | _ -> raise ParseError in rconstant term let rconstant gl term = if debug then Feedback.msg_debug ( Pp.str "rconstant: " ++ Printer.pr_leconstr_env gl.env gl.sigma term ++ fnl () ); let res = rconstant gl term in if debug then ( Printf.printf "rconstant -> %a\n" pp_Rcst res; flush stdout ); res let parse_qexpr gl = parse_expr gl qconstant (fun expr x -> let exp = zconstant gl x in match exp with | Mc.Zneg _ -> ( match expr with | Mc.PEc q -> Mc.PEc (Mc.qpower q exp) | _ -> raise ParseError ) | _ -> let exp = Mc.Z.to_N exp in Mc.PEpow (expr, exp)) qop_spec let parse_rexpr gl = parse_expr gl rconstant (fun expr x -> let exp = Mc.N.of_nat (parse_nat gl.sigma x) in Mc.PEpow (expr, exp)) rop_spec let parse_arith parse_op parse_expr env cstr gl = let sigma = gl.sigma in if debug then Feedback.msg_debug ( Pp.str "parse_arith: " ++ Printer.pr_leconstr_env gl.env sigma cstr ++ fnl () ); match EConstr.kind sigma cstr with | App (op, args) -> let op, lhs, rhs = parse_op gl (op, args) in let e1, env = parse_expr gl env lhs in let e2, env = parse_expr gl env rhs in ({Mc.flhs = e1; Mc.fop = op; Mc.frhs = e2}, env) | _ -> failwith "error : parse_arith(2)" let parse_zarith = parse_arith parse_zop parse_zexpr let parse_qarith = parse_arith parse_qop parse_qexpr let parse_rarith = parse_arith parse_rop parse_rexpr (* generic parsing of arithmetic expressions *) let mkC f1 f2 = Mc.Cj (f1, f2) let mkD f1 f2 = Mc.D (f1, f2) let mkIff f1 f2 = Mc.Cj (Mc.I (f1, None, f2), Mc.I (f2, None, f1)) let mkI f1 f2 = Mc.I (f1, None, f2) let mkformula_binary g term f1 f2 = match (f1, f2) with Mc.X _, Mc.X _ -> Mc.X term | _ -> g f1 f2 (** * This is the big generic function for formula parsers. *) let is_prop env sigma term = let sort = Retyping.get_sort_of env sigma term in Sorts.is_prop sort let parse_formula gl parse_atom env tg term = let sigma = gl.sigma in let is_prop term = is_prop gl.env gl.sigma term in let parse_atom env tg t = try let at, env = parse_atom env t gl in (Mc.A (at, (tg, t)), env, Tag.next tg) with ParseError -> if is_prop t then (Mc.X t, env, tg) else raise ParseError in let rec xparse_formula env tg term = match EConstr.kind sigma term with | App (l, rst) -> ( match rst with | [|a; b|] when EConstr.eq_constr sigma l (Lazy.force coq_and) -> let f, env, tg = xparse_formula env tg a in let g, env, tg = xparse_formula env tg b in (mkformula_binary mkC term f g, env, tg) | [|a; b|] when EConstr.eq_constr sigma l (Lazy.force coq_or) -> let f, env, tg = xparse_formula env tg a in let g, env, tg = xparse_formula env tg b in (mkformula_binary mkD term f g, env, tg) | [|a|] when EConstr.eq_constr sigma l (Lazy.force coq_not) -> let f, env, tg = xparse_formula env tg a in (Mc.N f, env, tg) | [|a; b|] when EConstr.eq_constr sigma l (Lazy.force coq_iff) -> let f, env, tg = xparse_formula env tg a in let g, env, tg = xparse_formula env tg b in (mkformula_binary mkIff term f g, env, tg) | _ -> parse_atom env tg term ) | Prod (typ, a, b) when typ.binder_name = Anonymous || EConstr.Vars.noccurn sigma 1 b -> let f, env, tg = xparse_formula env tg a in let g, env, tg = xparse_formula env tg b in (mkformula_binary mkI term f g, env, tg) | _ -> if EConstr.eq_constr sigma term (Lazy.force coq_True) then (Mc.TT, env, tg) else if EConstr.eq_constr sigma term (Lazy.force coq_False) then Mc.(FF, env, tg) else if is_prop term then (Mc.X term, env, tg) else raise ParseError in xparse_formula env tg (*Reductionops.whd_zeta*) term let dump_formula typ dump_atom f = let app_ctor c args = EConstr.mkApp ( Lazy.force c , Array.of_list ( typ :: EConstr.mkProp :: Lazy.force coq_unit :: Lazy.force coq_unit :: args ) ) in let rec xdump f = match f with | Mc.TT -> app_ctor coq_TT [] | Mc.FF -> app_ctor coq_FF [] | Mc.Cj (x, y) -> app_ctor coq_And [xdump x; xdump y] | Mc.D (x, y) -> app_ctor coq_Or [xdump x; xdump y] | Mc.I (x, _, y) -> app_ctor coq_Impl [ xdump x ; EConstr.mkApp (Lazy.force coq_None, [|Lazy.force coq_unit|]) ; xdump y ] | Mc.N x -> app_ctor coq_Neg [xdump x] | Mc.A (x, _) -> app_ctor coq_Atom [dump_atom x; Lazy.force coq_tt] | Mc.X t -> app_ctor coq_X [t] in xdump f let prop_env_of_formula gl form = Mc.( let rec doit env = function | TT | FF | A (_, _) -> env | X t -> fst (Env.compute_rank_add env t) | Cj (f1, f2) | D (f1, f2) | I (f1, _, f2) -> doit (doit env f1) f2 | N f -> doit env f in doit (Env.empty gl) form) let var_env_of_formula form = let rec vars_of_expr = function | Mc.PEX n -> ISet.singleton (CoqToCaml.positive n) | Mc.PEc z -> ISet.empty | Mc.PEadd (e1, e2) | Mc.PEmul (e1, e2) | Mc.PEsub (e1, e2) -> ISet.union (vars_of_expr e1) (vars_of_expr e2) | Mc.PEopp e | Mc.PEpow (e, _) -> vars_of_expr e in let vars_of_atom {Mc.flhs; Mc.fop; Mc.frhs} = ISet.union (vars_of_expr flhs) (vars_of_expr frhs) in Mc.( let rec doit = function | TT | FF | X _ -> ISet.empty | A (a, (t, c)) -> vars_of_atom a | Cj (f1, f2) | D (f1, f2) | I (f1, _, f2) -> ISet.union (doit f1) (doit f2) | N f -> doit f in doit form) type 'cst dump_expr = { (* 'cst is the type of the syntactic constants *) interp_typ : EConstr.constr ; dump_cst : 'cst -> EConstr.constr ; dump_add : EConstr.constr ; dump_sub : EConstr.constr ; dump_opp : EConstr.constr ; dump_mul : EConstr.constr ; dump_pow : EConstr.constr ; dump_pow_arg : Mc.n -> EConstr.constr ; dump_op : (Mc.op2 * EConstr.constr) list } let dump_zexpr = lazy { interp_typ = Lazy.force coq_Z ; dump_cst = dump_z ; dump_add = Lazy.force coq_Zplus ; dump_sub = Lazy.force coq_Zminus ; dump_opp = Lazy.force coq_Zopp ; dump_mul = Lazy.force coq_Zmult ; dump_pow = Lazy.force coq_Zpower ; dump_pow_arg = (fun n -> dump_z (CamlToCoq.z (CoqToCaml.n n))) ; dump_op = List.map (fun (x, y) -> (y, Lazy.force x)) zop_table } let dump_qexpr = lazy { interp_typ = Lazy.force coq_Q ; dump_cst = dump_q ; dump_add = Lazy.force coq_Qplus ; dump_sub = Lazy.force coq_Qminus ; dump_opp = Lazy.force coq_Qopp ; dump_mul = Lazy.force coq_Qmult ; dump_pow = Lazy.force coq_Qpower ; dump_pow_arg = (fun n -> dump_z (CamlToCoq.z (CoqToCaml.n n))) ; dump_op = List.map (fun (x, y) -> (y, Lazy.force x)) qop_table } let rec dump_Rcst_as_R cst = match cst with | Mc.C0 -> Lazy.force coq_R0 | Mc.C1 -> Lazy.force coq_R1 | Mc.CQ q -> EConstr.mkApp (Lazy.force coq_IQR, [|dump_q q|]) | Mc.CZ z -> EConstr.mkApp (Lazy.force coq_IZR, [|dump_z z|]) | Mc.CPlus (x, y) -> EConstr.mkApp (Lazy.force coq_Rplus, [|dump_Rcst_as_R x; dump_Rcst_as_R y|]) | Mc.CMinus (x, y) -> EConstr.mkApp (Lazy.force coq_Rminus, [|dump_Rcst_as_R x; dump_Rcst_as_R y|]) | Mc.CMult (x, y) -> EConstr.mkApp (Lazy.force coq_Rmult, [|dump_Rcst_as_R x; dump_Rcst_as_R y|]) | Mc.CPow (x, y) -> ( match y with | Mc.Inl z -> EConstr.mkApp (Lazy.force coq_powerZR, [|dump_Rcst_as_R x; dump_z z|]) | Mc.Inr n -> EConstr.mkApp (Lazy.force coq_Rpower, [|dump_Rcst_as_R x; dump_nat n|]) ) | Mc.CInv t -> EConstr.mkApp (Lazy.force coq_Rinv, [|dump_Rcst_as_R t|]) | Mc.COpp t -> EConstr.mkApp (Lazy.force coq_Ropp, [|dump_Rcst_as_R t|]) let dump_rexpr = lazy { interp_typ = Lazy.force coq_R ; dump_cst = dump_Rcst_as_R ; dump_add = Lazy.force coq_Rplus ; dump_sub = Lazy.force coq_Rminus ; dump_opp = Lazy.force coq_Ropp ; dump_mul = Lazy.force coq_Rmult ; dump_pow = Lazy.force coq_Rpower ; dump_pow_arg = (fun n -> dump_nat (CamlToCoq.nat (CoqToCaml.n n))) ; dump_op = List.map (fun (x, y) -> (y, Lazy.force x)) rop_table } let prodn n env b = let rec prodrec = function | 0, env, b -> b | n, (v, t) :: l, b -> prodrec (n - 1, l, EConstr.mkProd (make_annot v Sorts.Relevant, t, b)) | _ -> assert false in prodrec (n, env, b) (** [make_goal_of_formula depxr vars props form] where - vars is an environment for the arithmetic variables occurring in form - props is an environment for the propositions occurring in form @return a goal where all the variables and propositions of the formula are quantified *) let make_goal_of_formula gl dexpr form = let vars_idx = List.mapi (fun i v -> (v, i + 1)) (ISet.elements (var_env_of_formula form)) in (* List.iter (fun (v,i) -> Printf.fprintf stdout "var %i has index %i\n" v i) vars_idx ;*) let props = prop_env_of_formula gl form in let fresh_var str i = Names.Id.of_string (str ^ string_of_int i) in let fresh_prop str i = Names.Id.of_string (str ^ string_of_int i) in let vars_n = List.map (fun (_, i) -> (fresh_var "__x" i, dexpr.interp_typ)) vars_idx in let props_n = List.mapi (fun i _ -> (fresh_prop "__p" (i + 1), EConstr.mkProp)) (Env.elements props) in let var_name_pos = List.map2 (fun (idx, _) (id, _) -> (id, idx)) vars_idx vars_n in let dump_expr i e = let rec dump_expr = function | Mc.PEX n -> EConstr.mkRel (i + List.assoc (CoqToCaml.positive n) vars_idx) | Mc.PEc z -> dexpr.dump_cst z | Mc.PEadd (e1, e2) -> EConstr.mkApp (dexpr.dump_add, [|dump_expr e1; dump_expr e2|]) | Mc.PEsub (e1, e2) -> EConstr.mkApp (dexpr.dump_sub, [|dump_expr e1; dump_expr e2|]) | Mc.PEopp e -> EConstr.mkApp (dexpr.dump_opp, [|dump_expr e|]) | Mc.PEmul (e1, e2) -> EConstr.mkApp (dexpr.dump_mul, [|dump_expr e1; dump_expr e2|]) | Mc.PEpow (e, n) -> EConstr.mkApp (dexpr.dump_pow, [|dump_expr e; dexpr.dump_pow_arg n|]) in dump_expr e in let mkop op e1 e2 = try EConstr.mkApp (List.assoc op dexpr.dump_op, [|e1; e2|]) with Not_found -> EConstr.mkApp (Lazy.force coq_Eq, [|dexpr.interp_typ; e1; e2|]) in let dump_cstr i {Mc.flhs; Mc.fop; Mc.frhs} = mkop fop (dump_expr i flhs) (dump_expr i frhs) in let rec xdump pi xi f = match f with | Mc.TT -> Lazy.force coq_True | Mc.FF -> Lazy.force coq_False | Mc.Cj (x, y) -> EConstr.mkApp (Lazy.force coq_and, [|xdump pi xi x; xdump pi xi y|]) | Mc.D (x, y) -> EConstr.mkApp (Lazy.force coq_or, [|xdump pi xi x; xdump pi xi y|]) | Mc.I (x, _, y) -> EConstr.mkArrow (xdump pi xi x) Sorts.Relevant (xdump (pi + 1) (xi + 1) y) | Mc.N x -> EConstr.mkArrow (xdump pi xi x) Sorts.Relevant (Lazy.force coq_False) | Mc.A (x, _) -> dump_cstr xi x | Mc.X t -> let idx = Env.get_rank props t in EConstr.mkRel (pi + idx) in let nb_vars = List.length vars_n in let nb_props = List.length props_n in (* Printf.fprintf stdout "NBProps : %i\n" nb_props ;*) let subst_prop p = let idx = Env.get_rank props p in EConstr.mkVar (Names.Id.of_string (Printf.sprintf "__p%i" idx)) in let form' = Mc.mapX subst_prop form in ( prodn nb_props (List.map (fun (x, y) -> (Name.Name x, y)) props_n) (prodn nb_vars (List.map (fun (x, y) -> (Name.Name x, y)) vars_n) (xdump (List.length vars_n) 0 form)) , List.rev props_n , List.rev var_name_pos , form' ) (** * Given a conclusion and a list of affectations, rebuild a term prefixed by * the appropriate letins. * TODO: reverse the list of bindings! *) let set l concl = let rec xset acc = function | [] -> acc | e :: l -> let name, expr, typ = e in xset (EConstr.mkNamedLetIn (make_annot (Names.Id.of_string name) Sorts.Relevant) expr typ acc) l in xset concl l end open M let coq_Branch = lazy (gen_constant_in_modules "VarMap" [["Coq"; "micromega"; "VarMap"]; ["VarMap"]] "Branch") let coq_Elt = lazy (gen_constant_in_modules "VarMap" [["Coq"; "micromega"; "VarMap"]; ["VarMap"]] "Elt") let coq_Empty = lazy (gen_constant_in_modules "VarMap" [["Coq"; "micromega"; "VarMap"]; ["VarMap"]] "Empty") let coq_VarMap = lazy (gen_constant_in_modules "VarMap" [["Coq"; "micromega"; "VarMap"]; ["VarMap"]] "t") let rec dump_varmap typ m = match m with | Mc.Empty -> EConstr.mkApp (Lazy.force coq_Empty, [|typ|]) | Mc.Elt v -> EConstr.mkApp (Lazy.force coq_Elt, [|typ; v|]) | Mc.Branch (l, o, r) -> EConstr.mkApp (Lazy.force coq_Branch, [|typ; dump_varmap typ l; o; dump_varmap typ r|]) let vm_of_list env = match env with | [] -> Mc.Empty | (d, _) :: _ -> List.fold_left (fun vm (c, i) -> Mc.vm_add d (CamlToCoq.positive i) c vm) Mc.Empty env let rec dump_proof_term = function | Micromega.DoneProof -> Lazy.force coq_doneProof | Micromega.RatProof (cone, rst) -> EConstr.mkApp ( Lazy.force coq_ratProof , [|dump_psatz coq_Z dump_z cone; dump_proof_term rst|] ) | Micromega.CutProof (cone, prf) -> EConstr.mkApp ( Lazy.force coq_cutProof , [|dump_psatz coq_Z dump_z cone; dump_proof_term prf|] ) | Micromega.EnumProof (c1, c2, prfs) -> EConstr.mkApp ( Lazy.force coq_enumProof , [| dump_psatz coq_Z dump_z c1 ; dump_psatz coq_Z dump_z c2 ; dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |] ) | Micromega.ExProof (p, prf) -> EConstr.mkApp (Lazy.force coq_ExProof, [|dump_positive p; dump_proof_term prf|]) let rec size_of_psatz = function | Micromega.PsatzIn _ -> 1 | Micromega.PsatzSquare _ -> 1 | Micromega.PsatzMulC (_, p) -> 1 + size_of_psatz p | Micromega.PsatzMulE (p1, p2) | Micromega.PsatzAdd (p1, p2) -> size_of_psatz p1 + size_of_psatz p2 | Micromega.PsatzC _ -> 1 | Micromega.PsatzZ -> 1 let rec size_of_pf = function | Micromega.DoneProof -> 1 | Micromega.RatProof (p, a) -> size_of_pf a + size_of_psatz p | Micromega.CutProof (p, a) -> size_of_pf a + size_of_psatz p | Micromega.EnumProof (p1, p2, l) -> size_of_psatz p1 + size_of_psatz p2 + List.fold_left (fun acc p -> size_of_pf p + acc) 0 l | Micromega.ExProof (_, a) -> size_of_pf a + 1 let dump_proof_term t = if debug then Printf.printf "dump_proof_term %i\n" (size_of_pf t); dump_proof_term t let pp_q o q = Printf.fprintf o "%a/%a" pp_z q.Micromega.qnum pp_positive q.Micromega.qden let rec pp_proof_term o = function | Micromega.DoneProof -> Printf.fprintf o "D" | Micromega.RatProof (cone, rst) -> Printf.fprintf o "R[%a,%a]" (pp_psatz pp_z) cone pp_proof_term rst | Micromega.CutProof (cone, rst) -> Printf.fprintf o "C[%a,%a]" (pp_psatz pp_z) cone pp_proof_term rst | Micromega.EnumProof (c1, c2, rst) -> Printf.fprintf o "EP[%a,%a,%a]" (pp_psatz pp_z) c1 (pp_psatz pp_z) c2 (pp_list "[" "]" pp_proof_term) rst | Micromega.ExProof (p, prf) -> Printf.fprintf o "Ex[%a,%a]" pp_positive p pp_proof_term prf let rec parse_hyps gl parse_arith env tg hyps = match hyps with | [] -> ([], env, tg) | (i, t) :: l -> let lhyps, env, tg = parse_hyps gl parse_arith env tg l in if is_prop gl.env gl.sigma t then try let c, env, tg = parse_formula gl parse_arith env tg t in ((i, c) :: lhyps, env, tg) with ParseError -> (lhyps, env, tg) else (lhyps, env, tg) let parse_goal gl parse_arith (env : Env.t) hyps term = let f, env, tg = parse_formula gl parse_arith env (Tag.from 0) term in let lhyps, env, tg = parse_hyps gl parse_arith env tg hyps in (lhyps, f, env) type ('synt_c, 'prf) domain_spec = { typ : EConstr.constr ; (* is the type of the interpretation domain - Z, Q, R*) coeff : EConstr.constr ; (* is the type of the syntactic coeffs - Z , Q , Rcst *) dump_coeff : 'synt_c -> EConstr.constr ; proof_typ : EConstr.constr ; dump_proof : 'prf -> EConstr.constr } (** * The datastructures that aggregate theory-dependent proof values. *) let zz_domain_spec = lazy { typ = Lazy.force coq_Z ; coeff = Lazy.force coq_Z ; dump_coeff = dump_z ; proof_typ = Lazy.force coq_proofTerm ; dump_proof = dump_proof_term } let qq_domain_spec = lazy { typ = Lazy.force coq_Q ; coeff = Lazy.force coq_Q ; dump_coeff = dump_q ; proof_typ = Lazy.force coq_QWitness ; dump_proof = dump_psatz coq_Q dump_q } let max_tag f = 1 + Tag.to_int (Mc.foldA (fun t1 (t2, _) -> Tag.max t1 t2) f (Tag.from 0)) (** Naive topological sort of constr according to the subterm-ordering *) (* An element is minimal x is minimal w.r.t y if x <= y or (x and y are incomparable) *) (** * Instantiate the current Coq goal with a Micromega formula, a varmap, and a * witness. *) let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic*) = (* let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__v"^(string_of_int i)))) 0 env in *) let formula_typ = EConstr.mkApp (Lazy.force coq_Cstr, [|spec.coeff|]) in let ff = dump_formula formula_typ (dump_cstr spec.coeff spec.dump_coeff) ff in let vm = dump_varmap spec.typ (vm_of_list env) in (* todo : directly generate the proof term - or generalize before conversion? *) Proofview.Goal.enter (fun gl -> Tacticals.New.tclTHENLIST [ Tactics.change_concl (set [ ( "__ff" , ff , EConstr.mkApp (Lazy.force coq_Formula, [|formula_typ|]) ) ; ( "__varmap" , vm , EConstr.mkApp (Lazy.force coq_VarMap, [|spec.typ|]) ) ; ("__wit", cert, cert_typ) ] (Tacmach.New.pf_concl gl)) ]) (** * The datastructures that aggregate prover attributes. *) open Certificate type ('option, 'a, 'prf, 'model) prover = { name : string ; (* name of the prover *) get_option : unit -> 'option ; (* find the options of the prover *) prover : 'option * 'a list -> ('prf, 'model) Certificate.res ; (* the prover itself *) hyps : 'prf -> ISet.t ; (* extract the indexes of the hypotheses really used in the proof *) compact : 'prf -> (int -> int) -> 'prf ; (* remap the hyp indexes according to function *) pp_prf : out_channel -> 'prf -> unit ; (* pretting printing of proof *) pp_f : out_channel -> 'a -> unit (* pretty printing of the formulas (polynomials)*) } (** * Given a prover and a disjunction of atoms, find a proof of any of * the atoms. Returns an (optional) pair of a proof and a prover * datastructure. *) let find_witness p polys1 = let polys1 = List.map fst polys1 in match p.prover (p.get_option (), polys1) with | Model m -> Model m | Unknown -> Unknown | Prf prf -> Prf (prf, p) (** * Given a prover and a CNF, find a proof for each of the clauses. * Return the proofs as a list. *) let witness_list prover l = let rec xwitness_list l = match l with | [] -> Prf [] | e :: l -> ( match xwitness_list l with | Model (m, e) -> Model (m, e) | Unknown -> Unknown | Prf l -> ( match find_witness prover e with | Model m -> Model (m, e) | Unknown -> Unknown | Prf w -> Prf (w :: l) ) ) in xwitness_list l let witness_list_tags p g = witness_list p g (* let t1 = System.get_time () in let res = witness_list p g in let t2 = System.get_time () in Feedback.msg_info Pp.(str "Witness generation "++int (List.length g) ++ str " "++System.fmt_time_difference t1 t2) ; res *) (** * Prune the proof object, according to the 'diff' between two cnf formulas. *) let compact_proofs (cnf_ff : 'cst cnf) res (cnf_ff' : 'cst cnf) = let compact_proof (old_cl : 'cst clause) (prf, prover) (new_cl : 'cst clause) = let new_cl = List.mapi (fun i (f, _) -> (f, i)) new_cl in let remap i = let formula = try fst (List.nth old_cl i) with Failure _ -> failwith "bad old index" in List.assoc formula new_cl in (* if debug then begin Printf.printf "\ncompact_proof : %a %a %a" (pp_ml_list prover.pp_f) (List.map fst old_cl) prover.pp_prf prf (pp_ml_list prover.pp_f) (List.map fst new_cl) ; flush stdout end ; *) let res = try prover.compact prf remap with x when CErrors.noncritical x -> ( if debug then Printf.fprintf stdout "Proof compaction %s" (Printexc.to_string x); (* This should not happen -- this is the recovery plan... *) match prover.prover (prover.get_option (), List.map fst new_cl) with | Unknown | Model _ -> failwith "proof compaction error" | Prf p -> p ) in if debug then begin Printf.printf " -> %a\n" prover.pp_prf res; flush stdout end; res in let is_proof_compatible (old_cl : 'cst clause) (prf, prover) (new_cl : 'cst clause) = let hyps_idx = prover.hyps prf in let hyps = selecti hyps_idx old_cl in is_sublist ( = ) hyps new_cl in let cnf_res = List.combine cnf_ff res in (* we get pairs clause * proof *) if debug then begin Printf.printf "CNFRES\n"; flush stdout; Printf.printf "CNFOLD %a\n" pp_cnf_tag cnf_ff; List.iter (fun (cl, (prf, prover)) -> let hyps_idx = prover.hyps prf in let hyps = selecti hyps_idx cl in Printf.printf "\nProver %a -> %a\n" pp_clause_tag cl pp_clause_tag hyps; flush stdout) cnf_res; Printf.printf "CNFNEW %a\n" pp_cnf_tag cnf_ff' end; List.map (fun x -> let o, p = try List.find (fun (l, p) -> is_proof_compatible l p x) cnf_res with Not_found -> Printf.printf "ERROR: no compatible proof"; flush stdout; failwith "Cannot find compatible proof" in compact_proof o p x) cnf_ff' (** * "Hide out" tagged atoms of a formula by transforming them into generic * variables. See the Tag module in mutils.ml for more. *) let abstract_formula : TagSet.t -> 'a formula -> 'a formula = fun hyps f -> let to_constr = Mc. { mkTT = Lazy.force coq_True ; mkFF = Lazy.force coq_False ; mkA = (fun a (tg, t) -> t) ; mkCj = (let coq_and = Lazy.force coq_and in fun x y -> EConstr.mkApp (coq_and, [|x; y|])) ; mkD = (let coq_or = Lazy.force coq_or in fun x y -> EConstr.mkApp (coq_or, [|x; y|])) ; mkI = (fun x y -> EConstr.mkArrow x Sorts.Relevant y) ; mkN = (let coq_not = Lazy.force coq_not in fun x -> EConstr.mkApp (coq_not, [|x|])) } in Mc.abst_form to_constr (fun (t, _) -> TagSet.mem t hyps) true f (* [abstract_wrt_formula] is used in contexts whre f1 is already an abstraction of f2 *) let rec abstract_wrt_formula f1 f2 = Mc.( match (f1, f2) with | X c, _ -> X c | A _, A _ -> f2 | Cj (a, b), Cj (a', b') -> Cj (abstract_wrt_formula a a', abstract_wrt_formula b b') | D (a, b), D (a', b') -> D (abstract_wrt_formula a a', abstract_wrt_formula b b') | I (a, _, b), I (a', x, b') -> I (abstract_wrt_formula a a', x, abstract_wrt_formula b b') | FF, FF -> FF | TT, TT -> TT | N x, N y -> N (abstract_wrt_formula x y) | _ -> failwith "abstract_wrt_formula") (** * This exception is raised by really_call_csdpcert if Coq's configure didn't * find a CSDP executable. *) exception CsdpNotFound (** * This is the core of Micromega: apply the prover, analyze the result and * prune unused fomulas, and finally modify the proof state. *) let formula_hyps_concl hyps concl = List.fold_right (fun (id, f) (cc, ids) -> match f with | Mc.X _ -> (cc, ids) | _ -> (Mc.I (f, Some id, cc), id :: ids)) hyps (concl, []) (* let time str f x = let t1 = System.get_time () in let res = f x in let t2 = System.get_time () in Feedback.msg_info (Pp.str str ++ Pp.str " " ++ System.fmt_time_difference t1 t2) ; res *) let micromega_tauto pre_process cnf spec prover env (polys1 : (Names.Id.t * 'cst formula) list) (polys2 : 'cst formula) gl = (* Express the goal as one big implication *) let ff, ids = formula_hyps_concl polys1 polys2 in let mt = CamlToCoq.positive (max_tag ff) in (* Construction of cnf *) let pre_ff = pre_process mt (ff : 'a formula) in let cnf_ff, cnf_ff_tags = cnf pre_ff in match witness_list_tags prover cnf_ff with | Model m -> Model m | Unknown -> Unknown | Prf res -> (*Printf.printf "\nList %i" (List.length `res); *) let deps = List.fold_left (fun s (cl, (prf, p)) -> let tags = ISet.fold (fun i s -> let t = fst (snd (List.nth cl i)) in if debug then Printf.fprintf stdout "T : %i -> %a" i Tag.pp t; (*try*) TagSet.add t s (* with Invalid_argument _ -> s*)) (p.hyps prf) TagSet.empty in TagSet.union s tags) (List.fold_left (fun s (i, _) -> TagSet.add i s) TagSet.empty cnf_ff_tags) (List.combine cnf_ff res) in let ff' = abstract_formula deps ff in let pre_ff' = pre_process mt ff' in let cnf_ff', _ = cnf pre_ff' in if debug then begin output_string stdout "\n"; Printf.printf "TForm : %a\n" pp_formula ff; flush stdout; Printf.printf "CNF : %a\n" pp_cnf_tag cnf_ff; flush stdout; Printf.printf "TFormAbs : %a\n" pp_formula ff'; flush stdout; Printf.printf "TFormPre : %a\n" pp_formula pre_ff; flush stdout; Printf.printf "TFormPreAbs : %a\n" pp_formula pre_ff'; flush stdout; Printf.printf "CNF : %a\n" pp_cnf_tag cnf_ff'; flush stdout end; (* Even if it does not work, this does not mean it is not provable -- the prover is REALLY incomplete *) (* if debug then begin (* recompute the proofs *) match witness_list_tags prover cnf_ff' with | None -> failwith "abstraction is wrong" | Some res -> () end ; *) let res' = compact_proofs cnf_ff res cnf_ff' in let ff', res', ids = (ff', res', Mc.ids_of_formula ff') in let res' = dump_list spec.proof_typ spec.dump_proof res' in Prf (ids, ff', res') let micromega_tauto pre_process cnf spec prover env (polys1 : (Names.Id.t * 'cst formula) list) (polys2 : 'cst formula) gl = try micromega_tauto pre_process cnf spec prover env polys1 polys2 gl with Not_found -> Printexc.print_backtrace stdout; flush stdout; Unknown (** * Parse the proof environment, and call micromega_tauto *) let fresh_id avoid id gl = Tactics.fresh_id_in_env avoid id (Proofview.Goal.env gl) let clear_all_no_check = Proofview.Goal.enter (fun gl -> let concl = Tacmach.New.pf_concl gl in let env = Environ.reset_with_named_context Environ.empty_named_context_val (Tacmach.New.pf_env gl) in Refine.refine ~typecheck:false (fun sigma -> Evarutil.new_evar env sigma ~principal:true concl)) let micromega_gen parse_arith pre_process cnf spec dumpexpr prover tac = Proofview.Goal.enter (fun gl -> let sigma = Tacmach.New.project gl in let concl = Tacmach.New.pf_concl gl in let hyps = Tacmach.New.pf_hyps_types gl in try let gl0 = {env = Tacmach.New.pf_env gl; sigma} in let hyps, concl, env = parse_goal gl0 parse_arith (Env.empty gl0) hyps concl in let env = Env.elements env in let spec = Lazy.force spec in let dumpexpr = Lazy.force dumpexpr in if debug then Feedback.msg_debug (Pp.str "Env " ++ Env.pp gl0 env); match micromega_tauto pre_process cnf spec prover env hyps concl gl0 with | Unknown -> flush stdout; Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Model (m, e) -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Prf (ids, ff', res') -> let arith_goal, props, vars, ff_arith = make_goal_of_formula gl0 dumpexpr ff' in let intro (id, _) = Tactics.introduction id in let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in (* let ipat_of_name id = Some (CAst.make @@ IntroNaming (Namegen.IntroIdentifier id)) in*) let goal_name = fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id, i) -> (EConstr.mkVar id, i)) vars in let tac_arith = Tacticals.New.tclTHENLIST [ clear_all_no_check ; intro_props ; intro_vars ; micromega_order_change spec res' (EConstr.mkApp (Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in let goal_props = List.rev (Env.elements (prop_env_of_formula gl0 ff')) in let goal_vars = List.map (fun (_, i) -> List.nth env (i - 1)) vars in let arith_args = goal_props @ goal_vars in let kill_arith = Tacticals.New.tclTHEN tac_arith tac in (* (*tclABSTRACT fails in certain corner cases.*) Tacticals.New.tclTHEN clear_all_no_check (Abstract.tclABSTRACT ~opaque:false None (Tacticals.New.tclTHEN tac_arith tac)) in *) Tacticals.New.tclTHEN (Tactics.assert_by (Names.Name goal_name) arith_goal (*Proofview.tclTIME (Some "kill_arith")*) kill_arith) ((*Proofview.tclTIME (Some "apply_arith") *) Tactics.exact_check (EConstr.applist ( EConstr.mkVar goal_name , arith_args @ List.map EConstr.mkVar ids ))) with | Mfourier.TimeOut -> Tacticals.New.tclFAIL 0 (Pp.str "Timeout") | CsdpNotFound -> flush stdout; Tacticals.New.tclFAIL 0 (Pp.str ( " Skipping what remains of this tactic: the complexity of the \ goal requires " ^ "the use of a specialized external tool called csdp. \n\n" ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" \ executable in the path. \n\n" ^ "Csdp packages are provided by some OS distributions; binaries \ and source code can be downloaded from \ https://projects.coin-or.org/Csdp" )) | x -> if debug then Tacticals.New.tclFAIL 0 (Pp.str (Printexc.get_backtrace ())) else raise x) let micromega_order_changer cert env ff = (*let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__v"^(string_of_int i)))) 0 env in *) let coeff = Lazy.force coq_Rcst in let dump_coeff = dump_Rcst in let typ = Lazy.force coq_R in let cert_typ = EConstr.mkApp (Lazy.force coq_list, [|Lazy.force coq_QWitness|]) in let formula_typ = EConstr.mkApp (Lazy.force coq_Cstr, [|coeff|]) in let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in let vm = dump_varmap typ (vm_of_list env) in Proofview.Goal.enter (fun gl -> Tacticals.New.tclTHENLIST [ Tactics.change_concl (set [ ( "__ff" , ff , EConstr.mkApp (Lazy.force coq_Formula, [|formula_typ|]) ) ; ( "__varmap" , vm , EConstr.mkApp ( gen_constant_in_modules "VarMap" [["Coq"; "micromega"; "VarMap"]; ["VarMap"]] "t" , [|typ|] ) ) ; ("__wit", cert, cert_typ) ] (Tacmach.New.pf_concl gl)) (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*) ]) let micromega_genr prover tac = let parse_arith = parse_rarith in let spec = lazy { typ = Lazy.force coq_R ; coeff = Lazy.force coq_Rcst ; dump_coeff = dump_q ; proof_typ = Lazy.force coq_QWitness ; dump_proof = dump_psatz coq_Q dump_q } in Proofview.Goal.enter (fun gl -> let sigma = Tacmach.New.project gl in let concl = Tacmach.New.pf_concl gl in let hyps = Tacmach.New.pf_hyps_types gl in try let gl0 = {env = Tacmach.New.pf_env gl; sigma} in let hyps, concl, env = parse_goal gl0 parse_arith (Env.empty gl0) hyps concl in let env = Env.elements env in let spec = Lazy.force spec in let hyps' = List.map (fun (n, f) -> (n, Mc.map_bformula (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in let concl' = Mc.map_bformula (Micromega.map_Formula Micromega.q_of_Rcst) concl in match micromega_tauto (fun _ x -> x) Mc.cnfQ spec prover env hyps' concl' gl0 with | Unknown | Model _ -> flush stdout; Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Prf (ids, ff', res') -> let ff, ids = formula_hyps_concl (List.filter (fun (n, _) -> List.mem n ids) hyps) concl in let ff' = abstract_wrt_formula ff' ff in let arith_goal, props, vars, ff_arith = make_goal_of_formula gl0 (Lazy.force dump_rexpr) ff' in let intro (id, _) = Tactics.introduction id in let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in let ipat_of_name id = Some (CAst.make @@ IntroNaming (Namegen.IntroIdentifier id)) in let goal_name = fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id, i) -> (EConstr.mkVar id, i)) vars in let tac_arith = Tacticals.New.tclTHENLIST [ clear_all_no_check ; intro_props ; intro_vars ; micromega_order_changer res' env' ff_arith ] in let goal_props = List.rev (Env.elements (prop_env_of_formula gl0 ff')) in let goal_vars = List.map (fun (_, i) -> List.nth env (i - 1)) vars in let arith_args = goal_props @ goal_vars in let kill_arith = Tacticals.New.tclTHEN tac_arith tac in (* Tacticals.New.tclTHEN (Tactics.keep []) (Tactics.tclABSTRACT None*) Tacticals.New.tclTHENS (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal) [ kill_arith ; Tacticals.New.tclTHENLIST [ Tactics.generalize (List.map EConstr.mkVar ids) ; Tactics.exact_check (EConstr.applist (EConstr.mkVar goal_name, arith_args)) ] ] with | Mfourier.TimeOut -> Tacticals.New.tclFAIL 0 (Pp.str "Timeout") | CsdpNotFound -> flush stdout; Tacticals.New.tclFAIL 0 (Pp.str ( " Skipping what remains of this tactic: the complexity of the \ goal requires " ^ "the use of a specialized external tool called csdp. \n\n" ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" \ executable in the path. \n\n" ^ "Csdp packages are provided by some OS distributions; binaries \ and source code can be downloaded from \ https://projects.coin-or.org/Csdp" ))) let lift_ratproof prover l = match prover l with | Unknown | Model _ -> Unknown | Prf c -> Prf (Mc.RatProof (c, Mc.DoneProof)) type micromega_polys = (Micromega.q Mc.pol * Mc.op1) list [@@@ocaml.warning "-37"] type csdp_certificate = S of Sos_types.positivstellensatz option | F of string (* Used to read the result of the execution of csdpcert *) type provername = string * int option (** * The caching mechanism. *) open Persistent_cache module MakeCache (T : sig type prover_option type coeff val hash_prover_option : int -> prover_option -> int val hash_coeff : int -> coeff -> int val eq_prover_option : prover_option -> prover_option -> bool val eq_coeff : coeff -> coeff -> bool end) = struct module E = struct type t = T.prover_option * (T.coeff Mc.pol * Mc.op1) list let equal = Hash.( eq_pair T.eq_prover_option (CList.equal (eq_pair (eq_pol T.eq_coeff) Hash.eq_op1))) let hash = let hash_cstr = Hash.(hash_pair (hash_pol T.hash_coeff) hash_op1) in Hash.((hash_pair T.hash_prover_option (List.fold_left hash_cstr)) 0) end include PHashtable (E) let memo_opt use_cache cache_file f = let memof = memo cache_file f in fun x -> if !use_cache then memof x else f x end module CacheCsdp = MakeCache (struct type prover_option = provername type coeff = Mc.q let hash_prover_option = Hash.(hash_pair hash_string (hash_elt (Option.hash (fun x -> x)))) let eq_prover_option = Hash.(eq_pair String.equal (Option.equal Int.equal)) let hash_coeff = Hash.hash_q let eq_coeff = Hash.eq_q end) (** * Build the command to call csdpcert, and launch it. This in turn will call * the sos driver to the csdp executable. * Throw CsdpNotFound if Coq isn't aware of any csdp executable. *) let require_csdp = if System.is_in_system_path "csdp" then lazy () else lazy (raise CsdpNotFound) let really_call_csdpcert : provername -> micromega_polys -> Sos_types.positivstellensatz option = fun provername poly -> Lazy.force require_csdp; let cmdname = List.fold_left Filename.concat (Envars.coqlib ()) ["plugins"; "micromega"; "csdpcert" ^ Coq_config.exec_extension] in match (command cmdname [|cmdname|] (provername, poly) : csdp_certificate) with | F str -> if debug then Printf.fprintf stdout "really_call_csdpcert : %s\n" str; raise (failwith str) | S res -> res (** * Check the cache before calling the prover. *) let xcall_csdpcert = CacheCsdp.memo_opt use_csdp_cache ".csdp.cache" (fun (prover, pb) -> really_call_csdpcert prover pb) (** * Prover callback functions. *) let call_csdpcert prover pb = xcall_csdpcert (prover, pb) let rec z_to_q_pol e = match e with | Mc.Pc z -> Mc.Pc {Mc.qnum = z; Mc.qden = Mc.XH} | Mc.Pinj (p, pol) -> Mc.Pinj (p, z_to_q_pol pol) | Mc.PX (pol1, p, pol2) -> Mc.PX (z_to_q_pol pol1, p, z_to_q_pol pol2) let call_csdpcert_q provername poly = match call_csdpcert provername poly with | None -> Unknown | Some cert -> let cert = Certificate.q_cert_of_pos cert in if Mc.qWeakChecker poly cert then Prf cert else ( print_string "buggy certificate"; Unknown ) let call_csdpcert_z provername poly = let l = List.map (fun (e, o) -> (z_to_q_pol e, o)) poly in match call_csdpcert provername l with | None -> Unknown | Some cert -> let cert = Certificate.z_cert_of_pos cert in if Mc.zWeakChecker poly cert then Prf cert else ( print_string "buggy certificate"; flush stdout; Unknown ) let xhyps_of_cone base acc prf = let rec xtract e acc = match e with | Mc.PsatzC _ | Mc.PsatzZ | Mc.PsatzSquare _ -> acc | Mc.PsatzIn n -> let n = CoqToCaml.nat n in if n >= base then ISet.add (n - base) acc else acc | Mc.PsatzMulC (_, c) -> xtract c acc | Mc.PsatzAdd (e1, e2) | Mc.PsatzMulE (e1, e2) -> xtract e1 (xtract e2 acc) in xtract prf acc let hyps_of_cone prf = xhyps_of_cone 0 ISet.empty prf let compact_cone prf f = let np n = CamlToCoq.nat (f (CoqToCaml.nat n)) in let rec xinterp prf = match prf with | Mc.PsatzC _ | Mc.PsatzZ | Mc.PsatzSquare _ -> prf | Mc.PsatzIn n -> Mc.PsatzIn (np n) | Mc.PsatzMulC (e, c) -> Mc.PsatzMulC (e, xinterp c) | Mc.PsatzAdd (e1, e2) -> Mc.PsatzAdd (xinterp e1, xinterp e2) | Mc.PsatzMulE (e1, e2) -> Mc.PsatzMulE (xinterp e1, xinterp e2) in xinterp prf let hyps_of_pt pt = let rec xhyps base pt acc = match pt with | Mc.DoneProof -> acc | Mc.RatProof (c, pt) -> xhyps (base + 1) pt (xhyps_of_cone base acc c) | Mc.CutProof (c, pt) -> xhyps (base + 1) pt (xhyps_of_cone base acc c) | Mc.EnumProof (c1, c2, l) -> let s = xhyps_of_cone base (xhyps_of_cone base acc c2) c1 in List.fold_left (fun s x -> xhyps (base + 1) x s) s l | Mc.ExProof (_, pt) -> xhyps (base + 3) pt acc in xhyps 0 pt ISet.empty let compact_pt pt f = let translate ofset x = if x < ofset then x else f (x - ofset) + ofset in let rec compact_pt ofset pt = match pt with | Mc.DoneProof -> Mc.DoneProof | Mc.RatProof (c, pt) -> Mc.RatProof (compact_cone c (translate ofset), compact_pt (ofset + 1) pt) | Mc.CutProof (c, pt) -> Mc.CutProof (compact_cone c (translate ofset), compact_pt (ofset + 1) pt) | Mc.EnumProof (c1, c2, l) -> Mc.EnumProof ( compact_cone c1 (translate ofset) , compact_cone c2 (translate ofset) , Mc.map (fun x -> compact_pt (ofset + 1) x) l ) | Mc.ExProof (x, pt) -> Mc.ExProof (x, compact_pt (ofset + 3) pt) in compact_pt 0 pt (** * Definition of provers. * Instantiates the type ('a,'prf) prover defined above. *) let lift_pexpr_prover p l = p (List.map (fun (e, o) -> (Mc.denorm e, o)) l) module CacheZ = MakeCache (struct type prover_option = bool * bool * int type coeff = Mc.z let hash_prover_option : int -> prover_option -> int = Hash.hash_elt Hashtbl.hash let eq_prover_option : prover_option -> prover_option -> bool = ( = ) let eq_coeff = Hash.eq_z let hash_coeff = Hash.hash_z end) module CacheQ = MakeCache (struct type prover_option = int type coeff = Mc.q let hash_prover_option : int -> int -> int = Hash.hash_elt Hashtbl.hash let eq_prover_option = Int.equal let eq_coeff = Hash.eq_q let hash_coeff = Hash.hash_q end) let memo_lia = CacheZ.memo_opt use_lia_cache ".lia.cache" (fun ((_, ce, b), s) -> lift_pexpr_prover (Certificate.lia ce b) s) let memo_nlia = CacheZ.memo_opt use_nia_cache ".nia.cache" (fun ((_, ce, b), s) -> lift_pexpr_prover (Certificate.nlia ce b) s) let memo_nra = CacheQ.memo_opt use_nra_cache ".nra.cache" (fun (o, s) -> lift_pexpr_prover (Certificate.nlinear_prover o) s) let linear_prover_Q = { name = "linear prover" ; get_option = get_lra_option ; prover = (fun (o, l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o) l) ; hyps = hyps_of_cone ; compact = compact_cone ; pp_prf = pp_psatz pp_q ; pp_f = (fun o x -> pp_pol pp_q o (fst x)) } let linear_prover_R = { name = "linear prover" ; get_option = get_lra_option ; prover = (fun (o, l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o) l) ; hyps = hyps_of_cone ; compact = compact_cone ; pp_prf = pp_psatz pp_q ; pp_f = (fun o x -> pp_pol pp_q o (fst x)) } let nlinear_prover_R = { name = "nra" ; get_option = get_lra_option ; prover = memo_nra ; hyps = hyps_of_cone ; compact = compact_cone ; pp_prf = pp_psatz pp_q ; pp_f = (fun o x -> pp_pol pp_q o (fst x)) } let non_linear_prover_Q str o = { name = "real nonlinear prover" ; get_option = (fun () -> (str, o)) ; prover = (fun (o, l) -> call_csdpcert_q o l) ; hyps = hyps_of_cone ; compact = compact_cone ; pp_prf = pp_psatz pp_q ; pp_f = (fun o x -> pp_pol pp_q o (fst x)) } let non_linear_prover_R str o = { name = "real nonlinear prover" ; get_option = (fun () -> (str, o)) ; prover = (fun (o, l) -> call_csdpcert_q o l) ; hyps = hyps_of_cone ; compact = compact_cone ; pp_prf = pp_psatz pp_q ; pp_f = (fun o x -> pp_pol pp_q o (fst x)) } let non_linear_prover_Z str o = { name = "real nonlinear prover" ; get_option = (fun () -> (str, o)) ; prover = (fun (o, l) -> lift_ratproof (call_csdpcert_z o) l) ; hyps = hyps_of_pt ; compact = compact_pt ; pp_prf = pp_proof_term ; pp_f = (fun o x -> pp_pol pp_z o (fst x)) } let linear_Z = { name = "lia" ; get_option = get_lia_option ; prover = memo_lia ; hyps = hyps_of_pt ; compact = compact_pt ; pp_prf = pp_proof_term ; pp_f = (fun o x -> pp_pol pp_z o (fst x)) } let nlinear_Z = { name = "nlia" ; get_option = get_lia_option ; prover = memo_nlia ; hyps = hyps_of_pt ; compact = compact_pt ; pp_prf = pp_proof_term ; pp_f = (fun o x -> pp_pol pp_z o (fst x)) } (** * Functions instantiating micromega_gen with the appropriate theories and * solvers *) let exfalso_if_concl_not_Prop = Proofview.Goal.enter (fun gl -> Tacmach.New.( if is_prop (pf_env gl) (project gl) (pf_concl gl) then Tacticals.New.tclIDTAC else Tactics.elim_type (Lazy.force coq_False))) let micromega_gen parse_arith pre_process cnf spec dumpexpr prover tac = Tacticals.New.tclTHEN exfalso_if_concl_not_Prop (micromega_gen parse_arith pre_process cnf spec dumpexpr prover tac) let micromega_genr prover tac = Tacticals.New.tclTHEN exfalso_if_concl_not_Prop (micromega_genr prover tac) let lra_Q = micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr linear_prover_Q let psatz_Q i = micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr (non_linear_prover_Q "real_nonlinear_prover" (Some i)) let lra_R = micromega_genr linear_prover_R let psatz_R i = micromega_genr (non_linear_prover_R "real_nonlinear_prover" (Some i)) let psatz_Z i = micromega_gen parse_zarith (fun _ x -> x) Mc.cnfZ zz_domain_spec dump_zexpr (non_linear_prover_Z "real_nonlinear_prover" (Some i)) let sos_Z = micromega_gen parse_zarith (fun _ x -> x) Mc.cnfZ zz_domain_spec dump_zexpr (non_linear_prover_Z "pure_sos" None) let sos_Q = micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr (non_linear_prover_Q "pure_sos" None) let sos_R = micromega_genr (non_linear_prover_R "pure_sos" None) let xlia = micromega_gen parse_zarith (fun _ x -> x) Mc.cnfZ zz_domain_spec dump_zexpr linear_Z let xnlia = micromega_gen parse_zarith (fun _ x -> x) Mc.cnfZ zz_domain_spec dump_zexpr nlinear_Z let nra = micromega_genr nlinear_prover_R let nqa = micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr nlinear_prover_R (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/vect.ml0000644000175000017500000002126613612664533017214 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* true | [], _ -> false | _ :: _, [] -> false | (i1, n1) :: v1, (i2, n2) :: v2 -> Int.equal i1 i2 && n1 =/ n2 && equal v1 v2 let hash v = let rec hash i = function | [] -> i | (vr, vl) :: l -> hash (i + Hashtbl.hash (vr, float_of_num vl)) l in Hashtbl.hash (hash 0 v) let null = [] let is_null v = match v with [] | [(0, Int 0)] -> true | _ -> false let pp_var_num pp_var o (v, n) = if Int.equal v 0 then if eq_num (Int 0) n then () else Printf.fprintf o "%s" (string_of_num n) else match n with | Int 1 -> pp_var o v | Int -1 -> Printf.fprintf o "-%a" pp_var v | Int 0 -> () | _ -> Printf.fprintf o "%s*%a" (string_of_num n) pp_var v let pp_var_num_smt pp_var o (v, n) = if Int.equal v 0 then if eq_num (Int 0) n then () else Printf.fprintf o "%s" (string_of_num n) else match n with | Int 1 -> pp_var o v | Int -1 -> Printf.fprintf o "(- %a)" pp_var v | Int 0 -> () | _ -> Printf.fprintf o "(* %s %a)" (string_of_num n) pp_var v let rec pp_gen pp_var o v = match v with | [] -> output_string o "0" | [e] -> pp_var_num pp_var o e | e :: l -> Printf.fprintf o "%a + %a" (pp_var_num pp_var) e (pp_gen pp_var) l let pp_var o v = Printf.fprintf o "x%i" v let pp o v = pp_gen pp_var o v let pp_smt o v = let list o v = List.iter (fun e -> Printf.fprintf o "%a " (pp_var_num_smt pp_var) e) v in Printf.fprintf o "(+ %a)" list v let from_list (l : num list) = let rec xfrom_list i l = match l with | [] -> [] | e :: l -> if e <>/ Int 0 then (i, e) :: xfrom_list (i + 1) l else xfrom_list (i + 1) l in xfrom_list 0 l let zero_num = Int 0 let to_list m = let rec xto_list i l = match l with | [] -> [] | (x, v) :: l' -> if i = x then v :: xto_list (i + 1) l' else zero_num :: xto_list (i + 1) l in xto_list 0 m let cons i v rst = if v =/ Int 0 then rst else (i, v) :: rst let rec update i f t = match t with | [] -> cons i (f zero_num) [] | (k, v) :: l -> ( match Int.compare i k with | 0 -> cons k (f v) l | -1 -> cons i (f zero_num) t | 1 -> (k, v) :: update i f l | _ -> failwith "compare_num" ) let rec set i n t = match t with | [] -> cons i n [] | (k, v) :: l -> ( match Int.compare i k with | 0 -> cons k n l | -1 -> cons i n t | 1 -> (k, v) :: set i n l | _ -> failwith "compare_num" ) let cst n = if n =/ Int 0 then [] else [(0, n)] let mul z t = match z with | Int 0 -> [] | Int 1 -> t | _ -> List.map (fun (i, n) -> (i, mult_num z n)) t let div z t = if z <>/ Int 1 then List.map (fun (x, nx) -> (x, nx // z)) t else t let uminus t = List.map (fun (i, n) -> (i, minus_num n)) t let rec add (ve1 : t) (ve2 : t) = match (ve1, ve2) with | [], v | v, [] -> v | (v1, c1) :: l1, (v2, c2) :: l2 -> let cmp = Int.compare v1 v2 in if cmp == 0 then let s = add_num c1 c2 in if eq_num (Int 0) s then add l1 l2 else (v1, s) :: add l1 l2 else if cmp < 0 then (v1, c1) :: add l1 ve2 else (v2, c2) :: add l2 ve1 let rec xmul_add (n1 : num) (ve1 : t) (n2 : num) (ve2 : t) = match (ve1, ve2) with | [], _ -> mul n2 ve2 | _, [] -> mul n1 ve1 | (v1, c1) :: l1, (v2, c2) :: l2 -> let cmp = Int.compare v1 v2 in if cmp == 0 then let s = (n1 */ c1) +/ (n2 */ c2) in if eq_num (Int 0) s then xmul_add n1 l1 n2 l2 else (v1, s) :: xmul_add n1 l1 n2 l2 else if cmp < 0 then (v1, n1 */ c1) :: xmul_add n1 l1 n2 ve2 else (v2, n2 */ c2) :: xmul_add n1 ve1 n2 l2 let mul_add n1 ve1 n2 ve2 = if n1 =/ Int 1 && n2 =/ Int 1 then add ve1 ve2 else xmul_add n1 ve1 n2 ve2 let compare : t -> t -> int = Mutils.Cmp.compare_list (fun x y -> Mutils.Cmp.compare_lexical [ (fun () -> Int.compare (fst x) (fst y)) ; (fun () -> compare_num (snd x) (snd y)) ]) (** [tail v vect] returns - [None] if [v] is not a variable of the vector [vect] - [Some(vl,rst)] where [vl] is the value of [v] in vector [vect] and [rst] is the remaining of the vector We exploit that vectors are ordered lists *) let rec tail (v : var) (vect : t) = match vect with | [] -> None | (v', vl) :: vect' -> ( match Int.compare v' v with | 0 -> Some (vl, vect) (* Ok, found *) | -1 -> tail v vect' (* Might be in the tail *) | _ -> None ) (* Hopeless *) let get v vect = match tail v vect with None -> Int 0 | Some (vl, _) -> vl let is_constant v = match v with [] | [(0, _)] -> true | _ -> false let get_cst vect = match vect with (0, v) :: _ -> v | _ -> Int 0 let choose v = match v with [] -> None | (vr, vl) :: rst -> Some (vr, vl, rst) let rec fresh v = match v with [] -> 1 | [(v, _)] -> v + 1 | _ :: v -> fresh v let variables v = List.fold_left (fun acc (x, _) -> ISet.add x acc) ISet.empty v let decomp_cst v = match v with (0, vl) :: v -> (vl, v) | _ -> (Int 0, v) let rec decomp_at i v = match v with | [] -> (Int 0, null) | (vr, vl) :: r -> if i = vr then (vl, r) else if i < vr then (Int 0, v) else decomp_at i r let decomp_fst v = match v with [] -> ((0, Int 0), []) | x :: v -> (x, v) let rec subst (vr : int) (e : t) (v : t) = match v with | [] -> [] | (x, n) :: v' -> ( match Int.compare vr x with | 0 -> mul_add n e (Int 1) v' | -1 -> v | 1 -> add [(x, n)] (subst vr e v') | _ -> assert false ) let fold f acc v = List.fold_left (fun acc (v, i) -> f acc v i) acc v let fold_error f acc v = let rec fold acc v = match v with | [] -> Some acc | (x, i) :: v' -> ( match f acc x i with None -> None | Some acc' -> fold acc' v' ) in fold acc v let rec find p v = match v with | [] -> None | (v, n) :: v' -> ( match p v n with None -> find p v' | Some r -> Some r ) let for_all p l = List.for_all (fun (v, n) -> p v n) l let decr_var i v = List.map (fun (v, n) -> (v - i, n)) v let incr_var i v = List.map (fun (v, n) -> (v + i, n)) v open Big_int let gcd v = let res = fold (fun c _ n -> assert (Int.equal (compare_big_int (denominator n) unit_big_int) 0); gcd_big_int c (numerator n)) zero_big_int v in if Int.equal (compare_big_int res zero_big_int) 0 then unit_big_int else res let normalise v = let ppcm = fold (fun c _ n -> ppcm c (denominator n)) unit_big_int v in let gcd = let gcd = fold (fun c _ n -> gcd_big_int c (numerator n)) zero_big_int v in if Int.equal (compare_big_int gcd zero_big_int) 0 then unit_big_int else gcd in List.map (fun (x, v) -> (x, v */ Big_int ppcm // Big_int gcd)) v let rec exists2 p vect1 vect2 = match (vect1, vect2) with | _, [] | [], _ -> None | (v1, n1) :: vect1', (v2, n2) :: vect2' -> if Int.equal v1 v2 then if p n1 n2 then Some (v1, n1, n2) else exists2 p vect1' vect2' else if v1 < v2 then exists2 p vect1' vect2 else exists2 p vect1 vect2' let dotproduct v1 v2 = let rec dot acc v1 v2 = match (v1, v2) with | [], _ | _, [] -> acc | (x1, n1) :: v1', (x2, n2) :: v2' -> if x1 == x2 then dot (acc +/ (n1 */ n2)) v1' v2' else if x1 < x2 then dot acc v1' v2 else dot acc v1 v2' in dot (Int 0) v1 v2 let map f v = List.map (fun (x, v) -> f x v) v let abs_min_elt v = match v with | [] -> None | (v, vl) :: r -> Some (List.fold_left (fun (v1, vl1) (v2, vl2) -> if abs_num vl1 p vr vl) let mkvar x = set x (Int 1) null module Bound = struct type t = {cst : num; var : var; coeff : num} let of_vect (v : vector) = match v with | [(x, v)] -> if x = 0 then None else Some {cst = Int 0; var = x; coeff = v} | [(0, v); (x, v')] -> Some {cst = v; var = x; coeff = v'} | _ -> None end coq-8.11.0/plugins/micromega/Lia.v0000644000175000017500000000330513612664533016607 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a ; number_to_num : 'a -> num ; zero : 'a ; unit : 'a ; mult : 'a -> 'a -> 'a ; eqb : 'a -> 'a -> bool } let z_spec = { bigint_to_number = Ml2C.bigint ; number_to_num = (fun x -> Big_int (C2Ml.z_big_int x)) ; zero = Mc.Z0 ; unit = Mc.Zpos Mc.XH ; mult = Mc.Z.mul ; eqb = Mc.zeq_bool } let q_spec = { bigint_to_number = (fun x -> {Mc.qnum = Ml2C.bigint x; Mc.qden = Mc.XH}) ; number_to_num = C2Ml.q_to_num ; zero = {Mc.qnum = Mc.Z0; Mc.qden = Mc.XH} ; unit = {Mc.qnum = Mc.Zpos Mc.XH; Mc.qden = Mc.XH} ; mult = Mc.qmult ; eqb = Mc.qeq_bool } let dev_form n_spec p = let rec dev_form p = match p with | Mc.PEc z -> Poly.constant (n_spec.number_to_num z) | Mc.PEX v -> Poly.variable (C2Ml.positive v) | Mc.PEmul (p1, p2) -> let p1 = dev_form p1 in let p2 = dev_form p2 in Poly.product p1 p2 | Mc.PEadd (p1, p2) -> Poly.addition (dev_form p1) (dev_form p2) | Mc.PEopp p -> Poly.uminus (dev_form p) | Mc.PEsub (p1, p2) -> Poly.addition (dev_form p1) (Poly.uminus (dev_form p2)) | Mc.PEpow (p, n) -> let p = dev_form p in let n = C2Ml.n n in let rec pow n = if Int.equal n 0 then Poly.constant (n_spec.number_to_num n_spec.unit) else Poly.product p (pow (n - 1)) in pow n in dev_form p let rec fixpoint f x = let y' = f x in if y' = x then y' else fixpoint f y' let rec_simpl_cone n_spec e = let simpl_cone = Mc.simpl_cone n_spec.zero n_spec.unit n_spec.mult n_spec.eqb in let rec rec_simpl_cone = function | Mc.PsatzMulE (t1, t2) -> simpl_cone (Mc.PsatzMulE (rec_simpl_cone t1, rec_simpl_cone t2)) | Mc.PsatzAdd (t1, t2) -> simpl_cone (Mc.PsatzAdd (rec_simpl_cone t1, rec_simpl_cone t2)) | x -> simpl_cone x in rec_simpl_cone e let simplify_cone n_spec c = fixpoint (rec_simpl_cone n_spec) c (* The binding with Fourier might be a bit obsolete -- how does it handle equalities ? *) (* Certificates are elements of the cone such that P = 0 *) (* To begin with, we search for certificates of the form: a1.p1 + ... an.pn + b1.q1 +... + bn.qn + c = 0 where pi >= 0 qi > 0 ai >= 0 bi >= 0 Sum bi + c >= 1 This is a linear problem: each monomial is considered as a variable. Hence, we can use fourier. The variable c is at index 1 *) (* fold_left followed by a rev ! *) let constrain_variable v l = let coeffs = List.fold_left (fun acc p -> Vect.get v p.coeffs :: acc) [] l in { coeffs = Vect.from_list (Big_int zero_big_int :: Big_int zero_big_int :: List.rev coeffs) ; op = Eq ; cst = Big_int zero_big_int } let constrain_constant l = let coeffs = List.fold_left (fun acc p -> minus_num p.cst :: acc) [] l in { coeffs = Vect.from_list (Big_int zero_big_int :: Big_int unit_big_int :: List.rev coeffs) ; op = Eq ; cst = Big_int zero_big_int } let positivity l = let rec xpositivity i l = match l with | [] -> [] | c :: l -> ( match c.op with | Eq -> xpositivity (i + 1) l | _ -> { coeffs = Vect.update (i + 1) (fun _ -> Int 1) Vect.null ; op = Ge ; cst = Int 0 } :: xpositivity (i + 1) l ) in xpositivity 1 l let cstr_of_poly (p, o) = let c, l = Vect.decomp_cst p in {coeffs = l; op = o; cst = minus_num c} let variables_of_cstr c = Vect.variables c.coeffs (* If the certificate includes at least one strict inequality, the obtained polynomial can also be 0 *) let build_dual_linear_system l = let variables = List.fold_left (fun acc p -> ISet.union acc (variables_of_cstr p)) ISet.empty l in (* For each monomial, compute a constraint *) let s0 = ISet.fold (fun mn res -> constrain_variable mn l :: res) variables [] in let c = constrain_constant l in (* I need at least something strictly positive *) let strict = { coeffs = Vect.from_list ( Big_int zero_big_int :: Big_int unit_big_int :: List.map (fun c -> if is_strict c then Big_int unit_big_int else Big_int zero_big_int) l ) ; op = Ge ; cst = Big_int unit_big_int } in (* Add the positivity constraint *) { coeffs = Vect.from_list [Big_int zero_big_int; Big_int unit_big_int] ; op = Ge ; cst = Big_int zero_big_int } :: ((strict :: positivity l) @ (c :: s0)) open Util (** [direct_linear_prover l] does not handle strict inegalities *) let fourier_linear_prover l = match Mfourier.Fourier.find_point l with | Inr prf -> if debug then Printf.printf "AProof : %a\n" Mfourier.pp_proof prf; let cert = (*List.map (fun (x,n) -> x+1,n)*) fst (List.hd (Mfourier.Proof.mk_proof l prf)) in if debug then Printf.printf "CProof : %a" Vect.pp cert; (*Some (rats_to_ints (Vect.to_list cert))*) Some (Vect.normalise cert) | Inl _ -> None let direct_linear_prover l = if !use_simplex then Simplex.find_unsat_certificate l else fourier_linear_prover l let find_point l = if !use_simplex then Simplex.find_point l else match Mfourier.Fourier.find_point l with | Inr _ -> None | Inl cert -> Some cert let optimise v l = if !use_simplex then Simplex.optimise v l else Mfourier.Fourier.optimise v l let dual_raw_certificate l = if debug then begin Printf.printf "dual_raw_certificate\n"; List.iter (fun c -> Printf.fprintf stdout "%a\n" output_cstr c) l end; let sys = build_dual_linear_system l in if debug then begin Printf.printf "dual_system\n"; List.iter (fun c -> Printf.fprintf stdout "%a\n" output_cstr c) sys end; try match find_point sys with | None -> None | Some cert -> ( match Vect.choose cert with | None -> failwith "dual_raw_certificate: empty_certificate" | Some _ -> (*Some (rats_to_ints (Vect.to_list (Vect.decr_var 2 (Vect.set 1 (Int 0) cert))))*) Some (Vect.normalise (Vect.decr_var 2 (Vect.set 1 (Int 0) cert))) ) (* should not use rats_to_ints *) with x when CErrors.noncritical x -> if debug then ( Printf.printf "dual raw certificate %s" (Printexc.to_string x); flush stdout ); None let simple_linear_prover l = try direct_linear_prover l with Strict -> (* Fourier elimination should handle > *) dual_raw_certificate l let env_of_list l = snd (List.fold_left (fun (i, m) p -> (i + 1, IMap.add i p m)) (0, IMap.empty) l) let linear_prover_cstr sys = let sysi, prfi = List.split sys in match simple_linear_prover sysi with | None -> None | Some cert -> Some (ProofFormat.proof_of_farkas (env_of_list prfi) cert) let linear_prover_cstr = if debug then ( fun sys -> Printf.printf ""; flush stdout; res ) else linear_prover_cstr let compute_max_nb_cstr l d = let len = List.length l in max len (max d (len * d)) let develop_constraint z_spec (e, k) = ( dev_form z_spec e , match k with | Mc.NonStrict -> Ge | Mc.Equal -> Eq | Mc.Strict -> Gt | _ -> assert false ) (** A single constraint can be unsat for the following reasons: - 0 >= c for c a negative constant - 0 = c for c a non-zero constant - e = c when the coeffs of e are all integers and c is rational *) type checksat = | Tauto (* Tautology *) | Unsat of ProofFormat.prf_rule (* Unsatisfiable *) | Cut of cstr * ProofFormat.prf_rule (* Cutting plane *) | Normalise of cstr * ProofFormat.prf_rule (* Coefficients may be normalised i.e relatively prime *) exception FoundProof of ProofFormat.prf_rule (** [check_sat] - detects constraints that are not satisfiable; - normalises constraints and generate cuts. *) let check_int_sat (cstr, prf) = let {coeffs; op; cst} = cstr in match Vect.choose coeffs with | None -> if eval_op op (Int 0) cst then Tauto else Unsat prf | _ -> ( let gcdi = Vect.gcd coeffs in let gcd = Big_int gcdi in if eq_num gcd (Int 1) then Normalise (cstr, prf) else if Int.equal (sign_num (mod_num cst gcd)) 0 then begin (* We can really normalise *) assert (sign_num gcd >= 1); let cstr = {coeffs = Vect.div gcd coeffs; op; cst = cst // gcd} in Normalise (cstr, ProofFormat.Gcd (gcdi, prf)) (* Normalise(cstr,CutPrf prf)*) end else match op with | Eq -> Unsat (ProofFormat.CutPrf prf) | Ge -> let cstr = {coeffs = Vect.div gcd coeffs; op; cst = ceiling_num (cst // gcd)} in Cut (cstr, ProofFormat.CutPrf prf) | Gt -> failwith "check_sat : Unexpected operator" ) let apply_and_normalise check f psys = List.fold_left (fun acc pc' -> match f pc' with | None -> pc' :: acc | Some pc' -> ( match check pc' with | Tauto -> acc | Unsat prf -> raise (FoundProof prf) | Cut (c, p) -> (c, p) :: acc | Normalise (c, p) -> (c, p) :: acc )) [] psys let is_linear_for v pc = LinPoly.is_linear (fst (fst pc)) || LinPoly.is_linear_for v (fst (fst pc)) (*let non_linear_pivot sys pc v pc' = if LinPoly.is_linear (fst (fst pc')) then None (* There are other ways to deal with those *) else WithProof.linear_pivot sys pc v pc' *) let is_linear_substitution sys ((p, o), prf) = let pred v = v =/ Int 1 || v =/ Int (-1) in match o with | Eq -> ( match List.filter (fun v -> List.for_all (is_linear_for v) sys) (LinPoly.search_all_linear pred p) with | [] -> None | v :: _ -> Some v (* make a choice *) ) | _ -> None let elim_simple_linear_equality sys0 = let elim sys = let oeq, sys' = extract (is_linear_substitution sys) sys in match oeq with | None -> None | Some (v, pc) -> simplify (WithProof.linear_pivot sys0 pc v) sys' in iterate_until_stable elim sys0 let output_sys o sys = List.iter (fun s -> Printf.fprintf o "%a\n" WithProof.output s) sys let subst sys = let sys' = WithProof.subst sys in if debug then Printf.fprintf stdout "[subst:\n%a\n==>\n%a\n]" output_sys sys output_sys sys'; sys' (** [saturate_linear_equality sys] generate new constraints obtained by eliminating linear equalities by pivoting. For integers, the obtained constraints are sound but not complete. *) let saturate_by_linear_equalities sys0 = WithProof.saturate_subst false sys0 let saturate_by_linear_equalities sys = let sys' = saturate_by_linear_equalities sys in if debug then Printf.fprintf stdout "[saturate_by_linear_equalities:\n%a\n==>\n%a\n]" output_sys sys output_sys sys'; sys' (* let saturate_linear_equality_non_linear sys0 = let (l,_) = extract_all (is_substitution false) sys0 in let rec elim l acc = match l with | [] -> acc | (v,pc)::l' -> let nc = saturate (non_linear_pivot sys0 pc v) (sys0@acc) in elim l' (nc@acc) in elim l [] *) let bounded_vars (sys : WithProof.t list) = let l = fst (extract_all (fun ((p, o), prf) -> LinPoly.is_variable p) sys) in List.fold_left (fun acc (i, wp) -> IMap.add i wp acc) IMap.empty l let rec power n p = if n = 1 then p else WithProof.product p (power (n - 1) p) let bound_monomial mp m = if Monomial.is_var m || Monomial.is_const m then None else try Some (Monomial.fold (fun v i acc -> let wp = IMap.find v mp in WithProof.product (power i wp) acc) m (WithProof.const (Int 1))) with Not_found -> None let bound_monomials (sys : WithProof.t list) = let mp = bounded_vars sys in let m = List.fold_left (fun acc ((p, _), _) -> Vect.fold (fun acc v _ -> let m = LinPoly.MonT.retrieve v in match bound_monomial mp m with | None -> acc | Some r -> IMap.add v r acc) acc p) IMap.empty sys in IMap.fold (fun _ e acc -> e :: acc) m [] let develop_constraints prfdepth n_spec sys = LinPoly.MonT.clear (); max_nb_cstr := compute_max_nb_cstr sys prfdepth; let sys = List.map (develop_constraint n_spec) sys in List.mapi (fun i (p, o) -> ((LinPoly.linpol_of_pol p, o), ProofFormat.Hyp i)) sys let square_of_var i = let x = LinPoly.var i in ((LinPoly.product x x, Ge), ProofFormat.Square x) (** [nlinear_preprocess sys] augments the system [sys] by performing some limited non-linear reasoning. For instance, it asserts that the x² ≥0 but also that if c₁ ≥ 0 ∈ sys and c₂ ≥ 0 ∈ sys then c₁ × c₂ ≥ 0. The resulting system is linearised. *) let nlinear_preprocess (sys : WithProof.t list) = let is_linear = List.for_all (fun ((p, _), _) -> LinPoly.is_linear p) sys in if is_linear then sys else let collect_square = List.fold_left (fun acc ((p, _), _) -> MonMap.union (fun k e1 e2 -> Some e1) acc (LinPoly.collect_square p)) MonMap.empty sys in let sys = MonMap.fold (fun s m acc -> let s = LinPoly.of_monomial s in let m = LinPoly.of_monomial m in ((m, Ge), ProofFormat.Square s) :: acc) collect_square sys in let collect_vars = List.fold_left (fun acc p -> ISet.union acc (LinPoly.variables (fst (fst p)))) ISet.empty sys in let sys = ISet.fold (fun i acc -> square_of_var i :: acc) collect_vars sys in let sys = sys @ all_pairs WithProof.product sys in if debug then begin Printf.fprintf stdout "Preprocessed\n"; List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys end; List.map (WithProof.annot "P") sys let nlinear_prover prfdepth sys = let sys = develop_constraints prfdepth q_spec sys in let sys1 = elim_simple_linear_equality sys in let sys2 = saturate_by_linear_equalities sys1 in let sys = nlinear_preprocess sys1 @ sys2 in let sys = List.map (fun ((p, o), prf) -> (cstr_of_poly (p, o), prf)) sys in let id = List.fold_left (fun acc (_, r) -> max acc (ProofFormat.pr_rule_max_id r)) 0 sys in let env = CList.interval 0 id in match linear_prover_cstr sys with | None -> Unknown | Some cert -> Prf (ProofFormat.cmpl_prf_rule Mc.normQ CamlToCoq.q env cert) let linear_prover_with_cert prfdepth sys = let sys = develop_constraints prfdepth q_spec sys in (* let sys = nlinear_preprocess sys in *) let sys = List.map (fun (c, p) -> (cstr_of_poly c, p)) sys in match linear_prover_cstr sys with | None -> Unknown | Some cert -> Prf (ProofFormat.cmpl_prf_rule Mc.normQ CamlToCoq.q (List.mapi (fun i e -> i) sys) cert) (* The prover is (probably) incomplete -- only searching for naive cutting planes *) open Sos_types let rec scale_term t = match t with | Zero -> (unit_big_int, Zero) | Const n -> (denominator n, Const (Big_int (numerator n))) | Var n -> (unit_big_int, Var n) | Opp t -> let s, t = scale_term t in (s, Opp t) | Add (t1, t2) -> let s1, y1 = scale_term t1 and s2, y2 = scale_term t2 in let g = gcd_big_int s1 s2 in let s1' = div_big_int s1 g in let s2' = div_big_int s2 g in let e = mult_big_int g (mult_big_int s1' s2') in if Int.equal (compare_big_int e unit_big_int) 0 then (unit_big_int, Add (y1, y2)) else (e, Add (Mul (Const (Big_int s2'), y1), Mul (Const (Big_int s1'), y2))) | Sub _ -> failwith "scale term: not implemented" | Mul (y, z) -> let s1, y1 = scale_term y and s2, y2 = scale_term z in (mult_big_int s1 s2, Mul (y1, y2)) | Pow (t, n) -> let s, t = scale_term t in (power_big_int_positive_int s n, Pow (t, n)) let scale_term t = let s, t' = scale_term t in (s, t') let rec scale_certificate pos = match pos with | Axiom_eq i -> (unit_big_int, Axiom_eq i) | Axiom_le i -> (unit_big_int, Axiom_le i) | Axiom_lt i -> (unit_big_int, Axiom_lt i) | Monoid l -> (unit_big_int, Monoid l) | Rational_eq n -> (denominator n, Rational_eq (Big_int (numerator n))) | Rational_le n -> (denominator n, Rational_le (Big_int (numerator n))) | Rational_lt n -> (denominator n, Rational_lt (Big_int (numerator n))) | Square t -> let s, t' = scale_term t in (mult_big_int s s, Square t') | Eqmul (t, y) -> let s1, y1 = scale_term t and s2, y2 = scale_certificate y in (mult_big_int s1 s2, Eqmul (y1, y2)) | Sum (y, z) -> let s1, y1 = scale_certificate y and s2, y2 = scale_certificate z in let g = gcd_big_int s1 s2 in let s1' = div_big_int s1 g in let s2' = div_big_int s2 g in ( mult_big_int g (mult_big_int s1' s2') , Sum ( Product (Rational_le (Big_int s2'), y1) , Product (Rational_le (Big_int s1'), y2) ) ) | Product (y, z) -> let s1, y1 = scale_certificate y and s2, y2 = scale_certificate z in (mult_big_int s1 s2, Product (y1, y2)) open Micromega let rec term_to_q_expr = function | Const n -> PEc (Ml2C.q n) | Zero -> PEc (Ml2C.q (Int 0)) | Var s -> PEX (Ml2C.index (int_of_string (String.sub s 1 (String.length s - 1)))) | Mul (p1, p2) -> PEmul (term_to_q_expr p1, term_to_q_expr p2) | Add (p1, p2) -> PEadd (term_to_q_expr p1, term_to_q_expr p2) | Opp p -> PEopp (term_to_q_expr p) | Pow (t, n) -> PEpow (term_to_q_expr t, Ml2C.n n) | Sub (t1, t2) -> PEsub (term_to_q_expr t1, term_to_q_expr t2) let term_to_q_pol e = Mc.norm_aux (Ml2C.q (Int 0)) (Ml2C.q (Int 1)) Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool (term_to_q_expr e) let rec product l = match l with | [] -> Mc.PsatzZ | [i] -> Mc.PsatzIn (Ml2C.nat i) | i :: l -> Mc.PsatzMulE (Mc.PsatzIn (Ml2C.nat i), product l) let q_cert_of_pos pos = let rec _cert_of_pos = function | Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i) | Axiom_le i -> Mc.PsatzIn (Ml2C.nat i) | Axiom_lt i -> Mc.PsatzIn (Ml2C.nat i) | Monoid l -> product l | Rational_eq n | Rational_le n | Rational_lt n -> if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else Mc.PsatzC (Ml2C.q n) | Square t -> Mc.PsatzSquare (term_to_q_pol t) | Eqmul (t, y) -> Mc.PsatzMulC (term_to_q_pol t, _cert_of_pos y) | Sum (y, z) -> Mc.PsatzAdd (_cert_of_pos y, _cert_of_pos z) | Product (y, z) -> Mc.PsatzMulE (_cert_of_pos y, _cert_of_pos z) in simplify_cone q_spec (_cert_of_pos pos) let rec term_to_z_expr = function | Const n -> PEc (Ml2C.bigint (big_int_of_num n)) | Zero -> PEc Z0 | Var s -> PEX (Ml2C.index (int_of_string (String.sub s 1 (String.length s - 1)))) | Mul (p1, p2) -> PEmul (term_to_z_expr p1, term_to_z_expr p2) | Add (p1, p2) -> PEadd (term_to_z_expr p1, term_to_z_expr p2) | Opp p -> PEopp (term_to_z_expr p) | Pow (t, n) -> PEpow (term_to_z_expr t, Ml2C.n n) | Sub (t1, t2) -> PEsub (term_to_z_expr t1, term_to_z_expr t2) let term_to_z_pol e = Mc.norm_aux (Ml2C.z 0) (Ml2C.z 1) Mc.Z.add Mc.Z.mul Mc.Z.sub Mc.Z.opp Mc.zeq_bool (term_to_z_expr e) let z_cert_of_pos pos = let s, pos = scale_certificate pos in let rec _cert_of_pos = function | Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i) | Axiom_le i -> Mc.PsatzIn (Ml2C.nat i) | Axiom_lt i -> Mc.PsatzIn (Ml2C.nat i) | Monoid l -> product l | Rational_eq n | Rational_le n | Rational_lt n -> if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else Mc.PsatzC (Ml2C.bigint (big_int_of_num n)) | Square t -> Mc.PsatzSquare (term_to_z_pol t) | Eqmul (t, y) -> let is_unit = match t with Const n -> n =/ Int 1 | _ -> false in if is_unit then _cert_of_pos y else Mc.PsatzMulC (term_to_z_pol t, _cert_of_pos y) | Sum (y, z) -> Mc.PsatzAdd (_cert_of_pos y, _cert_of_pos z) | Product (y, z) -> Mc.PsatzMulE (_cert_of_pos y, _cert_of_pos z) in simplify_cone z_spec (_cert_of_pos pos) open Mutils (** All constraints (initial or derived) have an index and have a justification i.e., proof. Given a constraint, all the coefficients are always integers. *) open Num open Big_int open Polynomial type prf_sys = (cstr * ProofFormat.prf_rule) list (** Proof generating pivoting over variable v *) let pivot v (c1, p1) (c2, p2) = let {coeffs = v1; op = op1; cst = n1} = c1 and {coeffs = v2; op = op2; cst = n2} = c2 in (* Could factorise gcd... *) let xpivot cv1 cv2 = ( { coeffs = Vect.add (Vect.mul cv1 v1) (Vect.mul cv2 v2) ; op = opAdd op1 op2 ; cst = (n1 */ cv1) +/ (n2 */ cv2) } , ProofFormat.add_proof (ProofFormat.mul_cst_proof cv1 p1) (ProofFormat.mul_cst_proof cv2 p2) ) in match (Vect.get v v1, Vect.get v v2) with | Int 0, _ | _, Int 0 -> None | a, b -> if Int.equal (sign_num a * sign_num b) (-1) then let cv1 = abs_num b and cv2 = abs_num a in Some (xpivot cv1 cv2) else if op1 == Eq then let cv1 = minus_num (b */ Int (sign_num a)) and cv2 = abs_num a in Some (xpivot cv1 cv2) else if op2 == Eq then let cv1 = abs_num b and cv2 = minus_num (a */ Int (sign_num b)) in Some (xpivot cv1 cv2) else None (* op2 could be Eq ... this might happen *) let simpl_sys sys = List.fold_left (fun acc (c, p) -> match check_int_sat (c, p) with | Tauto -> acc | Unsat prf -> raise (FoundProof prf) | Cut (c, p) -> (c, p) :: acc | Normalise (c, p) -> (c, p) :: acc) [] sys (** [ext_gcd a b] is the extended Euclid algorithm. [ext_gcd a b = (x,y,g)] iff [ax+by=g] Source: http://en.wikipedia.org/wiki/Extended_Euclidean_algorithm *) let rec ext_gcd a b = if Int.equal (sign_big_int b) 0 then (unit_big_int, zero_big_int) else let q, r = quomod_big_int a b in let s, t = ext_gcd b r in (t, sub_big_int s (mult_big_int q t)) let extract_coprime (c1, p1) (c2, p2) = if c1.op == Eq && c2.op == Eq then Vect.exists2 (fun n1 n2 -> Int.equal (compare_big_int (gcd_big_int (numerator n1) (numerator n2)) unit_big_int) 0) c1.coeffs c2.coeffs else None let extract2 pred l = let rec xextract2 rl l = match l with | [] -> (None, rl) (* Did not find *) | e :: l -> ( match extract (pred e) l with | None, _ -> xextract2 (e :: rl) l | Some (r, e'), l' -> (Some (r, e, e'), List.rev_append rl l') ) in xextract2 [] l let extract_coprime_equation psys = extract2 extract_coprime psys let pivot_sys v pc psys = apply_and_normalise check_int_sat (pivot v pc) psys let reduce_coprime psys = let oeq, sys = extract_coprime_equation psys in match oeq with | None -> None (* Nothing to do *) | Some ((v, n1, n2), (c1, p1), (c2, p2)) -> let l1, l2 = ext_gcd (numerator n1) (numerator n2) in let l1' = Big_int l1 and l2' = Big_int l2 in let cstr = { coeffs = Vect.add (Vect.mul l1' c1.coeffs) (Vect.mul l2' c2.coeffs) ; op = Eq ; cst = (l1' */ c1.cst) +/ (l2' */ c2.cst) } in let prf = ProofFormat.add_proof (ProofFormat.mul_cst_proof l1' p1) (ProofFormat.mul_cst_proof l2' p2) in Some (pivot_sys v (cstr, prf) ((c1, p1) :: sys)) (** If there is an equation [eq] of the form 1.x + e = c, do a pivot over x with equation [eq] *) let reduce_unary psys = let is_unary_equation (cstr, prf) = if cstr.op == Eq then Vect.find (fun v n -> if n =/ Int 1 || n =/ Int (-1) then Some v else None) cstr.coeffs else None in let oeq, sys = extract is_unary_equation psys in match oeq with | None -> None (* Nothing to do *) | Some (v, pc) -> Some (pivot_sys v pc sys) let reduce_var_change psys = let rec rel_prime vect = match Vect.choose vect with | None -> None | Some (x, v, vect) -> ( let v = numerator v in match Vect.find (fun x' v' -> let v' = numerator v' in if eq_big_int (gcd_big_int v v') unit_big_int then Some (x', v') else None) vect with | Some (x', v') -> Some ((x, v), (x', v')) | None -> rel_prime vect ) in let rel_prime (cstr, prf) = if cstr.op == Eq then rel_prime cstr.coeffs else None in let oeq, sys = extract rel_prime psys in match oeq with | None -> None | Some (((x, v), (x', v')), (c, p)) -> let l1, l2 = ext_gcd v v' in let l1, l2 = (Big_int l1, Big_int l2) in let pivot_eq (c', p') = let {coeffs; op; cst} = c' in let vx = Vect.get x coeffs in let vx' = Vect.get x' coeffs in let m = minus_num ((vx */ l1) +/ (vx' */ l2)) in Some ( { coeffs = Vect.add (Vect.mul m c.coeffs) coeffs ; op ; cst = (m */ c.cst) +/ cst } , ProofFormat.add_proof (ProofFormat.mul_cst_proof m p) p' ) in Some (apply_and_normalise check_int_sat pivot_eq sys) let reduction_equations psys = iterate_until_stable (app_funs [reduce_unary; reduce_coprime; reduce_var_change (*; reduce_pivot*)]) psys (** [get_bound sys] returns upon success an interval (lb,e,ub) with proofs *) let get_bound sys = let is_small (v, i) = match Itv.range i with None -> false | Some i -> i <=/ Int 1 in let select_best (x1, i1) (x2, i2) = if Itv.smaller_itv i1 i2 then (x1, i1) else (x2, i2) in (* For lia, there are no equations => these precautions are not needed *) (* For nlia, there are equations => do not enumerate over equations! *) let all_planes sys = let eq, ineq = List.partition (fun c -> c.op == Eq) sys in match eq with | [] -> List.rev_map (fun c -> c.coeffs) ineq | _ -> List.fold_left (fun acc c -> if List.exists (fun c' -> Vect.equal c.coeffs c'.coeffs) eq then acc else c.coeffs :: acc) [] ineq in let smallest_interval = List.fold_left (fun acc vect -> if is_small acc then acc else match optimise vect sys with | None -> acc | Some i -> if debug then Printf.printf "Found a new bound %a in %a" Vect.pp vect Itv.pp i; select_best (vect, i) acc) (Vect.null, (None, None)) (all_planes sys) in let smallest_interval = match smallest_interval with | x, (Some i, Some j) -> Some (i, x, j) | x -> None (* This should not be possible *) in match smallest_interval with | Some (lb, e, ub) -> ( let lbn, lbd = (sub_big_int (numerator lb) unit_big_int, denominator lb) in let ubn, ubd = (add_big_int unit_big_int (numerator ub), denominator ub) in (* x <= ub -> x > ub *) match ( direct_linear_prover ( {coeffs = Vect.mul (Big_int ubd) e; op = Ge; cst = Big_int ubn} :: sys ) , (* lb <= x -> lb > x *) direct_linear_prover ( { coeffs = Vect.mul (minus_num (Big_int lbd)) e ; op = Ge ; cst = minus_num (Big_int lbn) } :: sys ) ) with | Some cub, Some clb -> Some (List.tl (Vect.to_list clb), (lb, e, ub), List.tl (Vect.to_list cub)) | _ -> failwith "Interval without proof" ) | None -> None let check_sys sys = List.for_all (fun (c, p) -> Vect.for_all (fun _ n -> sign_num n <> 0) c.coeffs) sys open ProofFormat let xlia (can_enum : bool) reduction_equations sys = let rec enum_proof (id : int) (sys : prf_sys) = if debug then ( Printf.printf "enum_proof\n"; flush stdout ); assert (check_sys sys); let nsys, prf = List.split sys in match get_bound nsys with | None -> Unknown (* Is the systeme really unbounded ? *) | Some (prf1, (lb, e, ub), prf2) -> ( if debug then Printf.printf "Found interval: %a in [%s;%s] -> " Vect.pp e (string_of_num lb) (string_of_num ub); match start_enum id e (ceiling_num lb) (floor_num ub) sys with | Prf prfl -> Prf (ProofFormat.Enum ( id , ProofFormat.proof_of_farkas (env_of_list prf) (Vect.from_list prf1) , e , ProofFormat.proof_of_farkas (env_of_list prf) (Vect.from_list prf2) , prfl )) | _ -> Unknown ) and start_enum id e clb cub sys = if clb >/ cub then Prf [] else let eq = {coeffs = e; op = Eq; cst = clb} in match aux_lia (id + 1) ((eq, ProofFormat.Def id) :: sys) with | Unknown | Model _ -> Unknown | Prf prf -> ( match start_enum id e (clb +/ Int 1) cub sys with | Prf l -> Prf (prf :: l) | _ -> Unknown ) and aux_lia (id : int) (sys : prf_sys) = assert (check_sys sys); if debug then Printf.printf "xlia: %a \n" (pp_list ";" (fun o (c, _) -> output_cstr o c)) sys; try let sys = reduction_equations sys in if debug then Printf.printf "after reduction: %a \n" (pp_list ";" (fun o (c, _) -> output_cstr o c)) sys; match linear_prover_cstr sys with | Some prf -> Prf (Step (id, prf, Done)) | None -> if can_enum then enum_proof id sys else Unknown with FoundProof prf -> (* [reduction_equations] can find a proof *) Prf (Step (id, prf, Done)) in (* let sys' = List.map (fun (p,o) -> Mc.norm0 p , o) sys in*) let id = 1 + List.fold_left (fun acc (_, r) -> max acc (ProofFormat.pr_rule_max_id r)) 0 sys in let orpf = try let sys = simpl_sys sys in aux_lia id sys with FoundProof pr -> Prf (Step (id, pr, Done)) in match orpf with | Unknown | Model _ -> Unknown | Prf prf -> let env = CList.interval 0 (id - 1) in if debug then begin Printf.fprintf stdout "direct proof %a\n" output_proof prf; flush stdout end; let prf = compile_proof env prf in (*try if Mc.zChecker sys' prf then Some prf else raise Certificate.BadCertificate with Failure s -> (Printf.printf "%s" s ; Some prf) *) Prf prf let xlia_simplex env red sys = let compile_prf sys prf = let id = 1 + List.fold_left (fun acc (_, r) -> max acc (ProofFormat.pr_rule_max_id r)) 0 sys in let env = CList.interval 0 (id - 1) in Prf (compile_proof env prf) in try let sys = red sys in match Simplex.integer_solver sys with | None -> Unknown | Some prf -> compile_prf sys prf with FoundProof prf -> compile_prf sys (Step (0, prf, Done)) let xlia env0 en red sys = if !use_simplex then xlia_simplex env0 red sys else xlia en red sys let dump_file = ref None let gen_bench (tac, prover) can_enum prfdepth sys = let res = prover can_enum prfdepth sys in ( match !dump_file with | None -> () | Some file -> let o = open_out (Filename.temp_file ~temp_dir:(Sys.getcwd ()) file ".v") in let sys = develop_constraints prfdepth z_spec sys in Printf.fprintf o "Require Import ZArith Lia. Open Scope Z_scope.\n"; Printf.fprintf o "Goal %a.\n" (LinPoly.pp_goal "Z") (List.map fst sys); begin match res with | Unknown | Model _ -> Printf.fprintf o "Proof.\n intros. Fail %s.\nAbort.\n" tac | Prf res -> Printf.fprintf o "Proof.\n intros. %s.\nQed.\n" tac end; flush o; close_out o ); res let lia (can_enum : bool) (prfdepth : int) sys = let sys = develop_constraints prfdepth z_spec sys in if debug then begin Printf.fprintf stdout "Input problem\n"; List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys; Printf.fprintf stdout "Input problem\n"; let string_of_op = function Eq -> "=" | Ge -> ">=" | Gt -> ">" in List.iter (fun ((p, op), _) -> Printf.fprintf stdout "(assert (%s %a))\n" (string_of_op op) Vect.pp_smt p) sys end; let sys = subst sys in let bnd = bound_monomials sys in (* To deal with non-linear monomials *) let sys = bnd @ saturate_by_linear_equalities sys @ sys in let sys' = List.map (fun ((p, o), prf) -> (cstr_of_poly (p, o), prf)) sys in xlia (List.map fst sys) can_enum reduction_equations sys' let make_cstr_system sys = List.map (fun ((p, o), prf) -> (cstr_of_poly (p, o), prf)) sys let nlia enum prfdepth sys = let sys = develop_constraints prfdepth z_spec sys in let is_linear = List.for_all (fun ((p, _), _) -> LinPoly.is_linear p) sys in if debug then begin Printf.fprintf stdout "Input problem\n"; List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys end; if is_linear then xlia (List.map fst sys) enum reduction_equations (make_cstr_system sys) else (* let sys1 = elim_every_substitution sys in No: if a wrong equation is chosen, the proof may fail. It would only be safe if the variable is linear... *) let sys1 = elim_simple_linear_equality sys in let sys2 = saturate_by_linear_equalities sys1 in let sys3 = nlinear_preprocess (sys1 @ sys2) in let sys4 = make_cstr_system (*sys2@*) sys3 in (* [reduction_equations] is too brutal - there should be some non-linear reasoning *) xlia (List.map fst sys) enum reduction_equations sys4 (* For regression testing, if bench = true generate a Coq goal *) let lia can_enum prfdepth sys = gen_bench ("lia", lia) can_enum prfdepth sys let nlia enum prfdepth sys = gen_bench ("nia", nlia) enum prfdepth sys (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/vect.mli0000644000175000017500000001455613612664533017371 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int val equal : t -> t -> bool val compare : t -> t -> int (** {1 Basic accessors and utility functions} *) val pp_gen : (out_channel -> var -> unit) -> out_channel -> t -> unit (** [pp_gen pp_var o v] prints the representation of the vector [v] over the channel [o] *) val pp : out_channel -> t -> unit (** [pp o v] prints the representation of the vector [v] over the channel [o] *) val pp_smt : out_channel -> t -> unit (** [pp_smt o v] prints the representation of the vector [v] over the channel [o] using SMTLIB conventions *) val variables : t -> ISet.t (** [variables v] returns the set of variables with non-zero coefficients *) val get_cst : t -> num (** [get_cst v] returns c i.e. the coefficient of the variable zero *) val decomp_cst : t -> num * t (** [decomp_cst v] returns the pair (c,a1.x1+...+an.xn) *) val decomp_at : int -> t -> num * t (** [decomp_cst v] returns the pair (ai, ai+1.xi+...+an.xn) *) val decomp_fst : t -> (var * num) * t val cst : num -> t (** [cst c] returns the vector v=c+0.x1+...+0.xn *) val is_constant : t -> bool (** [is_constant v] holds if [v] is a constant vector i.e. v=c+0.x1+...+0.xn *) val null : t (** [null] is the empty vector i.e. 0+0.x1+...+0.xn *) val is_null : t -> bool (** [is_null v] returns whether [v] is the [null] vector i.e [equal v null] *) val get : var -> t -> num (** [get xi v] returns the coefficient ai of the variable [xi]. [get] is also defined for the variable 0 *) val set : var -> num -> t -> t (** [set xi ai' v] returns the vector c+a1.x1+...ai'.xi+...+an.xn i.e. the coefficient of the variable xi is set to ai' *) val mkvar : var -> t (** [mkvar xi] returns 1.xi *) val update : var -> (num -> num) -> t -> t (** [update xi f v] returns c+a1.x1+...+f(ai).xi+...+an.xn *) val fresh : t -> int (** [fresh v] return the fresh variable with index 1+ max (variables v) *) val choose : t -> (var * num * t) option (** [choose v] decomposes a vector [v] depending on whether it is [null] or not. @return None if v is [null] @return Some(x,n,r) where v = r + n.x x is the smallest variable with non-zero coefficient n <> 0. *) val from_list : num list -> t (** [from_list l] returns the vector c+a1.x1...an.xn from the list of coefficient [l=c;a1;...;an] *) val to_list : t -> num list (** [to_list v] returns the list of all coefficient of the vector v i.e. [c;a1;...;an] The list representation is (obviously) not sparsed and therefore certain ai may be 0 *) val decr_var : int -> t -> t (** [decr_var i v] decrements the variables of the vector [v] by the amount [i]. Beware, it is only defined if all the variables of v are greater than i *) val incr_var : int -> t -> t (** [incr_var i v] increments the variables of the vector [v] by the amount [i]. *) val gcd : t -> Big_int.big_int (** [gcd v] returns gcd(num(c),num(a1),...,num(an)) where num extracts the numerator of a rational value. *) val normalise : t -> t (** [normalise v] returns a vector with only integer coefficients *) (** {1 Linear arithmetics} *) val add : t -> t -> t (** [add v1 v2] is vector addition. @param v1 is of the form c +a1.x1 +...+an.xn @param v2 is of the form c'+a1'.x1 +...+an'.xn @return c1+c1'+ (a1+a1').x1 + ... + (an+an').xn *) val mul : num -> t -> t (** [mul a v] is vector multiplication of vector [v] by a scalar [a]. @return a.v = a.c+a.a1.x1+...+a.an.xn *) val mul_add : num -> t -> num -> t -> t (** [mul_add c1 v1 c2 v2] returns the linear combination c1.v1+c2.v2 *) val subst : int -> t -> t -> t (** [subst x v v'] replaces x by v in vector v' *) val div : num -> t -> t (** [div c1 v1] returns the mutiplication by the inverse of c1 i.e (1/c1).v1 *) val uminus : t -> t (** [uminus v] @return -v the opposite vector of v i.e. (-1).v *) (** {1 Iterators} *) val fold : ('acc -> var -> num -> 'acc) -> 'acc -> t -> 'acc (** [fold f acc v] returns f (f (f acc 0 c ) x1 a1 ) ... xn an *) val fold_error : ('acc -> var -> num -> 'acc option) -> 'acc -> t -> 'acc option (** [fold_error f acc v] is the same as [fold (fun acc x i -> match acc with None -> None | Some acc' -> f acc' x i) (Some acc) v] but with early exit... *) val find : (var -> num -> 'c option) -> t -> 'c option (** [find f v] returns the first [f xi ai] such that [f xi ai <> None]. If no such xi ai exists, it returns None *) val for_all : (var -> num -> bool) -> t -> bool (** [for_all p v] returns /\_{i>=0} (f xi ai) *) val exists2 : (num -> num -> bool) -> t -> t -> (var * num * num) option (** [exists2 p v v'] returns Some(xi,ai,ai') if p(xi,ai,ai') holds and ai,ai' <> 0. It returns None if no such pair of coefficient exists. *) val dotproduct : t -> t -> num (** [dotproduct v1 v2] is the dot product of v1 and v2. *) val map : (var -> num -> 'a) -> t -> 'a list val abs_min_elt : t -> (var * num) option val partition : (var -> num -> bool) -> t -> t * t module Bound : sig type t = {cst : num; var : var; coeff : num} (** represents a0 + ai.xi *) val of_vect : vector -> t option end coq-8.11.0/plugins/micromega/mfourier.ml0000644000175000017500000006370113612664533020103 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* x.proof_idx) [c1;...;cn] - [pos] is the number of positive values of the vector - [neg] is the number of negative values of the vector ( [neg] + [pos] is therefore the length of the vector) [v] is an upper-bound of the set of variables which appear in [s]. *) exception SystemContradiction of proof (** To be thrown when a system has no solution *) (** Pretty printing *) let rec pp_proof o prf = match prf with | Assum i -> Printf.fprintf o "H%i" i | Elim (v, prf1, prf2) -> Printf.fprintf o "E(%i,%a,%a)" v pp_proof prf1 pp_proof prf2 | And (prf1, prf2) -> Printf.fprintf o "A(%a,%a)" pp_proof prf1 pp_proof prf2 let pp_cstr o (vect, bnd) = let l, r = bnd in ( match l with | None -> () | Some n -> Printf.fprintf o "%s <= " (string_of_num n) ); Vect.pp o vect; match r with | None -> output_string o "\n" | Some n -> Printf.fprintf o "<=%s\n" (string_of_num n) let pp_system o sys = System.iter (fun vect ibnd -> pp_cstr o (vect, !ibnd.bound)) sys (** [merge_cstr_info] takes: - the intersection of bounds and - the union of proofs - [pos] and [neg] fields should be identical *) let merge_cstr_info i1 i2 = let {pos = p1; neg = n1; bound = i1; prf = prf1} = i1 and {pos = p2; neg = n2; bound = i2; prf = prf2} = i2 in assert (Int.equal p1 p2 && Int.equal n1 n2); match inter i1 i2 with | None -> None (* Could directly raise a system contradiction exception *) | Some bnd -> Some {pos = p1; neg = n1; bound = bnd; prf = And (prf1, prf2)} (** [xadd_cstr vect cstr_info] loads an constraint into the system. The constraint is neither redundant nor contradictory. @raise SystemContradiction if [cstr_info] returns [None] *) let xadd_cstr vect cstr_info sys = try let info = System.find sys vect in match merge_cstr_info cstr_info !info with | None -> raise (SystemContradiction (And (cstr_info.prf, !info.prf))) | Some info' -> info := info' with Not_found -> System.replace sys vect (ref cstr_info) exception TimeOut let xadd_cstr vect cstr_info sys = if debug && Int.equal (System.length sys mod 1000) 0 then ( print_string "*"; flush stdout ); if System.length sys < !max_nb_cstr then xadd_cstr vect cstr_info sys else raise TimeOut type cstr_ext = | Contradiction (** The constraint is contradictory. Typically, a [SystemContradiction] exception will be raised. *) | Redundant (** The constrain is redundant. Typically, the constraint will be dropped *) | Cstr of vector * cstr_info (** Taken alone, the constraint is neither contradictory nor redundant. Typically, it will be added to the constraint system. *) (** [normalise_cstr] : vector -> cstr_info -> cstr_ext *) let normalise_cstr vect cinfo = match norm_itv cinfo.bound with | None -> Contradiction | Some (l, r) -> ( match Vect.choose vect with | None -> if Itv.in_bound (l, r) (Int 0) then Redundant else Contradiction | Some (_, n, _) -> Cstr ( Vect.div n vect , let divn x = x // n in if Int.equal (sign_num n) 1 then {cinfo with bound = (Option.map divn l, Option.map divn r)} else { cinfo with pos = cinfo.neg ; neg = cinfo.pos ; bound = (Option.map divn r, Option.map divn l) } ) ) (** For compatibility, there is an external representation of constraints *) let count v = Vect.fold (fun (n, p) _ vl -> let sg = sign_num vl in assert (sg <> 0); if Int.equal sg 1 then (n, p + 1) else (n + 1, p)) (0, 0) v let norm_cstr {coeffs = v; op = o; cst = c} idx = let n, p = count v in normalise_cstr v { pos = p ; neg = n ; bound = ( match o with | Eq -> (Some c, Some c) | Ge -> (Some c, None) | Gt -> raise Polynomial.Strict ) ; prf = Assum idx } (** [load_system l] takes a list of constraints of type [cstr_compat] @return a system of constraints @raise SystemContradiction if a contradiction is found *) let load_system l = let sys = System.create 1000 in let li = List.mapi (fun i e -> (e, i)) l in let vars = List.fold_left (fun vrs (cstr, i) -> match norm_cstr cstr i with | Contradiction -> raise (SystemContradiction (Assum i)) | Redundant -> vrs | Cstr (vect, info) -> xadd_cstr vect info sys; Vect.fold (fun s v _ -> ISet.add v s) vrs cstr.coeffs) ISet.empty li in {sys; vars} let system_list sys = let {sys = s; vars = v} = sys in System.fold (fun k bi l -> (k, !bi) :: l) s [] (** [add (v1,c1) (v2,c2) ] precondition: (c1 <>/ Int 0 && c2 <>/ Int 0) @return a pair [(v,ln)] such that [v] is the sum of vector [v1] divided by [c1] and vector [v2] divided by [c2] Note that the resulting vector is not normalised. *) let add (v1, c1) (v2, c2) = assert (c1 <>/ Int 0 && c2 <>/ Int 0); let res = mul_add (Int 1 // c1) v1 (Int 1 // c2) v2 in (res, count res) let add (v1, c1) (v2, c2) = let res = add (v1, c1) (v2, c2) in (* Printf.printf "add(%a,%s,%a,%s) -> %a\n" pp_vect v1 (string_of_num c1) pp_vect v2 (string_of_num c2) pp_vect (fst res) ;*) res (** To perform Fourier elimination, constraints are categorised depending on the sign of the variable to eliminate. *) (** [split x vect info (l,m,r)] @param v is the variable to eliminate @param l contains constraints such that (e + a*x) // a >= c / a @param r contains constraints such that (e + a*x) // - a >= c / -a @param m contains constraints which do not mention [x] *) let split x (vect : vector) info (l, m, r) = match get x vect with | Int 0 -> (* The constraint does not mention [x], store it in m *) (l, (vect, info) :: m, r) | vl -> (* otherwise *) let cons_bound lst bd = match bd with | None -> lst | Some bnd -> (vl, vect, {info with bound = (Some bnd, None)}) :: lst in let lb, rb = info.bound in if Int.equal (sign_num vl) 1 then (cons_bound l lb, m, cons_bound r rb) else (* sign_num vl = -1 *) (cons_bound l rb, m, cons_bound r lb) (** [project vr sys] projects system [sys] over the set of variables [ISet.remove vr sys.vars ]. This is a one step Fourier elimination. *) let project vr sys = let l, m, r = System.fold (fun vect rf l_m_r -> split vr vect !rf l_m_r) sys.sys ([], [], []) in let new_sys = System.create (System.length sys.sys) in (* Constraints in [m] belong to the projection - for those [vr] is already projected out *) List.iter (fun (vect, info) -> System.replace new_sys vect (ref info)) m; let elim (v1, vect1, info1) (v2, vect2, info2) = let {neg = n1; pos = p1; bound = bound1; prf = prf1} = info1 and {neg = n2; pos = p2; bound = bound2; prf = prf2} = info2 in let bnd1 = Option.get (fst bound1) and bnd2 = Option.get (fst bound2) in let bound = (bnd1 // v1) +/ (bnd2 // minus_num v2) in let vres, (n, p) = add (vect1, v1) (vect2, minus_num v2) in ( vres , { neg = n ; pos = p ; bound = (Some bound, None) ; prf = Elim (vr, info1.prf, info2.prf) } ) in List.iter (fun l_elem -> List.iter (fun r_elem -> let vect, info = elim l_elem r_elem in match normalise_cstr vect info with | Redundant -> () | Contradiction -> raise (SystemContradiction info.prf) | Cstr (vect, info) -> xadd_cstr vect info new_sys) r) l; {sys = new_sys; vars = ISet.remove vr sys.vars} (** [project_using_eq] performs elimination by pivoting using an equation. This is the counter_part of the [elim] sub-function of [!project]. @param vr is the variable to be used as pivot @param c is the coefficient of variable [vr] in vector [vect] @param len is the length of the equation @param bound is the bound of the equation @param prf is the proof of the equation *) let project_using_eq vr c vect bound prf (vect', info') = match get vr vect' with | Int 0 -> (vect', info') | c2 -> let c1 = if c2 >=/ Int 0 then minus_num c else c in let c2 = abs_num c2 in let vres, (n, p) = add (vect, c1) (vect', c2) in let cst = bound // c1 in let bndres = let f x = cst +/ (x // c2) in let l, r = info'.bound in (Option.map f l, Option.map f r) in (vres, {neg = n; pos = p; bound = bndres; prf = Elim (vr, prf, info'.prf)}) let elim_var_using_eq vr vect cst prf sys = let c = get vr vect in let elim_var = project_using_eq vr c vect cst prf in let new_sys = System.create (System.length sys.sys) in System.iter (fun vect iref -> let vect', info' = elim_var (vect, !iref) in match normalise_cstr vect' info' with | Redundant -> () | Contradiction -> raise (SystemContradiction info'.prf) | Cstr (vect, info') -> xadd_cstr vect info' new_sys) sys.sys; {sys = new_sys; vars = ISet.remove vr sys.vars} (** [size sys] computes the number of entries in the system of constraints *) let size sys = System.fold (fun v iref s -> s + !iref.neg + !iref.pos) sys 0 module IMap = CMap.Make (Int) (** [eval_vect map vect] evaluates vector [vect] using the values of [map]. If [map] binds all the variables of [vect], we get [eval_vect map [(x1,v1);...;(xn,vn)] = (IMap.find x1 map * v1) + ... + (IMap.find xn map) * vn , []] The function returns as second argument, a sub-vector consisting in the variables that are not in [map]. *) let eval_vect map vect = Vect.fold (fun (sum, rst) v vl -> try let val_v = IMap.find v map in (sum +/ (val_v */ vl), rst) with Not_found -> (sum, Vect.set v vl rst)) (Int 0, Vect.null) vect (** [restrict_bound n sum itv] returns the interval of [x] given that (fst itv) <= x * n + sum <= (snd itv) *) let restrict_bound n sum (itv : interval) = let f x = (x -/ sum) // n in let l, r = itv in match sign_num n with | 0 -> if in_bound itv sum then (None, None) (* redundant *) else failwith "SystemContradiction" | 1 -> (Option.map f l, Option.map f r) | _ -> (Option.map f r, Option.map f l) (** [bound_of_variable map v sys] computes the interval of [v] in [sys] given a mapping [map] binding all the other variables *) let bound_of_variable map v sys = System.fold (fun vect iref bnd -> let sum, rst = eval_vect map vect in let vl = Vect.get v rst in match inter bnd (restrict_bound vl sum !iref.bound) with | None -> Printf.fprintf stdout "bound_of_variable: eval_vecr %a = %s,%a\n" Vect.pp vect (Num.string_of_num sum) Vect.pp rst; Printf.fprintf stdout "current interval: %a\n" Itv.pp !iref.bound; failwith "bound_of_variable: impossible" | Some itv -> itv) sys (None, None) (** [pick_small_value bnd] picks a value being closed to zero within the interval *) let pick_small_value bnd = match bnd with | None, None -> Int 0 | None, Some i -> if Int 0 <=/ floor_num i then Int 0 else floor_num i | Some i, None -> if i <=/ Int 0 then Int 0 else ceiling_num i | Some i, Some j -> if i <=/ Int 0 && Int 0 <=/ j then Int 0 else if ceiling_num i <=/ floor_num j then ceiling_num i (* why not *) else i (** [solution s1 sys_l = Some(sn,\[(vn-1,sn-1);...; (v1,s1)\]\@sys_l)] then [sn] is a system which contains only [black_v] -- if it existed in [s1] and [sn+1] is obtained by projecting [vn] out of [sn] @raise SystemContradiction if system [s] has no solution *) let solve_sys black_v choose_eq choose_variable sys sys_l = let rec solve_sys sys sys_l = if debug then Printf.printf "S #%i size %i\n" (System.length sys.sys) (size sys.sys); if debug then Printf.printf "solve_sys :\n %a" pp_system sys.sys; let eqs = choose_eq sys in try let v, vect, cst, ln = fst (List.find (fun ((v, _, _, _), _) -> v <> black_v) eqs) in if debug then ( Printf.printf "\nE %a = %s variable %i\n" Vect.pp vect (string_of_num cst) v; flush stdout ); let sys' = elim_var_using_eq v vect cst ln sys in solve_sys sys' ((v, sys) :: sys_l) with Not_found -> ( let vars = choose_variable sys in try let v, est = List.find (fun (v, _) -> v <> black_v) vars in if debug then ( Printf.printf "\nV : %i estimate %f\n" v est; flush stdout ); let sys' = project v sys in solve_sys sys' ((v, sys) :: sys_l) with Not_found -> (* we are done *) Inl (sys, sys_l) ) in solve_sys sys sys_l let solve black_v choose_eq choose_variable cstrs = try let sys = load_system cstrs in if debug then Printf.printf "solve :\n %a" pp_system sys.sys; solve_sys black_v choose_eq choose_variable sys [] with SystemContradiction prf -> Inr prf (** The purpose of module [EstimateElimVar] is to try to estimate the cost of eliminating a variable. The output is an ordered list of (variable,cost). *) module EstimateElimVar = struct type sys_list = (vector * cstr_info) list let abstract_partition (v : int) (l : sys_list) = let rec xpart (l : sys_list) (ltl : sys_list) (n : int list) (z : int) (p : int list) = match l with | [] -> (ltl, n, z, p) | (l1, info) :: rl -> ( match Vect.choose l1 with | None -> xpart rl ((Vect.null, info) :: ltl) n (info.neg + info.pos + z) p | Some (vr, vl, rl1) -> if Int.equal v vr then let cons_bound lst bd = match bd with | None -> lst | Some bnd -> (info.neg + info.pos) :: lst in let lb, rb = info.bound in if Int.equal (sign_num vl) 1 then xpart rl ((rl1, info) :: ltl) (cons_bound n lb) z (cons_bound p rb) else xpart rl ((rl1, info) :: ltl) (cons_bound n rb) z (cons_bound p lb) else (* the variable is greater *) xpart rl ((l1, info) :: ltl) n (info.neg + info.pos + z) p ) in let sys', n, z, p = xpart l [] [] 0 [] in let ln = float_of_int (List.length n) in let sn = float_of_int (List.fold_left ( + ) 0 n) in let lp = float_of_int (List.length p) in let sp = float_of_int (List.fold_left ( + ) 0 p) in (sys', float_of_int z +. (lp *. sn) +. (ln *. sp) -. (lp *. ln)) let choose_variable sys = let {sys = s; vars = v} = sys in let sl = system_list sys in let evals = fst (ISet.fold (fun v (eval, s) -> let ts, vl = abstract_partition v s in ((v, vl) :: eval, ts)) v ([], sl)) in List.sort (fun x y -> compare_float (snd x) (snd y)) evals end open EstimateElimVar (** The module [EstimateElimEq] is similar to [EstimateElimVar] but it orders equations. *) module EstimateElimEq = struct let itv_point bnd = match bnd with Some a, Some b -> a =/ b | _ -> false let rec unroll_until v l = match Vect.choose l with | None -> (false, Vect.null) | Some (i, _, rl) -> if Int.equal i v then (true, rl) else if i < v then unroll_until v rl else (false, l) let rec choose_simple_equation eqs = match eqs with | [] -> None | (vect, a, prf, ln) :: eqs -> ( match Vect.choose vect with | Some (i, v, rst) -> if Vect.is_null rst then Some (i, vect, a, prf, ln) else choose_simple_equation eqs | _ -> choose_simple_equation eqs ) let choose_primal_equation eqs (sys_l : (Vect.t * cstr_info) list) = (* Counts the number of equations referring to variable [v] -- It looks like nb_cst is dead... *) let is_primal_equation_var v = List.fold_left (fun nb_eq (vect, info) -> if fst (unroll_until v vect) then if itv_point info.bound then nb_eq + 1 else nb_eq else nb_eq) 0 sys_l in let rec find_var vect = match Vect.choose vect with | None -> None | Some (i, _, vect) -> let nb_eq = is_primal_equation_var i in if Int.equal nb_eq 2 then Some i else find_var vect in let rec find_eq_var eqs = match eqs with | [] -> None | (vect, a, prf, ln) :: l -> ( match find_var vect with | None -> find_eq_var l | Some r -> Some (r, vect, a, prf, ln) ) in match choose_simple_equation eqs with | None -> find_eq_var eqs | Some res -> Some res let choose_equality_var sys = let sys_l = system_list sys in let equalities = List.fold_left (fun l (vect, info) -> match info.bound with | Some a, Some b -> if a =/ b then (* This an equation *) (vect, a, info.prf, info.neg + info.pos) :: l else l | _ -> l) [] sys_l in let rec estimate_cost v ct sysl acc tlsys = match sysl with | [] -> (acc, tlsys) | (l, info) :: rsys -> ( let ln = info.pos + info.neg in let b, l = unroll_until v l in match b with | true -> if itv_point info.bound then estimate_cost v ct rsys (acc + ln) ((l, info) :: tlsys) (* this is free *) else estimate_cost v ct rsys (acc + ln + ct) ((l, info) :: tlsys) (* should be more ? *) | false -> estimate_cost v ct rsys (acc + ln) ((l, info) :: tlsys) ) in match choose_primal_equation equalities sys_l with | None -> let cost_eq eq const prf ln acc_costs = let rec cost_eq eqr sysl costs = match Vect.choose eqr with | None -> costs | Some (v, _, eqr) -> let cst, tlsys = estimate_cost v (ln - 1) sysl 0 [] in cost_eq eqr tlsys (((v, eq, const, prf), cst) :: costs) in cost_eq eq sys_l acc_costs in let all_costs = List.fold_left (fun all_costs (vect, const, prf, ln) -> cost_eq vect const prf ln all_costs) [] equalities in (* pp_list (fun o ((v,eq,_,_),cst) -> Printf.fprintf o "((%i,%a),%i)\n" v pp_vect eq cst) stdout all_costs ; *) List.sort (fun x y -> Int.compare (snd x) (snd y)) all_costs | Some (v, vect, const, prf, _) -> [((v, vect, const, prf), 0)] end open EstimateElimEq module Fourier = struct let optimise vect l = (* We add a dummy (fresh) variable for vector *) let fresh = List.fold_left (fun fr c -> max fr (Vect.fresh c.coeffs)) 0 l in let cstr = {coeffs = Vect.set fresh (Int (-1)) vect; op = Eq; cst = Int 0} in match solve fresh choose_equality_var choose_variable (cstr :: l) with | Inr prf -> None (* This is an unsatisfiability proof *) | Inl (s, _) -> ( try Some (bound_of_variable IMap.empty fresh s.sys) with x when CErrors.noncritical x -> Printf.printf "optimise Exception : %s" (Printexc.to_string x); None ) let find_point cstrs = match solve max_int choose_equality_var choose_variable cstrs with | Inr prf -> Inr prf | Inl (_, l) -> let rec rebuild_solution l map = match l with | [] -> map | (v, e) :: l -> let itv = bound_of_variable map v e.sys in let map = IMap.add v (pick_small_value itv) map in rebuild_solution l map in let map = rebuild_solution l IMap.empty in let vect = IMap.fold (fun v i vect -> Vect.set v i vect) map Vect.null in if debug then Printf.printf "SOLUTION %a" Vect.pp vect; let res = Inl vect in res end module Proof = struct (** A proof term in the sense of a ZMicromega.RatProof is a positive combination of the hypotheses which leads to a contradiction. The proofs constructed by Fourier elimination are more like execution traces: - certain facts are recorded but are useless - certain inferences are implicit. The following code implements proof reconstruction. *) let add x y = fst (add x y) let forall_pairs f l1 l2 = List.fold_left (fun acc e1 -> List.fold_left (fun acc e2 -> match f e1 e2 with None -> acc | Some v -> v :: acc) acc l2) [] l1 let add_op x y = match (x, y) with Eq, Eq -> Eq | _ -> Ge let pivot v (p1, c1) (p2, c2) = let {coeffs = v1; op = op1; cst = n1} = c1 and {coeffs = v2; op = op2; cst = n2} = c2 in match (Vect.get v v1, Vect.get v v2) with | Int 0, _ | _, Int 0 -> None | a, b -> if Int.equal (sign_num a * sign_num b) (-1) then Some ( add (p1, abs_num a) (p2, abs_num b) , { coeffs = add (v1, abs_num a) (v2, abs_num b) ; op = add_op op1 op2 ; cst = (n1 // abs_num a) +/ (n2 // abs_num b) } ) else if op1 == Eq then Some ( add (p1, minus_num (a // b)) (p2, Int 1) , { coeffs = add (v1, minus_num (a // b)) (v2, Int 1) ; op = add_op op1 op2 ; cst = (n1 // minus_num (a // b)) +/ (n2 // Int 1) } ) else if op2 == Eq then Some ( add (p2, minus_num (b // a)) (p1, Int 1) , { coeffs = add (v2, minus_num (b // a)) (v1, Int 1) ; op = add_op op1 op2 ; cst = (n2 // minus_num (b // a)) +/ (n1 // Int 1) } ) else None (* op2 could be Eq ... this might happen *) let normalise_proofs l = List.fold_left (fun acc (prf, cstr) -> match acc with | Inr _ -> acc (* I already found a contradiction *) | Inl acc -> ( match norm_cstr cstr 0 with | Redundant -> Inl acc | Contradiction -> Inr (prf, cstr) | Cstr (v, info) -> Inl ((prf, cstr, v, info) :: acc) )) (Inl []) l type oproof = (vector * cstr * num) option let merge_proof (oleft : oproof) (prf, cstr, v, info) (oright : oproof) = let l, r = info.bound in let keep p ob bd = match (ob, bd) with | None, None -> None | None, Some b -> Some (prf, cstr, b) | Some _, None -> ob | Some (prfl, cstrl, bl), Some b -> if p bl b then Some (prf, cstr, b) else ob in let oleft = keep ( <=/ ) oleft l in let oright = keep ( >=/ ) oright r in (* Now, there might be a contradiction *) match (oleft, oright) with | None, _ | _, None -> Inl (oleft, oright) | Some (prfl, cstrl, l), Some (prfr, cstrr, r) -> ( if l <=/ r then Inl (oleft, oright) else (* There is a contradiction - it should show up by scaling up the vectors - any pivot should do*) match Vect.choose cstrr.coeffs with | None -> Inr (add (prfl, Int 1) (prfr, Int 1), cstrr) (* this is wrong *) | Some (v, _, _) -> ( match pivot v (prfl, cstrl) (prfr, cstrr) with | None -> failwith "merge_proof : pivot is not possible" | Some x -> Inr x ) ) let mk_proof hyps prf = (* I am keeping list - I might have a proof for the left bound and a proof for the right bound. If I perform aggressive elimination of redundancies, I expect the list to be of length at most 2. For each proof list, all the vectors should be of the form a.v for different constants a. *) let rec mk_proof prf = match prf with | Assum i -> [(Vect.set i (Int 1) Vect.null, List.nth hyps i)] | Elim (v, prf1, prf2) -> let prfsl = mk_proof prf1 and prfsr = mk_proof prf2 in (* I take only the pairs for which the elimination is meaningful *) forall_pairs (pivot v) prfsl prfsr | And (prf1, prf2) -> ( let prfsl1 = mk_proof prf1 and prfsl2 = mk_proof prf2 in (* detect trivial redundancies and contradictions *) match normalise_proofs (prfsl1 @ prfsl2) with | Inr x -> [x] (* This is a contradiction - this should be the end of the proof *) | Inl l -> ( (* All the vectors are the same *) let prfs = List.fold_left (fun acc e -> match acc with | Inr _ -> acc (* I have a contradiction *) | Inl (oleft, oright) -> merge_proof oleft e oright) (Inl (None, None)) l in match prfs with | Inr x -> [x] | Inl (oleft, oright) -> ( match (oleft, oright) with | None, None -> [] | None, Some (prf, cstr, _) | Some (prf, cstr, _), None -> [(prf, cstr)] | Some (prf1, cstr1, _), Some (prf2, cstr2, _) -> [(prf1, cstr1); (prf2, cstr2)] ) ) ) in mk_proof prf end coq-8.11.0/plugins/micromega/Lra.v0000644000175000017500000000425013612664533016620 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (sos_R rchecker) || (psatz_R d rchecker) | _ => fail "Unsupported domain" end in tac. Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n. Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1). (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/g_micromega.mli0000644000175000017500000000124313612664533020666 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Mc.q Mc.psatz (** [q_cert_of_pos prf] converts a Sos proof into a rational Coq proof *) val z_cert_of_pos : Sos_types.positivstellensatz -> Mc.z Mc.psatz (** [z_cert_of_pos prf] converts a Sos proof into an integer Coq proof *) val lia : bool -> int -> (Mc.z Mc.pExpr * Mc.op1) list -> zres (** [lia enum depth sys] generates an unsat proof for the linear constraints in [sys]. If the Simplex option is set, any failure to find a proof should be considered as a bug. *) val nlia : bool -> int -> (Mc.z Mc.pExpr * Mc.op1) list -> zres (** [nlia enum depth sys] generates an unsat proof for the non-linear constraints in [sys]. The solver is incomplete -- the problem is undecidable *) val linear_prover_with_cert : int -> (Mc.q Mc.pExpr * Mc.op1) list -> qres (** [linear_prover_with_cert depth sys] generates an unsat proof for the linear constraints in [sys]. Over the rationals, the solver is complete. *) val nlinear_prover : int -> (Mc.q Mc.pExpr * Mc.op1) list -> qres (** [nlinear depth sys] generates an unsat proof for the non-linear constraints in [sys]. The solver is incompete -- the problem is decidable. *) coq-8.11.0/plugins/micromega/Env.v0000644000175000017500000000626113612664533016636 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* D. Definition jump (j:positive) (e:Env) := fun x => e (x+j). Definition nth (n:positive) (e:Env) := e n. Definition hd (e:Env) := nth 1 e. Definition tail (e:Env) := jump 1 e. Lemma jump_add i j l x : jump (i + j) l x = jump i (jump j l) x. Proof. unfold jump. f_equal. apply Pos.add_assoc. Qed. Lemma jump_simpl p l x : jump p l x = match p with | xH => tail l x | xO p => jump p (jump p l) x | xI p => jump p (jump p (tail l)) x end. Proof. destruct p; unfold tail; rewrite <- ?jump_add; f_equal; now rewrite Pos.add_diag. Qed. Lemma jump_tl j l x : tail (jump j l) x = jump j (tail l) x. Proof. unfold tail. rewrite <- !jump_add. f_equal. apply Pos.add_comm. Qed. Lemma jump_succ j l x : jump (Pos.succ j) l x = jump 1 (jump j l) x. Proof. rewrite <- jump_add. f_equal. symmetry. apply Pos.add_1_l. Qed. Lemma jump_pred_double i l x : jump (Pos.pred_double i) (tail l) x = jump i (jump i l) x. Proof. unfold tail. rewrite <- !jump_add. f_equal. now rewrite Pos.add_1_r, Pos.succ_pred_double, Pos.add_diag. Qed. Lemma nth_spec p l : nth p l = match p with | xH => hd l | xO p => nth p (jump p l) | xI p => nth p (jump p (tail l)) end. Proof. unfold hd, nth, tail, jump. destruct p; f_equal; now rewrite Pos.add_diag. Qed. Lemma nth_jump p l : nth p (tail l) = hd (jump p l). Proof. unfold hd, nth, tail, jump. f_equal. apply Pos.add_comm. Qed. Lemma nth_pred_double p l : nth (Pos.pred_double p) (tail l) = nth p (jump p l). Proof. unfold nth, tail, jump. f_equal. now rewrite Pos.add_1_r, Pos.succ_pred_double, Pos.add_diag. Qed. End S. Ltac jump_simpl := repeat match goal with | |- context [jump xH] => rewrite (jump_simpl xH) | |- context [jump (xO ?p)] => rewrite (jump_simpl (xO p)) | |- context [jump (xI ?p)] => rewrite (jump_simpl (xI p)) end. coq-8.11.0/plugins/micromega/csdpcert.mli0000644000175000017500000000124313612664533020224 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* -1 | Eq => 0 | Gt => 1 end. Lemma Z_of_comparison_bound : forall x, -1 <= Z_of_comparison x <= 1. Proof. destruct x ; simpl; compute; intuition congruence. Qed. Instance Inj_comparison_Z : InjTyp comparison Z := { inj := Z_of_comparison ; pred :=(fun x => -1 <= x <= 1) ; cstr := Z_of_comparison_bound}. Add InjTyp Inj_comparison_Z. Definition ZcompareZ (x y : Z) := Z_of_comparison (Z.compare x y). Program Instance BinOp_Zcompare : BinOp Z.compare := { TBOp := ZcompareZ }. Add BinOp BinOp_Zcompare. Instance Op_eq_comparison : BinRel (@eq comparison) := {TR := @eq Z ; TRInj := ltac:(destruct n,m; simpl ; intuition congruence) }. Add BinRel Op_eq_comparison. Instance Op_Eq : CstOp Eq := { TCst := 0 ; TCstInj := eq_refl }. Add CstOp Op_Eq. Instance Op_Lt : CstOp Lt := { TCst := -1 ; TCstInj := eq_refl }. Add CstOp Op_Lt. Instance Op_Gt : CstOp Gt := { TCst := 1 ; TCstInj := eq_refl }. Add CstOp Op_Gt. Lemma Zcompare_spec : forall x y, (x = y -> ZcompareZ x y = 0) /\ (x > y -> ZcompareZ x y = 1) /\ (x < y -> ZcompareZ x y = -1). Proof. unfold ZcompareZ. intros. destruct (x ?= y) eqn:C; simpl. - rewrite Z.compare_eq_iff in C. intuition. - rewrite Z.compare_lt_iff in C. intuition. - rewrite Z.compare_gt_iff in C. intuition. Qed. Instance ZcompareSpec : BinOpSpec ZcompareZ := {| BPred := fun x y r => (x = y -> r = 0) /\ (x > y -> r = 1) /\ (x < y -> r = -1) ; BSpec := Zcompare_spec|}. Add Spec ZcompareSpec. coq-8.11.0/plugins/micromega/ZifyInst.v0000644000175000017500000003046513612664533017670 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* x -> y). Proof. constructor. intros. tauto. Defined. Add PropOp PropArrow. Instance PropIff : PropOp iff. Proof. constructor. intros. tauto. Defined. Add PropOp PropIff. Instance PropNot : PropUOp not. Proof. constructor. intros. tauto. Defined. Add PropUOp PropNot. Instance Inj_Z_Z : InjTyp Z Z := mkinj _ _ (fun x => x) (fun x => True ) (fun _ => I). Add InjTyp Inj_Z_Z. (** Support for nat *) Instance Inj_nat_Z : InjTyp nat Z := mkinj _ _ Z.of_nat (fun x => 0 <= x ) Nat2Z.is_nonneg. Add InjTyp Inj_nat_Z. (* zify_nat_rel *) Instance Op_ge : BinRel ge := {| TR := Z.ge; TRInj := Nat2Z.inj_ge |}. Add BinRel Op_ge. Instance Op_lt : BinRel lt := {| TR := Z.lt; TRInj := Nat2Z.inj_lt |}. Add BinRel Op_lt. Instance Op_gt : BinRel gt := {| TR := Z.gt; TRInj := Nat2Z.inj_gt |}. Add BinRel Op_gt. Instance Op_le : BinRel le := {| TR := Z.le; TRInj := Nat2Z.inj_le |}. Add BinRel Op_le. Instance Op_eq_nat : BinRel (@eq nat) := {| TR := @eq Z ; TRInj := fun x y : nat => iff_sym (Nat2Z.inj_iff x y) |}. Add BinRel Op_eq_nat. (* zify_nat_op *) Instance Op_plus : BinOp Nat.add := {| TBOp := Z.add; TBOpInj := Nat2Z.inj_add |}. Add BinOp Op_plus. Instance Op_sub : BinOp Nat.sub := {| TBOp := fun n m => Z.max 0 (n - m) ; TBOpInj := Nat2Z.inj_sub_max |}. Add BinOp Op_sub. Instance Op_mul : BinOp Nat.mul := {| TBOp := Z.mul ; TBOpInj := Nat2Z.inj_mul |}. Add BinOp Op_mul. Instance Op_min : BinOp Nat.min := {| TBOp := Z.min ; TBOpInj := Nat2Z.inj_min |}. Add BinOp Op_min. Instance Op_max : BinOp Nat.max := {| TBOp := Z.max ; TBOpInj := Nat2Z.inj_max |}. Add BinOp Op_max. Instance Op_pred : UnOp Nat.pred := {| TUOp := fun n => Z.max 0 (n - 1) ; TUOpInj := Nat2Z.inj_pred_max |}. Add UnOp Op_pred. Instance Op_S : UnOp S := {| TUOp := fun x => Z.add x 1 ; TUOpInj := Nat2Z.inj_succ |}. Add UnOp Op_S. Instance Op_O : CstOp O := {| TCst := Z0 ; TCstInj := eq_refl (Z.of_nat 0) |}. Add CstOp Op_O. Instance Op_Z_abs_nat : UnOp Z.abs_nat := { TUOp := Z.abs ; TUOpInj := Zabs2Nat.id_abs }. Add UnOp Op_Z_abs_nat. (** Support for positive *) Instance Inj_pos_Z : InjTyp positive Z := {| inj := Zpos ; pred := (fun x => 0 < x ) ; cstr := Pos2Z.pos_is_pos |}. Add InjTyp Inj_pos_Z. Instance Op_pos_to_nat : UnOp Pos.to_nat := {TUOp := (fun x => x); TUOpInj := positive_nat_Z}. Add UnOp Op_pos_to_nat. Instance Inj_N_Z : InjTyp N Z := mkinj _ _ Z.of_N (fun x => 0 <= x ) N2Z.is_nonneg. Add InjTyp Inj_N_Z. Instance Op_N_to_nat : UnOp N.to_nat := { TUOp := fun x => x ; TUOpInj := N_nat_Z }. Add UnOp Op_N_to_nat. (* zify_positive_rel *) Instance Op_pos_ge : BinRel Pos.ge := {| TR := Z.ge; TRInj := fun x y => iff_refl (Z.pos x >= Z.pos y) |}. Add BinRel Op_pos_ge. Instance Op_pos_lt : BinRel Pos.lt := {| TR := Z.lt; TRInj := fun x y => iff_refl (Z.pos x < Z.pos y) |}. Add BinRel Op_pos_lt. Instance Op_pos_gt : BinRel Pos.gt := {| TR := Z.gt; TRInj := fun x y => iff_refl (Z.pos x > Z.pos y) |}. Add BinRel Op_pos_gt. Instance Op_pos_le : BinRel Pos.le := {| TR := Z.le; TRInj := fun x y => iff_refl (Z.pos x <= Z.pos y) |}. Add BinRel Op_pos_le. Instance Op_eq_pos : BinRel (@eq positive) := {| TR := @eq Z ; TRInj := fun x y => iff_sym (Pos2Z.inj_iff x y) |}. Add BinRel Op_eq_pos. (* zify_positive_op *) Program Instance Op_Z_of_N : UnOp Z.of_N := { TUOp := (fun x => x) ; TUOpInj := fun x => eq_refl (Z.of_N x) }. Add UnOp Op_Z_of_N. Instance Op_Z_to_N : UnOp Z.to_N := { TUOp := fun x => Z.max 0 x ; TUOpInj := ltac:(now intro x; destruct x) }. Add UnOp Op_Z_to_N. Instance Op_Z_neg : UnOp Z.neg := { TUOp := Z.opp ; TUOpInj := (fun x => eq_refl (Zneg x))}. Add UnOp Op_Z_neg. Instance Op_Z_pos : UnOp Z.pos := { TUOp := (fun x => x) ; TUOpInj := (fun x => eq_refl (Z.pos x))}. Add UnOp Op_Z_pos. Instance Op_pos_succ : UnOp Pos.succ := { TUOp := fun x => x + 1; TUOpInj := Pos2Z.inj_succ }. Add UnOp Op_pos_succ. Instance Op_pos_pred : UnOp Pos.pred := { TUOp := fun x => Z.max 1 (x - 1) ; TUOpInj := ltac : (intros; rewrite <- Pos.sub_1_r; apply Pos2Z.inj_sub_max) }. Add UnOp Op_pos_pred. Instance Op_pos_of_succ_nat : UnOp Pos.of_succ_nat := { TUOp := fun x => x + 1 ; TUOpInj := Zpos_P_of_succ_nat }. Add UnOp Op_pos_of_succ_nat. Program Instance Op_pos_add : BinOp Pos.add := { TBOp := Z.add ; TBOpInj := ltac: (reflexivity) }. Add BinOp Op_pos_add. Instance Op_pos_sub : BinOp Pos.sub := { TBOp := fun n m => Z.max 1 (n - m) ;TBOpInj := Pos2Z.inj_sub_max }. Add BinOp Op_pos_sub. Instance Op_pos_mul : BinOp Pos.mul := { TBOp := Z.mul ; TBOpInj := ltac: (reflexivity) }. Add BinOp Op_pos_mul. Instance Op_pos_min : BinOp Pos.min := { TBOp := Z.min ; TBOpInj := Pos2Z.inj_min }. Add BinOp Op_pos_min. Instance Op_pos_max : BinOp Pos.max := { TBOp := Z.max ; TBOpInj := Pos2Z.inj_max }. Add BinOp Op_pos_max. Instance Op_xO : UnOp xO := { TUOp := fun x => 2 * x ; TUOpInj := ltac: (reflexivity) }. Add UnOp Op_xO. Instance Op_xI : UnOp xI := { TUOp := fun x => 2 * x + 1 ; TUOpInj := ltac: (reflexivity) }. Add UnOp Op_xI. Instance Op_xH : CstOp xH := { TCst := 1%Z ; TCstInj := ltac:(reflexivity)}. Add CstOp Op_xH. Instance Op_Z_of_nat : UnOp Z.of_nat:= { TUOp := fun x => x ; TUOpInj := ltac:(reflexivity) }. Add UnOp Op_Z_of_nat. (* zify_N_rel *) Instance Op_N_ge : BinRel N.ge := {| TR := Z.ge ; TRInj := N2Z.inj_ge |}. Add BinRel Op_N_ge. Instance Op_N_lt : BinRel N.lt := {| TR := Z.lt ; TRInj := N2Z.inj_lt |}. Add BinRel Op_N_lt. Instance Op_N_gt : BinRel N.gt := {| TR := Z.gt ; TRInj := N2Z.inj_gt |}. Add BinRel Op_N_gt. Instance Op_N_le : BinRel N.le := {| TR := Z.le ; TRInj := N2Z.inj_le |}. Add BinRel Op_N_le. Instance Op_eq_N : BinRel (@eq N) := {| TR := @eq Z ; TRInj := fun x y : N => iff_sym (N2Z.inj_iff x y) |}. Add BinRel Op_eq_N. (* zify_N_op *) Instance Op_N_of_nat : UnOp N.of_nat := { TUOp := fun x => x ; TUOpInj := nat_N_Z }. Add UnOp Op_N_of_nat. Instance Op_Z_abs_N : UnOp Z.abs_N := { TUOp := Z.abs ; TUOpInj := N2Z.inj_abs_N }. Add UnOp Op_Z_abs_N. Instance Op_N_pos : UnOp N.pos := { TUOp := fun x => x ; TUOpInj := ltac:(reflexivity)}. Add UnOp Op_N_pos. Instance Op_N_add : BinOp N.add := {| TBOp := Z.add ; TBOpInj := N2Z.inj_add |}. Add BinOp Op_N_add. Instance Op_N_min : BinOp N.min := {| TBOp := Z.min ; TBOpInj := N2Z.inj_min |}. Add BinOp Op_N_min. Instance Op_N_max : BinOp N.max := {| TBOp := Z.max ; TBOpInj := N2Z.inj_max |}. Add BinOp Op_N_max. Instance Op_N_mul : BinOp N.mul := {| TBOp := Z.mul ; TBOpInj := N2Z.inj_mul |}. Add BinOp Op_N_mul. Instance Op_N_sub : BinOp N.sub := {| TBOp := fun x y => Z.max 0 (x - y) ; TBOpInj := N2Z.inj_sub_max|}. Add BinOp Op_N_sub. Instance Op_N_div : BinOp N.div := {| TBOp := Z.div ; TBOpInj := N2Z.inj_div|}. Add BinOp Op_N_div. Instance Op_N_mod : BinOp N.modulo := {| TBOp := Z.rem ; TBOpInj := N2Z.inj_rem|}. Add BinOp Op_N_mod. Instance Op_N_pred : UnOp N.pred := { TUOp := fun x => Z.max 0 (x - 1) ; TUOpInj := ltac:(intros; rewrite N.pred_sub; apply N2Z.inj_sub_max) }. Add UnOp Op_N_pred. Instance Op_N_succ : UnOp N.succ := {| TUOp := fun x => x + 1 ; TUOpInj := N2Z.inj_succ |}. Add UnOp Op_N_succ. (** Support for Z - injected to itself *) (* zify_Z_rel *) Instance Op_Z_ge : BinRel Z.ge := {| TR := Z.ge ; TRInj := fun x y => iff_refl (x>= y)|}. Add BinRel Op_Z_ge. Instance Op_Z_lt : BinRel Z.lt := {| TR := Z.lt ; TRInj := fun x y => iff_refl (x < y)|}. Add BinRel Op_Z_lt. Instance Op_Z_gt : BinRel Z.gt := {| TR := Z.gt ;TRInj := fun x y => iff_refl (x > y)|}. Add BinRel Op_Z_gt. Instance Op_Z_le : BinRel Z.le := {| TR := Z.le ;TRInj := fun x y => iff_refl (x <= y)|}. Add BinRel Op_Z_le. Instance Op_eqZ : BinRel (@eq Z) := { TR := @eq Z ; TRInj := fun x y => iff_refl (x = y) }. Add BinRel Op_eqZ. Instance Op_Z_add : BinOp Z.add := { TBOp := Z.add ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_add. Instance Op_Z_min : BinOp Z.min := { TBOp := Z.min ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_min. Instance Op_Z_max : BinOp Z.max := { TBOp := Z.max ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_max. Instance Op_Z_mul : BinOp Z.mul := { TBOp := Z.mul ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_mul. Instance Op_Z_sub : BinOp Z.sub := { TBOp := Z.sub ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_sub. Instance Op_Z_div : BinOp Z.div := { TBOp := Z.div ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_div. Instance Op_Z_mod : BinOp Z.modulo := { TBOp := Z.modulo ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_mod. Instance Op_Z_rem : BinOp Z.rem := { TBOp := Z.rem ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_rem. Instance Op_Z_quot : BinOp Z.quot := { TBOp := Z.quot ; TBOpInj := ltac:(reflexivity) }. Add BinOp Op_Z_quot. Instance Op_Z_succ : UnOp Z.succ := { TUOp := fun x => x + 1 ; TUOpInj := ltac:(reflexivity) }. Add UnOp Op_Z_succ. Instance Op_Z_pred : UnOp Z.pred := { TUOp := fun x => x - 1 ; TUOpInj := ltac:(reflexivity) }. Add UnOp Op_Z_pred. Instance Op_Z_opp : UnOp Z.opp := { TUOp := Z.opp ; TUOpInj := ltac:(reflexivity) }. Add UnOp Op_Z_opp. Instance Op_Z_abs : UnOp Z.abs := { TUOp := Z.abs ; TUOpInj := ltac:(reflexivity) }. Add UnOp Op_Z_abs. Instance Op_Z_sgn : UnOp Z.sgn := { TUOp := Z.sgn ; TUOpInj := ltac:(reflexivity) }. Add UnOp Op_Z_sgn. Lemma of_nat_to_nat_eq : forall x, Z.of_nat (Z.to_nat x) = Z.max 0 x. Proof. destruct x. - reflexivity. - rewrite Z2Nat.id. reflexivity. compute. congruence. - reflexivity. Qed. Instance Op_Z_to_nat : UnOp Z.to_nat := { TUOp := fun x => Z.max 0 x ; TUOpInj := of_nat_to_nat_eq }. Add UnOp Op_Z_to_nat. (** Specification of derived operators over Z *) Lemma z_max_spec : forall n m, n <= Z.max n m /\ m <= Z.max n m /\ (Z.max n m = n \/ Z.max n m = m). Proof. intros. generalize (Z.le_max_l n m). generalize (Z.le_max_r n m). generalize (Z.max_spec_le n m). intuition idtac. Qed. Instance ZmaxSpec : BinOpSpec Z.max := {| BPred := fun n m r => n < m /\ r = m \/ m <= n /\ r = n ; BSpec := Z.max_spec|}. Add Spec ZmaxSpec. Lemma z_min_spec : forall n m, Z.min n m <= n /\ Z.min n m <= m /\ (Z.min n m = n \/ Z.min n m = m). Proof. intros. generalize (Z.le_min_l n m). generalize (Z.le_min_r n m). generalize (Z.min_spec_le n m). intuition idtac. Qed. Program Instance ZminSpec : BinOpSpec Z.min := {| BPred := fun n m r => n < m /\ r = n \/ m <= n /\ r = m ; BSpec := Z.min_spec |}. Add Spec ZminSpec. Instance ZsgnSpec : UnOpSpec Z.sgn := {| UPred := fun n r : Z => 0 < n /\ r = 1 \/ 0 = n /\ r = 0 \/ n < 0 /\ r = - (1) ; USpec := Z.sgn_spec|}. Add Spec ZsgnSpec. Instance ZabsSpec : UnOpSpec Z.abs := {| UPred := fun n r: Z => 0 <= n /\ r = n \/ n < 0 /\ r = - n ; USpec := Z.abs_spec|}. Add Spec ZabsSpec. (** Saturate positivity constraints *) Instance SatProd : Saturate Z.mul := {| PArg1 := fun x => 0 <= x; PArg2 := fun y => 0 <= y; PRes := fun r => 0 <= r; SatOk := Z.mul_nonneg_nonneg |}. Add Saturate SatProd. Instance SatProdPos : Saturate Z.mul := {| PArg1 := fun x => 0 < x; PArg2 := fun y => 0 < y; PRes := fun r => 0 < r; SatOk := Z.mul_pos_pos |}. Add Saturate SatProdPos. coq-8.11.0/plugins/micromega/sos_lib.ml0000644000175000017500000004544413612664533017711 0ustar treinentreinen(* ========================================================================= *) (* - This code originates from John Harrison's HOL LIGHT 2.30 *) (* (see file LICENSE.sos for license, copyright and disclaimer) *) (* This code is the HOL LIGHT library code used by sos.ml *) (* - Laurent Théry (thery@sophia.inria.fr) has isolated the HOL *) (* independent bits *) (* - Frédéric Besson (fbesson@irisa.fr) is using it to feed micromega *) (* ========================================================================= *) open Num (* ------------------------------------------------------------------------- *) (* Comparisons that are reflexive on NaN and also short-circuiting. *) (* ------------------------------------------------------------------------- *) (** FIXME *) let cmp = compare let ( =? ) x y = cmp x y = 0 let ( ? ) x y = cmp x y > 0 (* ------------------------------------------------------------------------- *) (* Combinators. *) (* ------------------------------------------------------------------------- *) let o f g x = f (g x) (* ------------------------------------------------------------------------- *) (* Some useful functions on "num" type. *) (* ------------------------------------------------------------------------- *) let num_0 = Int 0 and num_1 = Int 1 and num_2 = Int 2 and num_10 = Int 10 let pow2 n = power_num num_2 (Int n) let pow10 n = power_num num_10 (Int n) let numdom r = let r' = Ratio.normalize_ratio (ratio_of_num r) in ( num_of_big_int (Ratio.numerator_ratio r') , num_of_big_int (Ratio.denominator_ratio r') ) let numerator = o fst numdom and denominator = o snd numdom let gcd_num n1 n2 = num_of_big_int (Big_int.gcd_big_int (big_int_of_num n1) (big_int_of_num n2)) let lcm_num x y = if x =/ num_0 && y =/ num_0 then num_0 else abs_num (x */ y // gcd_num x y) (* ------------------------------------------------------------------------- *) (* Various versions of list iteration. *) (* ------------------------------------------------------------------------- *) let rec end_itlist f l = match l with | [] -> failwith "end_itlist" | [x] -> x | h :: t -> f h (end_itlist f t) (* ------------------------------------------------------------------------- *) (* All pairs arising from applying a function over two lists. *) (* ------------------------------------------------------------------------- *) let rec allpairs f l1 l2 = match l1 with | h1 :: t1 -> List.fold_right (fun x a -> f h1 x :: a) l2 (allpairs f t1 l2) | [] -> [] (* ------------------------------------------------------------------------- *) (* String operations (surely there is a better way...) *) (* ------------------------------------------------------------------------- *) let implode l = List.fold_right ( ^ ) l "" let explode s = let rec exap n l = if n < 0 then l else exap (n - 1) (String.sub s n 1 :: l) in exap (String.length s - 1) [] (* ------------------------------------------------------------------------- *) (* Repetition of a function. *) (* ------------------------------------------------------------------------- *) let rec funpow n f x = if n < 1 then x else funpow (n - 1) f (f x) (* ------------------------------------------------------------------------- *) (* Sequences. *) (* ------------------------------------------------------------------------- *) let rec ( -- ) m n = if m > n then [] else m :: (m + 1 -- n) (* ------------------------------------------------------------------------- *) (* Various useful list operations. *) (* ------------------------------------------------------------------------- *) let rec tryfind f l = match l with | [] -> failwith "tryfind" | h :: t -> ( try f h with Failure _ -> tryfind f t ) (* ------------------------------------------------------------------------- *) (* "Set" operations on lists. *) (* ------------------------------------------------------------------------- *) let rec mem x lis = match lis with [] -> false | h :: t -> x =? h || mem x t let insert x l = if mem x l then l else x :: l let union l1 l2 = List.fold_right insert l1 l2 let subtract l1 l2 = List.filter (fun x -> not (mem x l2)) l1 (* ------------------------------------------------------------------------- *) (* Common measure predicates to use with "sort". *) (* ------------------------------------------------------------------------- *) let increasing f x y = f x () | h :: t -> f h; do_list f t (* ------------------------------------------------------------------------- *) (* Sorting. *) (* ------------------------------------------------------------------------- *) let rec sort cmp lis = match lis with | [] -> [] | piv :: rest -> let r, l = List.partition (cmp piv) rest in sort cmp l @ (piv :: sort cmp r) (* ------------------------------------------------------------------------- *) (* Removing adjacent (NB!) equal elements from list. *) (* ------------------------------------------------------------------------- *) let rec uniq l = match l with | x :: (y :: _ as t) -> let t' = uniq t in if x =? y then t' else if t' == t then l else x :: t' | _ -> l (* ------------------------------------------------------------------------- *) (* Convert list into set by eliminating duplicates. *) (* ------------------------------------------------------------------------- *) let setify s = uniq (sort ( <=? ) s) (* ------------------------------------------------------------------------- *) (* Polymorphic finite partial functions via Patricia trees. *) (* *) (* The point of this strange representation is that it is canonical (equal *) (* functions have the same encoding) yet reasonably efficient on average. *) (* *) (* Idea due to Diego Olivier Fernandez Pons (OCaml list, 2003/11/10). *) (* ------------------------------------------------------------------------- *) type ('a, 'b) func = | Empty | Leaf of int * ('a * 'b) list | Branch of int * int * ('a, 'b) func * ('a, 'b) func (* ------------------------------------------------------------------------- *) (* Undefined function. *) (* ------------------------------------------------------------------------- *) let undefined = Empty (* ------------------------------------------------------------------------- *) (* In case of equality comparison worries, better use this. *) (* ------------------------------------------------------------------------- *) let is_undefined f = match f with Empty -> true | _ -> false (* ------------------------------------------------------------------------- *) (* Operation analogous to "map" for lists. *) (* ------------------------------------------------------------------------- *) let mapf = let rec map_list f l = match l with [] -> [] | (x, y) :: t -> (x, f y) :: map_list f t in let rec mapf f t = match t with | Empty -> Empty | Leaf (h, l) -> Leaf (h, map_list f l) | Branch (p, b, l, r) -> Branch (p, b, mapf f l, mapf f r) in mapf (* ------------------------------------------------------------------------- *) (* Operations analogous to "fold" for lists. *) (* ------------------------------------------------------------------------- *) let foldl = let rec foldl_list f a l = match l with [] -> a | (x, y) :: t -> foldl_list f (f a x y) t in let rec foldl f a t = match t with | Empty -> a | Leaf (h, l) -> foldl_list f a l | Branch (p, b, l, r) -> foldl f (foldl f a l) r in foldl let foldr = let rec foldr_list f l a = match l with [] -> a | (x, y) :: t -> f x y (foldr_list f t a) in let rec foldr f t a = match t with | Empty -> a | Leaf (h, l) -> foldr_list f l a | Branch (p, b, l, r) -> foldr f l (foldr f r a) in foldr (* ------------------------------------------------------------------------- *) (* Redefinition and combination. *) (* ------------------------------------------------------------------------- *) let ( |-> ), combine = let ldb x y = let z = x lxor y in z land -z in let newbranch p1 t1 p2 t2 = let b = ldb p1 p2 in let p = p1 land (b - 1) in if p1 land b = 0 then Branch (p, b, t1, t2) else Branch (p, b, t2, t1) in let rec define_list ((x, y) as xy) l = match l with | ((a, b) as ab) :: t -> if x =? a then xy :: t else if x [xy] and combine_list op z l1 l2 = match (l1, l2) with | [], _ -> l2 | _, [] -> l1 | ((x1, y1) as xy1) :: t1, ((x2, y2) as xy2) :: t2 -> if x1 ) x y = let k = Hashtbl.hash x in let rec upd t = match t with | Empty -> Leaf (k, [(x, y)]) | Leaf (h, l) -> if h = k then Leaf (h, define_list (x, y) l) else newbranch h t k (Leaf (k, [(x, y)])) | Branch (p, b, l, r) -> if k land (b - 1) <> p then newbranch p t k (Leaf (k, [(x, y)])) else if k land b = 0 then Branch (p, b, upd l, r) else Branch (p, b, l, upd r) in upd in let rec combine op z t1 t2 = match (t1, t2) with | Empty, _ -> t2 | _, Empty -> t1 | Leaf (h1, l1), Leaf (h2, l2) -> if h1 = h2 then let l = combine_list op z l1 l2 in if l = [] then Empty else Leaf (h1, l) else newbranch h1 t1 h2 t2 | (Leaf (k, lis) as lf), (Branch (p, b, l, r) as br) |(Branch (p, b, l, r) as br), (Leaf (k, lis) as lf) -> if k land (b - 1) = p then if k land b = 0 then let l' = combine op z lf l in if is_undefined l' then r else Branch (p, b, l', r) else let r' = combine op z lf r in if is_undefined r' then l else Branch (p, b, l, r') else newbranch k lf p br | Branch (p1, b1, l1, r1), Branch (p2, b2, l2, r2) -> if b1 < b2 then if p2 land (b1 - 1) <> p1 then newbranch p1 t1 p2 t2 else if p2 land b1 = 0 then let l = combine op z l1 t2 in if is_undefined l then r1 else Branch (p1, b1, l, r1) else let r = combine op z r1 t2 in if is_undefined r then l1 else Branch (p1, b1, l1, r) else if b2 < b1 then if p1 land (b2 - 1) <> p2 then newbranch p1 t1 p2 t2 else if p1 land b2 = 0 then let l = combine op z t1 l2 in if is_undefined l then r2 else Branch (p2, b2, l, r2) else let r = combine op z t1 r2 in if is_undefined r then l2 else Branch (p2, b2, l2, r) else if p1 = p2 then let l = combine op z l1 l2 and r = combine op z r1 r2 in if is_undefined l then r else if is_undefined r then l else Branch (p1, b1, l, r) else newbranch p1 t1 p2 t2 in (( |-> ), combine) (* ------------------------------------------------------------------------- *) (* Special case of point function. *) (* ------------------------------------------------------------------------- *) let ( |=> ) x y = (x |-> y) undefined (* ------------------------------------------------------------------------- *) (* Grab an arbitrary element. *) (* ------------------------------------------------------------------------- *) let rec choose t = match t with | Empty -> failwith "choose: completely undefined function" | Leaf (h, l) -> List.hd l | Branch (b, p, t1, t2) -> choose t1 (* ------------------------------------------------------------------------- *) (* Application. *) (* ------------------------------------------------------------------------- *) let applyd = let rec apply_listd l d x = match l with | (a, b) :: t -> if x =? a then b else if x >? a then apply_listd t d x else d x | [] -> d x in fun f d x -> let k = Hashtbl.hash x in let rec look t = match t with | Leaf (h, l) when h = k -> apply_listd l d x | Branch (p, b, l, r) -> look (if k land b = 0 then l else r) | _ -> d x in look f let apply f = applyd f (fun x -> failwith "apply") let tryapplyd f a d = applyd f (fun x -> d) a (* ------------------------------------------------------------------------- *) (* Undefinition. *) (* ------------------------------------------------------------------------- *) let undefine = let rec undefine_list x l = match l with | ((a, b) as ab) :: t -> if x =? a then t else if x [] in fun x -> let k = Hashtbl.hash x in let rec und t = match t with | Leaf (h, l) when h = k -> let l' = undefine_list x l in if l' == l then t else if l' = [] then Empty else Leaf (h, l') | Branch (p, b, l, r) when k land (b - 1) = p -> if k land b = 0 then let l' = und l in if l' == l then t else if is_undefined l' then r else Branch (p, b, l', r) else let r' = und r in if r' == r then t else if is_undefined r' then l else Branch (p, b, l, r') | _ -> t in und (* ------------------------------------------------------------------------- *) (* Mapping to sorted-list representation of the graph, domain and range. *) (* ------------------------------------------------------------------------- *) let graph f = setify (foldl (fun a x y -> (x, y) :: a) [] f) let dom f = setify (foldl (fun a x y -> x :: a) [] f) (* ------------------------------------------------------------------------- *) (* More parser basics. *) (* ------------------------------------------------------------------------- *) exception Noparse let isspace, isnum = let charcode s = Char.code s.[0] in let spaces = " \t\n\r" and separators = ",;" and brackets = "()[]{}" and symbs = "\\!@#$%^&*-+|\\<=>/?~.:" and alphas = "'abcdefghijklmnopqrstuvwxyz_ABCDEFGHIJKLMNOPQRSTUVWXYZ" and nums = "0123456789" in let allchars = spaces ^ separators ^ brackets ^ symbs ^ alphas ^ nums in let csetsize = List.fold_right (o max charcode) (explode allchars) 256 in let ctable = Array.make csetsize 0 in do_list (fun c -> ctable.(charcode c) <- 1) (explode spaces); do_list (fun c -> ctable.(charcode c) <- 2) (explode separators); do_list (fun c -> ctable.(charcode c) <- 4) (explode brackets); do_list (fun c -> ctable.(charcode c) <- 8) (explode symbs); do_list (fun c -> ctable.(charcode c) <- 16) (explode alphas); do_list (fun c -> ctable.(charcode c) <- 32) (explode nums); let isspace c = ctable.(charcode c) = 1 and isnum c = ctable.(charcode c) = 32 in (isspace, isnum) let parser_or parser1 parser2 input = try parser1 input with Noparse -> parser2 input let ( ++ ) parser1 parser2 input = let result1, rest1 = parser1 input in let result2, rest2 = parser2 rest1 in ((result1, result2), rest2) let rec many prs input = try let result, next = prs input in let results, rest = many prs next in (result :: results, rest) with Noparse -> ([], input) let ( >> ) prs treatment input = let result, rest = prs input in (treatment result, rest) let fix err prs input = try prs input with Noparse -> failwith (err ^ " expected") let listof prs sep err = prs ++ many (sep ++ fix err prs >> snd) >> fun (h, t) -> h :: t let possibly prs input = try let x, rest = prs input in ([x], rest) with Noparse -> ([], input) let some p = function | [] -> raise Noparse | h :: t -> if p h then (h, t) else raise Noparse let a tok = some (fun item -> item = tok) let rec atleast n prs i = ( if n <= 0 then many prs else prs ++ atleast (n - 1) prs >> fun (h, t) -> h :: t ) i (* ------------------------------------------------------------------------- *) let temp_path = Filename.get_temp_dir_name () (* ------------------------------------------------------------------------- *) (* Convenient conversion between files and (lists of) strings. *) (* ------------------------------------------------------------------------- *) let strings_of_file filename = let fd = try open_in filename with Sys_error _ -> failwith ("strings_of_file: can't open " ^ filename) in let rec suck_lines acc = try let l = input_line fd in suck_lines (l :: acc) with End_of_file -> List.rev acc in let data = suck_lines [] in close_in fd; data let string_of_file filename = String.concat "\n" (strings_of_file filename) let file_of_string filename s = let fd = open_out filename in output_string fd s; close_out fd (* ------------------------------------------------------------------------- *) (* Iterative deepening. *) (* ------------------------------------------------------------------------- *) let rec deepen f n = try (*print_string "Searching with depth limit "; print_int n; print_newline();*) f n with Failure _ -> deepen f (n + 1) exception TooDeep let deepen_until limit f n = match compare limit 0 with | 0 -> raise TooDeep | -1 -> deepen f n | _ -> let rec d_until f n = try (* if !debugging then (print_string "Searching with depth limit "; print_int n; print_newline()) ;*) f n with Failure x -> (*if !debugging then (Printf.printf "solver error : %s\n" x) ; *) if n = limit then raise TooDeep else d_until f (n + 1) in d_until f n coq-8.11.0/plugins/micromega/micromega.ml0000644000175000017500000017506113612664533020221 0ustar treinentreinen type __ = Obj.t type unit0 = | Tt (** val negb : bool -> bool **) let negb = function | true -> false | false -> true type nat = | O | S of nat type ('a, 'b) sum = | Inl of 'a | Inr of 'b (** val fst : ('a1 * 'a2) -> 'a1 **) let fst = function | x,_ -> x (** val snd : ('a1 * 'a2) -> 'a2 **) let snd = function | _,y -> y (** val app : 'a1 list -> 'a1 list -> 'a1 list **) let rec app l m = match l with | [] -> m | a::l1 -> a::(app l1 m) type comparison = | Eq | Lt | Gt (** val compOpp : comparison -> comparison **) let compOpp = function | Eq -> Eq | Lt -> Gt | Gt -> Lt module Coq__1 = struct (** val add : nat -> nat -> nat **) let rec add n0 m = match n0 with | O -> m | S p -> S (add p m) end include Coq__1 (** val nth : nat -> 'a1 list -> 'a1 -> 'a1 **) let rec nth n0 l default = match n0 with | O -> (match l with | [] -> default | x::_ -> x) | S m -> (match l with | [] -> default | _::t0 -> nth m t0 default) (** val rev_append : 'a1 list -> 'a1 list -> 'a1 list **) let rec rev_append l l' = match l with | [] -> l' | a::l0 -> rev_append l0 (a::l') (** val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list **) let rec map f = function | [] -> [] | a::t0 -> (f a)::(map f t0) (** val fold_left : ('a1 -> 'a2 -> 'a1) -> 'a2 list -> 'a1 -> 'a1 **) let rec fold_left f l a0 = match l with | [] -> a0 | b::t0 -> fold_left f t0 (f a0 b) (** val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1 **) let rec fold_right f a0 = function | [] -> a0 | b::t0 -> f b (fold_right f a0 t0) type positive = | XI of positive | XO of positive | XH type n = | N0 | Npos of positive type z = | Z0 | Zpos of positive | Zneg of positive module Pos = struct type mask = | IsNul | IsPos of positive | IsNeg end module Coq_Pos = struct (** val succ : positive -> positive **) let rec succ = function | XI p -> XO (succ p) | XO p -> XI p | XH -> XO XH (** val add : positive -> positive -> positive **) let rec add x y = match x with | XI p -> (match y with | XI q0 -> XO (add_carry p q0) | XO q0 -> XI (add p q0) | XH -> XO (succ p)) | XO p -> (match y with | XI q0 -> XI (add p q0) | XO q0 -> XO (add p q0) | XH -> XI p) | XH -> (match y with | XI q0 -> XO (succ q0) | XO q0 -> XI q0 | XH -> XO XH) (** val add_carry : positive -> positive -> positive **) and add_carry x y = match x with | XI p -> (match y with | XI q0 -> XI (add_carry p q0) | XO q0 -> XO (add_carry p q0) | XH -> XI (succ p)) | XO p -> (match y with | XI q0 -> XO (add_carry p q0) | XO q0 -> XI (add p q0) | XH -> XO (succ p)) | XH -> (match y with | XI q0 -> XI (succ q0) | XO q0 -> XO (succ q0) | XH -> XI XH) (** val pred_double : positive -> positive **) let rec pred_double = function | XI p -> XI (XO p) | XO p -> XI (pred_double p) | XH -> XH type mask = Pos.mask = | IsNul | IsPos of positive | IsNeg (** val succ_double_mask : mask -> mask **) let succ_double_mask = function | IsNul -> IsPos XH | IsPos p -> IsPos (XI p) | IsNeg -> IsNeg (** val double_mask : mask -> mask **) let double_mask = function | IsPos p -> IsPos (XO p) | x0 -> x0 (** val double_pred_mask : positive -> mask **) let double_pred_mask = function | XI p -> IsPos (XO (XO p)) | XO p -> IsPos (XO (pred_double p)) | XH -> IsNul (** val sub_mask : positive -> positive -> mask **) let rec sub_mask x y = match x with | XI p -> (match y with | XI q0 -> double_mask (sub_mask p q0) | XO q0 -> succ_double_mask (sub_mask p q0) | XH -> IsPos (XO p)) | XO p -> (match y with | XI q0 -> succ_double_mask (sub_mask_carry p q0) | XO q0 -> double_mask (sub_mask p q0) | XH -> IsPos (pred_double p)) | XH -> (match y with | XH -> IsNul | _ -> IsNeg) (** val sub_mask_carry : positive -> positive -> mask **) and sub_mask_carry x y = match x with | XI p -> (match y with | XI q0 -> succ_double_mask (sub_mask_carry p q0) | XO q0 -> double_mask (sub_mask p q0) | XH -> IsPos (pred_double p)) | XO p -> (match y with | XI q0 -> double_mask (sub_mask_carry p q0) | XO q0 -> succ_double_mask (sub_mask_carry p q0) | XH -> double_pred_mask p) | XH -> IsNeg (** val sub : positive -> positive -> positive **) let sub x y = match sub_mask x y with | IsPos z0 -> z0 | _ -> XH (** val mul : positive -> positive -> positive **) let rec mul x y = match x with | XI p -> add y (XO (mul p y)) | XO p -> XO (mul p y) | XH -> y (** val iter : ('a1 -> 'a1) -> 'a1 -> positive -> 'a1 **) let rec iter f x = function | XI n' -> f (iter f (iter f x n') n') | XO n' -> iter f (iter f x n') n' | XH -> f x (** val size_nat : positive -> nat **) let rec size_nat = function | XI p2 -> S (size_nat p2) | XO p2 -> S (size_nat p2) | XH -> S O (** val compare_cont : comparison -> positive -> positive -> comparison **) let rec compare_cont r x y = match x with | XI p -> (match y with | XI q0 -> compare_cont r p q0 | XO q0 -> compare_cont Gt p q0 | XH -> Gt) | XO p -> (match y with | XI q0 -> compare_cont Lt p q0 | XO q0 -> compare_cont r p q0 | XH -> Gt) | XH -> (match y with | XH -> r | _ -> Lt) (** val compare : positive -> positive -> comparison **) let compare = compare_cont Eq (** val max : positive -> positive -> positive **) let max p p' = match compare p p' with | Gt -> p | _ -> p' (** val leb : positive -> positive -> bool **) let leb x y = match compare x y with | Gt -> false | _ -> true (** val gcdn : nat -> positive -> positive -> positive **) let rec gcdn n0 a b = match n0 with | O -> XH | S n1 -> (match a with | XI a' -> (match b with | XI b' -> (match compare a' b' with | Eq -> a | Lt -> gcdn n1 (sub b' a') a | Gt -> gcdn n1 (sub a' b') b) | XO b0 -> gcdn n1 a b0 | XH -> XH) | XO a0 -> (match b with | XI _ -> gcdn n1 a0 b | XO b0 -> XO (gcdn n1 a0 b0) | XH -> XH) | XH -> XH) (** val gcd : positive -> positive -> positive **) let gcd a b = gcdn (Coq__1.add (size_nat a) (size_nat b)) a b (** val of_succ_nat : nat -> positive **) let rec of_succ_nat = function | O -> XH | S x -> succ (of_succ_nat x) end module N = struct (** val of_nat : nat -> n **) let of_nat = function | O -> N0 | S n' -> Npos (Coq_Pos.of_succ_nat n') end (** val pow_pos : ('a1 -> 'a1 -> 'a1) -> 'a1 -> positive -> 'a1 **) let rec pow_pos rmul x = function | XI i0 -> let p = pow_pos rmul x i0 in rmul x (rmul p p) | XO i0 -> let p = pow_pos rmul x i0 in rmul p p | XH -> x module Z = struct (** val double : z -> z **) let double = function | Z0 -> Z0 | Zpos p -> Zpos (XO p) | Zneg p -> Zneg (XO p) (** val succ_double : z -> z **) let succ_double = function | Z0 -> Zpos XH | Zpos p -> Zpos (XI p) | Zneg p -> Zneg (Coq_Pos.pred_double p) (** val pred_double : z -> z **) let pred_double = function | Z0 -> Zneg XH | Zpos p -> Zpos (Coq_Pos.pred_double p) | Zneg p -> Zneg (XI p) (** val pos_sub : positive -> positive -> z **) let rec pos_sub x y = match x with | XI p -> (match y with | XI q0 -> double (pos_sub p q0) | XO q0 -> succ_double (pos_sub p q0) | XH -> Zpos (XO p)) | XO p -> (match y with | XI q0 -> pred_double (pos_sub p q0) | XO q0 -> double (pos_sub p q0) | XH -> Zpos (Coq_Pos.pred_double p)) | XH -> (match y with | XI q0 -> Zneg (XO q0) | XO q0 -> Zneg (Coq_Pos.pred_double q0) | XH -> Z0) (** val add : z -> z -> z **) let add x y = match x with | Z0 -> y | Zpos x' -> (match y with | Z0 -> x | Zpos y' -> Zpos (Coq_Pos.add x' y') | Zneg y' -> pos_sub x' y') | Zneg x' -> (match y with | Z0 -> x | Zpos y' -> pos_sub y' x' | Zneg y' -> Zneg (Coq_Pos.add x' y')) (** val opp : z -> z **) let opp = function | Z0 -> Z0 | Zpos x0 -> Zneg x0 | Zneg x0 -> Zpos x0 (** val sub : z -> z -> z **) let sub m n0 = add m (opp n0) (** val mul : z -> z -> z **) let mul x y = match x with | Z0 -> Z0 | Zpos x' -> (match y with | Z0 -> Z0 | Zpos y' -> Zpos (Coq_Pos.mul x' y') | Zneg y' -> Zneg (Coq_Pos.mul x' y')) | Zneg x' -> (match y with | Z0 -> Z0 | Zpos y' -> Zneg (Coq_Pos.mul x' y') | Zneg y' -> Zpos (Coq_Pos.mul x' y')) (** val pow_pos : z -> positive -> z **) let pow_pos z0 = Coq_Pos.iter (mul z0) (Zpos XH) (** val pow : z -> z -> z **) let pow x = function | Z0 -> Zpos XH | Zpos p -> pow_pos x p | Zneg _ -> Z0 (** val compare : z -> z -> comparison **) let compare x y = match x with | Z0 -> (match y with | Z0 -> Eq | Zpos _ -> Lt | Zneg _ -> Gt) | Zpos x' -> (match y with | Zpos y' -> Coq_Pos.compare x' y' | _ -> Gt) | Zneg x' -> (match y with | Zneg y' -> compOpp (Coq_Pos.compare x' y') | _ -> Lt) (** val leb : z -> z -> bool **) let leb x y = match compare x y with | Gt -> false | _ -> true (** val ltb : z -> z -> bool **) let ltb x y = match compare x y with | Lt -> true | _ -> false (** val gtb : z -> z -> bool **) let gtb x y = match compare x y with | Gt -> true | _ -> false (** val max : z -> z -> z **) let max n0 m = match compare n0 m with | Lt -> m | _ -> n0 (** val abs : z -> z **) let abs = function | Zneg p -> Zpos p | x -> x (** val to_N : z -> n **) let to_N = function | Zpos p -> Npos p | _ -> N0 (** val of_nat : nat -> z **) let of_nat = function | O -> Z0 | S n1 -> Zpos (Coq_Pos.of_succ_nat n1) (** val of_N : n -> z **) let of_N = function | N0 -> Z0 | Npos p -> Zpos p (** val pos_div_eucl : positive -> z -> z * z **) let rec pos_div_eucl a b = match a with | XI a' -> let q0,r = pos_div_eucl a' b in let r' = add (mul (Zpos (XO XH)) r) (Zpos XH) in if ltb r' b then (mul (Zpos (XO XH)) q0),r' else (add (mul (Zpos (XO XH)) q0) (Zpos XH)),(sub r' b) | XO a' -> let q0,r = pos_div_eucl a' b in let r' = mul (Zpos (XO XH)) r in if ltb r' b then (mul (Zpos (XO XH)) q0),r' else (add (mul (Zpos (XO XH)) q0) (Zpos XH)),(sub r' b) | XH -> if leb (Zpos (XO XH)) b then Z0,(Zpos XH) else (Zpos XH),Z0 (** val div_eucl : z -> z -> z * z **) let div_eucl a b = match a with | Z0 -> Z0,Z0 | Zpos a' -> (match b with | Z0 -> Z0,Z0 | Zpos _ -> pos_div_eucl a' b | Zneg b' -> let q0,r = pos_div_eucl a' (Zpos b') in (match r with | Z0 -> (opp q0),Z0 | _ -> (opp (add q0 (Zpos XH))),(add b r))) | Zneg a' -> (match b with | Z0 -> Z0,Z0 | Zpos _ -> let q0,r = pos_div_eucl a' b in (match r with | Z0 -> (opp q0),Z0 | _ -> (opp (add q0 (Zpos XH))),(sub b r)) | Zneg b' -> let q0,r = pos_div_eucl a' (Zpos b') in q0,(opp r)) (** val div : z -> z -> z **) let div a b = let q0,_ = div_eucl a b in q0 (** val gcd : z -> z -> z **) let gcd a b = match a with | Z0 -> abs b | Zpos a0 -> (match b with | Z0 -> abs a | Zpos b0 -> Zpos (Coq_Pos.gcd a0 b0) | Zneg b0 -> Zpos (Coq_Pos.gcd a0 b0)) | Zneg a0 -> (match b with | Z0 -> abs a | Zpos b0 -> Zpos (Coq_Pos.gcd a0 b0) | Zneg b0 -> Zpos (Coq_Pos.gcd a0 b0)) end (** val zeq_bool : z -> z -> bool **) let zeq_bool x y = match Z.compare x y with | Eq -> true | _ -> false type 'c pExpr = | PEc of 'c | PEX of positive | PEadd of 'c pExpr * 'c pExpr | PEsub of 'c pExpr * 'c pExpr | PEmul of 'c pExpr * 'c pExpr | PEopp of 'c pExpr | PEpow of 'c pExpr * n type 'c pol = | Pc of 'c | Pinj of positive * 'c pol | PX of 'c pol * positive * 'c pol (** val p0 : 'a1 -> 'a1 pol **) let p0 cO = Pc cO (** val p1 : 'a1 -> 'a1 pol **) let p1 cI = Pc cI (** val peq : ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> bool **) let rec peq ceqb p p' = match p with | Pc c -> (match p' with | Pc c' -> ceqb c c' | _ -> false) | Pinj (j, q0) -> (match p' with | Pinj (j', q') -> (match Coq_Pos.compare j j' with | Eq -> peq ceqb q0 q' | _ -> false) | _ -> false) | PX (p2, i, q0) -> (match p' with | PX (p'0, i', q') -> (match Coq_Pos.compare i i' with | Eq -> if peq ceqb p2 p'0 then peq ceqb q0 q' else false | _ -> false) | _ -> false) (** val mkPinj : positive -> 'a1 pol -> 'a1 pol **) let mkPinj j p = match p with | Pc _ -> p | Pinj (j', q0) -> Pinj ((Coq_Pos.add j j'), q0) | PX (_, _, _) -> Pinj (j, p) (** val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol **) let mkPinj_pred j p = match j with | XI j0 -> Pinj ((XO j0), p) | XO j0 -> Pinj ((Coq_Pos.pred_double j0), p) | XH -> p (** val mkPX : 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) let mkPX cO ceqb p i q0 = match p with | Pc c -> if ceqb c cO then mkPinj XH q0 else PX (p, i, q0) | Pinj (_, _) -> PX (p, i, q0) | PX (p', i', q') -> if peq ceqb q' (p0 cO) then PX (p', (Coq_Pos.add i' i), q0) else PX (p, i, q0) (** val mkXi : 'a1 -> 'a1 -> positive -> 'a1 pol **) let mkXi cO cI i = PX ((p1 cI), i, (p0 cO)) (** val mkX : 'a1 -> 'a1 -> 'a1 pol **) let mkX cO cI = mkXi cO cI XH (** val popp : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol **) let rec popp copp = function | Pc c -> Pc (copp c) | Pinj (j, q0) -> Pinj (j, (popp copp q0)) | PX (p2, i, q0) -> PX ((popp copp p2), i, (popp copp q0)) (** val paddC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol **) let rec paddC cadd p c = match p with | Pc c1 -> Pc (cadd c1 c) | Pinj (j, q0) -> Pinj (j, (paddC cadd q0 c)) | PX (p2, i, q0) -> PX (p2, i, (paddC cadd q0 c)) (** val psubC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol **) let rec psubC csub p c = match p with | Pc c1 -> Pc (csub c1 c) | Pinj (j, q0) -> Pinj (j, (psubC csub q0 c)) | PX (p2, i, q0) -> PX (p2, i, (psubC csub q0 c)) (** val paddI : ('a1 -> 'a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) let rec paddI cadd pop q0 j = function | Pc c -> mkPinj j (paddC cadd q0 c) | Pinj (j', q') -> (match Z.pos_sub j' j with | Z0 -> mkPinj j (pop q' q0) | Zpos k -> mkPinj j (pop (Pinj (k, q')) q0) | Zneg k -> mkPinj j' (paddI cadd pop q0 k q')) | PX (p2, i, q') -> (match j with | XI j0 -> PX (p2, i, (paddI cadd pop q0 (XO j0) q')) | XO j0 -> PX (p2, i, (paddI cadd pop q0 (Coq_Pos.pred_double j0) q')) | XH -> PX (p2, i, (pop q' q0))) (** val psubI : ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) let rec psubI cadd copp pop q0 j = function | Pc c -> mkPinj j (paddC cadd (popp copp q0) c) | Pinj (j', q') -> (match Z.pos_sub j' j with | Z0 -> mkPinj j (pop q' q0) | Zpos k -> mkPinj j (pop (Pinj (k, q')) q0) | Zneg k -> mkPinj j' (psubI cadd copp pop q0 k q')) | PX (p2, i, q') -> (match j with | XI j0 -> PX (p2, i, (psubI cadd copp pop q0 (XO j0) q')) | XO j0 -> PX (p2, i, (psubI cadd copp pop q0 (Coq_Pos.pred_double j0) q')) | XH -> PX (p2, i, (pop q' q0))) (** val paddX : 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) let rec paddX cO ceqb pop p' i' p = match p with | Pc _ -> PX (p', i', p) | Pinj (j, q') -> (match j with | XI j0 -> PX (p', i', (Pinj ((XO j0), q'))) | XO j0 -> PX (p', i', (Pinj ((Coq_Pos.pred_double j0), q'))) | XH -> PX (p', i', q')) | PX (p2, i, q') -> (match Z.pos_sub i i' with | Z0 -> mkPX cO ceqb (pop p2 p') i q' | Zpos k -> mkPX cO ceqb (pop (PX (p2, k, (p0 cO))) p') i' q' | Zneg k -> mkPX cO ceqb (paddX cO ceqb pop p' k p2) i q') (** val psubX : 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) let rec psubX cO copp ceqb pop p' i' p = match p with | Pc _ -> PX ((popp copp p'), i', p) | Pinj (j, q') -> (match j with | XI j0 -> PX ((popp copp p'), i', (Pinj ((XO j0), q'))) | XO j0 -> PX ((popp copp p'), i', (Pinj ((Coq_Pos.pred_double j0), q'))) | XH -> PX ((popp copp p'), i', q')) | PX (p2, i, q') -> (match Z.pos_sub i i' with | Z0 -> mkPX cO ceqb (pop p2 p') i q' | Zpos k -> mkPX cO ceqb (pop (PX (p2, k, (p0 cO))) p') i' q' | Zneg k -> mkPX cO ceqb (psubX cO copp ceqb pop p' k p2) i q') (** val padd : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) let rec padd cO cadd ceqb p = function | Pc c' -> paddC cadd p c' | Pinj (j', q') -> paddI cadd (padd cO cadd ceqb) q' j' p | PX (p'0, i', q') -> (match p with | Pc c -> PX (p'0, i', (paddC cadd q' c)) | Pinj (j, q0) -> (match j with | XI j0 -> PX (p'0, i', (padd cO cadd ceqb (Pinj ((XO j0), q0)) q')) | XO j0 -> PX (p'0, i', (padd cO cadd ceqb (Pinj ((Coq_Pos.pred_double j0), q0)) q')) | XH -> PX (p'0, i', (padd cO cadd ceqb q0 q'))) | PX (p2, i, q0) -> (match Z.pos_sub i i' with | Z0 -> mkPX cO ceqb (padd cO cadd ceqb p2 p'0) i (padd cO cadd ceqb q0 q') | Zpos k -> mkPX cO ceqb (padd cO cadd ceqb (PX (p2, k, (p0 cO))) p'0) i' (padd cO cadd ceqb q0 q') | Zneg k -> mkPX cO ceqb (paddX cO ceqb (padd cO cadd ceqb) p'0 k p2) i (padd cO cadd ceqb q0 q'))) (** val psub : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) let rec psub cO cadd csub copp ceqb p = function | Pc c' -> psubC csub p c' | Pinj (j', q') -> psubI cadd copp (psub cO cadd csub copp ceqb) q' j' p | PX (p'0, i', q') -> (match p with | Pc c -> PX ((popp copp p'0), i', (paddC cadd (popp copp q') c)) | Pinj (j, q0) -> (match j with | XI j0 -> PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb (Pinj ((XO j0), q0)) q')) | XO j0 -> PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb (Pinj ((Coq_Pos.pred_double j0), q0)) q')) | XH -> PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb q0 q'))) | PX (p2, i, q0) -> (match Z.pos_sub i i' with | Z0 -> mkPX cO ceqb (psub cO cadd csub copp ceqb p2 p'0) i (psub cO cadd csub copp ceqb q0 q') | Zpos k -> mkPX cO ceqb (psub cO cadd csub copp ceqb (PX (p2, k, (p0 cO))) p'0) i' (psub cO cadd csub copp ceqb q0 q') | Zneg k -> mkPX cO ceqb (psubX cO copp ceqb (psub cO cadd csub copp ceqb) p'0 k p2) i (psub cO cadd csub copp ceqb q0 q'))) (** val pmulC_aux : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 pol **) let rec pmulC_aux cO cmul ceqb p c = match p with | Pc c' -> Pc (cmul c' c) | Pinj (j, q0) -> mkPinj j (pmulC_aux cO cmul ceqb q0 c) | PX (p2, i, q0) -> mkPX cO ceqb (pmulC_aux cO cmul ceqb p2 c) i (pmulC_aux cO cmul ceqb q0 c) (** val pmulC : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 pol **) let pmulC cO cI cmul ceqb p c = if ceqb c cO then p0 cO else if ceqb c cI then p else pmulC_aux cO cmul ceqb p c (** val pmulI : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) let rec pmulI cO cI cmul ceqb pmul0 q0 j = function | Pc c -> mkPinj j (pmulC cO cI cmul ceqb q0 c) | Pinj (j', q') -> (match Z.pos_sub j' j with | Z0 -> mkPinj j (pmul0 q' q0) | Zpos k -> mkPinj j (pmul0 (Pinj (k, q')) q0) | Zneg k -> mkPinj j' (pmulI cO cI cmul ceqb pmul0 q0 k q')) | PX (p', i', q') -> (match j with | XI j' -> mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 j p') i' (pmulI cO cI cmul ceqb pmul0 q0 (XO j') q') | XO j' -> mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 j p') i' (pmulI cO cI cmul ceqb pmul0 q0 (Coq_Pos.pred_double j') q') | XH -> mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 XH p') i' (pmul0 q' q0)) (** val pmul : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) let rec pmul cO cI cadd cmul ceqb p p'' = match p'' with | Pc c -> pmulC cO cI cmul ceqb p c | Pinj (j', q') -> pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' j' p | PX (p', i', q') -> (match p with | Pc c -> pmulC cO cI cmul ceqb p'' c | Pinj (j, q0) -> let qQ' = match j with | XI j0 -> pmul cO cI cadd cmul ceqb (Pinj ((XO j0), q0)) q' | XO j0 -> pmul cO cI cadd cmul ceqb (Pinj ((Coq_Pos.pred_double j0), q0)) q' | XH -> pmul cO cI cadd cmul ceqb q0 q' in mkPX cO ceqb (pmul cO cI cadd cmul ceqb p p') i' qQ' | PX (p2, i, q0) -> let qQ' = pmul cO cI cadd cmul ceqb q0 q' in let pQ' = pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' XH p2 in let qP' = pmul cO cI cadd cmul ceqb (mkPinj XH q0) p' in let pP' = pmul cO cI cadd cmul ceqb p2 p' in padd cO cadd ceqb (mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb pP' i (p0 cO)) qP') i' (p0 cO)) (mkPX cO ceqb pQ' i qQ')) (** val psquare : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol **) let rec psquare cO cI cadd cmul ceqb = function | Pc c -> Pc (cmul c c) | Pinj (j, q0) -> Pinj (j, (psquare cO cI cadd cmul ceqb q0)) | PX (p2, i, q0) -> let twoPQ = pmul cO cI cadd cmul ceqb p2 (mkPinj XH (pmulC cO cI cmul ceqb q0 (cadd cI cI))) in let q2 = psquare cO cI cadd cmul ceqb q0 in let p3 = psquare cO cI cadd cmul ceqb p2 in mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb p3 i (p0 cO)) twoPQ) i q2 (** val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol **) let mk_X cO cI j = mkPinj_pred j (mkX cO cI) (** val ppow_pos : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 pol **) let rec ppow_pos cO cI cadd cmul ceqb subst_l res p = function | XI p3 -> subst_l (pmul cO cI cadd cmul ceqb (ppow_pos cO cI cadd cmul ceqb subst_l (ppow_pos cO cI cadd cmul ceqb subst_l res p p3) p p3) p) | XO p3 -> ppow_pos cO cI cadd cmul ceqb subst_l (ppow_pos cO cI cadd cmul ceqb subst_l res p p3) p p3 | XH -> subst_l (pmul cO cI cadd cmul ceqb res p) (** val ppow_N : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol **) let ppow_N cO cI cadd cmul ceqb subst_l p = function | N0 -> p1 cI | Npos p2 -> ppow_pos cO cI cadd cmul ceqb subst_l (p1 cI) p p2 (** val norm_aux : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **) let rec norm_aux cO cI cadd cmul csub copp ceqb = function | PEc c -> Pc c | PEX j -> mk_X cO cI j | PEadd (pe1, pe2) -> (match pe1 with | PEopp pe3 -> psub cO cadd csub copp ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe2) (norm_aux cO cI cadd cmul csub copp ceqb pe3) | _ -> (match pe2 with | PEopp pe3 -> psub cO cadd csub copp ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) (norm_aux cO cI cadd cmul csub copp ceqb pe3) | _ -> padd cO cadd ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) (norm_aux cO cI cadd cmul csub copp ceqb pe2))) | PEsub (pe1, pe2) -> psub cO cadd csub copp ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) (norm_aux cO cI cadd cmul csub copp ceqb pe2) | PEmul (pe1, pe2) -> pmul cO cI cadd cmul ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) (norm_aux cO cI cadd cmul csub copp ceqb pe2) | PEopp pe1 -> popp copp (norm_aux cO cI cadd cmul csub copp ceqb pe1) | PEpow (pe1, n0) -> ppow_N cO cI cadd cmul ceqb (fun p -> p) (norm_aux cO cI cadd cmul csub copp ceqb pe1) n0 type ('tA, 'tX, 'aA, 'aF) gFormula = | TT | FF | X of 'tX | A of 'tA * 'aA | Cj of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula | D of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula | N of ('tA, 'tX, 'aA, 'aF) gFormula | I of ('tA, 'tX, 'aA, 'aF) gFormula * 'aF option * ('tA, 'tX, 'aA, 'aF) gFormula (** val mapX : ('a2 -> 'a2) -> ('a1, 'a2, 'a3, 'a4) gFormula -> ('a1, 'a2, 'a3, 'a4) gFormula **) let rec mapX f = function | X x -> X (f x) | Cj (f1, f2) -> Cj ((mapX f f1), (mapX f f2)) | D (f1, f2) -> D ((mapX f f1), (mapX f f2)) | N f1 -> N (mapX f f1) | I (f1, o, f2) -> I ((mapX f f1), o, (mapX f f2)) | x -> x (** val foldA : ('a5 -> 'a3 -> 'a5) -> ('a1, 'a2, 'a3, 'a4) gFormula -> 'a5 -> 'a5 **) let rec foldA f f0 acc = match f0 with | A (_, an) -> f acc an | Cj (f1, f2) -> foldA f f1 (foldA f f2 acc) | D (f1, f2) -> foldA f f1 (foldA f f2 acc) | N f1 -> foldA f f1 acc | I (f1, _, f2) -> foldA f f1 (foldA f f2 acc) | _ -> acc (** val cons_id : 'a1 option -> 'a1 list -> 'a1 list **) let cons_id id l = match id with | Some id0 -> id0::l | None -> l (** val ids_of_formula : ('a1, 'a2, 'a3, 'a4) gFormula -> 'a4 list **) let rec ids_of_formula = function | I (_, id, f') -> cons_id id (ids_of_formula f') | _ -> [] (** val collect_annot : ('a1, 'a2, 'a3, 'a4) gFormula -> 'a3 list **) let rec collect_annot = function | A (_, a) -> a::[] | Cj (f1, f2) -> app (collect_annot f1) (collect_annot f2) | D (f1, f2) -> app (collect_annot f1) (collect_annot f2) | N f0 -> collect_annot f0 | I (f1, _, f2) -> app (collect_annot f1) (collect_annot f2) | _ -> [] type 'a bFormula = ('a, __, unit0, unit0) gFormula (** val map_bformula : ('a1 -> 'a2) -> ('a1, 'a3, 'a4, 'a5) gFormula -> ('a2, 'a3, 'a4, 'a5) gFormula **) let rec map_bformula fct = function | TT -> TT | FF -> FF | X p -> X p | A (a, t0) -> A ((fct a), t0) | Cj (f1, f2) -> Cj ((map_bformula fct f1), (map_bformula fct f2)) | D (f1, f2) -> D ((map_bformula fct f1), (map_bformula fct f2)) | N f0 -> N (map_bformula fct f0) | I (f1, a, f2) -> I ((map_bformula fct f1), a, (map_bformula fct f2)) type ('x, 'annot) clause = ('x * 'annot) list type ('x, 'annot) cnf = ('x, 'annot) clause list (** val cnf_tt : ('a1, 'a2) cnf **) let cnf_tt = [] (** val cnf_ff : ('a1, 'a2) cnf **) let cnf_ff = []::[] (** val add_term : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) -> ('a1, 'a2) clause -> ('a1, 'a2) clause option **) let rec add_term unsat deduce t0 = function | [] -> (match deduce (fst t0) (fst t0) with | Some u -> if unsat u then None else Some (t0::[]) | None -> Some (t0::[])) | t'::cl0 -> (match deduce (fst t0) (fst t') with | Some u -> if unsat u then None else (match add_term unsat deduce t0 cl0 with | Some cl' -> Some (t'::cl') | None -> None) | None -> (match add_term unsat deduce t0 cl0 with | Some cl' -> Some (t'::cl') | None -> None)) (** val or_clause : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) clause -> ('a1, 'a2) clause option **) let rec or_clause unsat deduce cl1 cl2 = match cl1 with | [] -> Some cl2 | t0::cl -> (match add_term unsat deduce t0 cl2 with | Some cl' -> or_clause unsat deduce cl cl' | None -> None) (** val xor_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) let xor_clause_cnf unsat deduce t0 f = fold_left (fun acc e -> match or_clause unsat deduce t0 e with | Some cl -> cl::acc | None -> acc) f [] (** val or_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) let or_clause_cnf unsat deduce t0 f = match t0 with | [] -> f | _::_ -> xor_clause_cnf unsat deduce t0 f (** val or_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) let rec or_cnf unsat deduce f f' = match f with | [] -> cnf_tt | e::rst -> rev_append (or_cnf unsat deduce rst f') (or_clause_cnf unsat deduce e f') (** val and_cnf : ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) let and_cnf = rev_append type ('term, 'annot, 'tX, 'aF) tFormula = ('term, 'tX, 'annot, 'aF) gFormula (** val is_cnf_tt : ('a1, 'a2) cnf -> bool **) let is_cnf_tt = function | [] -> true | _::_ -> false (** val is_cnf_ff : ('a1, 'a2) cnf -> bool **) let is_cnf_ff = function | [] -> false | c0::l -> (match c0 with | [] -> (match l with | [] -> true | _::_ -> false) | _::_ -> false) (** val and_cnf_opt : ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) let and_cnf_opt f1 f2 = if if is_cnf_ff f1 then true else is_cnf_ff f2 then cnf_ff else and_cnf f1 f2 (** val or_cnf_opt : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) let or_cnf_opt unsat deduce f1 f2 = if if is_cnf_tt f1 then true else is_cnf_tt f2 then cnf_tt else if is_cnf_ff f2 then f1 else or_cnf unsat deduce f1 f2 (** val xcnf : ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> bool -> ('a1, 'a3, 'a4, 'a5) tFormula -> ('a2, 'a3) cnf **) let rec xcnf unsat deduce normalise1 negate0 pol0 = function | TT -> if pol0 then cnf_tt else cnf_ff | FF -> if pol0 then cnf_ff else cnf_tt | X _ -> cnf_ff | A (x, t0) -> if pol0 then normalise1 x t0 else negate0 x t0 | Cj (e1, e2) -> if pol0 then and_cnf_opt (xcnf unsat deduce normalise1 negate0 pol0 e1) (xcnf unsat deduce normalise1 negate0 pol0 e2) else or_cnf_opt unsat deduce (xcnf unsat deduce normalise1 negate0 pol0 e1) (xcnf unsat deduce normalise1 negate0 pol0 e2) | D (e1, e2) -> if pol0 then or_cnf_opt unsat deduce (xcnf unsat deduce normalise1 negate0 pol0 e1) (xcnf unsat deduce normalise1 negate0 pol0 e2) else and_cnf_opt (xcnf unsat deduce normalise1 negate0 pol0 e1) (xcnf unsat deduce normalise1 negate0 pol0 e2) | N e -> xcnf unsat deduce normalise1 negate0 (negb pol0) e | I (e1, _, e2) -> if pol0 then or_cnf_opt unsat deduce (xcnf unsat deduce normalise1 negate0 (negb pol0) e1) (xcnf unsat deduce normalise1 negate0 pol0 e2) else and_cnf_opt (xcnf unsat deduce normalise1 negate0 (negb pol0) e1) (xcnf unsat deduce normalise1 negate0 pol0 e2) (** val radd_term : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) -> ('a1, 'a2) clause -> (('a1, 'a2) clause, 'a2 list) sum **) let rec radd_term unsat deduce t0 = function | [] -> (match deduce (fst t0) (fst t0) with | Some u -> if unsat u then Inr ((snd t0)::[]) else Inl (t0::[]) | None -> Inl (t0::[])) | t'::cl0 -> (match deduce (fst t0) (fst t') with | Some u -> if unsat u then Inr ((snd t0)::((snd t')::[])) else (match radd_term unsat deduce t0 cl0 with | Inl cl' -> Inl (t'::cl') | Inr l -> Inr l) | None -> (match radd_term unsat deduce t0 cl0 with | Inl cl' -> Inl (t'::cl') | Inr l -> Inr l)) (** val ror_clause : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause -> (('a1, 'a2) clause, 'a2 list) sum **) let rec ror_clause unsat deduce cl1 cl2 = match cl1 with | [] -> Inl cl2 | t0::cl -> (match radd_term unsat deduce t0 cl2 with | Inl cl' -> ror_clause unsat deduce cl cl' | Inr l -> Inr l) (** val xror_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list * 'a2 list **) let xror_clause_cnf unsat deduce t0 f = fold_left (fun pat e -> let acc,tg = pat in (match ror_clause unsat deduce t0 e with | Inl cl -> (cl::acc),tg | Inr l -> acc,(rev_append tg l))) f ([],[]) (** val ror_clause_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list * 'a2 list **) let ror_clause_cnf unsat deduce t0 f = match t0 with | [] -> f,[] | _::_ -> xror_clause_cnf unsat deduce t0 f (** val ror_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list -> ('a1, 'a2) cnf * 'a2 list **) let rec ror_cnf unsat deduce f f' = match f with | [] -> cnf_tt,[] | e::rst -> let rst_f',t0 = ror_cnf unsat deduce rst f' in let e_f',t' = ror_clause_cnf unsat deduce e f' in (rev_append rst_f' e_f'),(rev_append t0 t') (** val ror_cnf_opt : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf * 'a2 list **) let ror_cnf_opt unsat deduce f1 f2 = if is_cnf_tt f1 then cnf_tt,[] else if is_cnf_tt f2 then cnf_tt,[] else if is_cnf_ff f2 then f1,[] else ror_cnf unsat deduce f1 f2 (** val ratom : ('a1, 'a2) cnf -> 'a2 -> ('a1, 'a2) cnf * 'a2 list **) let ratom c a = if if is_cnf_ff c then true else is_cnf_tt c then c,(a::[]) else c,[] (** val rxcnf : ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> bool -> ('a1, 'a3, 'a4, 'a5) tFormula -> ('a2, 'a3) cnf * 'a3 list **) let rec rxcnf unsat deduce normalise1 negate0 polarity = function | TT -> if polarity then cnf_tt,[] else cnf_ff,[] | FF -> if polarity then cnf_ff,[] else cnf_tt,[] | X _ -> cnf_ff,[] | A (x, t0) -> ratom (if polarity then normalise1 x t0 else negate0 x t0) t0 | Cj (e1, e2) -> let e3,t1 = rxcnf unsat deduce normalise1 negate0 polarity e1 in let e4,t2 = rxcnf unsat deduce normalise1 negate0 polarity e2 in if polarity then (and_cnf_opt e3 e4),(rev_append t1 t2) else let f',t' = ror_cnf_opt unsat deduce e3 e4 in f',(rev_append t1 (rev_append t2 t')) | D (e1, e2) -> let e3,t1 = rxcnf unsat deduce normalise1 negate0 polarity e1 in let e4,t2 = rxcnf unsat deduce normalise1 negate0 polarity e2 in if polarity then let f',t' = ror_cnf_opt unsat deduce e3 e4 in f',(rev_append t1 (rev_append t2 t')) else (and_cnf_opt e3 e4),(rev_append t1 t2) | N e -> rxcnf unsat deduce normalise1 negate0 (negb polarity) e | I (e1, _, e2) -> let e3,t1 = rxcnf unsat deduce normalise1 negate0 (negb polarity) e1 in if polarity then if is_cnf_ff e3 then rxcnf unsat deduce normalise1 negate0 polarity e2 else let e4,t2 = rxcnf unsat deduce normalise1 negate0 polarity e2 in let f',t' = ror_cnf_opt unsat deduce e3 e4 in f',(rev_append t1 (rev_append t2 t')) else let e4,t2 = rxcnf unsat deduce normalise1 negate0 polarity e2 in (and_cnf_opt e3 e4),(rev_append t1 t2) type ('term, 'annot, 'tX) to_constrT = { mkTT : 'tX; mkFF : 'tX; mkA : ('term -> 'annot -> 'tX); mkCj : ('tX -> 'tX -> 'tX); mkD : ('tX -> 'tX -> 'tX); mkI : ('tX -> 'tX -> 'tX); mkN : ('tX -> 'tX) } (** val aformula : ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> 'a3 **) let rec aformula to_constr = function | TT -> to_constr.mkTT | FF -> to_constr.mkFF | X p -> p | A (x, t0) -> to_constr.mkA x t0 | Cj (f1, f2) -> to_constr.mkCj (aformula to_constr f1) (aformula to_constr f2) | D (f1, f2) -> to_constr.mkD (aformula to_constr f1) (aformula to_constr f2) | N f0 -> to_constr.mkN (aformula to_constr f0) | I (f1, _, f2) -> to_constr.mkI (aformula to_constr f1) (aformula to_constr f2) (** val is_X : ('a1, 'a2, 'a3, 'a4) tFormula -> 'a3 option **) let is_X = function | X p -> Some p | _ -> None (** val abs_and : ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> (('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula) -> ('a1, 'a3, 'a2, 'a4) gFormula **) let abs_and to_constr f1 f2 c = match is_X f1 with | Some _ -> X (aformula to_constr (c f1 f2)) | None -> (match is_X f2 with | Some _ -> X (aformula to_constr (c f1 f2)) | None -> c f1 f2) (** val abs_or : ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> (('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula) -> ('a1, 'a3, 'a2, 'a4) gFormula **) let abs_or to_constr f1 f2 c = match is_X f1 with | Some _ -> (match is_X f2 with | Some _ -> X (aformula to_constr (c f1 f2)) | None -> c f1 f2) | None -> c f1 f2 (** val mk_arrow : 'a4 option -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula **) let mk_arrow o f1 f2 = match o with | Some _ -> (match is_X f1 with | Some _ -> f2 | None -> I (f1, o, f2)) | None -> I (f1, None, f2) (** val abst_form : ('a1, 'a2, 'a3) to_constrT -> ('a2 -> bool) -> bool -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a3, 'a2, 'a4) gFormula **) let rec abst_form to_constr needA pol0 = function | TT -> if pol0 then TT else X to_constr.mkTT | FF -> if pol0 then X to_constr.mkFF else FF | X p -> X p | A (x, t0) -> if needA t0 then A (x, t0) else X (to_constr.mkA x t0) | Cj (f1, f2) -> let f3 = abst_form to_constr needA pol0 f1 in let f4 = abst_form to_constr needA pol0 f2 in if pol0 then abs_and to_constr f3 f4 (fun x x0 -> Cj (x, x0)) else abs_or to_constr f3 f4 (fun x x0 -> Cj (x, x0)) | D (f1, f2) -> let f3 = abst_form to_constr needA pol0 f1 in let f4 = abst_form to_constr needA pol0 f2 in if pol0 then abs_or to_constr f3 f4 (fun x x0 -> D (x, x0)) else abs_and to_constr f3 f4 (fun x x0 -> D (x, x0)) | N f0 -> let f1 = abst_form to_constr needA (negb pol0) f0 in (match is_X f1 with | Some a -> X (to_constr.mkN a) | None -> N f1) | I (f1, o, f2) -> let f3 = abst_form to_constr needA (negb pol0) f1 in let f4 = abst_form to_constr needA pol0 f2 in if pol0 then abs_or to_constr f3 f4 (mk_arrow o) else abs_and to_constr f3 f4 (mk_arrow o) (** val cnf_checker : (('a1 * 'a2) list -> 'a3 -> bool) -> ('a1, 'a2) cnf -> 'a3 list -> bool **) let rec cnf_checker checker f l = match f with | [] -> true | e::f0 -> (match l with | [] -> false | c::l0 -> if checker e c then cnf_checker checker f0 l0 else false) (** val tauto_checker : ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> (('a2 * 'a3) list -> 'a4 -> bool) -> ('a1, __, 'a3, unit0) gFormula -> 'a4 list -> bool **) let tauto_checker unsat deduce normalise1 negate0 checker f w = cnf_checker checker (xcnf unsat deduce normalise1 negate0 true f) w (** val cneqb : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool **) let cneqb ceqb x y = negb (ceqb x y) (** val cltb : ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool **) let cltb ceqb cleb x y = (&&) (cleb x y) (cneqb ceqb x y) type 'c polC = 'c pol type op1 = | Equal | NonEqual | Strict | NonStrict type 'c nFormula = 'c polC * op1 (** val opMult : op1 -> op1 -> op1 option **) let opMult o o' = match o with | Equal -> Some Equal | NonEqual -> (match o' with | Equal -> Some Equal | NonEqual -> Some NonEqual | _ -> None) | Strict -> (match o' with | NonEqual -> None | _ -> Some o') | NonStrict -> (match o' with | Equal -> Some Equal | NonEqual -> None | _ -> Some NonStrict) (** val opAdd : op1 -> op1 -> op1 option **) let opAdd o o' = match o with | Equal -> Some o' | NonEqual -> (match o' with | Equal -> Some NonEqual | _ -> None) | Strict -> (match o' with | NonEqual -> None | _ -> Some Strict) | NonStrict -> (match o' with | Equal -> Some NonStrict | NonEqual -> None | x -> Some x) type 'c psatz = | PsatzIn of nat | PsatzSquare of 'c polC | PsatzMulC of 'c polC * 'c psatz | PsatzMulE of 'c psatz * 'c psatz | PsatzAdd of 'c psatz * 'c psatz | PsatzC of 'c | PsatzZ (** val map_option : ('a1 -> 'a2 option) -> 'a1 option -> 'a2 option **) let map_option f = function | Some x -> f x | None -> None (** val map_option2 : ('a1 -> 'a2 -> 'a3 option) -> 'a1 option -> 'a2 option -> 'a3 option **) let map_option2 f o o' = match o with | Some x -> (match o' with | Some x' -> f x x' | None -> None) | None -> None (** val pexpr_times_nformula : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option **) let pexpr_times_nformula cO cI cplus ctimes ceqb e = function | ef,o -> (match o with | Equal -> Some ((pmul cO cI cplus ctimes ceqb e ef),Equal) | _ -> None) (** val nformula_times_nformula : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option **) let nformula_times_nformula cO cI cplus ctimes ceqb f1 f2 = let e1,o1 = f1 in let e2,o2 = f2 in map_option (fun x -> Some ((pmul cO cI cplus ctimes ceqb e1 e2),x)) (opMult o1 o2) (** val nformula_plus_nformula : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option **) let nformula_plus_nformula cO cplus ceqb f1 f2 = let e1,o1 = f1 in let e2,o2 = f2 in map_option (fun x -> Some ((padd cO cplus ceqb e1 e2),x)) (opAdd o1 o2) (** val eval_Psatz : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1 nFormula option **) let rec eval_Psatz cO cI cplus ctimes ceqb cleb l = function | PsatzIn n0 -> Some (nth n0 l ((Pc cO),Equal)) | PsatzSquare e0 -> Some ((psquare cO cI cplus ctimes ceqb e0),NonStrict) | PsatzMulC (re, e0) -> map_option (pexpr_times_nformula cO cI cplus ctimes ceqb re) (eval_Psatz cO cI cplus ctimes ceqb cleb l e0) | PsatzMulE (f1, f2) -> map_option2 (nformula_times_nformula cO cI cplus ctimes ceqb) (eval_Psatz cO cI cplus ctimes ceqb cleb l f1) (eval_Psatz cO cI cplus ctimes ceqb cleb l f2) | PsatzAdd (f1, f2) -> map_option2 (nformula_plus_nformula cO cplus ceqb) (eval_Psatz cO cI cplus ctimes ceqb cleb l f1) (eval_Psatz cO cI cplus ctimes ceqb cleb l f2) | PsatzC c -> if cltb ceqb cleb cO c then Some ((Pc c),Strict) else None | PsatzZ -> Some ((Pc cO),Equal) (** val check_inconsistent : 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> bool **) let check_inconsistent cO ceqb cleb = function | e,op -> (match e with | Pc c -> (match op with | Equal -> cneqb ceqb c cO | NonEqual -> ceqb c cO | Strict -> cleb c cO | NonStrict -> cltb ceqb cleb c cO) | _ -> false) (** val check_normalised_formulas : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool **) let check_normalised_formulas cO cI cplus ctimes ceqb cleb l cm = match eval_Psatz cO cI cplus ctimes ceqb cleb l cm with | Some f -> check_inconsistent cO ceqb cleb f | None -> false type op2 = | OpEq | OpNEq | OpLe | OpGe | OpLt | OpGt type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr } (** val norm : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **) let norm = norm_aux (** val psub0 : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) let psub0 = psub (** val padd0 : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) let padd0 = padd (** val popp0 : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol **) let popp0 = popp (** val normalise : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula **) let normalise cO cI cplus ctimes cminus copp ceqb f = let { flhs = lhs; fop = op; frhs = rhs } = f in let lhs0 = norm cO cI cplus ctimes cminus copp ceqb lhs in let rhs0 = norm cO cI cplus ctimes cminus copp ceqb rhs in (match op with | OpEq -> (psub0 cO cplus cminus copp ceqb lhs0 rhs0),Equal | OpNEq -> (psub0 cO cplus cminus copp ceqb lhs0 rhs0),NonEqual | OpLe -> (psub0 cO cplus cminus copp ceqb rhs0 lhs0),NonStrict | OpGe -> (psub0 cO cplus cminus copp ceqb lhs0 rhs0),NonStrict | OpLt -> (psub0 cO cplus cminus copp ceqb rhs0 lhs0),Strict | OpGt -> (psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict) (** val xnormalise : ('a1 -> 'a1) -> 'a1 nFormula -> 'a1 nFormula list **) let xnormalise copp = function | e,o -> (match o with | Equal -> (e,Strict)::(((popp0 copp e),Strict)::[]) | NonEqual -> (e,Equal)::[] | Strict -> ((popp0 copp e),NonStrict)::[] | NonStrict -> ((popp0 copp e),Strict)::[]) (** val xnegate : ('a1 -> 'a1) -> 'a1 nFormula -> 'a1 nFormula list **) let xnegate copp = function | e,o -> (match o with | NonEqual -> (e,Strict)::(((popp0 copp e),Strict)::[]) | x -> (e,x)::[]) (** val cnf_of_list : 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a2 -> ('a1 nFormula, 'a2) cnf **) let cnf_of_list cO ceqb cleb l tg = fold_right (fun x acc -> if check_inconsistent cO ceqb cleb x then acc else ((x,tg)::[])::acc) cnf_tt l (** val cnf_normalise : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula, 'a2) cnf **) let cnf_normalise cO cI cplus ctimes cminus copp ceqb cleb t0 tg = let f = normalise cO cI cplus ctimes cminus copp ceqb t0 in if check_inconsistent cO ceqb cleb f then cnf_ff else cnf_of_list cO ceqb cleb (xnormalise copp f) tg (** val cnf_negate : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula, 'a2) cnf **) let cnf_negate cO cI cplus ctimes cminus copp ceqb cleb t0 tg = let f = normalise cO cI cplus ctimes cminus copp ceqb t0 in if check_inconsistent cO ceqb cleb f then cnf_tt else cnf_of_list cO ceqb cleb (xnegate copp f) tg (** val xdenorm : positive -> 'a1 pol -> 'a1 pExpr **) let rec xdenorm jmp = function | Pc c -> PEc c | Pinj (j, p2) -> xdenorm (Coq_Pos.add j jmp) p2 | PX (p2, j, q0) -> PEadd ((PEmul ((xdenorm jmp p2), (PEpow ((PEX jmp), (Npos j))))), (xdenorm (Coq_Pos.succ jmp) q0)) (** val denorm : 'a1 pol -> 'a1 pExpr **) let denorm p = xdenorm XH p (** val map_PExpr : ('a2 -> 'a1) -> 'a2 pExpr -> 'a1 pExpr **) let rec map_PExpr c_of_S = function | PEc c -> PEc (c_of_S c) | PEX p -> PEX p | PEadd (e1, e2) -> PEadd ((map_PExpr c_of_S e1), (map_PExpr c_of_S e2)) | PEsub (e1, e2) -> PEsub ((map_PExpr c_of_S e1), (map_PExpr c_of_S e2)) | PEmul (e1, e2) -> PEmul ((map_PExpr c_of_S e1), (map_PExpr c_of_S e2)) | PEopp e0 -> PEopp (map_PExpr c_of_S e0) | PEpow (e0, n0) -> PEpow ((map_PExpr c_of_S e0), n0) (** val map_Formula : ('a2 -> 'a1) -> 'a2 formula -> 'a1 formula **) let map_Formula c_of_S f = let { flhs = l; fop = o; frhs = r } = f in { flhs = (map_PExpr c_of_S l); fop = o; frhs = (map_PExpr c_of_S r) } (** val simpl_cone : 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz -> 'a1 psatz **) let simpl_cone cO cI ctimes ceqb e = match e with | PsatzSquare t0 -> (match t0 with | Pc c -> if ceqb cO c then PsatzZ else PsatzC (ctimes c c) | _ -> PsatzSquare t0) | PsatzMulE (t1, t2) -> (match t1 with | PsatzMulE (x, x0) -> (match x with | PsatzC p2 -> (match t2 with | PsatzC c -> PsatzMulE ((PsatzC (ctimes c p2)), x0) | PsatzZ -> PsatzZ | _ -> e) | _ -> (match x0 with | PsatzC p2 -> (match t2 with | PsatzC c -> PsatzMulE ((PsatzC (ctimes c p2)), x) | PsatzZ -> PsatzZ | _ -> e) | _ -> (match t2 with | PsatzC c -> if ceqb cI c then t1 else PsatzMulE (t1, t2) | PsatzZ -> PsatzZ | _ -> e))) | PsatzC c -> (match t2 with | PsatzMulE (x, x0) -> (match x with | PsatzC p2 -> PsatzMulE ((PsatzC (ctimes c p2)), x0) | _ -> (match x0 with | PsatzC p2 -> PsatzMulE ((PsatzC (ctimes c p2)), x) | _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2))) | PsatzAdd (y, z0) -> PsatzAdd ((PsatzMulE ((PsatzC c), y)), (PsatzMulE ((PsatzC c), z0))) | PsatzC c0 -> PsatzC (ctimes c c0) | PsatzZ -> PsatzZ | _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2)) | PsatzZ -> PsatzZ | _ -> (match t2 with | PsatzC c -> if ceqb cI c then t1 else PsatzMulE (t1, t2) | PsatzZ -> PsatzZ | _ -> e)) | PsatzAdd (t1, t2) -> (match t1 with | PsatzZ -> t2 | _ -> (match t2 with | PsatzZ -> t1 | _ -> PsatzAdd (t1, t2))) | _ -> e type q = { qnum : z; qden : positive } (** val qeq_bool : q -> q -> bool **) let qeq_bool x y = zeq_bool (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)) (** val qle_bool : q -> q -> bool **) let qle_bool x y = Z.leb (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)) (** val qplus : q -> q -> q **) let qplus x y = { qnum = (Z.add (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden))); qden = (Coq_Pos.mul x.qden y.qden) } (** val qmult : q -> q -> q **) let qmult x y = { qnum = (Z.mul x.qnum y.qnum); qden = (Coq_Pos.mul x.qden y.qden) } (** val qopp : q -> q **) let qopp x = { qnum = (Z.opp x.qnum); qden = x.qden } (** val qminus : q -> q -> q **) let qminus x y = qplus x (qopp y) (** val qinv : q -> q **) let qinv x = match x.qnum with | Z0 -> { qnum = Z0; qden = XH } | Zpos p -> { qnum = (Zpos x.qden); qden = p } | Zneg p -> { qnum = (Zneg x.qden); qden = p } (** val qpower_positive : q -> positive -> q **) let qpower_positive = pow_pos qmult (** val qpower : q -> z -> q **) let qpower q0 = function | Z0 -> { qnum = (Zpos XH); qden = XH } | Zpos p -> qpower_positive q0 p | Zneg p -> qinv (qpower_positive q0 p) type 'a t = | Empty | Elt of 'a | Branch of 'a t * 'a * 'a t (** val find : 'a1 -> 'a1 t -> positive -> 'a1 **) let rec find default vm p = match vm with | Empty -> default | Elt i -> i | Branch (l, e, r) -> (match p with | XI p2 -> find default r p2 | XO p2 -> find default l p2 | XH -> e) (** val singleton : 'a1 -> positive -> 'a1 -> 'a1 t **) let rec singleton default x v = match x with | XI p -> Branch (Empty, default, (singleton default p v)) | XO p -> Branch ((singleton default p v), default, Empty) | XH -> Elt v (** val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t **) let rec vm_add default x v = function | Empty -> singleton default x v | Elt vl -> (match x with | XI p -> Branch (Empty, vl, (singleton default p v)) | XO p -> Branch ((singleton default p v), vl, Empty) | XH -> Elt v) | Branch (l, o, r) -> (match x with | XI p -> Branch (l, o, (vm_add default p v r)) | XO p -> Branch ((vm_add default p v l), o, r) | XH -> Branch (l, v, r)) (** val zeval_const : z pExpr -> z option **) let rec zeval_const = function | PEc c -> Some c | PEX _ -> None | PEadd (e1, e2) -> map_option2 (fun x y -> Some (Z.add x y)) (zeval_const e1) (zeval_const e2) | PEsub (e1, e2) -> map_option2 (fun x y -> Some (Z.sub x y)) (zeval_const e1) (zeval_const e2) | PEmul (e1, e2) -> map_option2 (fun x y -> Some (Z.mul x y)) (zeval_const e1) (zeval_const e2) | PEopp e0 -> map_option (fun x -> Some (Z.opp x)) (zeval_const e0) | PEpow (e1, n0) -> map_option (fun x -> Some (Z.pow x (Z.of_N n0))) (zeval_const e1) type zWitness = z psatz (** val zWeakChecker : z nFormula list -> z psatz -> bool **) let zWeakChecker = check_normalised_formulas Z0 (Zpos XH) Z.add Z.mul zeq_bool Z.leb (** val psub1 : z pol -> z pol -> z pol **) let psub1 = psub0 Z0 Z.add Z.sub Z.opp zeq_bool (** val padd1 : z pol -> z pol -> z pol **) let padd1 = padd0 Z0 Z.add zeq_bool (** val normZ : z pExpr -> z pol **) let normZ = norm Z0 (Zpos XH) Z.add Z.mul Z.sub Z.opp zeq_bool (** val zunsat : z nFormula -> bool **) let zunsat = check_inconsistent Z0 zeq_bool Z.leb (** val zdeduce : z nFormula -> z nFormula -> z nFormula option **) let zdeduce = nformula_plus_nformula Z0 Z.add zeq_bool (** val xnnormalise : z formula -> z nFormula **) let xnnormalise t0 = let { flhs = lhs; fop = o; frhs = rhs } = t0 in let lhs0 = normZ lhs in let rhs0 = normZ rhs in (match o with | OpEq -> (psub1 rhs0 lhs0),Equal | OpNEq -> (psub1 rhs0 lhs0),NonEqual | OpLe -> (psub1 rhs0 lhs0),NonStrict | OpGe -> (psub1 lhs0 rhs0),NonStrict | OpLt -> (psub1 rhs0 lhs0),Strict | OpGt -> (psub1 lhs0 rhs0),Strict) (** val xnormalise0 : z nFormula -> z nFormula list **) let xnormalise0 = function | e,o -> (match o with | Equal -> ((psub1 e (Pc (Zpos XH))),NonStrict)::(((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) | NonEqual -> (e,Equal)::[] | Strict -> ((psub1 (Pc Z0) e),NonStrict)::[] | NonStrict -> ((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) (** val cnf_of_list0 : 'a1 -> z nFormula list -> (z nFormula * 'a1) list list **) let cnf_of_list0 tg l = fold_right (fun x acc -> if zunsat x then acc else ((x,tg)::[])::acc) cnf_tt l (** val normalise0 : z formula -> 'a1 -> (z nFormula, 'a1) cnf **) let normalise0 t0 tg = let f = xnnormalise t0 in if zunsat f then cnf_ff else cnf_of_list0 tg (xnormalise0 f) (** val xnegate0 : z nFormula -> z nFormula list **) let xnegate0 = function | e,o -> (match o with | NonEqual -> ((psub1 e (Pc (Zpos XH))),NonStrict)::(((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) | Strict -> ((psub1 e (Pc (Zpos XH))),NonStrict)::[] | x -> (e,x)::[]) (** val negate : z formula -> 'a1 -> (z nFormula, 'a1) cnf **) let negate t0 tg = let f = xnnormalise t0 in if zunsat f then cnf_tt else cnf_of_list0 tg (xnegate0 f) (** val cnfZ : (z formula, 'a1, 'a2, 'a3) tFormula -> (z nFormula, 'a1) cnf * 'a1 list **) let cnfZ f = rxcnf zunsat zdeduce normalise0 negate true f (** val ceiling : z -> z -> z **) let ceiling a b = let q0,r = Z.div_eucl a b in (match r with | Z0 -> q0 | _ -> Z.add q0 (Zpos XH)) type zArithProof = | DoneProof | RatProof of zWitness * zArithProof | CutProof of zWitness * zArithProof | EnumProof of zWitness * zWitness * zArithProof list | ExProof of positive * zArithProof (** val zgcdM : z -> z -> z **) let zgcdM x y = Z.max (Z.gcd x y) (Zpos XH) (** val zgcd_pol : z polC -> z * z **) let rec zgcd_pol = function | Pc c -> Z0,c | Pinj (_, p2) -> zgcd_pol p2 | PX (p2, _, q0) -> let g1,c1 = zgcd_pol p2 in let g2,c2 = zgcd_pol q0 in (zgcdM (zgcdM g1 c1) g2),c2 (** val zdiv_pol : z polC -> z -> z polC **) let rec zdiv_pol p x = match p with | Pc c -> Pc (Z.div c x) | Pinj (j, p2) -> Pinj (j, (zdiv_pol p2 x)) | PX (p2, j, q0) -> PX ((zdiv_pol p2 x), j, (zdiv_pol q0 x)) (** val makeCuttingPlane : z polC -> z polC * z **) let makeCuttingPlane p = let g,c = zgcd_pol p in if Z.gtb g Z0 then (zdiv_pol (psubC Z.sub p c) g),(Z.opp (ceiling (Z.opp c) g)) else p,Z0 (** val genCuttingPlane : z nFormula -> ((z polC * z) * op1) option **) let genCuttingPlane = function | e,op -> (match op with | Equal -> let g,c = zgcd_pol e in if (&&) (Z.gtb g Z0) ((&&) (negb (zeq_bool c Z0)) (negb (zeq_bool (Z.gcd g c) g))) then None else Some ((makeCuttingPlane e),Equal) | NonEqual -> Some ((e,Z0),op) | Strict -> Some ((makeCuttingPlane (psubC Z.sub e (Zpos XH))),NonStrict) | NonStrict -> Some ((makeCuttingPlane e),NonStrict)) (** val nformula_of_cutting_plane : ((z polC * z) * op1) -> z nFormula **) let nformula_of_cutting_plane = function | e_z,o -> let e,z0 = e_z in (padd1 e (Pc z0)),o (** val is_pol_Z0 : z polC -> bool **) let is_pol_Z0 = function | Pc z0 -> (match z0 with | Z0 -> true | _ -> false) | _ -> false (** val eval_Psatz0 : z nFormula list -> zWitness -> z nFormula option **) let eval_Psatz0 = eval_Psatz Z0 (Zpos XH) Z.add Z.mul zeq_bool Z.leb (** val valid_cut_sign : op1 -> bool **) let valid_cut_sign = function | Equal -> true | NonStrict -> true | _ -> false (** val bound_var : positive -> z formula **) let bound_var v = { flhs = (PEX v); fop = OpGe; frhs = (PEc Z0) } (** val mk_eq_pos : positive -> positive -> positive -> z formula **) let mk_eq_pos x y t0 = { flhs = (PEX x); fop = OpEq; frhs = (PEsub ((PEX y), (PEX t0))) } (** val max_var : positive -> z pol -> positive **) let rec max_var jmp = function | Pc _ -> jmp | Pinj (j, p2) -> max_var (Coq_Pos.add j jmp) p2 | PX (p2, _, q0) -> Coq_Pos.max (max_var jmp p2) (max_var (Coq_Pos.succ jmp) q0) (** val max_var_nformulae : z nFormula list -> positive **) let max_var_nformulae l = fold_left (fun acc f -> Coq_Pos.max acc (max_var XH (fst f))) l XH (** val zChecker : z nFormula list -> zArithProof -> bool **) let rec zChecker l = function | DoneProof -> false | RatProof (w, pf0) -> (match eval_Psatz0 l w with | Some f -> if zunsat f then true else zChecker (f::l) pf0 | None -> false) | CutProof (w, pf0) -> (match eval_Psatz0 l w with | Some f -> (match genCuttingPlane f with | Some cp -> zChecker ((nformula_of_cutting_plane cp)::l) pf0 | None -> true) | None -> false) | EnumProof (w1, w2, pf0) -> (match eval_Psatz0 l w1 with | Some f1 -> (match eval_Psatz0 l w2 with | Some f2 -> (match genCuttingPlane f1 with | Some p -> let p2,op3 = p in let e1,z1 = p2 in (match genCuttingPlane f2 with | Some p3 -> let p4,op4 = p3 in let e2,z2 = p4 in if (&&) ((&&) (valid_cut_sign op3) (valid_cut_sign op4)) (is_pol_Z0 (padd1 e1 e2)) then let rec label pfs lb ub = match pfs with | [] -> Z.gtb lb ub | pf1::rsr -> (&&) (zChecker (((psub1 e1 (Pc lb)),Equal)::l) pf1) (label rsr (Z.add lb (Zpos XH)) ub) in label pf0 (Z.opp z1) z2 else false | None -> true) | None -> true) | None -> false) | None -> false) | ExProof (x, prf) -> let fr = max_var_nformulae l in if Coq_Pos.leb x fr then let z0 = Coq_Pos.succ fr in let t0 = Coq_Pos.succ z0 in let nfx = xnnormalise (mk_eq_pos x z0 t0) in let posz = xnnormalise (bound_var z0) in let post = xnnormalise (bound_var t0) in zChecker (nfx::(posz::(post::l))) prf else false (** val zTautoChecker : z formula bFormula -> zArithProof list -> bool **) let zTautoChecker f w = tauto_checker zunsat zdeduce normalise0 negate (fun cl -> zChecker (map fst cl)) f w type qWitness = q psatz (** val qWeakChecker : q nFormula list -> q psatz -> bool **) let qWeakChecker = check_normalised_formulas { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus qmult qeq_bool qle_bool (** val qnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) let qnormalise t0 tg = cnf_normalise { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus qmult qminus qopp qeq_bool qle_bool t0 tg (** val qnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) let qnegate t0 tg = cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus qmult qminus qopp qeq_bool qle_bool t0 tg (** val qunsat : q nFormula -> bool **) let qunsat = check_inconsistent { qnum = Z0; qden = XH } qeq_bool qle_bool (** val qdeduce : q nFormula -> q nFormula -> q nFormula option **) let qdeduce = nformula_plus_nformula { qnum = Z0; qden = XH } qplus qeq_bool (** val normQ : q pExpr -> q pol **) let normQ = norm { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus qmult qminus qopp qeq_bool (** val cnfQ : (q formula, 'a1, 'a2, 'a3) tFormula -> (q nFormula, 'a1) cnf * 'a1 list **) let cnfQ f = rxcnf qunsat qdeduce qnormalise qnegate true f (** val qTautoChecker : q formula bFormula -> qWitness list -> bool **) let qTautoChecker f w = tauto_checker qunsat qdeduce qnormalise qnegate (fun cl -> qWeakChecker (map fst cl)) f w type rcst = | C0 | C1 | CQ of q | CZ of z | CPlus of rcst * rcst | CMinus of rcst * rcst | CMult of rcst * rcst | CPow of rcst * (z, nat) sum | CInv of rcst | COpp of rcst (** val z_of_exp : (z, nat) sum -> z **) let z_of_exp = function | Inl z1 -> z1 | Inr n0 -> Z.of_nat n0 (** val q_of_Rcst : rcst -> q **) let rec q_of_Rcst = function | C0 -> { qnum = Z0; qden = XH } | C1 -> { qnum = (Zpos XH); qden = XH } | CQ q0 -> q0 | CZ z0 -> { qnum = z0; qden = XH } | CPlus (r1, r2) -> qplus (q_of_Rcst r1) (q_of_Rcst r2) | CMinus (r1, r2) -> qminus (q_of_Rcst r1) (q_of_Rcst r2) | CMult (r1, r2) -> qmult (q_of_Rcst r1) (q_of_Rcst r2) | CPow (r1, z0) -> qpower (q_of_Rcst r1) (z_of_exp z0) | CInv r0 -> qinv (q_of_Rcst r0) | COpp r0 -> qopp (q_of_Rcst r0) type rWitness = q psatz (** val rWeakChecker : q nFormula list -> q psatz -> bool **) let rWeakChecker = check_normalised_formulas { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus qmult qeq_bool qle_bool (** val rnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) let rnormalise t0 tg = cnf_normalise { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus qmult qminus qopp qeq_bool qle_bool t0 tg (** val rnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) let rnegate t0 tg = cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus qmult qminus qopp qeq_bool qle_bool t0 tg (** val runsat : q nFormula -> bool **) let runsat = check_inconsistent { qnum = Z0; qden = XH } qeq_bool qle_bool (** val rdeduce : q nFormula -> q nFormula -> q nFormula option **) let rdeduce = nformula_plus_nformula { qnum = Z0; qden = XH } qplus qeq_bool (** val rTautoChecker : rcst formula bFormula -> rWitness list -> bool **) let rTautoChecker f w = tauto_checker runsat rdeduce rnormalise rnegate (fun cl -> rWeakChecker (map fst cl)) (map_bformula (map_Formula q_of_Rcst) f) w coq-8.11.0/plugins/micromega/zify.mli0000644000175000017500000000210213612664533017371 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit val print : unit -> unit end module InjTable : S module UnOp : S module BinOp : S module CstOp : S module BinRel : S module PropOp : S module PropUnOp : S module Spec : S module Saturate : S val zify_tac : unit Proofview.tactic val saturate : unit Proofview.tactic val iter_specs : Ltac_plugin.Tacinterp.Value.t -> unit Proofview.tactic coq-8.11.0/plugins/micromega/coq_micromega.mli0000644000175000017500000000320613612664533021223 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit Proofview.tactic*) val psatz_Z : int -> unit Proofview.tactic -> unit Proofview.tactic val psatz_Q : int -> unit Proofview.tactic -> unit Proofview.tactic val psatz_R : int -> unit Proofview.tactic -> unit Proofview.tactic val xlia : unit Proofview.tactic -> unit Proofview.tactic val xnlia : unit Proofview.tactic -> unit Proofview.tactic val nra : unit Proofview.tactic -> unit Proofview.tactic val nqa : unit Proofview.tactic -> unit Proofview.tactic val sos_Z : unit Proofview.tactic -> unit Proofview.tactic val sos_Q : unit Proofview.tactic -> unit Proofview.tactic val sos_R : unit Proofview.tactic -> unit Proofview.tactic val lra_Q : unit Proofview.tactic -> unit Proofview.tactic val lra_R : unit Proofview.tactic -> unit Proofview.tactic (** {5 Use Micromega independently from tactics. } *) val dump_proof_term : Micromega.zArithProof -> EConstr.t (** [dump_proof_term] generates the Coq representation of a Micromega proof witness *) coq-8.11.0/plugins/micromega/MExtraction.v0000644000175000017500000000543413612664533020344 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* "( * )" [ "(,)" ]. Extract Inductive list => list [ "[]" "(::)" ]. Extract Inductive bool => bool [ true false ]. Extract Inductive sumbool => bool [ true false ]. Extract Inductive option => option [ Some None ]. Extract Inductive sumor => option [ Some None ]. (** Then, in a ternary alternative { }+{ }+{ }, - leftmost choice (Inleft Left) is (Some true), - middle choice (Inleft Right) is (Some false), - rightmost choice (Inright) is (None) *) (** To preserve its laziness, andb is normally expanded. Let's rather use the ocaml && *) Extract Inlined Constant andb => "(&&)". Import Reals.Rdefinitions. Extract Constant R => "int". Extract Constant R0 => "0". Extract Constant R1 => "1". Extract Constant Rplus => "( + )". Extract Constant Rmult => "( * )". Extract Constant Ropp => "fun x -> - x". Extract Constant Rinv => "fun x -> 1 / x". (** In order to avoid annoying build dependencies the actual extraction is only performed as a test in the test suite. *) (*Extraction "micromega.ml" Tauto.mapX Tauto.foldA Tauto.collect_annot Tauto.ids_of_formula Tauto.map_bformula Tauto.abst_form ZMicromega.cnfZ ZMicromega.Zeval_const QMicromega.cnfQ List.map simpl_cone (*map_cone indexes*) denorm Qpower vm_add normZ normQ normQ n_of_Z N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find. *) (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/sos_types.mli0000644000175000017500000000257513612664533020456 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* term -> unit type positivstellensatz = | Axiom_eq of int | Axiom_le of int | Axiom_lt of int | Rational_eq of Num.num | Rational_le of Num.num | Rational_lt of Num.num | Square of term | Monoid of int list | Eqmul of term * positivstellensatz | Sum of positivstellensatz * positivstellensatz | Product of positivstellensatz * positivstellensatz val output_psatz : out_channel -> positivstellensatz -> unit coq-8.11.0/plugins/micromega/polynomial.ml0000644000175000017500000010002413612664533020424 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ) = add_num let ( <*> ) = mult_num module Monomial : sig type t val const : t val is_const : t -> bool val var : var -> t val is_var : t -> bool val get_var : t -> var option val prod : t -> t -> t val exp : t -> int -> t val div : t -> t -> t * int val compare : t -> t -> int val pp : out_channel -> t -> unit val fold : (var -> int -> 'a -> 'a) -> t -> 'a -> 'a val sqrt : t -> t option val variables : t -> ISet.t val degree : t -> int end = struct (* A monomial is represented by a multiset of variables *) module Map = Map.Make (Int) open Map type t = int Map.t let degree m = Map.fold (fun _ i d -> i + d) m 0 let is_singleton m = try let k, v = choose m in let l, e, r = split k m in if is_empty l && is_empty r then Some (k, v) else None with Not_found -> None let pp o m = let pp_elt o (k, v) = if v = 1 then Printf.fprintf o "x%i" k else Printf.fprintf o "x%i^%i" k v in let rec pp_list o l = match l with | [] -> () | [e] -> pp_elt o e | e :: l -> Printf.fprintf o "%a*%a" pp_elt e pp_list l in pp_list o (Map.bindings m) (* The monomial that corresponds to a constant *) let const = Map.empty let sum_degree m = Map.fold (fun _ n s -> s + n) m 0 (* Total ordering of monomials *) let compare : t -> t -> int = fun m1 m2 -> let s1 = sum_degree m1 and s2 = sum_degree m2 in if Int.equal s1 s2 then Map.compare Int.compare m1 m2 else Int.compare s1 s2 let is_const m = m = Map.empty (* The monomial 'x' *) let var x = Map.add x 1 Map.empty let is_var m = match is_singleton m with None -> false | Some (_, i) -> i = 1 let get_var m = match is_singleton m with | None -> None | Some (k, i) -> if i = 1 then Some k else None let sqrt m = if is_const m then None else try Some (Map.fold (fun v i acc -> let i' = i / 2 in if i mod 2 = 0 then add v i' acc else raise Not_found) m const) with Not_found -> None (* Get the degre of a variable in a monomial *) let find x m = try find x m with Not_found -> 0 (* Product of monomials *) let prod m1 m2 = Map.fold (fun k d m -> add k (find k m + d) m) m1 m2 let exp m n = let rec exp acc n = if n = 0 then acc else exp (prod acc m) (n - 1) in exp const n (* [div m1 m2 = mr,n] such that mr * (m2)^n = m1 *) let div m1 m2 = let n = fold (fun x i n -> let i' = find x m1 in let nx = i' / i in min n nx) m2 max_int in let mr = fold (fun x i' m -> let i = find x m2 in let ir = i' - (i * n) in if ir = 0 then m else add x ir m) m1 empty in (mr, n) let variables m = fold (fun v i acc -> ISet.add v acc) m ISet.empty let fold = fold end module MonMap = struct include Map.Make (Monomial) let union f = merge (fun x v1 v2 -> match (v1, v2) with | None, None -> None | Some v, None | None, Some v -> Some v | Some v1, Some v2 -> f x v1 v2) end let pp_mon o (m, i) = if Monomial.is_const m then if eq_num (Int 0) i then () else Printf.fprintf o "%s" (string_of_num i) else match i with | Int 1 -> Monomial.pp o m | Int -1 -> Printf.fprintf o "-%a" Monomial.pp m | Int 0 -> () | _ -> Printf.fprintf o "%s*%a" (string_of_num i) Monomial.pp m module Poly : (* A polynomial is a map of monomials *) (* This is probably a naive implementation (expected to be fast enough - Coq is probably the bottleneck) *The new ring contribution is using a sparse Horner representation. *) sig type t val pp : out_channel -> t -> unit val get : Monomial.t -> t -> num val variable : var -> t val add : Monomial.t -> num -> t -> t val constant : num -> t val product : t -> t -> t val addition : t -> t -> t val uminus : t -> t val fold : (Monomial.t -> num -> 'a -> 'a) -> t -> 'a -> 'a val factorise : var -> t -> t * t end = struct (*normalisation bug : 0*x ... *) module P = Map.Make (Monomial) open P type t = num P.t let pp o p = P.iter (fun mn i -> Printf.fprintf o "%a + " pp_mon (mn, i)) p (* Get the coefficient of monomial mn *) let get : Monomial.t -> t -> num = fun mn p -> try find mn p with Not_found -> Int 0 (* The polynomial 1.x *) let variable : var -> t = fun x -> add (Monomial.var x) (Int 1) empty (*The constant polynomial *) let constant : num -> t = fun c -> add Monomial.const c empty (* The addition of a monomial *) let add : Monomial.t -> num -> t -> t = fun mn v p -> if sign_num v = 0 then p else let vl = get mn p <+> v in if sign_num vl = 0 then remove mn p else add mn vl p (** Design choice: empty is not a polynomial I do not remember why .... **) (* The product by a monomial *) let mult : Monomial.t -> num -> t -> t = fun mn v p -> if sign_num v = 0 then constant (Int 0) else fold (fun mn' v' res -> P.add (Monomial.prod mn mn') (v <*> v') res) p empty let addition : t -> t -> t = fun p1 p2 -> fold (fun mn v p -> add mn v p) p1 p2 let product : t -> t -> t = fun p1 p2 -> fold (fun mn v res -> addition (mult mn v p2) res) p1 empty let uminus : t -> t = fun p -> map (fun v -> minus_num v) p let fold = P.fold let factorise x p = let x = Monomial.var x in P.fold (fun m v (px, cx) -> let m1, i = Monomial.div m x in if i = 0 then (px, add m v cx) else let mx = Monomial.prod m1 (Monomial.exp x (i - 1)) in (add mx v px, cx)) p (constant (Int 0), constant (Int 0)) end type vector = Vect.t type cstr = {coeffs : vector; op : op; cst : num} and op = Eq | Ge | Gt exception Strict let is_strict c = c.op = Gt let eval_op = function Eq -> ( =/ ) | Ge -> ( >=/ ) | Gt -> ( >/ ) let string_of_op = function Eq -> "=" | Ge -> ">=" | Gt -> ">" let output_cstr o {coeffs; op; cst} = Printf.fprintf o "%a %s %s" Vect.pp coeffs (string_of_op op) (string_of_num cst) let opMult o1 o2 = match (o1, o2) with Eq, _ | _, Eq -> Eq | Ge, _ | _, Ge -> Ge | Gt, Gt -> Gt let opAdd o1 o2 = match (o1, o2) with Eq, x | x, Eq -> x | Gt, x | x, Gt -> Gt | Ge, Ge -> Ge module LinPoly = struct (** A linear polynomial a0 + a1.x1 + ... + an.xn By convention, the constant a0 is the coefficient of the variable 0. *) type t = Vect.t module MonT = struct module MonoMap = Map.Make (Monomial) module IntMap = Map.Make (Int) (** A hash table might be preferable but requires a hash function. *) let (index_of_monomial : int MonoMap.t ref) = ref MonoMap.empty let (monomial_of_index : Monomial.t IntMap.t ref) = ref IntMap.empty let fresh = ref 0 let reserve vr = if !fresh > vr then failwith (Printf.sprintf "Cannot reserve %i" vr) else fresh := vr + 1 let get_fresh () = !fresh let register m = try MonoMap.find m !index_of_monomial with Not_found -> let res = !fresh in index_of_monomial := MonoMap.add m res !index_of_monomial; monomial_of_index := IntMap.add res m !monomial_of_index; incr fresh; res let retrieve i = IntMap.find i !monomial_of_index let clear () = index_of_monomial := MonoMap.empty; monomial_of_index := IntMap.empty; fresh := 0; ignore (register Monomial.const) let _ = register Monomial.const end let var v = Vect.set (MonT.register (Monomial.var v)) (Int 1) Vect.null let of_monomial m = let v = MonT.register m in Vect.set v (Int 1) Vect.null let linpol_of_pol p = Poly.fold (fun mon num vct -> let vr = MonT.register mon in Vect.set vr num vct) p Vect.null let pol_of_linpol v = Vect.fold (fun p vr n -> Poly.add (MonT.retrieve vr) n p) (Poly.constant (Int 0)) v let coq_poly_of_linpol cst p = let pol_of_mon m = Monomial.fold (fun x v p -> Mc.PEmul (Mc.PEpow (Mc.PEX (CamlToCoq.positive x), CamlToCoq.n v), p)) m (Mc.PEc (cst (Int 1))) in Vect.fold (fun acc x v -> let mn = MonT.retrieve x in Mc.PEadd (Mc.PEmul (Mc.PEc (cst v), pol_of_mon mn), acc)) (Mc.PEc (cst (Int 0))) p let pp_var o vr = try Monomial.pp o (MonT.retrieve vr) (* this is a non-linear variable *) with Not_found -> Printf.fprintf o "v%i" vr let pp o p = Vect.pp_gen pp_var o p let constant c = if sign_num c = 0 then Vect.null else Vect.set 0 c Vect.null let is_linear p = Vect.for_all (fun v _ -> let mn = MonT.retrieve v in Monomial.is_var mn || Monomial.is_const mn) p let is_variable p = let (x, v), r = Vect.decomp_fst p in if Vect.is_null r && v >/ Int 0 then Monomial.get_var (MonT.retrieve x) else None let factorise x p = let px, cx = Poly.factorise x (pol_of_linpol p) in (linpol_of_pol px, linpol_of_pol cx) let is_linear_for x p = let a, b = factorise x p in Vect.is_constant a let search_all_linear p l = Vect.fold (fun acc x v -> if p v then let x' = MonT.retrieve x in match Monomial.get_var x' with | None -> acc | Some x -> if is_linear_for x l then x :: acc else acc else acc) [] l let min_list (l : int list) = match l with [] -> None | e :: l -> Some (List.fold_left min e l) let search_linear p l = min_list (search_all_linear p l) let product p1 p2 = linpol_of_pol (Poly.product (pol_of_linpol p1) (pol_of_linpol p2)) let addition p1 p2 = Vect.add p1 p2 let of_vect v = Vect.fold (fun acc v vl -> addition (product (var v) (constant vl)) acc) Vect.null v let variables p = Vect.fold (fun acc v _ -> ISet.union (Monomial.variables (MonT.retrieve v)) acc) ISet.empty p let monomials p = Vect.fold (fun acc v _ -> ISet.add v acc) ISet.empty p let degree v = Vect.fold (fun acc v vl -> max acc (Monomial.degree (MonT.retrieve v))) 0 v let pp_goal typ o l = let vars = List.fold_left (fun acc p -> ISet.union acc (variables (fst p))) ISet.empty l in let pp_vars o i = ISet.iter (fun v -> Printf.fprintf o "(x%i : %s) " v typ) vars in Printf.fprintf o "forall %a\n" pp_vars vars; List.iteri (fun i (p, op) -> Printf.fprintf o "(H%i : %a %s 0)\n" i pp p (string_of_op op)) l; Printf.fprintf o ", False\n" let collect_square p = Vect.fold (fun acc v _ -> let m = MonT.retrieve v in match Monomial.sqrt m with None -> acc | Some s -> MonMap.add s m acc) MonMap.empty p end module ProofFormat = struct open Big_int type prf_rule = | Annot of string * prf_rule | Hyp of int | Def of int | Cst of Num.num | Zero | Square of Vect.t | MulC of Vect.t * prf_rule | Gcd of Big_int.big_int * prf_rule | MulPrf of prf_rule * prf_rule | AddPrf of prf_rule * prf_rule | CutPrf of prf_rule type proof = | Done | Step of int * prf_rule * proof | Enum of int * prf_rule * Vect.t * prf_rule * proof list | ExProof of int * int * int * var * var * var * proof (* x = z - t, z >= 0, t >= 0 *) let rec output_prf_rule o = function | Annot (s, p) -> Printf.fprintf o "(%a)@%s" output_prf_rule p s | Hyp i -> Printf.fprintf o "Hyp %i" i | Def i -> Printf.fprintf o "Def %i" i | Cst c -> Printf.fprintf o "Cst %s" (string_of_num c) | Zero -> Printf.fprintf o "Zero" | Square s -> Printf.fprintf o "(%a)^2" Poly.pp (LinPoly.pol_of_linpol s) | MulC (p, pr) -> Printf.fprintf o "(%a) * (%a)" Poly.pp (LinPoly.pol_of_linpol p) output_prf_rule pr | MulPrf (p1, p2) -> Printf.fprintf o "(%a) * (%a)" output_prf_rule p1 output_prf_rule p2 | AddPrf (p1, p2) -> Printf.fprintf o "%a + %a" output_prf_rule p1 output_prf_rule p2 | CutPrf p -> Printf.fprintf o "[%a]" output_prf_rule p | Gcd (c, p) -> Printf.fprintf o "(%a)/%s" output_prf_rule p (string_of_big_int c) let rec output_proof o = function | Done -> Printf.fprintf o "." | Step (i, p, pf) -> Printf.fprintf o "%i:= %a ; %a" i output_prf_rule p output_proof pf | Enum (i, p1, v, p2, pl) -> Printf.fprintf o "%i{%a<=%a<=%a}%a" i output_prf_rule p1 Vect.pp v output_prf_rule p2 (pp_list ";" output_proof) pl | ExProof (i, j, k, x, z, t, pr) -> Printf.fprintf o "%i := %i = %i - %i ; %i := %i >= 0 ; %i := %i >= 0 ; %a" i x z t j z k t output_proof pr let rec pr_size = function | Annot (_, p) -> pr_size p | Zero | Square _ -> Int 0 | Hyp _ -> Int 1 | Def _ -> Int 1 | Cst n -> n | Gcd (i, p) -> pr_size p // Big_int i | MulPrf (p1, p2) | AddPrf (p1, p2) -> pr_size p1 +/ pr_size p2 | CutPrf p -> pr_size p | MulC (v, p) -> pr_size p let rec pr_rule_max_id = function | Annot (_, p) -> pr_rule_max_id p | Hyp i | Def i -> i | Cst _ | Zero | Square _ -> -1 | MulC (_, p) | CutPrf p | Gcd (_, p) -> pr_rule_max_id p | MulPrf (p1, p2) | AddPrf (p1, p2) -> max (pr_rule_max_id p1) (pr_rule_max_id p2) let rec proof_max_id = function | Done -> -1 | Step (i, pr, prf) -> max i (max (pr_rule_max_id pr) (proof_max_id prf)) | Enum (i, p1, _, p2, l) -> let m = max (pr_rule_max_id p1) (pr_rule_max_id p2) in List.fold_left (fun i prf -> max i (proof_max_id prf)) (max i m) l | ExProof (i, j, k, _, _, _, prf) -> max (max (max i j) k) (proof_max_id prf) let rec pr_rule_def_cut id = function | Annot (_, p) -> pr_rule_def_cut id p | MulC (p, prf) -> let bds, id', prf' = pr_rule_def_cut id prf in (bds, id', MulC (p, prf')) | MulPrf (p1, p2) -> let bds1, id, p1 = pr_rule_def_cut id p1 in let bds2, id, p2 = pr_rule_def_cut id p2 in (bds2 @ bds1, id, MulPrf (p1, p2)) | AddPrf (p1, p2) -> let bds1, id, p1 = pr_rule_def_cut id p1 in let bds2, id, p2 = pr_rule_def_cut id p2 in (bds2 @ bds1, id, AddPrf (p1, p2)) | CutPrf p -> let bds, id, p = pr_rule_def_cut id p in ((id, p) :: bds, id + 1, Def id) | Gcd (c, p) -> let bds, id, p = pr_rule_def_cut id p in ((id, p) :: bds, id + 1, Def id) | (Square _ | Cst _ | Def _ | Hyp _ | Zero) as x -> ([], id, x) (* Do not define top-level cuts *) let pr_rule_def_cut id = function | CutPrf p -> let bds, ids, p' = pr_rule_def_cut id p in (bds, ids, CutPrf p') | p -> pr_rule_def_cut id p let rec implicit_cut p = match p with CutPrf p -> implicit_cut p | _ -> p let rec pr_rule_collect_hyps pr = match pr with | Annot (_, pr) -> pr_rule_collect_hyps pr | Hyp i | Def i -> ISet.add i ISet.empty | Cst _ | Zero | Square _ -> ISet.empty | MulC (_, pr) | Gcd (_, pr) | CutPrf pr -> pr_rule_collect_hyps pr | MulPrf (p1, p2) | AddPrf (p1, p2) -> ISet.union (pr_rule_collect_hyps p1) (pr_rule_collect_hyps p2) let simplify_proof p = let rec simplify_proof p = match p with | Done -> (Done, ISet.empty) | Step (i, pr, Done) -> (p, ISet.add i (pr_rule_collect_hyps pr)) | Step (i, pr, prf) -> let prf', hyps = simplify_proof prf in if not (ISet.mem i hyps) then (prf', hyps) else ( Step (i, pr, prf') , ISet.add i (ISet.union (pr_rule_collect_hyps pr) hyps) ) | Enum (i, p1, v, p2, pl) -> let pl, hl = List.split (List.map simplify_proof pl) in let hyps = List.fold_left ISet.union ISet.empty hl in ( Enum (i, p1, v, p2, pl) , ISet.add i (ISet.union (ISet.union (pr_rule_collect_hyps p1) (pr_rule_collect_hyps p2)) hyps) ) | ExProof (i, j, k, x, z, t, prf) -> let prf', hyps = simplify_proof prf in if (not (ISet.mem i hyps)) && (not (ISet.mem j hyps)) && not (ISet.mem k hyps) then (prf', hyps) else ( ExProof (i, j, k, x, z, t, prf') , ISet.add i (ISet.add j (ISet.add k hyps)) ) in fst (simplify_proof p) let rec normalise_proof id prf = match prf with | Done -> (id, Done) | Step (i, Gcd (c, p), Done) -> normalise_proof id (Step (i, p, Done)) | Step (i, p, prf) -> let bds, id, p' = pr_rule_def_cut id p in let id, prf = normalise_proof id prf in let prf = List.fold_left (fun acc (i, p) -> Step (i, CutPrf p, acc)) (Step (i, p', prf)) bds in (id, prf) | ExProof (i, j, k, x, z, t, prf) -> let id, prf = normalise_proof id prf in (id, ExProof (i, j, k, x, z, t, prf)) | Enum (i, p1, v, p2, pl) -> (* Why do I have top-level cuts ? *) (* let p1 = implicit_cut p1 in let p2 = implicit_cut p2 in let (ids,prfs) = List.split (List.map (normalise_proof id) pl) in (List.fold_left max 0 ids , Enum(i,p1,v,p2,prfs)) *) let bds1, id, p1' = pr_rule_def_cut id (implicit_cut p1) in let bds2, id, p2' = pr_rule_def_cut id (implicit_cut p2) in let ids, prfs = List.split (List.map (normalise_proof id) pl) in ( List.fold_left max 0 ids , List.fold_left (fun acc (i, p) -> Step (i, CutPrf p, acc)) (Enum (i, p1', v, p2', prfs)) (bds2 @ bds1) ) let normalise_proof id prf = let prf = simplify_proof prf in let res = normalise_proof id prf in if debug then Printf.printf "normalise_proof %a -> %a" output_proof prf output_proof (snd res); res module OrdPrfRule = struct type t = prf_rule let id_of_constr = function | Annot _ -> 0 | Hyp _ -> 1 | Def _ -> 2 | Cst _ -> 3 | Zero -> 4 | Square _ -> 5 | MulC _ -> 6 | Gcd _ -> 7 | MulPrf _ -> 8 | AddPrf _ -> 9 | CutPrf _ -> 10 let cmp_pair c1 c2 (x1, x2) (y1, y2) = match c1 x1 y1 with 0 -> c2 x2 y2 | i -> i let rec compare p1 p2 = match (p1, p2) with | Annot (s1, p1), Annot (s2, p2) -> if s1 = s2 then compare p1 p2 else String.compare s1 s2 | Hyp i, Hyp j -> Int.compare i j | Def i, Def j -> Int.compare i j | Cst n, Cst m -> Num.compare_num n m | Zero, Zero -> 0 | Square v1, Square v2 -> Vect.compare v1 v2 | MulC (v1, p1), MulC (v2, p2) -> cmp_pair Vect.compare compare (v1, p1) (v2, p2) | Gcd (b1, p1), Gcd (b2, p2) -> cmp_pair Big_int.compare_big_int compare (b1, p1) (b2, p2) | MulPrf (p1, q1), MulPrf (p2, q2) -> cmp_pair compare compare (p1, q1) (p2, q2) | AddPrf (p1, q1), MulPrf (p2, q2) -> cmp_pair compare compare (p1, q1) (p2, q2) | CutPrf p, CutPrf p' -> compare p p' | _, _ -> Int.compare (id_of_constr p1) (id_of_constr p2) end let add_proof x y = match (x, y) with Zero, p | p, Zero -> p | _ -> AddPrf (x, y) let rec mul_cst_proof c p = match p with | Annot (s, p) -> Annot (s, mul_cst_proof c p) | MulC (v, p') -> MulC (Vect.mul c v, p') | _ -> ( match sign_num c with | 0 -> Zero (* This is likely to be a bug *) | -1 -> MulC (LinPoly.constant c, p) (* [p] should represent an equality *) | 1 -> if eq_num (Int 1) c then p else MulPrf (Cst c, p) | _ -> assert false ) let sMulC v p = let c, v' = Vect.decomp_cst v in if Vect.is_null v' then mul_cst_proof c p else MulC (v, p) let mul_proof p1 p2 = match (p1, p2) with | Zero, _ | _, Zero -> Zero | Cst c, p | p, Cst c -> mul_cst_proof c p | _, _ -> MulPrf (p1, p2) module PrfRuleMap = Map.Make (OrdPrfRule) let prf_rule_of_map m = PrfRuleMap.fold (fun k v acc -> add_proof (sMulC v k) acc) m Zero let rec dev_prf_rule p = match p with | Annot (s, p) -> dev_prf_rule p | Hyp _ | Def _ | Cst _ | Zero | Square _ -> PrfRuleMap.singleton p (LinPoly.constant (Int 1)) | MulC (v, p) -> PrfRuleMap.map (fun v1 -> LinPoly.product v v1) (dev_prf_rule p) | AddPrf (p1, p2) -> PrfRuleMap.merge (fun k o1 o2 -> match (o1, o2) with | None, None -> None | None, Some v | Some v, None -> Some v | Some v1, Some v2 -> Some (LinPoly.addition v1 v2)) (dev_prf_rule p1) (dev_prf_rule p2) | MulPrf (p1, p2) -> ( let p1' = dev_prf_rule p1 in let p2' = dev_prf_rule p2 in let p1'' = prf_rule_of_map p1' in let p2'' = prf_rule_of_map p2' in match p1'' with | Cst c -> PrfRuleMap.map (fun v1 -> Vect.mul c v1) p2' | _ -> PrfRuleMap.singleton (MulPrf (p1'', p2'')) (LinPoly.constant (Int 1)) ) | _ -> PrfRuleMap.singleton p (LinPoly.constant (Int 1)) let simplify_prf_rule p = prf_rule_of_map (dev_prf_rule p) (* let mul_proof p1 p2 = let res = mul_proof p1 p2 in Printf.printf "mul_proof %a %a = %a\n" output_prf_rule p1 output_prf_rule p2 output_prf_rule res; res let add_proof p1 p2 = let res = add_proof p1 p2 in Printf.printf "add_proof %a %a = %a\n" output_prf_rule p1 output_prf_rule p2 output_prf_rule res; res let sMulC v p = let res = sMulC v p in Printf.printf "sMulC %a %a = %a\n" Vect.pp v output_prf_rule p output_prf_rule res ; res let mul_cst_proof c p = let res = mul_cst_proof c p in Printf.printf "mul_cst_proof %s %a = %a\n" (Num.string_of_num c) output_prf_rule p output_prf_rule res ; res *) let proof_of_farkas env vect = Vect.fold (fun prf x n -> add_proof (mul_cst_proof n (IMap.find x env)) prf) Zero vect module Env = struct let rec string_of_int_list l = match l with | [] -> "" | i :: l -> Printf.sprintf "%i,%s" i (string_of_int_list l) let id_of_hyp hyp l = let rec xid_of_hyp i l' = match l' with | [] -> failwith (Printf.sprintf "id_of_hyp %i %s" hyp (string_of_int_list l)) | hyp' :: l' -> if hyp = hyp' then i else xid_of_hyp (i + 1) l' in xid_of_hyp 0 l end let cmpl_prf_rule norm (cst : num -> 'a) env prf = let rec cmpl = function | Annot (s, p) -> cmpl p | Hyp i | Def i -> Mc.PsatzIn (CamlToCoq.nat (Env.id_of_hyp i env)) | Cst i -> Mc.PsatzC (cst i) | Zero -> Mc.PsatzZ | MulPrf (p1, p2) -> Mc.PsatzMulE (cmpl p1, cmpl p2) | AddPrf (p1, p2) -> Mc.PsatzAdd (cmpl p1, cmpl p2) | MulC (lp, p) -> let lp = norm (LinPoly.coq_poly_of_linpol cst lp) in Mc.PsatzMulC (lp, cmpl p) | Square lp -> Mc.PsatzSquare (norm (LinPoly.coq_poly_of_linpol cst lp)) | _ -> failwith "Cuts should already be compiled" in cmpl prf let cmpl_prf_rule_z env r = cmpl_prf_rule Mc.normZ (fun x -> CamlToCoq.bigint (numerator x)) env r let rec cmpl_proof env = function | Done -> Mc.DoneProof | Step (i, p, prf) -> ( match p with | CutPrf p' -> Mc.CutProof (cmpl_prf_rule_z env p', cmpl_proof (i :: env) prf) | _ -> Mc.RatProof (cmpl_prf_rule_z env p, cmpl_proof (i :: env) prf) ) | Enum (i, p1, _, p2, l) -> Mc.EnumProof ( cmpl_prf_rule_z env p1 , cmpl_prf_rule_z env p2 , List.map (cmpl_proof (i :: env)) l ) | ExProof (i, j, k, x, _, _, prf) -> Mc.ExProof (CamlToCoq.positive x, cmpl_proof (i :: j :: k :: env) prf) let compile_proof env prf = let id = 1 + proof_max_id prf in let _, prf = normalise_proof id prf in cmpl_proof env prf let rec eval_prf_rule env = function | Annot (s, p) -> eval_prf_rule env p | Hyp i | Def i -> env i | Cst n -> ( ( Vect.set 0 n Vect.null , match Num.compare_num n (Int 0) with | 0 -> Ge | 1 -> Gt | _ -> failwith "eval_prf_rule : negative constant" ) ) | Zero -> (Vect.null, Ge) | Square v -> (LinPoly.product v v, Ge) | MulC (v, p) -> ( let p1, o = eval_prf_rule env p in match o with | Eq -> (LinPoly.product v p1, Eq) | _ -> Printf.fprintf stdout "MulC(%a,%a) invalid 2d arg %a %s" Vect.pp v output_prf_rule p Vect.pp p1 (string_of_op o); failwith "eval_prf_rule : not an equality" ) | Gcd (g, p) -> let v, op = eval_prf_rule env p in (Vect.div (Big_int g) v, op) | MulPrf (p1, p2) -> let v1, o1 = eval_prf_rule env p1 in let v2, o2 = eval_prf_rule env p2 in (LinPoly.product v1 v2, opMult o1 o2) | AddPrf (p1, p2) -> let v1, o1 = eval_prf_rule env p1 in let v2, o2 = eval_prf_rule env p2 in (LinPoly.addition v1 v2, opAdd o1 o2) | CutPrf p -> eval_prf_rule env p let is_unsat (p, o) = let c, r = Vect.decomp_cst p in if Vect.is_null r then not (eval_op o c (Int 0)) else false let rec eval_proof env p = match p with | Done -> failwith "Proof is not finished" | Step (i, prf, rst) -> let p, o = eval_prf_rule (fun i -> IMap.find i env) prf in if is_unsat (p, o) then true else if rst = Done then begin Printf.fprintf stdout "Last inference %a %s\n" LinPoly.pp p (string_of_op o); false end else eval_proof (IMap.add i (p, o) env) rst | Enum (i, r1, v, r2, l) -> let _ = eval_prf_rule (fun i -> IMap.find i env) r1 in let _ = eval_prf_rule (fun i -> IMap.find i env) r2 in (* Should check bounds *) failwith "Not implemented" | ExProof _ -> failwith "Not implemented" end module WithProof = struct type t = (LinPoly.t * op) * ProofFormat.prf_rule let annot s (p, prf) = (p, ProofFormat.Annot (s, prf)) let output o ((lp, op), prf) = Printf.fprintf o "%a %s 0 by %a\n" LinPoly.pp lp (string_of_op op) ProofFormat.output_prf_rule prf let output_sys o l = List.iter (Printf.fprintf o "%a\n" output) l exception InvalidProof let zero = ((Vect.null, Eq), ProofFormat.Zero) let const n = ((LinPoly.constant n, Ge), ProofFormat.Cst n) let of_cstr (c, prf) = ((Vect.set 0 (Num.minus_num c.cst) c.coeffs, c.op), prf) let product : t -> t -> t = fun ((p1, o1), prf1) ((p2, o2), prf2) -> ((LinPoly.product p1 p2, opMult o1 o2), ProofFormat.mul_proof prf1 prf2) let addition : t -> t -> t = fun ((p1, o1), prf1) ((p2, o2), prf2) -> ((Vect.add p1 p2, opAdd o1 o2), ProofFormat.add_proof prf1 prf2) let mult p ((p1, o1), prf1) = match o1 with | Eq -> ((LinPoly.product p p1, o1), ProofFormat.sMulC p prf1) | Gt | Ge -> let n, r = Vect.decomp_cst p in if Vect.is_null r && n >/ Int 0 then ((LinPoly.product p p1, o1), ProofFormat.mul_cst_proof n prf1) else ( Printf.printf "mult_error %a [*] %a\n" LinPoly.pp p output ((p1, o1), prf1); raise InvalidProof ) let cutting_plane ((p, o), prf) = let c, p' = Vect.decomp_cst p in let g = Vect.gcd p' in if Big_int.eq_big_int Big_int.unit_big_int g || c =/ Int 0 || not (Big_int.eq_big_int (denominator c) Big_int.unit_big_int) then None (* Nothing to do *) else let c1 = c // Big_int g in let c1' = Num.floor_num c1 in if c1 =/ c1' then None else match o with | Eq -> Some ((Vect.set 0 (Int (-1)) Vect.null, Eq), ProofFormat.Gcd (g, prf)) | Gt -> failwith "cutting_plane ignore strict constraints" | Ge -> (* This is a non-trivial common divisor *) Some ( (Vect.set 0 c1' (Vect.div (Big_int g) p), o) , ProofFormat.Gcd (g, prf) ) let construct_sign p = let c, p' = Vect.decomp_cst p in if Vect.is_null p' then Some ( match sign_num c with | 0 -> (true, Eq, ProofFormat.Zero) | 1 -> (true, Gt, ProofFormat.Cst c) | _ (*-1*) -> (false, Gt, ProofFormat.Cst (minus_num c)) ) else None let get_sign l p = match construct_sign p with | None -> ( try let (p', o), prf = List.find (fun ((p', o), prf) -> Vect.equal p p') l in Some (true, o, prf) with Not_found -> ( let p = Vect.uminus p in try let (p', o), prf = List.find (fun ((p', o), prf) -> Vect.equal p p') l in Some (false, o, prf) with Not_found -> None ) ) | Some s -> Some s let mult_sign : bool -> t -> t = fun b ((p, o), prf) -> if b then ((p, o), prf) else ((Vect.uminus p, o), prf) let rec linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) = (* lp1 = a1.x + b1 *) let a1, b1 = LinPoly.factorise x lp1 in (* lp2 = a2.x + b2 *) let a2, b2 = LinPoly.factorise x lp2 in if Vect.is_null a2 then (* We are done *) Some ((lp2, op2), prf2) else match (op1, op2) with | Eq, (Ge | Gt) -> ( match get_sign sys a1 with | None -> None (* Impossible to pivot without sign information *) | Some (b, o, prf) -> let sa1 = mult_sign b ((a1, o), prf) in let sa2 = if b then Vect.uminus a2 else a2 in let (lp2, op2), prf2 = addition (product sa1 ((lp2, op2), prf2)) (mult sa2 ((lp1, op1), prf1)) in linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) ) | Eq, Eq -> let (lp2, op2), prf2 = addition (mult a1 ((lp2, op2), prf2)) (mult (Vect.uminus a2) ((lp1, op1), prf1)) in linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) | (Ge | Gt), (Ge | Gt) -> ( match (get_sign sys a1, get_sign sys a2) with | Some (b1, o1, p1), Some (b2, o2, p2) -> if b1 <> b2 then let (lp2, op2), prf2 = addition (product (mult_sign b1 ((a1, o1), p1)) ((lp2, op2), prf2)) (product (mult_sign b2 ((a2, o2), p2)) ((lp1, op1), prf1)) in linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) else None | _ -> None ) | (Ge | Gt), Eq -> failwith "pivot: equality as second argument" let linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) = match linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) with | None -> None | Some (c, p) -> Some (c, ProofFormat.simplify_prf_rule p) let is_substitution strict ((p, o), prf) = let pred v = if strict then v =/ Int 1 || v =/ Int (-1) else true in match o with Eq -> LinPoly.search_linear pred p | _ -> None let subst1 sys0 = let oeq, sys' = extract (is_substitution true) sys0 in match oeq with | None -> sys0 | Some (v, pc) -> ( match simplify (linear_pivot sys0 pc v) sys' with | None -> sys0 | Some sys' -> sys' ) let subst sys0 = let elim sys = let oeq, sys' = extract (is_substitution true) sys in match oeq with | None -> None | Some (v, pc) -> simplify (linear_pivot sys0 pc v) sys' in iterate_until_stable elim sys0 let saturate_subst b sys0 = let select = is_substitution b in let gen (v, pc) ((c, op), prf) = if ISet.mem v (LinPoly.variables c) then linear_pivot sys0 pc v ((c, op), prf) else None in saturate select gen sys0 end (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/simplex.ml0000644000175000017500000005363013612664533017734 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* = than a threshold rst are restricted.*) module Restricted = struct type t = { base : int (** All variables above [base] are restricted *) ; exc : int option (** Except [exc] which is currently optimised *) } let pp o {base; exc} = Printf.fprintf o ">= %a " LinPoly.pp_var base; match exc with | None -> Printf.fprintf o "-" | Some x -> Printf.fprintf o "-%a" LinPoly.pp_var base let is_exception (x : var) (r : t) = match r.exc with None -> false | Some x' -> x = x' let restrict x rst = if is_exception x rst then {base = rst.base; exc = None} else failwith (Printf.sprintf "Cannot restrict %i" x) let is_restricted x r0 = x >= r0.base && not (is_exception x r0) let make x = {base = x; exc = None} let set_exc x rst = {base = rst.base; exc = Some x} let fold rst f m acc = IMap.fold (fun k v acc -> if is_exception k rst then acc else f k v acc) (IMap.from rst.base m) acc end let pp_row o v = LinPoly.pp o v let output_tableau o t = IMap.iter (fun k v -> Printf.fprintf o "%a = %a\n" LinPoly.pp_var k pp_row v) t let output_env o t = IMap.iter (fun k v -> Printf.fprintf o "%a : %a\n" LinPoly.pp_var k WithProof.output v) t let output_vars o m = IMap.iter (fun k _ -> Printf.fprintf o "%a " LinPoly.pp_var k) m (** A tableau is feasible iff for every basic restricted variable xi, we have ci>=0. When all the non-basic variables are set to 0, the value of a basic variable xi is necessarily ci. If xi is restricted, it is feasible if ci>=0. *) let unfeasible (rst : Restricted.t) tbl = Restricted.fold rst (fun k v m -> if Vect.get_cst v >=/ Int 0 then m else IMap.add k () m) tbl IMap.empty let is_feasible rst tb = IMap.is_empty (unfeasible rst tb) (** Let a1.x1+...+an.xn be a vector of non-basic variables. It is maximised if all the xi are restricted and the ai are negative. If xi>= 0 (restricted) and ai is negative, the maximum for ai.xi is obtained for xi = 0 Otherwise, it is possible to make ai.xi arbitrarily big: - if xi is not restricted, take +/- oo depending on the sign of ai - if ai is positive, take +oo *) let is_maximised_vect rst v = Vect.for_all (fun xi ai -> if ai >/ Int 0 then false else Restricted.is_restricted xi rst) v (** [is_maximised rst v] @return None if the variable is not maximised @return Some v where v is the maximal value *) let is_maximised rst v = try let vl, v = Vect.decomp_cst v in if is_maximised_vect rst v then Some vl else None with Not_found -> None (** A variable xi is unbounded if for every equation xj= ...ai.xi ... if ai < 0 then xj is not restricted. As a result, even if we increase the value of xi, it is always possible to adjust the value of xj without violating a restriction. *) type result = | Max of num (** Maximum is reached *) | Ubnd of var (** Problem is unbounded *) | Feas (** Problem is feasible *) type pivot = Done of result | Pivot of int * int * num type simplex = Opt of tableau * result (** For a row, x = ao.xo+...+ai.xi a valid pivot variable is such that it can improve the value of xi. it is the case, if xi is unrestricted (increase if ai> 0, decrease if ai < 0) xi is restricted but ai > 0 This is the entering variable. *) let rec find_pivot_column (rst : Restricted.t) (r : Vect.t) = match Vect.choose r with | None -> failwith "find_pivot_column" | Some (xi, ai, r') -> if ai Some (i1, sc1) | Some (i0, sc0) -> if sc0 let aij = Vect.get j v in if Int sgn */ aij if debug then Printf.fprintf stdout "safe_find %s x%i %a\n" err x output_tableau t; failwith err (** [find_pivot vr t] aims at improving the objective function of the basic variable vr *) let find_pivot vr (rst : Restricted.t) tbl = (* Get the objective of the basic variable vr *) let v = safe_find "find_pivot" vr tbl in match is_maximised rst v with | Some mx -> Done (Max mx) (* Maximum is reached; we are done *) | None -> ( (* Extract the vector *) let _, v = Vect.decomp_cst v in let j', sgn = find_pivot_column rst v in match find_pivot_row rst (IMap.remove vr tbl) j' sgn with | None -> Done (Ubnd j') | Some (i', sc) -> Pivot (i', j', sc) ) (** [solve_column c r e] @param c is a non-basic variable @param r is a basic variable @param e is a vector such that r = e and e is of the form ai.c+e' @return the vector (-r + e').-1/ai i.e c = (r - e')/ai *) let solve_column (c : var) (r : var) (e : Vect.t) : Vect.t = let a = Vect.get c e in if a =/ Int 0 then failwith "Cannot solve column" else let a' = Int (-1) // a in Vect.mul a' (Vect.set r (Int (-1)) (Vect.set c (Int 0) e)) (** [pivot_row r c e] @param c is such that c = e @param r is a vector r = g.c + r' @return g.e+r' *) let pivot_row (row : Vect.t) (c : var) (e : Vect.t) : Vect.t = let g = Vect.get c row in if g =/ Int 0 then row else Vect.mul_add g e (Int 1) (Vect.set c (Int 0) row) let pivot_with (m : tableau) (v : var) (p : Vect.t) = IMap.map (fun (r : Vect.t) -> pivot_row r v p) m let pivot (m : tableau) (r : var) (c : var) = let row = safe_find "pivot" r m in let piv = solve_column c r row in IMap.add c piv (pivot_with (IMap.remove r m) c piv) let adapt_unbounded vr x rst tbl = if Vect.get_cst (IMap.find vr tbl) >=/ Int 0 then tbl else pivot tbl vr x module BaseSet = Set.Make (struct type t = iset let compare = IMap.compare (fun x y -> 0) end) let get_base tbl = IMap.mapi (fun k _ -> ()) tbl let simplex opt vr rst tbl = let b = ref BaseSet.empty in let rec simplex opt vr rst tbl = ( if debug then let base = get_base tbl in if BaseSet.mem base !b then Printf.fprintf stdout "Cycling detected\n" else b := BaseSet.add base !b ); if debug && not (is_feasible rst tbl) then begin let m = unfeasible rst tbl in Printf.fprintf stdout "Simplex error\n"; Printf.fprintf stdout "The current tableau is not feasible\n"; Printf.fprintf stdout "Restricted >= %a\n" Restricted.pp rst; output_tableau stdout tbl; Printf.fprintf stdout "Error for variables %a\n" output_vars m end; if (not opt) && Vect.get_cst (IMap.find vr tbl) >=/ Int 0 then Opt (tbl, Feas) else match find_pivot vr rst tbl with | Done r -> ( match r with | Max _ -> Opt (tbl, r) | Ubnd x -> let t' = adapt_unbounded vr x rst tbl in Opt (t', r) | Feas -> raise (Invalid_argument "find_pivot") ) | Pivot (i, j, s) -> if debug then begin Printf.fprintf stdout "Find pivot for x%i(%s)\n" vr (string_of_num s); Printf.fprintf stdout "Leaving variable x%i\n" i; Printf.fprintf stdout "Entering variable x%i\n" j end; let m' = pivot tbl i j in simplex opt vr rst m' in simplex opt vr rst tbl type certificate = Unsat of Vect.t | Sat of tableau * var option (** [normalise_row t v] @return a row obtained by pivoting the basic variables of the vector v *) let normalise_row (t : tableau) (v : Vect.t) = Vect.fold (fun acc vr ai -> try let e = IMap.find vr t in Vect.add (Vect.mul ai e) acc with Not_found -> Vect.add (Vect.set vr ai Vect.null) acc) Vect.null v let normalise_row (t : tableau) (v : Vect.t) = let v' = normalise_row t v in if debug then Printf.fprintf stdout "Normalised Optimising %a\n" LinPoly.pp v'; v' let add_row (nw : var) (t : tableau) (v : Vect.t) : tableau = IMap.add nw (normalise_row t v) t (** [push_real] performs reasoning over the rationals *) let push_real (opt : bool) (nw : var) (v : Vect.t) (rst : Restricted.t) (t : tableau) : certificate = if debug then begin Printf.fprintf stdout "Current Tableau\n%a\n" output_tableau t; Printf.fprintf stdout "Optimising %a=%a\n" LinPoly.pp_var nw LinPoly.pp v end; match simplex opt nw rst (add_row nw t v) with | Opt (t', r) -> ( (* Look at the optimal *) match r with | Ubnd x -> if debug then Printf.printf "The objective is unbounded (variable %a)\n" LinPoly.pp_var x; Sat (t', Some x) (* This is sat and we can extract a value *) | Feas -> Sat (t', None) | Max n -> if debug then begin Printf.printf "The objective is maximised %s\n" (string_of_num n); Printf.printf "%a = %a\n" LinPoly.pp_var nw pp_row (IMap.find nw t') end; if n >=/ Int 0 then Sat (t', None) else let v' = safe_find "push_real" nw t' in Unsat (Vect.set nw (Int 1) (Vect.set 0 (Int 0) (Vect.mul (Int (-1)) v'))) ) open Mutils (** One complication is that equalities needs some pre-processing. *) open Polynomial (*type varmap = (int * bool) IMap.t*) let make_certificate vm l = Vect.normalise (Vect.fold (fun acc x n -> let x', b = IMap.find x vm in Vect.set x' (if b then n else Num.minus_num n) acc) Vect.null l) (** [eliminate_equalities vr0 l] represents an equality e = 0 of index idx in the list l by 2 constraints (vr:e >= 0) and (vr+1:-e >= 0) The mapping vm maps vr to idx *) let eliminate_equalities (vr0 : var) (l : Polynomial.cstr list) = let rec elim idx vr vm l acc = match l with | [] -> (vr, vm, acc) | c :: l -> ( match c.op with | Ge -> let v = Vect.set 0 (minus_num c.cst) c.coeffs in elim (idx + 1) (vr + 1) (IMap.add vr (idx, true) vm) l ((vr, v) :: acc) | Eq -> let v1 = Vect.set 0 (minus_num c.cst) c.coeffs in let v2 = Vect.mul (Int (-1)) v1 in let vm = IMap.add vr (idx, true) (IMap.add (vr + 1) (idx, false) vm) in elim (idx + 1) (vr + 2) vm l ((vr, v1) :: (vr + 1, v2) :: acc) | Gt -> raise Strict ) in elim 0 vr0 IMap.empty l [] let find_solution rst tbl = IMap.fold (fun vr v res -> if Restricted.is_restricted vr rst then res else Vect.set vr (Vect.get_cst v) res) tbl Vect.null let find_full_solution rst tbl = IMap.fold (fun vr v res -> Vect.set vr (Vect.get_cst v) res) tbl Vect.null let choose_conflict (sol : Vect.t) (l : (var * Vect.t) list) = let esol = Vect.set 0 (Int 1) sol in let rec most_violating l e (x, v) rst = match l with | [] -> Some ((x, v), rst) | (x', v') :: l -> let e' = Vect.dotproduct esol v' in if e' <=/ e then most_violating l e' (x', v') ((x, v) :: rst) else most_violating l e (x, v) ((x', v') :: rst) in match l with | [] -> None | (x, v) :: l -> let e = Vect.dotproduct esol v in most_violating l e (x, v) [] let rec solve opt l (rst : Restricted.t) (t : tableau) = let sol = find_solution rst t in match choose_conflict sol l with | None -> Inl (rst, t, None) | Some ((vr, v), l) -> ( match push_real opt vr v (Restricted.set_exc vr rst) t with | Sat (t', x) -> ( (* let t' = remove_redundant rst t' in*) match l with | [] -> Inl (rst, t', x) | _ -> solve opt l rst t' ) | Unsat c -> Inr c ) let find_unsat_certificate (l : Polynomial.cstr list) = let vr = LinPoly.MonT.get_fresh () in let _, vm, l' = eliminate_equalities vr l in match solve false l' (Restricted.make vr) IMap.empty with | Inr c -> Some (make_certificate vm c) | Inl _ -> None let fresh_var l = 1 + try ISet.max_elt (List.fold_left (fun acc c -> ISet.union acc (Vect.variables c.coeffs)) ISet.empty l) with Not_found -> 0 let find_point (l : Polynomial.cstr list) = let vr = fresh_var l in let _, vm, l' = eliminate_equalities vr l in match solve false l' (Restricted.make vr) IMap.empty with | Inl (rst, t, _) -> Some (find_solution rst t) | _ -> None let optimise obj l = let vr0 = LinPoly.MonT.get_fresh () in let _, vm, l' = eliminate_equalities (vr0 + 1) l in let bound pos res = match res with | Opt (_, Max n) -> Some (if pos then n else minus_num n) | Opt (_, Ubnd _) -> None | Opt (_, Feas) -> None in match solve false l' (Restricted.make vr0) IMap.empty with | Inl (rst, t, _) -> Some ( bound false (simplex true vr0 rst (add_row vr0 t (Vect.uminus obj))) , bound true (simplex true vr0 rst (add_row vr0 t obj)) ) | _ -> None open Polynomial let env_of_list l = List.fold_left (fun (i, m) l -> (i + 1, IMap.add i l m)) (0, IMap.empty) l open ProofFormat let make_farkas_certificate (env : WithProof.t IMap.t) vm v = Vect.fold (fun acc x n -> add_proof acc begin try let x', b = IMap.find x vm in mul_cst_proof (if b then n else Num.minus_num n) (snd (IMap.find x' env)) with Not_found -> (* This is an introduced hypothesis *) mul_cst_proof n (snd (IMap.find x env)) end) Zero v let make_farkas_proof (env : WithProof.t IMap.t) vm v = Vect.fold (fun wp x n -> WithProof.addition wp begin try let x', b = IMap.find x vm in let n = if b then n else Num.minus_num n in WithProof.mult (Vect.cst n) (IMap.find x' env) with Not_found -> WithProof.mult (Vect.cst n) (IMap.find x env) end) WithProof.zero v let frac_num n = n -/ Num.floor_num n type ('a, 'b) hitkind = | Forget (* Not interesting *) | Hit of 'a (* Yes, we have a positive result *) | Keep of 'b let cut env rmin sol vm (rst : Restricted.t) tbl (x, v) = let n, r = Vect.decomp_cst v in let f = frac_num n in if f =/ Int 0 then Forget (* The solution is integral *) else (* This is potentially a cut *) let t = if f if Restricted.is_restricted x rst then Vect.set x (ccoeff n) acc else acc) Vect.null r in let lcut = List.map (fun cv -> Vect.normalise (cut_vector cv)) [cut_coeff1; cut_coeff2] in let lcut = List.map (make_farkas_proof env vm) lcut in let check_cutting_plane c = match WithProof.cutting_plane c with | None -> if debug then Printf.printf "This is not a cutting plane for %a\n%a:" LinPoly.pp_var x WithProof.output c; None | Some (v, prf) -> if debug then ( Printf.printf "This is a cutting plane for %a:" LinPoly.pp_var x; Printf.printf " %a\n" WithProof.output (v, prf) ); if snd v = Eq then (* Unsat *) Some (x, (v, prf)) else let vl = Vect.dotproduct (fst v) (Vect.set 0 (Int 1) sol) in if eval_op Ge vl (Int 0) then ( if debug then Printf.printf "The cut is feasible %s >= 0 \n" (Num.string_of_num vl); None ) else Some (x, (v, prf)) in match find_some check_cutting_plane lcut with | Some r -> Hit r | None -> Keep (x, v) let merge_result_old oldr f x = match oldr with | Hit v -> Hit v | Forget -> ( match f x with Forget -> Forget | Hit v -> Hit v | Keep v -> Keep v ) | Keep v -> ( match f x with Forget -> Keep v | Keep v' -> Keep v | Hit v -> Hit v ) let merge_best lt oldr newr = match (oldr, newr) with | x, Forget -> x | Hit v, Hit v' -> if lt v v' then Hit v else Hit v' | _, Hit v | Hit v, _ -> Hit v | Forget, Keep v -> Keep v | Keep v, Keep v' -> Keep v' let find_cut nb env u sol vm rst tbl = if nb = 0 then IMap.fold (fun x v acc -> merge_result_old acc (cut env u sol vm rst tbl) (x, v)) tbl Forget else let lt (_, (_, p1)) (_, (_, p2)) = ProofFormat.pr_size p1 merge_best lt acc (cut env u sol vm rst tbl (x, v))) tbl Forget let var_of_vect v = fst (fst (Vect.decomp_fst v)) let eliminate_variable (bounded, vr, env, tbl) x = if debug then Printf.printf "Eliminating variable %a from tableau\n%a\n" LinPoly.pp_var x output_tableau tbl; (* We identify the new variables with the constraint. *) LinPoly.MonT.reserve vr; let z = LinPoly.var (vr + 1) in let zv = var_of_vect z in let t = LinPoly.var (vr + 2) in let tv = var_of_vect t in (* x = z - t *) let xdef = Vect.add z (Vect.uminus t) in let xp = ((Vect.set x (Int 1) (Vect.uminus xdef), Eq), Def vr) in let zp = ((z, Ge), Def zv) in let tp = ((t, Ge), Def tv) in (* Pivot the current tableau using xdef *) let tbl = IMap.map (fun v -> Vect.subst x xdef v) tbl in (* Pivot the environment *) let env = IMap.map (fun lp -> let (v, o), p = lp in let ai = Vect.get x v in if ai =/ Int 0 then lp else WithProof.addition (WithProof.mult (Vect.cst (Num.minus_num ai)) xp) lp) env in (* Add the variables to the environment *) let env = IMap.add vr xp (IMap.add zv zp (IMap.add tv tp env)) in (* Remember the mapping *) let bounded = IMap.add x (vr, zv, tv) bounded in if debug then ( Printf.printf "Tableau without\n %a\n" output_tableau tbl; Printf.printf "Environment\n %a\n" output_env env ); (bounded, vr + 3, env, tbl) let integer_solver lp = let l, _ = List.split lp in let vr0 = 3 * LinPoly.MonT.get_fresh () in let vr, vm, l' = eliminate_equalities vr0 l in let _, env = env_of_list (List.map WithProof.of_cstr lp) in let insert_row vr v rst tbl = match push_real true vr v rst tbl with | Sat (t', x) -> Inl (Restricted.restrict vr rst, t', x) | Unsat c -> Inr c in let nb = ref 0 in let rec isolve env cr vr res = incr nb; match res with | Inr c -> Some (Step (vr, make_farkas_certificate env vm (Vect.normalise c), Done)) | Inl (rst, tbl, x) -> ( if debug then begin Printf.fprintf stdout "Looking for a cut\n"; Printf.fprintf stdout "Restricted %a ...\n" Restricted.pp rst; Printf.fprintf stdout "Current Tableau\n%a\n" output_tableau tbl; flush stdout (* Printf.fprintf stdout "Bounding box\n%a\n" output_box (bounding_box (vr+1) rst tbl l')*) end; let sol = find_full_solution rst tbl in match find_cut (!nb mod 2) env cr (*x*) sol vm rst tbl with | Forget -> None (* There is no hope, there should be an integer solution *) | Hit (cr, ((v, op), cut)) -> if op = Eq then (* This is a contradiction *) Some (Step (vr, CutPrf cut, Done)) else ( LinPoly.MonT.reserve vr; let res = insert_row vr v (Restricted.set_exc vr rst) tbl in let prf = isolve (IMap.add vr ((v, op), Def vr) env) (Some cr) (vr + 1) res in match prf with | None -> None | Some p -> Some (Step (vr, CutPrf cut, p)) ) | Keep (x, v) -> ( if debug then Printf.fprintf stdout "Remove %a from Tableau\n" LinPoly.pp_var x; let bounded, vr, env, tbl = Vect.fold (fun acc x n -> if x <> 0 && not (Restricted.is_restricted x rst) then eliminate_variable acc x else acc) (IMap.empty, vr, env, tbl) v in let prf = isolve env cr vr (Inl (rst, tbl, None)) in match prf with | None -> None | Some pf -> Some (IMap.fold (fun x (vr, zv, tv) acc -> ExProof (vr, zv, tv, x, zv, tv, acc)) bounded pf) ) ) in let res = solve true l' (Restricted.make vr0) IMap.empty in isolve env None vr res let integer_solver lp = if debug then Printf.printf "Input integer solver\n%a\n" WithProof.output_sys (List.map WithProof.of_cstr lp); match integer_solver lp with | None -> None | Some prf -> if debug then Printf.fprintf stdout "Proof %a\n" ProofFormat.output_proof prf; Some prf coq-8.11.0/plugins/micromega/persistent_cache.ml0000644000175000017500000001241513612664533021572 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a t val find : 'a t -> key -> 'a val add : 'a t -> key -> 'a -> unit val memo : string -> (key -> 'a) -> key -> 'a val memo_cond : string -> (key -> bool) -> (key -> 'a) -> key -> 'a end open Hashtbl module PHashtable (Key : HashedType) : PHashtable with type key = Key.t = struct open Unix type key = Key.t module Table = Hashtbl.Make (Key) exception InvalidTableFormat exception UnboundTable type mode = Closed | Open type 'a t = {outch : out_channel; mutable status : mode; htbl : 'a Table.t} let finally f rst = try let res = f () in rst (); res with reraise -> (try rst () with any -> raise reraise); raise reraise let read_key_elem inch = try Some (Marshal.from_channel inch) with | End_of_file -> None | e when CErrors.noncritical e -> raise InvalidTableFormat (** We used to only lock/unlock regions. Is-it more robust/portable to lock/unlock a fixed region e.g. [0;1]? In case of locking failure, the cache is not used. **) type lock_kind = Read | Write let lock kd fd = let pos = lseek fd 0 SEEK_CUR in let success = try ignore (lseek fd 0 SEEK_SET); let lk = match kd with Read -> F_RLOCK | Write -> F_LOCK in lockf fd lk 1; true with Unix.Unix_error (_, _, _) -> false in ignore (lseek fd pos SEEK_SET); success let unlock fd = let pos = lseek fd 0 SEEK_CUR in try ignore (lseek fd 0 SEEK_SET); lockf fd F_ULOCK 1 with Unix.Unix_error (_, _, _) -> () (* Here, this is really bad news -- there is a pending lock which could cause a deadlock. Should it be an anomaly or produce a warning ? *); ignore (lseek fd pos SEEK_SET) (* We make the assumption that an acquired lock can always be released *) let do_under_lock kd fd f = if lock kd fd then finally f (fun () -> unlock fd) else f () let open_in f = let flags = [O_RDONLY; O_CREAT] in let finch = openfile f flags 0o666 in let inch = in_channel_of_descr finch in let htbl = Table.create 100 in let rec xload () = match read_key_elem inch with | None -> () | Some (key, elem) -> Table.add htbl key elem; xload () in try (* Locking of the (whole) file while reading *) do_under_lock Read finch xload; close_in_noerr inch; { outch = out_channel_of_descr (openfile f [O_WRONLY; O_APPEND; O_CREAT] 0o666) ; status = Open ; htbl } with InvalidTableFormat -> (* The file is corrupted *) close_in_noerr inch; let flags = [O_WRONLY; O_TRUNC; O_CREAT] in let out = openfile f flags 0o666 in let outch = out_channel_of_descr out in do_under_lock Write out (fun () -> Table.iter (fun k e -> Marshal.to_channel outch (k, e) [Marshal.No_sharing]) htbl; flush outch); {outch; status = Open; htbl} let add t k e = let {outch; status; htbl = tbl} = t in if status == Closed then raise UnboundTable else let fd = descr_of_out_channel outch in Table.add tbl k e; do_under_lock Write fd (fun _ -> Marshal.to_channel outch (k, e) [Marshal.No_sharing]; flush outch) let find t k = let {outch; status; htbl = tbl} = t in if status == Closed then raise UnboundTable else let res = Table.find tbl k in res let memo cache f = let tbl = lazy (try Some (open_in cache) with _ -> None) in fun x -> match Lazy.force tbl with | None -> f x | Some tbl -> ( try find tbl x with Not_found -> let res = f x in add tbl x res; res ) let memo_cond cache cond f = let tbl = lazy (try Some (open_in cache) with _ -> None) in fun x -> match Lazy.force tbl with | None -> f x | Some tbl -> if cond x then begin try find tbl x with Not_found -> let res = f x in add tbl x res; res end else f x end (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/Psatz.v0000644000175000017500000000424013612664533017202 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (sos_Z Lia.zchecker) || (psatz_Z d Lia.zchecker) | R => (sos_R Lra.rchecker) || (psatz_R d Lra.rchecker) | Q => (sos_Q Lqa.rchecker) || (psatz_Q d Lqa.rchecker) | _ => fail "Unsupported domain" end in tac. Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n. Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1). Ltac psatzl dom := let tac := lazymatch dom with | Z => Lia.lia | Q => Lqa.lra | R => Lra.lra | _ => fail "Unsupported domain" end in tac. Ltac lra := first [ psatzl R | psatzl Q ]. Ltac nra := first [ Lra.nra | Lqa.nra ]. (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/mutils.mli0000644000175000017500000000755413612664533017745 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int -> int val equal : int -> int -> bool end module ISet : sig include Set.S with type elt = int val pp : out_channel -> t -> unit end module IMap : sig include Map.S with type key = int val from : key -> 'elt t -> 'elt t (** [from k m] returns the submap of [m] with keys greater or equal k *) end val numerator : Num.num -> Big_int.big_int val denominator : Num.num -> Big_int.big_int module Cmp : sig val compare_list : ('a -> 'b -> int) -> 'a list -> 'b list -> int val compare_lexical : (unit -> int) list -> int end module Tag : sig type t val pp : out_channel -> t -> unit val next : t -> t val max : t -> t -> t val from : int -> t val to_int : t -> int end module TagSet : CSig.SetS with type elt = Tag.t val pp_list : string -> (out_channel -> 'a -> unit) -> out_channel -> 'a list -> unit module CamlToCoq : sig val positive : int -> Micromega.positive val bigint : Big_int.big_int -> Micromega.z val n : int -> Micromega.n val nat : int -> Micromega.nat val q : Num.num -> Micromega.q val index : int -> Micromega.positive val z : int -> Micromega.z val positive_big_int : Big_int.big_int -> Micromega.positive end module CoqToCaml : sig val z_big_int : Micromega.z -> Big_int.big_int val z : Micromega.z -> int val q_to_num : Micromega.q -> Num.num val positive : Micromega.positive -> int val n : Micromega.n -> int val nat : Micromega.nat -> int val index : Micromega.positive -> int end module Hash : sig val eq_op1 : Micromega.op1 -> Micromega.op1 -> bool val eq_positive : Micromega.positive -> Micromega.positive -> bool val eq_z : Micromega.z -> Micromega.z -> bool val eq_q : Micromega.q -> Micromega.q -> bool val eq_pol : ('a -> 'a -> bool) -> 'a Micromega.pol -> 'a Micromega.pol -> bool val eq_pair : ('a -> 'a -> bool) -> ('b -> 'b -> bool) -> 'a * 'b -> 'a * 'b -> bool val hash_op1 : int -> Micromega.op1 -> int val hash_pol : (int -> 'a -> int) -> int -> 'a Micromega.pol -> int val hash_pair : (int -> 'a -> int) -> (int -> 'b -> int) -> int -> 'a * 'b -> int val hash_z : int -> Micromega.z -> int val hash_q : int -> Micromega.q -> int val hash_string : int -> string -> int val hash_elt : ('a -> int) -> int -> 'a -> int end val ppcm : Big_int.big_int -> Big_int.big_int -> Big_int.big_int val all_pairs : ('a -> 'a -> 'b) -> 'a list -> 'b list val try_any : (('a -> 'b option) * 'c) list -> 'a -> 'b option val is_sublist : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool val extract : ('a -> 'b option) -> 'a list -> ('b * 'a) option * 'a list val extract_all : ('a -> 'b option) -> 'a list -> ('b * 'a) list * 'a list val extract_best : ('a -> 'b option) -> ('b -> 'b -> bool) -> 'a list -> ('b * 'a) option * 'a list val find_some : ('a -> 'b option) -> 'a list -> 'b option val iterate_until_stable : ('a -> 'a option) -> 'a -> 'a val simplify : ('a -> 'a option) -> 'a list -> 'a list option val saturate : ('a -> 'b option) -> ('b * 'a -> 'a -> 'a option) -> 'a list -> 'a list val generate : ('a -> 'b option) -> 'a list -> 'b list val app_funs : ('a -> 'b option) list -> 'a -> 'b option val command : string -> string array -> 'a -> 'b coq-8.11.0/plugins/micromega/Tauto.v0000644000175000017500000011510313612664533017176 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* GFormula | A : TA -> AA -> GFormula | Cj : GFormula -> GFormula -> GFormula | D : GFormula -> GFormula -> GFormula | N : GFormula -> GFormula | I : GFormula -> option AF -> GFormula -> GFormula. Section MAPX. Variable F : TX -> TX. Fixpoint mapX (f : GFormula) : GFormula := match f with | TT => TT | FF => FF | X x => X (F x) | A a an => A a an | Cj f1 f2 => Cj (mapX f1) (mapX f2) | D f1 f2 => D (mapX f1) (mapX f2) | N f => N (mapX f) | I f1 o f2 => I (mapX f1) o (mapX f2) end. End MAPX. Section FOLDANNOT. Variable ACC : Type. Variable F : ACC -> AA -> ACC. Fixpoint foldA (f : GFormula) (acc : ACC) : ACC := match f with | TT => acc | FF => acc | X x => acc | A a an => F acc an | Cj f1 f2 | D f1 f2 | I f1 _ f2 => foldA f1 (foldA f2 acc) | N f => foldA f acc end. End FOLDANNOT. Definition cons_id (id : option AF) (l : list AF) := match id with | None => l | Some id => id :: l end. Fixpoint ids_of_formula f := match f with | I f id f' => cons_id id (ids_of_formula f') | _ => nil end. Fixpoint collect_annot (f : GFormula) : list AA := match f with | TT | FF | X _ => nil | A _ a => a ::nil | Cj f1 f2 | D f1 f2 | I f1 _ f2 => collect_annot f1 ++ collect_annot f2 | N f => collect_annot f end. Variable ex : TX -> Prop. (* [ex] will be the identity *) Section EVAL. Variable ea : TA -> Prop. Fixpoint eval_f (f:GFormula) {struct f}: Prop := match f with | TT => True | FF => False | A a _ => ea a | X p => ex p | Cj e1 e2 => (eval_f e1) /\ (eval_f e2) | D e1 e2 => (eval_f e1) \/ (eval_f e2) | N e => ~ (eval_f e) | I f1 _ f2 => (eval_f f1) -> (eval_f f2) end. End EVAL. Lemma eval_f_morph : forall (ev ev' : TA -> Prop) (f : GFormula), (forall a, ev a <-> ev' a) -> (eval_f ev f <-> eval_f ev' f). Proof. induction f ; simpl ; try tauto. intros. apply H. Qed. End S. (** Typical boolean formulae *) Definition BFormula (A : Type) := @GFormula A Prop unit unit. Section MAPATOMS. Context {TA TA':Type}. Context {TX : Type}. Context {AA : Type}. Context {AF : Type}. Fixpoint map_bformula (fct : TA -> TA') (f : @GFormula TA TX AA AF ) : @GFormula TA' TX AA AF := match f with | TT => TT | FF => FF | X p => X p | A a t => A (fct a) t | Cj f1 f2 => Cj (map_bformula fct f1) (map_bformula fct f2) | D f1 f2 => D (map_bformula fct f1) (map_bformula fct f2) | N f => N (map_bformula fct f) | I f1 a f2 => I (map_bformula fct f1) a (map_bformula fct f2) end. End MAPATOMS. Lemma map_simpl : forall A B f l, @map A B f l = match l with | nil => nil | a :: l=> (f a) :: (@map A B f l) end. Proof. destruct l ; reflexivity. Qed. Section S. (** A cnf tracking annotations of atoms. *) (** Type parameters *) Variable Env : Type. Variable Term : Type. Variable Term' : Type. Variable Annot : Type. Variable unsat : Term' -> bool. (* see [unsat_prop] *) Variable deduce : Term' -> Term' -> option Term'. (* see [deduce_prop] *) Definition clause := list (Term' * Annot). Definition cnf := list clause. Variable normalise : Term -> Annot -> cnf. Variable negate : Term -> Annot -> cnf. Definition cnf_tt : cnf := @nil clause. Definition cnf_ff : cnf := cons (@nil (Term' * Annot)) nil. (** Our cnf is optimised and detects contradictions on the fly. *) Fixpoint add_term (t: Term' * Annot) (cl : clause) : option clause := match cl with | nil => match deduce (fst t) (fst t) with | None => Some (t ::nil) | Some u => if unsat u then None else Some (t::nil) end | t'::cl => match deduce (fst t) (fst t') with | None => match add_term t cl with | None => None | Some cl' => Some (t' :: cl') end | Some u => if unsat u then None else match add_term t cl with | None => None | Some cl' => Some (t' :: cl') end end end. Fixpoint or_clause (cl1 cl2 : clause) : option clause := match cl1 with | nil => Some cl2 | t::cl => match add_term t cl2 with | None => None | Some cl' => or_clause cl cl' end end. Definition xor_clause_cnf (t:clause) (f:cnf) : cnf := List.fold_left (fun acc e => match or_clause t e with | None => acc | Some cl => cl :: acc end) f nil . Definition or_clause_cnf (t: clause) (f:cnf) : cnf := match t with | nil => f | _ => xor_clause_cnf t f end. Fixpoint or_cnf (f : cnf) (f' : cnf) {struct f}: cnf := match f with | nil => cnf_tt | e :: rst => (or_cnf rst f') +++ (or_clause_cnf e f') end. Definition and_cnf (f1 : cnf) (f2 : cnf) : cnf := f1 +++ f2. (** TX is Prop in Coq and EConstr.constr in Ocaml. AF i s unit in Coq and Names.Id.t in Ocaml *) Definition TFormula (TX: Type) (AF: Type) := @GFormula Term TX Annot AF. Definition is_cnf_tt (c : cnf) : bool := match c with | nil => true | _ => false end. Definition is_cnf_ff (c : cnf) : bool := match c with | nil::nil => true | _ => false end. Definition and_cnf_opt (f1 : cnf) (f2 : cnf) : cnf := if is_cnf_ff f1 || is_cnf_ff f2 then cnf_ff else and_cnf f1 f2. Definition or_cnf_opt (f1 : cnf) (f2 : cnf) : cnf := if is_cnf_tt f1 || is_cnf_tt f2 then cnf_tt else if is_cnf_ff f2 then f1 else or_cnf f1 f2. Fixpoint xcnf {TX AF: Type} (pol : bool) (f : TFormula TX AF) {struct f}: cnf := match f with | TT => if pol then cnf_tt else cnf_ff | FF => if pol then cnf_ff else cnf_tt | X p => if pol then cnf_ff else cnf_ff (* This is not complete - cannot negate any proposition *) | A x t => if pol then normalise x t else negate x t | N e => xcnf (negb pol) e | Cj e1 e2 => (if pol then and_cnf_opt else or_cnf_opt) (xcnf pol e1) (xcnf pol e2) | D e1 e2 => (if pol then or_cnf_opt else and_cnf_opt) (xcnf pol e1) (xcnf pol e2) | I e1 _ e2 => (if pol then or_cnf_opt else and_cnf_opt) (xcnf (negb pol) e1) (xcnf pol e2) end. Section CNFAnnot. (** Records annotations used to optimise the cnf. Those need to be kept when pruning the formula. For efficiency, this is a separate function. *) Fixpoint radd_term (t : Term' * Annot) (cl : clause) : clause + list Annot := match cl with | nil => (* if t is unsat, the clause is empty BUT t is needed. *) match deduce (fst t) (fst t) with | Some u => if unsat u then inr ((snd t)::nil) else inl (t::nil) | None => inl (t::nil) end | t'::cl => (* if t /\ t' is unsat, the clause is empty BUT t & t' are needed *) match deduce (fst t) (fst t') with | Some u => if unsat u then inr ((snd t)::(snd t')::nil) else match radd_term t cl with | inl cl' => inl (t'::cl') | inr l => inr l end | None => match radd_term t cl with | inl cl' => inl (t'::cl') | inr l => inr l end end end. Fixpoint ror_clause cl1 cl2 := match cl1 with | nil => inl cl2 | t::cl => match radd_term t cl2 with | inl cl' => ror_clause cl cl' | inr l => inr l end end. Definition xror_clause_cnf t f := List.fold_left (fun '(acc,tg) e => match ror_clause t e with | inl cl => (cl :: acc,tg) | inr l => (acc,tg+++l) end) f (nil,nil). Definition ror_clause_cnf t f := match t with | nil => (f,nil) | _ => xror_clause_cnf t f end. Fixpoint ror_cnf (f f':list clause) := match f with | nil => (cnf_tt,nil) | e :: rst => let (rst_f',t) := ror_cnf rst f' in let (e_f', t') := ror_clause_cnf e f' in (rst_f' +++ e_f', t +++ t') end. Definition annot_of_clause (l : clause) : list Annot := List.map snd l. Definition annot_of_cnf (f : cnf) : list Annot := List.fold_left (fun acc e => annot_of_clause e +++ acc ) f nil. Definition ror_cnf_opt f1 f2 := if is_cnf_tt f1 then (cnf_tt , nil) else if is_cnf_tt f2 then (cnf_tt, nil) else if is_cnf_ff f2 then (f1,nil) else ror_cnf f1 f2. Definition ocons {A : Type} (o : option A) (l : list A) : list A := match o with | None => l | Some e => e ::l end. Definition ratom (c : cnf) (a : Annot) : cnf * list Annot := if is_cnf_ff c || is_cnf_tt c then (c,a::nil) else (c,nil). (* t is embedded in c *) Fixpoint rxcnf {TX AF: Type}(polarity : bool) (f : TFormula TX AF) : cnf * list Annot := match f with | TT => if polarity then (cnf_tt,nil) else (cnf_ff,nil) | FF => if polarity then (cnf_ff,nil) else (cnf_tt,nil) | X p => if polarity then (cnf_ff,nil) else (cnf_ff,nil) | A x t => ratom (if polarity then normalise x t else negate x t) t | N e => rxcnf (negb polarity) e | Cj e1 e2 => let '(e1,t1) := rxcnf polarity e1 in let '(e2,t2) := rxcnf polarity e2 in if polarity then (and_cnf_opt e1 e2, t1 +++ t2) else let (f',t') := ror_cnf_opt e1 e2 in (f', t1 +++ t2 +++ t') | D e1 e2 => let '(e1,t1) := rxcnf polarity e1 in let '(e2,t2) := rxcnf polarity e2 in if polarity then let (f',t') := ror_cnf_opt e1 e2 in (f', t1 +++ t2 +++ t') else (and_cnf_opt e1 e2, t1 +++ t2) | I e1 a e2 => let '(e1 , t1) := (rxcnf (negb polarity) e1) in if polarity then if is_cnf_ff e1 then rxcnf polarity e2 else (* compute disjunction *) let '(e2 , t2) := (rxcnf polarity e2) in let (f',t') := ror_cnf_opt e1 e2 in (f', t1 +++ t2 +++ t') (* record the hypothesis *) else let '(e2 , t2) := (rxcnf polarity e2) in (and_cnf_opt e1 e2, t1 +++ t2) end. Section Abstraction. Variable TX : Type. Variable AF : Type. Class to_constrT : Type := { mkTT : TX; mkFF : TX; mkA : Term -> Annot -> TX; mkCj : TX -> TX -> TX; mkD : TX -> TX -> TX; mkI : TX -> TX -> TX; mkN : TX -> TX }. Context {to_constr : to_constrT}. Fixpoint aformula (f : TFormula TX AF) : TX := match f with | TT => mkTT | FF => mkFF | X p => p | A x t => mkA x t | Cj f1 f2 => mkCj (aformula f1) (aformula f2) | D f1 f2 => mkD (aformula f1) (aformula f2) | I f1 o f2 => mkI (aformula f1) (aformula f2) | N f => mkN (aformula f) end. Definition is_X (f : TFormula TX AF) : option TX := match f with | X p => Some p | _ => None end. Definition is_X_inv : forall f x, is_X f = Some x -> f = X x. Proof. destruct f ; simpl ; congruence. Qed. Variable needA : Annot -> bool. Definition abs_and (f1 f2 : TFormula TX AF) (c : TFormula TX AF -> TFormula TX AF -> TFormula TX AF) := match is_X f1 , is_X f2 with | Some _ , _ | _ , Some _ => X (aformula (c f1 f2)) | _ , _ => c f1 f2 end. Definition abs_or (f1 f2 : TFormula TX AF) (c : TFormula TX AF -> TFormula TX AF -> TFormula TX AF) := match is_X f1 , is_X f2 with | Some _ , Some _ => X (aformula (c f1 f2)) | _ , _ => c f1 f2 end. Definition mk_arrow (o : option AF) (f1 f2: TFormula TX AF) := match o with | None => I f1 None f2 | Some _ => if is_X f1 then f2 else I f1 o f2 end. Fixpoint abst_form (pol : bool) (f : TFormula TX AF) := match f with | TT => if pol then TT else X mkTT | FF => if pol then X mkFF else FF | X p => X p | A x t => if needA t then A x t else X (mkA x t) | Cj f1 f2 => let f1 := abst_form pol f1 in let f2 := abst_form pol f2 in if pol then abs_and f1 f2 Cj else abs_or f1 f2 Cj | D f1 f2 => let f1 := abst_form pol f1 in let f2 := abst_form pol f2 in if pol then abs_or f1 f2 D else abs_and f1 f2 D | I f1 o f2 => let f1 := abst_form (negb pol) f1 in let f2 := abst_form pol f2 in if pol then abs_or f1 f2 (mk_arrow o) else abs_and f1 f2 (mk_arrow o) | N f => let f := abst_form (negb pol) f in match is_X f with | Some a => X (mkN a) | _ => N f end end. Lemma if_same : forall {A: Type} (b:bool) (t:A), (if b then t else t) = t. Proof. destruct b ; reflexivity. Qed. Lemma is_cnf_tt_cnf_ff : is_cnf_tt cnf_ff = false. Proof. reflexivity. Qed. Lemma is_cnf_ff_cnf_ff : is_cnf_ff cnf_ff = true. Proof. reflexivity. Qed. Lemma is_cnf_tt_inv : forall f1, is_cnf_tt f1 = true -> f1 = cnf_tt. Proof. unfold cnf_tt. destruct f1 ; simpl ; try congruence. Qed. Lemma is_cnf_ff_inv : forall f1, is_cnf_ff f1 = true -> f1 = cnf_ff. Proof. unfold cnf_ff. destruct f1 ; simpl ; try congruence. destruct c ; simpl ; try congruence. destruct f1 ; try congruence. reflexivity. Qed. Lemma if_cnf_tt : forall f, (if is_cnf_tt f then cnf_tt else f) = f. Proof. intros. destruct (is_cnf_tt f) eqn:EQ. apply is_cnf_tt_inv in EQ;auto. reflexivity. Qed. Lemma or_cnf_opt_cnf_ff : forall f, or_cnf_opt cnf_ff f = f. Proof. intros. unfold or_cnf_opt. rewrite is_cnf_tt_cnf_ff. simpl. destruct (is_cnf_tt f) eqn:EQ. apply is_cnf_tt_inv in EQ. congruence. destruct (is_cnf_ff f) eqn:EQ1. apply is_cnf_ff_inv in EQ1. congruence. reflexivity. Qed. Lemma abs_and_pol : forall f1 f2 pol, and_cnf_opt (xcnf pol f1) (xcnf pol f2) = xcnf pol (abs_and f1 f2 (if pol then Cj else D)). Proof. unfold abs_and; intros. destruct (is_X f1) eqn:EQ1. apply is_X_inv in EQ1. subst. simpl. rewrite if_same. reflexivity. destruct (is_X f2) eqn:EQ2. apply is_X_inv in EQ2. subst. simpl. rewrite if_same. unfold and_cnf_opt. rewrite orb_comm. reflexivity. destruct pol ; simpl; auto. Qed. Lemma abs_or_pol : forall f1 f2 pol, or_cnf_opt (xcnf pol f1) (xcnf pol f2) = xcnf pol (abs_or f1 f2 (if pol then D else Cj)). Proof. unfold abs_or; intros. destruct (is_X f1) eqn:EQ1. apply is_X_inv in EQ1. subst. destruct (is_X f2) eqn:EQ2. apply is_X_inv in EQ2. subst. simpl. rewrite if_same. reflexivity. simpl. rewrite if_same. destruct pol ; simpl; auto. destruct pol ; simpl ; auto. Qed. Variable needA_all : forall a, needA a = true. Lemma xcnf_true_mk_arrow_l : forall o t f, xcnf true (mk_arrow o (X t) f) = xcnf true f. Proof. destruct o ; simpl; auto. intros. rewrite or_cnf_opt_cnf_ff. reflexivity. Qed. Lemma or_cnf_opt_cnf_ff_r : forall f, or_cnf_opt f cnf_ff = f. Proof. unfold or_cnf_opt. intros. rewrite is_cnf_tt_cnf_ff. rewrite orb_comm. simpl. apply if_cnf_tt. Qed. Lemma xcnf_true_mk_arrow_r : forall o t f, xcnf true (mk_arrow o f (X t)) = xcnf false f. Proof. destruct o ; simpl; auto. - intros. destruct (is_X f) eqn:EQ. apply is_X_inv in EQ. subst. reflexivity. simpl. apply or_cnf_opt_cnf_ff_r. - intros. apply or_cnf_opt_cnf_ff_r. Qed. Lemma abst_form_correct : forall f pol, xcnf pol f = xcnf pol (abst_form pol f). Proof. induction f;intros. - simpl. destruct pol ; reflexivity. - simpl. destruct pol ; reflexivity. - simpl. reflexivity. - simpl. rewrite needA_all. reflexivity. - simpl. specialize (IHf1 pol). specialize (IHf2 pol). rewrite IHf1. rewrite IHf2. destruct pol. + apply abs_and_pol; auto. + apply abs_or_pol; auto. - simpl. specialize (IHf1 pol). specialize (IHf2 pol). rewrite IHf1. rewrite IHf2. destruct pol. + apply abs_or_pol; auto. + apply abs_and_pol; auto. - simpl. specialize (IHf (negb pol)). destruct (is_X (abst_form (negb pol) f)) eqn:EQ1. + apply is_X_inv in EQ1. rewrite EQ1 in *. simpl in *. destruct pol ; auto. + simpl. congruence. - simpl. specialize (IHf1 (negb pol)). specialize (IHf2 pol). destruct pol. + simpl in *. unfold abs_or. destruct (is_X (abst_form false f1)) eqn:EQ1; destruct (is_X (abst_form true f2)) eqn:EQ2 ; simpl. * apply is_X_inv in EQ1. apply is_X_inv in EQ2. rewrite EQ1 in *. rewrite EQ2 in *. rewrite IHf1. rewrite IHf2. simpl. reflexivity. * apply is_X_inv in EQ1. rewrite EQ1 in *. rewrite IHf1. simpl. rewrite xcnf_true_mk_arrow_l. rewrite or_cnf_opt_cnf_ff. congruence. * apply is_X_inv in EQ2. rewrite EQ2 in *. rewrite IHf2. simpl. rewrite xcnf_true_mk_arrow_r. rewrite or_cnf_opt_cnf_ff_r. congruence. * destruct o ; simpl ; try congruence. rewrite EQ1. simpl. congruence. + simpl in *. unfold abs_and. destruct (is_X (abst_form true f1)) eqn:EQ1; destruct (is_X (abst_form false f2)) eqn:EQ2 ; simpl. * apply is_X_inv in EQ1. apply is_X_inv in EQ2. rewrite EQ1 in *. rewrite EQ2 in *. rewrite IHf1. rewrite IHf2. simpl. reflexivity. * apply is_X_inv in EQ1. rewrite EQ1 in *. rewrite IHf1. simpl. reflexivity. * apply is_X_inv in EQ2. rewrite EQ2 in *. rewrite IHf2. simpl. unfold and_cnf_opt. rewrite orb_comm. reflexivity. * destruct o; simpl. rewrite EQ1. simpl. congruence. congruence. Qed. End Abstraction. End CNFAnnot. Lemma radd_term_term : forall a' a cl, radd_term a a' = inl cl -> add_term a a' = Some cl. Proof. induction a' ; simpl. - intros. destruct (deduce (fst a) (fst a)). destruct (unsat t). congruence. inversion H. reflexivity. inversion H ;reflexivity. - intros. destruct (deduce (fst a0) (fst a)). destruct (unsat t). congruence. destruct (radd_term a0 a') eqn:RADD; try congruence. inversion H. subst. apply IHa' in RADD. rewrite RADD. reflexivity. destruct (radd_term a0 a') eqn:RADD; try congruence. inversion H. subst. apply IHa' in RADD. rewrite RADD. reflexivity. Qed. Lemma radd_term_term' : forall a' a cl, add_term a a' = Some cl -> radd_term a a' = inl cl. Proof. induction a' ; simpl. - intros. destruct (deduce (fst a) (fst a)). destruct (unsat t). congruence. inversion H. reflexivity. inversion H ;reflexivity. - intros. destruct (deduce (fst a0) (fst a)). destruct (unsat t). congruence. destruct (add_term a0 a') eqn:RADD; try congruence. inversion H. subst. apply IHa' in RADD. rewrite RADD. reflexivity. destruct (add_term a0 a') eqn:RADD; try congruence. inversion H. subst. apply IHa' in RADD. rewrite RADD. reflexivity. Qed. Lemma xror_clause_clause : forall a f, fst (xror_clause_cnf a f) = xor_clause_cnf a f. Proof. unfold xror_clause_cnf. unfold xor_clause_cnf. assert (ACC: fst (@nil clause,@nil Annot) = nil). reflexivity. intros. set (F1:= (fun '(acc, tg) (e : clause) => match ror_clause a e with | inl cl => (cl :: acc, tg) | inr l => (acc, tg +++ l) end)). set (F2:= (fun (acc : list clause) (e : clause) => match or_clause a e with | Some cl => cl :: acc | None => acc end)). revert ACC. generalize (@nil clause,@nil Annot). generalize (@nil clause). induction f ; simpl ; auto. intros. apply IHf. unfold F1 , F2. destruct p ; simpl in * ; subst. clear. revert a0. induction a; simpl; auto. intros. destruct (radd_term a a1) eqn:RADD. apply radd_term_term in RADD. rewrite RADD. auto. destruct (add_term a a1) eqn:RADD'. apply radd_term_term' in RADD'. congruence. reflexivity. Qed. Lemma ror_clause_clause : forall a f, fst (ror_clause_cnf a f) = or_clause_cnf a f. Proof. unfold ror_clause_cnf,or_clause_cnf. destruct a ; auto. apply xror_clause_clause. Qed. Lemma ror_cnf_cnf : forall f1 f2, fst (ror_cnf f1 f2) = or_cnf f1 f2. Proof. induction f1 ; simpl ; auto. intros. specialize (IHf1 f2). destruct(ror_cnf f1 f2). rewrite <- ror_clause_clause. destruct(ror_clause_cnf a f2). simpl. rewrite <- IHf1. reflexivity. Qed. Lemma ror_opt_cnf_cnf : forall f1 f2, fst (ror_cnf_opt f1 f2) = or_cnf_opt f1 f2. Proof. unfold ror_cnf_opt, or_cnf_opt. intros. destruct (is_cnf_tt f1). - simpl ; auto. - simpl. destruct (is_cnf_tt f2) ; simpl ; auto. destruct (is_cnf_ff f2) eqn:EQ. reflexivity. apply ror_cnf_cnf. Qed. Lemma ratom_cnf : forall f a, fst (ratom f a) = f. Proof. unfold ratom. intros. destruct (is_cnf_ff f || is_cnf_tt f); auto. Qed. Lemma rxcnf_xcnf : forall {TX AF:Type} (f:TFormula TX AF) b, fst (rxcnf b f) = xcnf b f. Proof. induction f ; simpl ; auto. - destruct b; simpl ; auto. - destruct b; simpl ; auto. - destruct b ; simpl ; auto. - intros. rewrite ratom_cnf. reflexivity. - intros. specialize (IHf1 b). specialize (IHf2 b). destruct (rxcnf b f1). destruct (rxcnf b f2). simpl in *. subst. destruct b ; auto. rewrite <- ror_opt_cnf_cnf. destruct (ror_cnf_opt (xcnf false f1) (xcnf false f2)). reflexivity. - intros. specialize (IHf1 b). specialize (IHf2 b). rewrite <- IHf1. rewrite <- IHf2. destruct (rxcnf b f1). destruct (rxcnf b f2). simpl in *. subst. destruct b ; auto. rewrite <- ror_opt_cnf_cnf. destruct (ror_cnf_opt (xcnf true f1) (xcnf true f2)). reflexivity. - intros. specialize (IHf1 (negb b)). specialize (IHf2 b). rewrite <- IHf1. rewrite <- IHf2. destruct (rxcnf (negb b) f1). destruct (rxcnf b f2). simpl in *. subst. destruct b;auto. generalize (is_cnf_ff_inv (xcnf (negb true) f1)). destruct (is_cnf_ff (xcnf (negb true) f1)). + intros. rewrite H by auto. unfold or_cnf_opt. simpl. destruct (is_cnf_tt (xcnf true f2)) eqn:EQ;auto. apply is_cnf_tt_inv in EQ; auto. destruct (is_cnf_ff (xcnf true f2)) eqn:EQ1. apply is_cnf_ff_inv in EQ1. congruence. reflexivity. + rewrite <- ror_opt_cnf_cnf. destruct (ror_cnf_opt (xcnf (negb true) f1) (xcnf true f2)). intros. reflexivity. Qed. Variable eval' : Env -> Term' -> Prop. Variable no_middle_eval' : forall env d, (eval' env d) \/ ~ (eval' env d). Variable unsat_prop : forall t, unsat t = true -> forall env, eval' env t -> False. Variable deduce_prop : forall t t' u, deduce t t' = Some u -> forall env, eval' env t -> eval' env t' -> eval' env u. Definition eval_tt (env : Env) (tt : Term' * Annot) := eval' env (fst tt). Definition eval_clause (env : Env) (cl : clause) := ~ make_conj (eval_tt env) cl. Definition eval_cnf (env : Env) (f:cnf) := make_conj (eval_clause env) f. Lemma eval_cnf_app : forall env x y, eval_cnf env (x+++y) <-> eval_cnf env x /\ eval_cnf env y. Proof. unfold eval_cnf. intros. rewrite make_conj_rapp. rewrite make_conj_app ; auto. tauto. Qed. Lemma eval_cnf_ff : forall env, eval_cnf env cnf_ff <-> False. Proof. unfold cnf_ff, eval_cnf,eval_clause. simpl. tauto. Qed. Lemma eval_cnf_tt : forall env, eval_cnf env cnf_tt <-> True. Proof. unfold cnf_tt, eval_cnf,eval_clause. simpl. tauto. Qed. Lemma eval_cnf_and_opt : forall env x y, eval_cnf env (and_cnf_opt x y) <-> eval_cnf env (and_cnf x y). Proof. unfold and_cnf_opt. intros. destruct (is_cnf_ff x) eqn:F1. { apply is_cnf_ff_inv in F1. simpl. subst. unfold and_cnf. rewrite eval_cnf_app. rewrite eval_cnf_ff. tauto. } simpl. destruct (is_cnf_ff y) eqn:F2. { apply is_cnf_ff_inv in F2. simpl. subst. unfold and_cnf. rewrite eval_cnf_app. rewrite eval_cnf_ff. tauto. } tauto. Qed. Definition eval_opt_clause (env : Env) (cl: option clause) := match cl with | None => True | Some cl => eval_clause env cl end. Lemma add_term_correct : forall env t cl , eval_opt_clause env (add_term t cl) <-> eval_clause env (t::cl). Proof. induction cl. - (* BC *) simpl. case_eq (deduce (fst t) (fst t)) ; try tauto. intros. generalize (@deduce_prop _ _ _ H env). case_eq (unsat t0) ; try tauto. { intros. generalize (@unsat_prop _ H0 env). unfold eval_clause. rewrite make_conj_cons. simpl; intros. tauto. } - (* IC *) simpl. case_eq (deduce (fst t) (fst a)); intros. generalize (@deduce_prop _ _ _ H env). case_eq (unsat t0); intros. { generalize (@unsat_prop _ H0 env). simpl. unfold eval_clause. repeat rewrite make_conj_cons. tauto. } destruct (add_term t cl) ; simpl in * ; try tauto. { intros. unfold eval_clause in *. repeat rewrite make_conj_cons in *. tauto. } { unfold eval_clause in *. repeat rewrite make_conj_cons in *. tauto. } destruct (add_term t cl) ; simpl in *; unfold eval_clause in * ; repeat rewrite make_conj_cons in *; tauto. Qed. Lemma no_middle_eval_tt : forall env a, eval_tt env a \/ ~ eval_tt env a. Proof. unfold eval_tt. auto. Qed. Hint Resolve no_middle_eval_tt : tauto. Lemma or_clause_correct : forall cl cl' env, eval_opt_clause env (or_clause cl cl') <-> eval_clause env cl \/ eval_clause env cl'. Proof. induction cl. - simpl. unfold eval_clause at 2. simpl. tauto. - intros *. simpl. assert (HH := add_term_correct env a cl'). assert (eval_tt env a \/ ~ eval_tt env a) by (apply no_middle_eval'). destruct (add_term a cl'); simpl in *. + rewrite IHcl. unfold eval_clause in *. rewrite !make_conj_cons in *. tauto. + unfold eval_clause in *. repeat rewrite make_conj_cons in *. tauto. Qed. Lemma or_clause_cnf_correct : forall env t f, eval_cnf env (or_clause_cnf t f) <-> (eval_clause env t) \/ (eval_cnf env f). Proof. unfold eval_cnf. unfold or_clause_cnf. intros until t. set (F := (fun (acc : list clause) (e : clause) => match or_clause t e with | Some cl => cl :: acc | None => acc end)). intro f. assert ( make_conj (eval_clause env) (fold_left F f nil) <-> (eval_clause env t \/ make_conj (eval_clause env) f) /\ make_conj (eval_clause env) nil). { generalize (@nil clause) as acc. induction f. - simpl. intros ; tauto. - intros. simpl fold_left. rewrite IHf. rewrite make_conj_cons. unfold F in *; clear F. generalize (or_clause_correct t a env). destruct (or_clause t a). + rewrite make_conj_cons. simpl. tauto. + simpl. tauto. } destruct t ; auto. - unfold eval_clause ; simpl. tauto. - unfold xor_clause_cnf. unfold F in H. rewrite H. unfold make_conj at 2. tauto. Qed. Lemma eval_cnf_cons : forall env a f, (~ make_conj (eval_tt env) a /\ eval_cnf env f) <-> eval_cnf env (a::f). Proof. intros. unfold eval_cnf in *. rewrite make_conj_cons ; eauto. unfold eval_clause at 2. tauto. Qed. Lemma eval_cnf_cons_iff : forall env a f, ((~ make_conj (eval_tt env) a) /\ eval_cnf env f) <-> eval_cnf env (a::f). Proof. intros. unfold eval_cnf in *. rewrite make_conj_cons ; eauto. unfold eval_clause. tauto. Qed. Lemma or_cnf_correct : forall env f f', eval_cnf env (or_cnf f f') <-> (eval_cnf env f) \/ (eval_cnf env f'). Proof. induction f. unfold eval_cnf. simpl. tauto. (**) intros. simpl. rewrite eval_cnf_app. rewrite <- eval_cnf_cons_iff. rewrite IHf. rewrite or_clause_cnf_correct. unfold eval_clause. tauto. Qed. Lemma or_cnf_opt_correct : forall env f f', eval_cnf env (or_cnf_opt f f') <-> eval_cnf env (or_cnf f f'). Proof. unfold or_cnf_opt. intros. destruct (is_cnf_tt f) eqn:TF. { simpl. apply is_cnf_tt_inv in TF. subst. rewrite or_cnf_correct. rewrite eval_cnf_tt. tauto. } destruct (is_cnf_tt f') eqn:TF'. { simpl. apply is_cnf_tt_inv in TF'. subst. rewrite or_cnf_correct. rewrite eval_cnf_tt. tauto. } { simpl. destruct (is_cnf_ff f') eqn:EQ. apply is_cnf_ff_inv in EQ. subst. rewrite or_cnf_correct. rewrite eval_cnf_ff. tauto. tauto. } Qed. Variable eval : Env -> Term -> Prop. Variable normalise_correct : forall env t tg, eval_cnf env (normalise t tg) -> eval env t. Variable negate_correct : forall env t tg, eval_cnf env (negate t tg) -> ~ eval env t. Lemma xcnf_correct : forall (f : @GFormula Term Prop Annot unit) pol env, eval_cnf env (xcnf pol f) -> eval_f (fun x => x) (eval env) (if pol then f else N f). Proof. induction f. - (* TT *) unfold eval_cnf. simpl. destruct pol ; simpl ; auto. - (* FF *) unfold eval_cnf. destruct pol; simpl ; auto. unfold eval_clause ; simpl. tauto. - (* P *) simpl. destruct pol ; intros ;simpl. unfold eval_cnf in H. (* Here I have to drop the proposition *) simpl in H. unfold eval_clause in H ; simpl in H. tauto. (* Here, I could store P in the clause *) unfold eval_cnf in H;simpl in H. unfold eval_clause in H ; simpl in H. tauto. - (* A *) simpl. destruct pol ; simpl. intros. eapply normalise_correct ; eauto. (* A 2 *) intros. eapply negate_correct ; eauto. - (* Cj *) destruct pol ; simpl. + (* pol = true *) intros. rewrite eval_cnf_and_opt in H. unfold and_cnf in H. rewrite eval_cnf_app in H. destruct H. split. apply (IHf1 _ _ H). apply (IHf2 _ _ H0). + (* pol = false *) intros. rewrite or_cnf_opt_correct in H. rewrite or_cnf_correct in H. destruct H as [H | H]. generalize (IHf1 false env H). simpl. tauto. generalize (IHf2 false env H). simpl. tauto. - (* D *) simpl. destruct pol. + (* pol = true *) intros. rewrite or_cnf_opt_correct in H. rewrite or_cnf_correct in H. destruct H as [H | H]. generalize (IHf1 _ env H). simpl. tauto. generalize (IHf2 _ env H). simpl. tauto. + (* pol = true *) intros. rewrite eval_cnf_and_opt in H. unfold and_cnf. rewrite eval_cnf_app in H. destruct H as [H0 H1]. simpl. generalize (IHf1 _ _ H0). generalize (IHf2 _ _ H1). simpl. tauto. - (**) simpl. destruct pol ; simpl. intros. apply (IHf false) ; auto. intros. generalize (IHf _ _ H). tauto. - (* I *) simpl; intros. destruct pol. + simpl. intro. rewrite or_cnf_opt_correct in H. rewrite or_cnf_correct in H. destruct H as [H | H]. generalize (IHf1 _ _ H). simpl in *. tauto. generalize (IHf2 _ _ H). auto. + (* pol = false *) rewrite eval_cnf_and_opt in H. unfold and_cnf in H. simpl in H. rewrite eval_cnf_app in H. destruct H as [H0 H1]. generalize (IHf1 _ _ H0). generalize (IHf2 _ _ H1). simpl. tauto. Qed. Variable Witness : Type. Variable checker : list (Term'*Annot) -> Witness -> bool. Variable checker_sound : forall t w, checker t w = true -> forall env, make_impl (eval_tt env) t False. Fixpoint cnf_checker (f : cnf) (l : list Witness) {struct f}: bool := match f with | nil => true | e::f => match l with | nil => false | c::l => match checker e c with | true => cnf_checker f l | _ => false end end end. Lemma cnf_checker_sound : forall t w, cnf_checker t w = true -> forall env, eval_cnf env t. Proof. unfold eval_cnf. induction t. (* bc *) simpl. auto. (* ic *) simpl. destruct w. intros ; discriminate. case_eq (checker a w) ; intros ; try discriminate. generalize (@checker_sound _ _ H env). generalize (IHt _ H0 env) ; intros. destruct t. red ; intro. rewrite <- make_conj_impl in H2. tauto. rewrite <- make_conj_impl in H2. tauto. Qed. Definition tauto_checker (f:@GFormula Term Prop Annot unit) (w:list Witness) : bool := cnf_checker (xcnf true f) w. Lemma tauto_checker_sound : forall t w, tauto_checker t w = true -> forall env, eval_f (fun x => x) (eval env) t. Proof. unfold tauto_checker. intros. change (eval_f (fun x => x) (eval env) t) with (eval_f (fun x => x) (eval env) (if true then t else TT)). apply (xcnf_correct t true). eapply cnf_checker_sound ; eauto. Qed. Definition eval_bf {A : Type} (ea : A -> Prop) (f: BFormula A) := eval_f (fun x => x) ea f. Lemma eval_bf_map : forall T U (fct: T-> U) env f , eval_bf env (map_bformula fct f) = eval_bf (fun x => env (fct x)) f. Proof. induction f ; simpl ; try (rewrite IHf1 ; rewrite IHf2) ; auto. rewrite <- IHf. auto. Qed. End S. (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/ZifyClasses.v0000644000175000017500000002102513612664533020340 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* S -> S]. Another limitation is that our injection theorems e.g. [TBOpInj], are using Leibniz equality; the payoff is that there is no need for morphisms... *) (** An injection [InjTyp S T] declares an injection from source type S to target type T. *) Class InjTyp (S : Type) (T : Type) := mkinj { (* [inj] is the injection function *) inj : S -> T; pred : T -> Prop; (* [cstr] states that [pred] holds for any injected element. [cstr (inj x)] is introduced in the goal for any leaf term of the form [inj x] *) cstr : forall x, pred (inj x) }. (** [BinOp Op] declares a source operator [Op: S1 -> S2 -> S3]. *) Class BinOp {S1 S2 S3:Type} {T:Type} (Op : S1 -> S2 -> S3) {I1 : InjTyp S1 T} {I2 : InjTyp S2 T} {I3 : InjTyp S3 T} := mkbop { (* [TBOp] is the target operator after injection of operands. *) TBOp : T -> T -> T; (* [TBOpInj] states the correctness of the injection. *) TBOpInj : forall (n:S1) (m:S2), inj (Op n m) = TBOp (inj n) (inj m) }. (** [Unop Op] declares a source operator [Op : S1 -> S2]. *) Class UnOp {S1 S2 T:Type} (Op : S1 -> S2) {I1 : InjTyp S1 T} {I2 : InjTyp S2 T} := mkuop { (* [TUOp] is the target operator after injection of operands. *) TUOp : T -> T; (* [TUOpInj] states the correctness of the injection. *) TUOpInj : forall (x:S1), inj (Op x) = TUOp (inj x) }. (** [CstOp Op] declares a source constant [Op : S]. *) Class CstOp {S T:Type} (Op : S) {I : InjTyp S T} := mkcst { (* [TCst] is the target constant. *) TCst : T; (* [TCstInj] states the correctness of the injection. *) TCstInj : inj Op = TCst }. (** In the framework, [Prop] is mapped to [Prop] and the injection is phrased in terms of [=] instead of [<->]. *) (** [BinRel R] declares the injection of a binary relation. *) Class BinRel {S:Type} {T:Type} (R : S -> S -> Prop) {I : InjTyp S T} := mkbrel { TR : T -> T -> Prop; TRInj : forall n m : S, R n m <-> TR (@inj _ _ I n) (inj m) }. (** [PropOp Op] declares morphisms for [<->]. This will be used to deal with e.g. [and], [or],... *) Class PropOp (Op : Prop -> Prop -> Prop) := mkprop { op_iff : forall (p1 p2 q1 q2:Prop), (p1 <-> q1) -> (p2 <-> q2) -> (Op p1 p2 <-> Op q1 q2) }. Class PropUOp (Op : Prop -> Prop) := mkuprop { uop_iff : forall (p1 q1 :Prop), (p1 <-> q1) -> (Op p1 <-> Op q1) }. (** Once the term is injected, terms can be replaced by their specification. NB1: The Ltac code is currently limited to (Op: Z -> Z -> Z) NB2: This is not sufficient to cope with [Z.div] or [Z.mod] *) Class BinOpSpec {S T: Type} (Op : T -> T -> T) {I : InjTyp S T} := mkbspec { BPred : T -> T -> T -> Prop; BSpec : forall x y, BPred x y (Op x y) }. Class UnOpSpec {S T: Type} (Op : T -> T) {I : InjTyp S T} := mkuspec { UPred : T -> T -> Prop; USpec : forall x, UPred x (Op x) }. (** After injections, e.g. nat -> Z, the fact that Z.of_nat x * Z.of_nat y is positive is lost. This information can be recovered using instance of the [Saturate] class. *) Class Saturate {T: Type} (Op : T -> T -> T) := mksat { (** Given [Op x y], - [PArg1] is the pre-condition of x - [PArg2] is the pre-condition of y - [PRes] is the pos-condition of (Op x y) *) PArg1 : T -> Prop; PArg2 : T -> Prop; PRes : T -> Prop; (** [SatOk] states the correctness of the reasoning *) SatOk : forall x y, PArg1 x -> PArg2 y -> PRes (Op x y) }. (* The [ZifyInst.saturate] iterates over all the instances and for every pattern of the form [H1 : PArg1 ?x , H2 : PArg2 ?y , T : context[Op ?x ?y] |- _ ] [H1 : PArg1 ?x , H2 : PArg2 ?y |- context[Op ?x ?y] ] asserts (SatOK x y H1 H2) *) (** The rest of the file is for internal use by the ML tactic. There are data-structures and lemmas used to inductively construct the injected terms. *) (** The data-structures [injterm] and [injected_prop] are used to store source and target expressions together with a correctness proof. *) Record injterm {S T: Type} {I : S -> T} := mkinjterm { source : S ; target : T ; inj_ok : I source = target}. Record injprop := mkinjprop { source_prop : Prop ; target_prop : Prop ; injprop_ok : source_prop <-> target_prop}. (** Lemmas for building [injterm] and [injprop]. *) Definition mkprop_op (Op : Prop -> Prop -> Prop) (POp : PropOp Op) (p1 :injprop) (p2: injprop) : injprop := {| source_prop := (Op (source_prop p1) (source_prop p2)) ; target_prop := (Op (target_prop p1) (target_prop p2)) ; injprop_ok := (op_iff (source_prop p1) (source_prop p2) (target_prop p1) (target_prop p2) (injprop_ok p1) (injprop_ok p2)) |}. Definition mkuprop_op (Op : Prop -> Prop) (POp : PropUOp Op) (p1 :injprop) : injprop := {| source_prop := (Op (source_prop p1)) ; target_prop := (Op (target_prop p1)) ; injprop_ok := (uop_iff (source_prop p1) (target_prop p1) (injprop_ok p1)) |}. Lemma mkapp2 (S1 S2 S3 T : Type) (Op : S1 -> S2 -> S3) {I1 : InjTyp S1 T} {I2 : InjTyp S2 T} {I3 : InjTyp S3 T} (B : @BinOp S1 S2 S3 T Op I1 I2 I3) (t1 : @injterm S1 T inj) (t2 : @injterm S2 T inj) : @injterm S3 T inj. Proof. apply (mkinjterm _ _ inj (Op (source t1) (source t2)) (TBOp (target t1) (target t2))). (rewrite <- inj_ok; rewrite <- inj_ok; apply TBOpInj). Defined. Lemma mkapp (S1 S2 T : Type) (Op : S1 -> S2) {I1 : InjTyp S1 T} {I2 : InjTyp S2 T} (B : @UnOp S1 S2 T Op I1 I2 ) (t1 : @injterm S1 T inj) : @injterm S2 T inj. Proof. apply (mkinjterm _ _ inj (Op (source t1)) (TUOp (target t1))). (rewrite <- inj_ok; apply TUOpInj). Defined. Lemma mkapp0 (S T : Type) (Op : S) {I : InjTyp S T} (B : @CstOp S T Op I) : @injterm S T inj. Proof. apply (mkinjterm _ _ inj Op TCst). (apply TCstInj). Defined. Lemma mkrel (S T : Type) (R : S -> S -> Prop) {Inj : InjTyp S T} (B : @BinRel S T R Inj) (t1 : @injterm S T inj) (t2 : @injterm S T inj) : @injprop. Proof. apply (mkinjprop (R (source t1) (source t2)) (TR (target t1) (target t2))). (rewrite <- inj_ok; rewrite <- inj_ok;apply TRInj). Defined. (** Registering constants for use by the plugin *) Register target_prop as ZifyClasses.target_prop. Register mkrel as ZifyClasses.mkrel. Register target as ZifyClasses.target. Register mkapp2 as ZifyClasses.mkapp2. Register mkapp as ZifyClasses.mkapp. Register mkapp0 as ZifyClasses.mkapp0. Register op_iff as ZifyClasses.op_iff. Register uop_iff as ZifyClasses.uop_iff. Register TR as ZifyClasses.TR. Register TBOp as ZifyClasses.TBOp. Register TUOp as ZifyClasses.TUOp. Register TCst as ZifyClasses.TCst. Register mkprop_op as ZifyClasses.mkprop_op. Register mkuprop_op as ZifyClasses.mkuprop_op. Register injprop_ok as ZifyClasses.injprop_ok. Register inj_ok as ZifyClasses.inj_ok. Register source as ZifyClasses.source. Register source_prop as ZifyClasses.source_prop. Register inj as ZifyClasses.inj. Register TRInj as ZifyClasses.TRInj. Register TUOpInj as ZifyClasses.TUOpInj. Register not as ZifyClasses.not. Register mkinjterm as ZifyClasses.mkinjterm. Register eq_refl as ZifyClasses.eq_refl. Register mkinjprop as ZifyClasses.mkinjprop. Register iff_refl as ZifyClasses.iff_refl. Register source_prop as ZifyClasses.source_prop. Register injprop_ok as ZifyClasses.injprop_ok. Register iff as ZifyClasses.iff. coq-8.11.0/plugins/micromega/LICENSE.sos0000644000175000017500000000305713612664533017527 0ustar treinentreinen HOL Light copyright notice, licence and disclaimer (c) University of Cambridge 1998 (c) Copyright, John Harrison 1998-2006 HOL Light version 2.20, hereinafter referred to as "the software", is a computer theorem proving system written by John Harrison. Much of the software was developed at the University of Cambridge Computer Laboratory, New Museums Site, Pembroke Street, Cambridge, CB2 3QG, England. The software is copyright, University of Cambridge 1998 and John Harrison 1998-2006. Permission to use, copy, modify, and distribute the software and its documentation for any purpose and without fee is hereby granted. In the case of further distribution of the software the present text, including copyright notice, licence and disclaimer of warranty, must be included in full and unmodified form in any release. Distribution of derivative software obtained by modifying the software, or incorporating it into other software, is permitted, provided the inclusion of the software is acknowledged and that any changes made to the software are clearly documented. John Harrison and the University of Cambridge disclaim all warranties with regard to the software, including all implied warranties of merchantability and fitness. In no event shall John Harrison or the University of Cambridge be liable for any special, indirect, incidental or consequential damages or any damages whatsoever, including, but not limited to, those arising from computer failure or malfunction, work stoppage, loss of profit or loss of contracts. coq-8.11.0/plugins/micromega/sos.ml0000644000175000017500000013106513612664533017056 0ustar treinentreinen(* ========================================================================= *) (* - This code originates from John Harrison's HOL LIGHT 2.30 *) (* (see file LICENSE.sos for license, copyright and disclaimer) *) (* - Laurent Théry (thery@sophia.inria.fr) has isolated the HOL *) (* independent bits *) (* - Frédéric Besson (fbesson@irisa.fr) is using it to feed micromega *) (* ========================================================================= *) (* ========================================================================= *) (* Nonlinear universal reals procedure using SOS decomposition. *) (* ========================================================================= *) open Num open Sos_types open Sos_lib (* prioritize_real();; *) let debugging = ref false exception Sanity (* ------------------------------------------------------------------------- *) (* Turn a rational into a decimal string with d sig digits. *) (* ------------------------------------------------------------------------- *) let decimalize = let rec normalize y = if abs_num y =/ Int 1 then normalize (y // Int 10) + 1 else 0 in fun d x -> if x =/ Int 0 then "0.0" else let y = abs_num x in let e = normalize y in let z = (pow10 (-e) */ y) +/ Int 1 in let k = round_num (pow10 d */ z) in (if x a | h :: t -> itern (k + 1) t f (f h k a) let rec iter (m, n) f a = if n < m then a else iter (m + 1, n) f (f m a) (* ------------------------------------------------------------------------- *) (* The main types. *) (* ------------------------------------------------------------------------- *) type vector = int * (int, num) func type matrix = (int * int) * (int * int, num) func type monomial = (vname, int) func type poly = (monomial, num) func (* ------------------------------------------------------------------------- *) (* Assignment avoiding zeros. *) (* ------------------------------------------------------------------------- *) let ( |--> ) x y a = if y =/ Int 0 then a else (x |-> y) a (* ------------------------------------------------------------------------- *) (* This can be generic. *) (* ------------------------------------------------------------------------- *) let element (d, v) i = tryapplyd v i (Int 0) let mapa f (d, v) = (d, foldl (fun a i c -> (i |--> f c) a) undefined v) let is_zero (d, v) = match v with Empty -> true | _ -> false (* ------------------------------------------------------------------------- *) (* Vectors. Conventionally indexed 1..n. *) (* ------------------------------------------------------------------------- *) let vector_0 n = ((n, undefined) : vector) let dim (v : vector) = fst v let vector_const c n = if c =/ Int 0 then vector_0 n else ((n, List.fold_right (fun k -> k |-> c) (1 -- n) undefined) : vector) let vector_cmul c (v : vector) = let n = dim v in if c =/ Int 0 then vector_0 n else (n, mapf (fun x -> c */ x) (snd v)) let vector_of_list l = let n = List.length l in ((n, List.fold_right2 ( |-> ) (1 -- n) l undefined) : vector) (* ------------------------------------------------------------------------- *) (* Matrices; again rows and columns indexed from 1. *) (* ------------------------------------------------------------------------- *) let matrix_0 (m, n) = (((m, n), undefined) : matrix) let dimensions (m : matrix) = fst m let matrix_cmul c (m : matrix) = let i, j = dimensions m in if c =/ Int 0 then matrix_0 (i, j) else ((i, j), mapf (fun x -> c */ x) (snd m)) let matrix_neg (m : matrix) = ((dimensions m, mapf minus_num (snd m)) : matrix) let matrix_add (m1 : matrix) (m2 : matrix) = let d1 = dimensions m1 and d2 = dimensions m2 in if d1 <> d2 then failwith "matrix_add: incompatible dimensions" else ((d1, combine ( +/ ) (fun x -> x =/ Int 0) (snd m1) (snd m2)) : matrix) let row k (m : matrix) = let i, j = dimensions m in ( ( j , foldl (fun a (i, j) c -> if i = k then (j |-> c) a else a) undefined (snd m) ) : vector ) let column k (m : matrix) = let i, j = dimensions m in ( ( i , foldl (fun a (i, j) c -> if j = k then (i |-> c) a else a) undefined (snd m) ) : vector ) let diagonal (v : vector) = let n = dim v in (((n, n), foldl (fun a i c -> ((i, i) |-> c) a) undefined (snd v)) : matrix) (* ------------------------------------------------------------------------- *) (* Monomials. *) (* ------------------------------------------------------------------------- *) let monomial_1 = (undefined : monomial) let monomial_var x = (x |=> 1 : monomial) let (monomial_mul : monomial -> monomial -> monomial) = combine ( + ) (fun x -> false) let monomial_degree x (m : monomial) = tryapplyd m x 0 let monomial_multidegree (m : monomial) = foldl (fun a x k -> k + a) 0 m let monomial_variables m = dom m (* ------------------------------------------------------------------------- *) (* Polynomials. *) (* ------------------------------------------------------------------------- *) let poly_0 = (undefined : poly) let poly_isconst (p : poly) = foldl (fun a m c -> m = monomial_1 && a) true p let poly_var x = (monomial_var x |=> Int 1 : poly) let poly_const c = if c =/ Int 0 then poly_0 else monomial_1 |=> c let poly_cmul c (p : poly) = if c =/ Int 0 then poly_0 else mapf (fun x -> c */ x) p let poly_neg (p : poly) = (mapf minus_num p : poly) let poly_add (p1 : poly) (p2 : poly) = (combine ( +/ ) (fun x -> x =/ Int 0) p1 p2 : poly) let poly_sub p1 p2 = poly_add p1 (poly_neg p2) let poly_cmmul (c, m) (p : poly) = if c =/ Int 0 then poly_0 else if m = monomial_1 then mapf (fun d -> c */ d) p else foldl (fun a m' d -> (monomial_mul m m' |-> c */ d) a) poly_0 p let poly_mul (p1 : poly) (p2 : poly) = foldl (fun a m c -> poly_add (poly_cmmul (c, m) p2) a) poly_0 p1 let poly_square p = poly_mul p p let rec poly_pow p k = if k = 0 then poly_const (Int 1) else if k = 1 then p else let q = poly_square (poly_pow p (k / 2)) in if k mod 2 = 1 then poly_mul p q else q let degree x (p : poly) = foldl (fun a m c -> max (monomial_degree x m) a) 0 p let multidegree (p : poly) = foldl (fun a m c -> max (monomial_multidegree m) a) 0 p let poly_variables (p : poly) = foldr (fun m c -> union (monomial_variables m)) p [] (* ------------------------------------------------------------------------- *) (* Order monomials for human presentation. *) (* ------------------------------------------------------------------------- *) let humanorder_varpow (x1, k1) (x2, k2) = x1 < x2 || (x1 = x2 && k1 > k2) let humanorder_monomial = let rec ord l1 l2 = match (l1, l2) with | _, [] -> true | [], _ -> false | h1 :: t1, h2 :: t2 -> humanorder_varpow h1 h2 || (h1 = h2 && ord t1 t2) in fun m1 m2 -> m1 = m2 || ord (sort humanorder_varpow (graph m1)) (sort humanorder_varpow (graph m2)) (* ------------------------------------------------------------------------- *) (* Conversions to strings. *) (* ------------------------------------------------------------------------- *) let string_of_vname (v : vname) : string = (v : string) let string_of_varpow x k = if k = 1 then string_of_vname x else string_of_vname x ^ "^" ^ string_of_int k let string_of_monomial m = if m = monomial_1 then "1" else let vps = List.fold_right (fun (x, k) a -> string_of_varpow x k :: a) (sort humanorder_varpow (graph m)) [] in String.concat "*" vps let string_of_cmonomial (c, m) = if m = monomial_1 then string_of_num c else if c =/ Int 1 then string_of_monomial m else string_of_num c ^ "*" ^ string_of_monomial m let string_of_poly (p : poly) = if p = poly_0 then "<<0>>" else let cms = sort (fun (m1, _) (m2, _) -> humanorder_monomial m1 m2) (graph p) in let s = List.fold_left (fun a (m, c) -> if c >" (* ------------------------------------------------------------------------- *) (* Printers. *) (* ------------------------------------------------------------------------- *) (* let print_vector v = Format.print_string(string_of_vector 0 20 v);; let print_matrix m = Format.print_string(string_of_matrix 20 m);; let print_monomial m = Format.print_string(string_of_monomial m);; let print_poly m = Format.print_string(string_of_poly m);; #install_printer print_vector;; #install_printer print_matrix;; #install_printer print_monomial;; #install_printer print_poly;; *) (* ------------------------------------------------------------------------- *) (* Conversion from term. *) (* ------------------------------------------------------------------------- *) let rec poly_of_term t = match t with | Zero -> poly_0 | Const n -> poly_const n | Var x -> poly_var x | Opp t1 -> poly_neg (poly_of_term t1) | Add (l, r) -> poly_add (poly_of_term l) (poly_of_term r) | Sub (l, r) -> poly_sub (poly_of_term l) (poly_of_term r) | Mul (l, r) -> poly_mul (poly_of_term l) (poly_of_term r) | Pow (t, n) -> poly_pow (poly_of_term t) n (* ------------------------------------------------------------------------- *) (* String of vector (just a list of space-separated numbers). *) (* ------------------------------------------------------------------------- *) let sdpa_of_vector (v : vector) = let n = dim v in let strs = List.map (o (decimalize 20) (element v)) (1 -- n) in String.concat " " strs ^ "\n" (* ------------------------------------------------------------------------- *) (* String for a matrix numbered k, in SDPA sparse format. *) (* ------------------------------------------------------------------------- *) let sdpa_of_matrix k (m : matrix) = let pfx = string_of_int k ^ " 1 " in let ms = foldr (fun (i, j) c a -> if i > j then a else ((i, j), c) :: a) (snd m) [] in let mss = sort (increasing fst) ms in List.fold_right (fun ((i, j), c) a -> pfx ^ string_of_int i ^ " " ^ string_of_int j ^ " " ^ decimalize 20 c ^ "\n" ^ a) mss "" (* ------------------------------------------------------------------------- *) (* String in SDPA sparse format for standard SDP problem: *) (* *) (* X = v_1 * [M_1] + ... + v_m * [M_m] - [M_0] must be PSD *) (* Minimize obj_1 * v_1 + ... obj_m * v_m *) (* ------------------------------------------------------------------------- *) let sdpa_of_problem comment obj mats = let m = List.length mats - 1 and n, _ = dimensions (List.hd mats) in "\"" ^ comment ^ "\"\n" ^ string_of_int m ^ "\n" ^ "1\n" ^ string_of_int n ^ "\n" ^ sdpa_of_vector obj ^ List.fold_right2 (fun k m a -> sdpa_of_matrix (k - 1) m ^ a) (1 -- List.length mats) mats "" (* ------------------------------------------------------------------------- *) (* More parser basics. *) (* ------------------------------------------------------------------------- *) let word s = end_itlist (fun p1 p2 -> p1 ++ p2 >> fun (s, t) -> s ^ t) (List.map a (explode s)) let token s = many (some isspace) ++ word s ++ many (some isspace) >> fun ((_, t), _) -> t let decimal = let ( || ) = parser_or in let numeral = some isnum in let decimalint = atleast 1 numeral >> o Num.num_of_string implode in let decimalfrac = atleast 1 numeral >> fun s -> Num.num_of_string (implode s) // pow10 (List.length s) in let decimalsig = decimalint ++ possibly (a "." ++ decimalfrac >> snd) >> function h, [x] -> h +/ x | h, _ -> h in let signed prs = a "-" ++ prs >> o minus_num snd || a "+" ++ prs >> snd || prs in let exponent = (a "e" || a "E") ++ signed decimalint >> snd in signed decimalsig ++ possibly exponent >> function h, [x] -> h */ power_num (Int 10) x | h, _ -> h let mkparser p s = let x, rst = p (explode s) in if rst = [] then x else failwith "mkparser: unparsed input" (* ------------------------------------------------------------------------- *) (* Parse back a vector. *) (* ------------------------------------------------------------------------- *) let _parse_sdpaoutput, parse_csdpoutput = let ( || ) = parser_or in let vector = token "{" ++ listof decimal (token ",") "decimal" ++ token "}" >> fun ((_, v), _) -> vector_of_list v in let rec skipupto dscr prs inp = (dscr ++ prs >> snd || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in let ignore inp = ((), []) in let sdpaoutput = skipupto (word "xVec" ++ token "=") (vector ++ ignore >> fst) in let csdpoutput = (decimal ++ many (a " " ++ decimal >> snd) >> fun (h, t) -> h :: t) ++ (a " " ++ a "\n" ++ ignore) >> o vector_of_list fst in (mkparser sdpaoutput, mkparser csdpoutput) (* ------------------------------------------------------------------------- *) (* The default parameters. Unfortunately this goes to a fixed file. *) (* ------------------------------------------------------------------------- *) let _sdpa_default_parameters = "100 unsigned int maxIteration;\n\ 1.0E-7 double 0.0 < epsilonStar;\n\ 1.0E2 double 0.0 < lambdaStar;\n\ 2.0 double 1.0 < omegaStar;\n\ -1.0E5 double lowerBound;\n\ 1.0E5 double upperBound;\n\ 0.1 double 0.0 <= betaStar < 1.0;\n\ 0.2 double 0.0 <= betaBar < 1.0, betaStar <= betaBar;\n\ 0.9 double 0.0 < gammaStar < 1.0;\n\ 1.0E-7 double 0.0 < epsilonDash;\n" (* ------------------------------------------------------------------------- *) (* These were suggested by Makoto Yamashita for problems where we are *) (* right at the edge of the semidefinite cone, as sometimes happens. *) (* ------------------------------------------------------------------------- *) let sdpa_alt_parameters = "1000 unsigned int maxIteration;\n\ 1.0E-7 double 0.0 < epsilonStar;\n\ 1.0E4 double 0.0 < lambdaStar;\n\ 2.0 double 1.0 < omegaStar;\n\ -1.0E5 double lowerBound;\n\ 1.0E5 double upperBound;\n\ 0.1 double 0.0 <= betaStar < 1.0;\n\ 0.2 double 0.0 <= betaBar < 1.0, betaStar <= betaBar;\n\ 0.9 double 0.0 < gammaStar < 1.0;\n\ 1.0E-7 double 0.0 < epsilonDash;\n" let _sdpa_params = sdpa_alt_parameters (* ------------------------------------------------------------------------- *) (* CSDP parameters; so far I'm sticking with the defaults. *) (* ------------------------------------------------------------------------- *) let csdp_default_parameters = "axtol=1.0e-8\n\ atytol=1.0e-8\n\ objtol=1.0e-8\n\ pinftol=1.0e8\n\ dinftol=1.0e8\n\ maxiter=100\n\ minstepfrac=0.9\n\ maxstepfrac=0.97\n\ minstepp=1.0e-8\n\ minstepd=1.0e-8\n\ usexzgap=1\n\ tweakgap=0\n\ affine=0\n\ printlevel=1\n" let csdp_params = csdp_default_parameters (* ------------------------------------------------------------------------- *) (* Now call CSDP on a problem and parse back the output. *) (* ------------------------------------------------------------------------- *) let run_csdp dbg obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_problem "" obj mats); file_of_string params_file csdp_params; let rv = Sys.command ( "cd " ^ temp_path ^ "; csdp " ^ input_file ^ " " ^ output_file ^ if dbg then "" else "> /dev/null" ) in let op = string_of_file output_file in let res = parse_csdpoutput op in if dbg then () else (Sys.remove input_file; Sys.remove output_file); (rv, res) (* ------------------------------------------------------------------------- *) (* Try some apparently sensible scaling first. Note that this is purely to *) (* get a cleaner translation to floating-point, and doesn't affect any of *) (* the results, in principle. In practice it seems a lot better when there *) (* are extreme numbers in the original problem. *) (* ------------------------------------------------------------------------- *) let scale_then = let common_denominator amat acc = foldl (fun a m c -> lcm_num (denominator c) a) acc amat and maximal_element amat acc = foldl (fun maxa m c -> max_num maxa (abs_num c)) acc amat in fun solver obj mats -> let cd1 = List.fold_right common_denominator mats (Int 1) and cd2 = common_denominator (snd obj) (Int 1) in let mats' = List.map (mapf (fun x -> cd1 */ x)) mats and obj' = vector_cmul cd2 obj in let max1 = List.fold_right maximal_element mats' (Int 0) and max2 = maximal_element (snd obj') (Int 0) in let scal1 = pow2 (20 - int_of_float (log (float_of_num max1) /. log 2.0)) and scal2 = pow2 (20 - int_of_float (log (float_of_num max2) /. log 2.0)) in let mats'' = List.map (mapf (fun x -> x */ scal1)) mats' and obj'' = vector_cmul scal2 obj' in solver obj'' mats'' (* ------------------------------------------------------------------------- *) (* Round a vector to "nice" rationals. *) (* ------------------------------------------------------------------------- *) let nice_rational n x = round_num (n */ x) // n let nice_vector n = mapa (nice_rational n) (* ------------------------------------------------------------------------- *) (* Reduce linear program to SDP (diagonal matrices) and test with CSDP. This *) (* one tests A [-1;x1;..;xn] >= 0 (i.e. left column is negated constants). *) (* ------------------------------------------------------------------------- *) let linear_program_basic a = let m, n = dimensions a in let mats = List.map (fun j -> diagonal (column j a)) (1 -- n) and obj = vector_const (Int 1) m in let rv, res = run_csdp false obj mats in if rv = 1 || rv = 2 then false else if rv = 0 then true else failwith "linear_program: An error occurred in the SDP solver" (* ------------------------------------------------------------------------- *) (* Test whether a point is in the convex hull of others. Rather than use *) (* computational geometry, express as linear inequalities and call CSDP. *) (* This is a bit lazy of me, but it's easy and not such a bottleneck so far. *) (* ------------------------------------------------------------------------- *) let in_convex_hull pts pt = let pts1 = (1 :: pt) :: List.map (fun x -> 1 :: x) pts in let pts2 = List.map (fun p -> List.map (fun x -> -x) p @ p) pts1 in let n = List.length pts + 1 and v = 2 * (List.length pt + 1) in let m = v + n - 1 in let mat = ( (m, n) , itern 1 pts2 (fun pts j -> itern 1 pts (fun x i -> (i, j) |-> Int x)) (iter (1, n) (fun i -> (v + i, i + 1) |-> Int 1) undefined) ) in linear_program_basic mat (* ------------------------------------------------------------------------- *) (* Filter down a set of points to a minimal set with the same convex hull. *) (* ------------------------------------------------------------------------- *) let minimal_convex_hull = let augment1 = function | [] -> assert false | m :: ms -> if in_convex_hull ms m then ms else ms @ [m] in let augment m ms = funpow 3 augment1 (m :: ms) in fun mons -> let mons' = List.fold_right augment (List.tl mons) [List.hd mons] in funpow (List.length mons') augment1 mons' (* ------------------------------------------------------------------------- *) (* Stuff for "equations" (generic A->num functions). *) (* ------------------------------------------------------------------------- *) let equation_cmul c eq = if c =/ Int 0 then Empty else mapf (fun d -> c */ d) eq let equation_add eq1 eq2 = combine ( +/ ) (fun x -> x =/ Int 0) eq1 eq2 let equation_eval assig eq = let value v = apply assig v in foldl (fun a v c -> a +/ (value v */ c)) (Int 0) eq (* ------------------------------------------------------------------------- *) (* Eliminate all variables, in an essentially arbitrary order. *) (* ------------------------------------------------------------------------- *) let eliminate_all_equations one = let choose_variable eq = let v, _ = choose eq in if v = one then let eq' = undefine v eq in if is_undefined eq' then failwith "choose_variable" else let w, _ = choose eq' in w else v in let rec eliminate dun eqs = match eqs with | [] -> dun | eq :: oeqs -> if is_undefined eq then eliminate dun oeqs else let v = choose_variable eq in let a = apply eq v in let eq' = equation_cmul (Int (-1) // a) (undefine v eq) in let elim e = let b = tryapplyd e v (Int 0) in if b =/ Int 0 then e else equation_add e (equation_cmul (minus_num b // a) eq) in eliminate ((v |-> eq') (mapf elim dun)) (List.map elim oeqs) in fun eqs -> let assig = eliminate undefined eqs in let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in (setify vs, assig) (* ------------------------------------------------------------------------- *) (* Hence produce the "relevant" monomials: those whose squares lie in the *) (* Newton polytope of the monomials in the input. (This is enough according *) (* to Reznik: "Extremal PSD forms with few terms", Duke Math. Journal, *) (* vol 45, pp. 363--374, 1978. *) (* *) (* These are ordered in sort of decreasing degree. In particular the *) (* constant monomial is last; this gives an order in diagonalization of the *) (* quadratic form that will tend to display constants. *) (* ------------------------------------------------------------------------- *) let newton_polytope pol = let vars = poly_variables pol in let mons = List.map (fun m -> List.map (fun x -> monomial_degree x m) vars) (dom pol) and ds = List.map (fun x -> (degree x pol + 1) / 2) vars in let all = List.fold_right (fun n -> allpairs (fun h t -> h :: t) (0 -- n)) ds [[]] and mons' = minimal_convex_hull mons in let all' = List.filter (fun m -> in_convex_hull mons' (List.map (fun x -> 2 * x) m)) all in List.map (fun m -> List.fold_right2 (fun v i a -> if i = 0 then a else (v |-> i) a) vars m monomial_1) (List.rev all') (* ------------------------------------------------------------------------- *) (* Diagonalize (Cholesky/LDU) the matrix corresponding to a quadratic form. *) (* ------------------------------------------------------------------------- *) let diag m = let nn = dimensions m in let n = fst nn in if snd nn <> n then failwith "diagonalize: non-square matrix" else let rec diagonalize i m = if is_zero m then [] else let a11 = element m (i, i) in if a11 a1k // a11) v in let m' = ( (n, n) , iter (i + 1, n) (fun j -> iter (i + 1, n) (fun k -> (j, k) |--> element m (j, k) -/ (element v j */ element v' k))) undefined ) in (a11, v') :: diagonalize (i + 1) m' in diagonalize 1 m (* ------------------------------------------------------------------------- *) (* Adjust a diagonalization to collect rationals at the start. *) (* ------------------------------------------------------------------------- *) let deration d = if d = [] then (Int 0, d) else let adj (c, l) = let a = foldl (fun a i c -> lcm_num a (denominator c)) (Int 1) (snd l) // foldl (fun a i c -> gcd_num a (numerator c)) (Int 0) (snd l) in (c // (a */ a), mapa (fun x -> a */ x) l) in let d' = List.map adj d in let a = List.fold_right (o lcm_num (o denominator fst)) d' (Int 1) // List.fold_right (o gcd_num (o numerator fst)) d' (Int 0) in (Int 1 // a, List.map (fun (c, l) -> (a */ c, l)) d') (* ------------------------------------------------------------------------- *) (* Enumeration of monomials with given multidegree bound. *) (* ------------------------------------------------------------------------- *) let rec enumerate_monomials d vars = if d < 0 then [] else if d = 0 then [undefined] else if vars = [] then [monomial_1] else let alts = List.map (fun k -> let oths = enumerate_monomials (d - k) (List.tl vars) in List.map (fun ks -> if k = 0 then ks else (List.hd vars |-> k) ks) oths) (0 -- d) in end_itlist ( @ ) alts (* ------------------------------------------------------------------------- *) (* Enumerate products of distinct input polys with degree <= d. *) (* We ignore any constant input polynomials. *) (* Give the output polynomial and a record of how it was derived. *) (* ------------------------------------------------------------------------- *) let rec enumerate_products d pols = if d = 0 then [(poly_const num_1, Rational_lt num_1)] else if d < 0 then [] else match pols with | [] -> [(poly_const num_1, Rational_lt num_1)] | (p, b) :: ps -> let e = multidegree p in if e = 0 then enumerate_products d ps else enumerate_products d ps @ List.map (fun (q, c) -> (poly_mul p q, Product (b, c))) (enumerate_products (d - e) ps) (* ------------------------------------------------------------------------- *) (* Multiply equation-parametrized poly by regular poly and add accumulator. *) (* ------------------------------------------------------------------------- *) let epoly_pmul p q acc = foldl (fun a m1 c -> foldl (fun b m2 e -> let m = monomial_mul m1 m2 in let es = tryapplyd b m undefined in (m |-> equation_add (equation_cmul c e) es) b) a q) acc p (* ------------------------------------------------------------------------- *) (* Convert regular polynomial. Note that we treat (0,0,0) as -1. *) (* ------------------------------------------------------------------------- *) let epoly_of_poly p = foldl (fun a m c -> (m |-> ((0, 0, 0) |=> minus_num c)) a) undefined p (* ------------------------------------------------------------------------- *) (* String for block diagonal matrix numbered k. *) (* ------------------------------------------------------------------------- *) let sdpa_of_blockdiagonal k m = let pfx = string_of_int k ^ " " in let ents = foldl (fun a (b, i, j) c -> if i > j then a else ((b, i, j), c) :: a) [] m in let entss = sort (increasing fst) ents in List.fold_right (fun ((b, i, j), c) a -> pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ " " ^ decimalize 20 c ^ "\n" ^ a) entss "" (* ------------------------------------------------------------------------- *) (* SDPA for problem using block diagonal (i.e. multiple SDPs) *) (* ------------------------------------------------------------------------- *) let sdpa_of_blockproblem comment nblocks blocksizes obj mats = let m = List.length mats - 1 in "\"" ^ comment ^ "\"\n" ^ string_of_int m ^ "\n" ^ string_of_int nblocks ^ "\n" ^ String.concat " " (List.map string_of_int blocksizes) ^ "\n" ^ sdpa_of_vector obj ^ List.fold_right2 (fun k m a -> sdpa_of_blockdiagonal (k - 1) m ^ a) (1 -- List.length mats) mats "" (* ------------------------------------------------------------------------- *) (* Hence run CSDP on a problem in block diagonal form. *) (* ------------------------------------------------------------------------- *) let run_csdp dbg nblocks blocksizes obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_blockproblem "" nblocks blocksizes obj mats); file_of_string params_file csdp_params; let rv = Sys.command ( "cd " ^ temp_path ^ "; csdp " ^ input_file ^ " " ^ output_file ^ if dbg then "" else "> /dev/null" ) in let op = string_of_file output_file in let res = parse_csdpoutput op in if dbg then () else (Sys.remove input_file; Sys.remove output_file); (rv, res) let csdp nblocks blocksizes obj mats = let rv, res = run_csdp !debugging nblocks blocksizes obj mats in if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (*Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline() *) else if rv <> 0 then failwith ("csdp: error " ^ string_of_int rv) else (); res (* ------------------------------------------------------------------------- *) (* 3D versions of matrix operations to consider blocks separately. *) (* ------------------------------------------------------------------------- *) let bmatrix_add = combine ( +/ ) (fun x -> x =/ Int 0) let bmatrix_cmul c bm = if c =/ Int 0 then undefined else mapf (fun x -> c */ x) bm let bmatrix_neg = bmatrix_cmul (Int (-1)) (* ------------------------------------------------------------------------- *) (* Smash a block matrix into components. *) (* ------------------------------------------------------------------------- *) let blocks blocksizes bm = List.map (fun (bs, b0) -> let m = foldl (fun a (b, i, j) c -> if b = b0 then ((i, j) |-> c) a else a) undefined bm in (((bs, bs), m) : matrix)) (List.combine blocksizes (1 -- List.length blocksizes)) (* ------------------------------------------------------------------------- *) (* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *) (* ------------------------------------------------------------------------- *) let real_positivnullstellensatz_general linf d eqs leqs pol = let vars = List.fold_right (o union poly_variables) ((pol :: eqs) @ List.map fst leqs) [] in let monoid = if linf then (poly_const num_1, Rational_lt num_1) :: List.filter (fun (p, c) -> multidegree p <= d) leqs else enumerate_products d leqs in let nblocks = List.length monoid in let mk_idmultiplier k p = let e = d - multidegree p in let mons = enumerate_monomials e vars in let nons = List.combine mons (1 -- List.length mons) in ( mons , List.fold_right (fun (m, n) -> m |-> ((-k, -n, n) |=> Int 1)) nons undefined ) in let mk_sqmultiplier k (p, c) = let e = (d - multidegree p) / 2 in let mons = enumerate_monomials e vars in let nons = List.combine mons (1 -- List.length mons) in ( mons , List.fold_right (fun (m1, n1) -> List.fold_right (fun (m2, n2) a -> let m = monomial_mul m1 m2 in if n1 > n2 then a else let c = if n1 = n2 then Int 1 else Int 2 in let e = tryapplyd a m undefined in (m |-> equation_add ((k, n1, n2) |=> c) e) a) nons) nons undefined ) in let sqmonlist, sqs = List.split (List.map2 mk_sqmultiplier (1 -- List.length monoid) monoid) and idmonlist, ids = List.split (List.map2 mk_idmultiplier (1 -- List.length eqs) eqs) in let blocksizes = List.map List.length sqmonlist in let bigsum = List.fold_right2 (fun p q a -> epoly_pmul p q a) eqs ids (List.fold_right2 (fun (p, c) s a -> epoly_pmul p s a) monoid sqs (epoly_of_poly (poly_neg pol))) in let eqns = foldl (fun a m e -> e :: a) [] bigsum in let pvs, assig = eliminate_all_equations (0, 0, 0) eqns in let qvars = (0, 0, 0) :: pvs in let allassig = List.fold_right (fun v -> v |-> (v |=> Int 1)) pvs assig in let mk_matrix v = foldl (fun m (b, i, j) ass -> if b < 0 then m else let c = tryapplyd ass v (Int 0) in if c =/ Int 0 then m else ((b, j, i) |-> c) (((b, i, j) |-> c) m)) undefined allassig in let diagents = foldl (fun a (b, i, j) e -> if b > 0 && i = j then equation_add e a else a) undefined allassig in let mats = List.map mk_matrix qvars and obj = ( List.length pvs , itern 1 pvs (fun v i -> i |--> tryapplyd diagents v (Int 0)) undefined ) in let raw_vec = if pvs = [] then vector_0 0 else scale_then (csdp nblocks blocksizes) obj mats in let find_rounding d = if !debugging then ( Format.print_string ("Trying rounding with limit " ^ string_of_num d); Format.print_newline () ) else (); let vec = nice_vector d raw_vec in let blockmat = iter (1, dim vec) (fun i a -> bmatrix_add (bmatrix_cmul (element vec i) (List.nth mats i)) a) (bmatrix_neg (List.nth mats 0)) in let allmats = blocks blocksizes blockmat in (vec, List.map diag allmats) in let vec, ratdias = if pvs = [] then find_rounding num_1 else tryfind find_rounding (List.map Num.num_of_int (1 -- 31) @ List.map pow2 (5 -- 66)) in let newassigs = List.fold_right (fun k -> List.nth pvs (k - 1) |-> element vec k) (1 -- dim vec) ((0, 0, 0) |=> Int (-1)) in let finalassigs = foldl (fun a v e -> (v |-> equation_eval newassigs e) a) newassigs allassig in let poly_of_epoly p = foldl (fun a v e -> (v |--> equation_eval finalassigs e) a) undefined p in let mk_sos mons = let mk_sq (c, m) = ( c , List.fold_right (fun k a -> (List.nth mons (k - 1) |--> element m k) a) (1 -- List.length mons) undefined ) in List.map mk_sq in let sqs = List.map2 mk_sos sqmonlist ratdias and cfs = List.map poly_of_epoly ids in let msq = List.filter (fun (a, b) -> b <> []) (List.map2 (fun a b -> (a, b)) monoid sqs) in let eval_sq sqs = List.fold_right (fun (c, q) -> poly_add (poly_cmul c (poly_mul q q))) sqs poly_0 in let sanity = List.fold_right (fun ((p, c), s) -> poly_add (poly_mul p (eval_sq s))) msq (List.fold_right2 (fun p q -> poly_add (poly_mul p q)) cfs eqs (poly_neg pol)) in if not (is_undefined sanity) then raise Sanity else (cfs, List.map (fun (a, b) -> (snd a, b)) msq) (* ------------------------------------------------------------------------- *) (* The ordering so we can create canonical HOL polynomials. *) (* ------------------------------------------------------------------------- *) let dest_monomial mon = sort (increasing fst) (graph mon) let monomial_order = let rec lexorder l1 l2 = match (l1, l2) with | [], [] -> true | vps, [] -> false | [], vps -> true | (x1, n1) :: vs1, (x2, n2) :: vs2 -> if x1 < x2 then true else if x2 < x1 then false else if n1 < n2 then false else if n2 < n1 then true else lexorder vs1 vs2 in fun m1 m2 -> if m2 = monomial_1 then true else if m1 = monomial_1 then false else let mon1 = dest_monomial m1 and mon2 = dest_monomial m2 in let deg1 = List.fold_right (o ( + ) snd) mon1 0 and deg2 = List.fold_right (o ( + ) snd) mon2 0 in if deg1 < deg2 then false else if deg1 > deg2 then true else lexorder mon1 mon2 (* ------------------------------------------------------------------------- *) (* Map back polynomials and their composites to HOL. *) (* ------------------------------------------------------------------------- *) let term_of_varpow x k = if k = 1 then Var x else Pow (Var x, k) let term_of_monomial m = if m = monomial_1 then Const num_1 else let m' = dest_monomial m in let vps = List.fold_right (fun (x, k) a -> term_of_varpow x k :: a) m' [] in end_itlist (fun s t -> Mul (s, t)) vps let term_of_cmonomial (m, c) = if m = monomial_1 then Const c else if c =/ num_1 then term_of_monomial m else Mul (Const c, term_of_monomial m) let term_of_poly p = if p = poly_0 then Zero else let cms = List.map term_of_cmonomial (sort (fun (m1, _) (m2, _) -> monomial_order m1 m2) (graph p)) in end_itlist (fun t1 t2 -> Add (t1, t2)) cms let term_of_sqterm (c, p) = Product (Rational_lt c, Square (term_of_poly p)) let term_of_sos (pr, sqs) = if sqs = [] then pr else Product (pr, end_itlist (fun a b -> Sum (a, b)) (List.map term_of_sqterm sqs)) (* ------------------------------------------------------------------------- *) (* Some combinatorial helper functions. *) (* ------------------------------------------------------------------------- *) let rec allpermutations l = if l = [] then [[]] else List.fold_right (fun h acc -> List.map (fun t -> h :: t) (allpermutations (subtract l [h])) @ acc) l [] let changevariables_monomial zoln (m : monomial) = foldl (fun a x k -> (List.assoc x zoln |-> k) a) monomial_1 m let changevariables zoln pol = foldl (fun a m c -> (changevariables_monomial zoln m |-> c) a) poly_0 pol (* ------------------------------------------------------------------------- *) (* Return to original non-block matrices. *) (* ------------------------------------------------------------------------- *) let sdpa_of_vector (v : vector) = let n = dim v in let strs = List.map (o (decimalize 20) (element v)) (1 -- n) in String.concat " " strs ^ "\n" let sdpa_of_matrix k (m : matrix) = let pfx = string_of_int k ^ " 1 " in let ms = foldr (fun (i, j) c a -> if i > j then a else ((i, j), c) :: a) (snd m) [] in let mss = sort (increasing fst) ms in List.fold_right (fun ((i, j), c) a -> pfx ^ string_of_int i ^ " " ^ string_of_int j ^ " " ^ decimalize 20 c ^ "\n" ^ a) mss "" let sdpa_of_problem comment obj mats = let m = List.length mats - 1 and n, _ = dimensions (List.hd mats) in "\"" ^ comment ^ "\"\n" ^ string_of_int m ^ "\n" ^ "1\n" ^ string_of_int n ^ "\n" ^ sdpa_of_vector obj ^ List.fold_right2 (fun k m a -> sdpa_of_matrix (k - 1) m ^ a) (1 -- List.length mats) mats "" let run_csdp dbg obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_problem "" obj mats); file_of_string params_file csdp_params; let rv = Sys.command ( "cd " ^ temp_path ^ "; csdp " ^ input_file ^ " " ^ output_file ^ if dbg then "" else "> /dev/null" ) in let op = string_of_file output_file in let res = parse_csdpoutput op in if dbg then () else (Sys.remove input_file; Sys.remove output_file); (rv, res) let csdp obj mats = let rv, res = run_csdp !debugging obj mats in if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (* (Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline()) *) else if rv <> 0 then failwith ("csdp: error " ^ string_of_int rv) else (); res (* ------------------------------------------------------------------------- *) (* Sum-of-squares function with some lowbrow symmetry reductions. *) (* ------------------------------------------------------------------------- *) let sumofsquares_general_symmetry tool pol = let vars = poly_variables pol and lpps = newton_polytope pol in let n = List.length lpps in let sym_eqs = let invariants = List.filter (fun vars' -> is_undefined (poly_sub pol (changevariables (List.combine vars vars') pol))) (allpermutations vars) in let lpns = List.combine lpps (1 -- List.length lpps) in let lppcs = List.filter (fun (m, (n1, n2)) -> n1 <= n2) (allpairs (fun (m1, n1) (m2, n2) -> ((m1, m2), (n1, n2))) lpns lpns) in let clppcs = end_itlist ( @ ) (List.map (fun ((m1, m2), (n1, n2)) -> List.map (fun vars' -> ( ( changevariables_monomial (List.combine vars vars') m1 , changevariables_monomial (List.combine vars vars') m2 ) , (n1, n2) )) invariants) lppcs) in let clppcs_dom = setify (List.map fst clppcs) in let clppcs_cls = List.map (fun d -> List.filter (fun (e, _) -> e = d) clppcs) clppcs_dom in let eqvcls = List.map (o setify (List.map snd)) clppcs_cls in let mk_eq cls acc = match cls with | [] -> raise Sanity | [h] -> acc | h :: t -> List.map (fun k -> (k |-> Int (-1)) (h |=> Int 1)) t @ acc in List.fold_right mk_eq eqvcls [] in let eqs = foldl (fun a x y -> y :: a) [] (itern 1 lpps (fun m1 n1 -> itern 1 lpps (fun m2 n2 f -> let m = monomial_mul m1 m2 in if n1 > n2 then f else let c = if n1 = n2 then Int 1 else Int 2 in (m |-> ((n1, n2) |-> c) (tryapplyd f m undefined)) f)) (foldl (fun a m c -> (m |-> ((0, 0) |=> c)) a) undefined pol)) @ sym_eqs in let pvs, assig = eliminate_all_equations (0, 0) eqs in let allassig = List.fold_right (fun v -> v |-> (v |=> Int 1)) pvs assig in let qvars = (0, 0) :: pvs in let diagents = end_itlist equation_add (List.map (fun i -> apply allassig (i, i)) (1 -- n)) in let mk_matrix v = ( ( (n, n) , foldl (fun m (i, j) ass -> let c = tryapplyd ass v (Int 0) in if c =/ Int 0 then m else ((j, i) |-> c) (((i, j) |-> c) m)) undefined allassig ) : matrix ) in let mats = List.map mk_matrix qvars and obj = ( List.length pvs , itern 1 pvs (fun v i -> i |--> tryapplyd diagents v (Int 0)) undefined ) in let raw_vec = if pvs = [] then vector_0 0 else tool obj mats in let find_rounding d = if !debugging then ( Format.print_string ("Trying rounding with limit " ^ string_of_num d); Format.print_newline () ) else (); let vec = nice_vector d raw_vec in let mat = iter (1, dim vec) (fun i a -> matrix_add (matrix_cmul (element vec i) (List.nth mats i)) a) (matrix_neg (List.nth mats 0)) in deration (diag mat) in let rat, dia = if pvs = [] then let mat = matrix_neg (List.nth mats 0) in deration (diag mat) else tryfind find_rounding (List.map Num.num_of_int (1 -- 31) @ List.map pow2 (5 -- 66)) in let poly_of_lin (d, v) = (d, foldl (fun a i c -> (List.nth lpps (i - 1) |-> c) a) undefined (snd v)) in let lins = List.map poly_of_lin dia in let sqs = List.map (fun (d, l) -> poly_mul (poly_const d) (poly_pow l 2)) lins in let sos = poly_cmul rat (end_itlist poly_add sqs) in if is_undefined (poly_sub sos pol) then (rat, lins) else raise Sanity let sumofsquares = sumofsquares_general_symmetry csdp coq-8.11.0/plugins/micromega/mutils.ml0000644000175000017500000002753213612664533017572 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int -> int = compare let equal : int -> int -> bool = ( = ) end module ISet = struct include Set.Make (Int) let pp o s = iter (fun i -> Printf.fprintf o "%i " i) s end module IMap = struct include Map.Make (Int) let from k m = let _, _, r = split (k - 1) m in r end let rec pp_list s f o l = match l with | [] -> () | [e] -> f o e | e :: l -> f o e; output_string o s; pp_list s f o l let finally f rst = try let res = f () in rst (); res with reraise -> (try rst () with any -> raise reraise); raise reraise let rec try_any l x = match l with | [] -> None | (f, s) :: l -> ( match f x with None -> try_any l x | x -> x ) let all_pairs f l = let pair_with acc e l = List.fold_left (fun acc x -> f e x :: acc) acc l in let rec xpairs acc l = match l with [] -> acc | e :: lx -> xpairs (pair_with acc e l) lx in xpairs [] l let rec is_sublist f l1 l2 = match (l1, l2) with | [], _ -> true | e :: l1', [] -> false | e :: l1', e' :: l2' -> if f e e' then is_sublist f l1' l2' else is_sublist f l1 l2' let extract pred l = List.fold_left (fun (fd, sys) e -> match fd with | None -> ( match pred e with None -> (fd, e :: sys) | Some v -> (Some (v, e), sys) ) | _ -> (fd, e :: sys)) (None, []) l let extract_best red lt l = let rec extractb c e rst l = match l with | [] -> (Some (c, e), rst) | e' :: l' -> ( match red e' with | None -> extractb c e (e' :: rst) l' | Some c' -> if lt c' c then extractb c' e' (e :: rst) l' else extractb c e (e' :: rst) l' ) in match extract red l with | None, _ -> (None, l) | Some (c, e), rst -> extractb c e [] rst let rec find_option pred l = match l with | [] -> raise Not_found | e :: l -> ( match pred e with Some r -> r | None -> find_option pred l ) let find_some pred l = try Some (find_option pred l) with Not_found -> None let extract_all pred l = List.fold_left (fun (s1, s2) e -> match pred e with None -> (s1, e :: s2) | Some v -> ((v, e) :: s1, s2)) ([], []) l let simplify f sys = let sys', b = List.fold_left (fun (sys', b) c -> match f c with None -> (c :: sys', b) | Some c' -> (c' :: sys', true)) ([], false) sys in if b then Some sys' else None let generate_acc f acc sys = List.fold_left (fun sys' c -> match f c with None -> sys' | Some c' -> c' :: sys') acc sys let generate f sys = generate_acc f [] sys let saturate p f sys = let rec sat acc l = match extract p l with | None, _ -> acc | Some r, l' -> let n = generate (f r) (l' @ acc) in sat (n @ acc) l' in try sat [] sys with x -> Printexc.print_backtrace stdout; raise x open Num open Big_int let ppcm x y = let g = gcd_big_int x y in let x' = div_big_int x g in let y' = div_big_int y g in mult_big_int g (mult_big_int x' y') let denominator = function | Int _ | Big_int _ -> unit_big_int | Ratio r -> Ratio.denominator_ratio r let numerator = function | Ratio r -> Ratio.numerator_ratio r | Int i -> Big_int.big_int_of_int i | Big_int i -> i let iterate_until_stable f x = let rec iter x = match f x with None -> x | Some x' -> iter x' in iter x let rec app_funs l x = match l with | [] -> None | f :: fl -> ( match f x with None -> app_funs fl x | Some x' -> Some x' ) (** * MODULE: Coq to Caml data-structure mappings *) module CoqToCaml = struct open Micromega let rec nat = function O -> 0 | S n -> nat n + 1 let rec positive p = match p with | XH -> 1 | XI p -> 1 + (2 * positive p) | XO p -> 2 * positive p let n nt = match nt with N0 -> 0 | Npos p -> positive p let rec index i = (* Swap left-right ? *) match i with XH -> 1 | XI i -> 1 + (2 * index i) | XO i -> 2 * index i open Big_int let rec positive_big_int p = match p with | XH -> unit_big_int | XI p -> add_int_big_int 1 (mult_int_big_int 2 (positive_big_int p)) | XO p -> mult_int_big_int 2 (positive_big_int p) let z_big_int x = match x with | Z0 -> zero_big_int | Zpos p -> positive_big_int p | Zneg p -> minus_big_int (positive_big_int p) let z x = match x with Z0 -> 0 | Zpos p -> index p | Zneg p -> -index p let q_to_num {qnum = x; qden = y} = Big_int (z_big_int x) // Big_int (z_big_int (Zpos y)) end (** * MODULE: Caml to Coq data-structure mappings *) module CamlToCoq = struct open Micromega let rec nat = function 0 -> O | n -> S (nat (n - 1)) let rec positive n = if Int.equal n 1 then XH else if Int.equal (n land 1) 1 then XI (positive (n lsr 1)) else XO (positive (n lsr 1)) let n nt = if nt < 0 then assert false else if Int.equal nt 0 then N0 else Npos (positive nt) let rec index n = if Int.equal n 1 then XH else if Int.equal (n land 1) 1 then XI (index (n lsr 1)) else XO (index (n lsr 1)) let z x = match compare x 0 with | 0 -> Z0 | 1 -> Zpos (positive x) | _ -> (* this should be -1 *) Zneg (positive (-x)) open Big_int let positive_big_int n = let two = big_int_of_int 2 in let rec _pos n = if eq_big_int n unit_big_int then XH else let q, m = quomod_big_int n two in if eq_big_int unit_big_int m then XI (_pos q) else XO (_pos q) in _pos n let bigint x = match sign_big_int x with | 0 -> Z0 | 1 -> Zpos (positive_big_int x) | _ -> Zneg (positive_big_int (minus_big_int x)) let q n = { Micromega.qnum = bigint (numerator n) ; Micromega.qden = positive_big_int (denominator n) } end (** * MODULE: Comparisons on lists: by evaluating the elements in a single list, * between two lists given an ordering, and using a hash computation *) module Cmp = struct let rec compare_lexical l = match l with | [] -> 0 (* Equal *) | f :: l -> let cmp = f () in if Int.equal cmp 0 then compare_lexical l else cmp let rec compare_list cmp l1 l2 = match (l1, l2) with | [], [] -> 0 | [], _ -> -1 | _, [] -> 1 | e1 :: l1, e2 :: l2 -> let c = cmp e1 e2 in if Int.equal c 0 then compare_list cmp l1 l2 else c end (** * MODULE: Labels for atoms in propositional formulas. * Tags are used to identify unused atoms in CNFs, and propagate them back to * the original formula. The translation back to Coq then ignores these * superfluous items, which speeds the translation up a bit. *) module type Tag = sig type t = int val from : int -> t val next : t -> t val pp : out_channel -> t -> unit val compare : t -> t -> int val max : t -> t -> t val to_int : t -> int end module Tag : Tag = struct type t = int let from i = i let next i = i + 1 let max : int -> int -> int = max let pp o i = output_string o (string_of_int i) let compare : int -> int -> int = Int.compare let to_int x = x end (** * MODULE: Ordered sets of tags. *) module TagSet = struct include Set.Make (Tag) end (** As for Unix.close_process, our Unix.waipid will ignore all EINTR *) let rec waitpid_non_intr pid = try snd (Unix.waitpid [] pid) with Unix.Unix_error (Unix.EINTR, _, _) -> waitpid_non_intr pid (** * Forking routine, plumbing the appropriate pipes where needed. *) let command exe_path args vl = (* creating pipes for stdin, stdout, stderr *) let stdin_read, stdin_write = Unix.pipe () and stdout_read, stdout_write = Unix.pipe () and stderr_read, stderr_write = Unix.pipe () in (* Create the process *) let pid = Unix.create_process exe_path args stdin_read stdout_write stderr_write in (* Write the data on the stdin of the created process *) let outch = Unix.out_channel_of_descr stdin_write in output_value outch vl; flush outch; (* Wait for its completion *) let status = waitpid_non_intr pid in finally (* Recover the result *) (fun () -> match status with | Unix.WEXITED 0 -> ( let inch = Unix.in_channel_of_descr stdout_read in try Marshal.from_channel inch with any -> failwith (Printf.sprintf "command \"%s\" exited %s" exe_path (Printexc.to_string any)) ) | Unix.WEXITED i -> failwith (Printf.sprintf "command \"%s\" exited %i" exe_path i) | Unix.WSIGNALED i -> failwith (Printf.sprintf "command \"%s\" killed %i" exe_path i) | Unix.WSTOPPED i -> failwith (Printf.sprintf "command \"%s\" stopped %i" exe_path i)) (* Cleanup *) (fun () -> List.iter (fun x -> try Unix.close x with any -> ()) [ stdin_read ; stdin_write ; stdout_read ; stdout_write ; stderr_read ; stderr_write ]) (** Hashing utilities *) module Hash = struct module Mc = Micromega open Hashset.Combine let int_of_eq_op1 = Mc.(function Equal -> 0 | NonEqual -> 1 | Strict -> 2 | NonStrict -> 3) let eq_op1 o1 o2 = int_of_eq_op1 o1 = int_of_eq_op1 o2 let hash_op1 h o = combine h (int_of_eq_op1 o) let rec eq_positive p1 p2 = match (p1, p2) with | Mc.XH, Mc.XH -> true | Mc.XI p1, Mc.XI p2 -> eq_positive p1 p2 | Mc.XO p1, Mc.XO p2 -> eq_positive p1 p2 | _, _ -> false let eq_z z1 z2 = match (z1, z2) with | Mc.Z0, Mc.Z0 -> true | Mc.Zpos p1, Mc.Zpos p2 | Mc.Zneg p1, Mc.Zneg p2 -> eq_positive p1 p2 | _, _ -> false let eq_q {Mc.qnum = qn1; Mc.qden = qd1} {Mc.qnum = qn2; Mc.qden = qd2} = eq_z qn1 qn2 && eq_positive qd1 qd2 let rec eq_pol eq p1 p2 = match (p1, p2) with | Mc.Pc c1, Mc.Pc c2 -> eq c1 c2 | Mc.Pinj (i1, p1), Mc.Pinj (i2, p2) -> eq_positive i1 i2 && eq_pol eq p1 p2 | Mc.PX (p1, i1, p1'), Mc.PX (p2, i2, p2') -> eq_pol eq p1 p2 && eq_positive i1 i2 && eq_pol eq p1' p2' | _, _ -> false let eq_pair eq1 eq2 (x1, y1) (x2, y2) = eq1 x1 x2 && eq2 y1 y2 let hash_pol helt = let rec hash acc = function | Mc.Pc c -> helt (combine acc 1) c | Mc.Pinj (p, c) -> hash (combine (combine acc 1) (CoqToCaml.index p)) c | Mc.PX (p1, i, p2) -> hash (hash (combine (combine acc 2) (CoqToCaml.index i)) p1) p2 in hash let hash_pair h1 h2 h (e1, e2) = h2 (h1 h e1) e2 let hash_elt f h e = combine h (f e) let hash_string h (e : string) = hash_elt Hashtbl.hash h e let hash_z = hash_elt CoqToCaml.z let hash_q = hash_elt (fun q -> Hashtbl.hash (CoqToCaml.q_to_num q)) end (* Local Variables: *) (* coding: utf-8 *) (* End: *) coq-8.11.0/plugins/micromega/Zify.v0000644000175000017500000000613413612664533017026 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let u := eval compute in (t a) in change (t a) with u in * | _ => zify_unop_var_or_term t thm a end. Ltac zify_unop_nored t thm a := (* in this version, we don't try to reduce the unop (that can be (Z.add x)) *) let isz := isZcst a in match isz with | true => zify_unop_core t thm a | _ => zify_unop_var_or_term t thm a end. Ltac zify_binop t thm a b:= (* works as zify_unop, except that we should be careful when dealing with b, since it can be equal to a *) let isza := isZcst a in match isza with | true => zify_unop (t a) (thm a) b | _ => let za := fresh "z" in (rename a into za; rename za into a; zify_unop_nored (t a) (thm a) b) || (remember a as za; match goal with | H : za = b |- _ => zify_unop_nored (t za) (thm za) za | _ => zify_unop_nored (t za) (thm za) b end) end. (* end from PreOmega *) Ltac applySpec S := let t := type of S in match t with | @BinOpSpec _ _ ?Op _ => let Spec := (eval unfold S, BSpec in (@BSpec _ _ Op _ S)) in repeat match goal with | H : context[Op ?X ?Y] |- _ => zify_binop Op Spec X Y | |- context[Op ?X ?Y] => zify_binop Op Spec X Y end | @UnOpSpec _ _ ?Op _ => let Spec := (eval unfold S, USpec in (@USpec _ _ Op _ S)) in repeat match goal with | H : context[Op ?X] |- _ => zify_unop Op Spec X | |- context[Op ?X ] => zify_unop Op Spec X end end. (** [zify_post_hook] is there to be redefined. *) Ltac zify_post_hook := idtac. Ltac zify := zify_op ; (iter_specs applySpec) ; zify_post_hook. coq-8.11.0/plugins/micromega/g_zify.mlg0000644000175000017500000000403513612664533017704 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Zify.InjTable.register t } | ["Add" "BinOp" constr(t) ] -> { Zify.BinOp.register t } | ["Add" "UnOp" constr(t) ] -> { Zify.UnOp.register t } | ["Add" "CstOp" constr(t) ] -> { Zify.CstOp.register t } | ["Add" "BinRel" constr(t) ] -> { Zify.BinRel.register t } | ["Add" "PropOp" constr(t) ] -> { Zify.PropOp.register t } | ["Add" "PropUOp" constr(t) ] -> { Zify.PropUnOp.register t } | ["Add" "Spec" constr(t) ] -> { Zify.Spec.register t } | ["Add" "BinOpSpec" constr(t) ] -> { Zify.Spec.register t } | ["Add" "UnOpSpec" constr(t) ] -> { Zify.Spec.register t } | ["Add" "Saturate" constr(t) ] -> { Zify.Saturate.register t } END TACTIC EXTEND ITER | [ "iter_specs" tactic(t)] -> { Zify.iter_specs t } END TACTIC EXTEND TRANS | [ "zify_op" ] -> { Zify.zify_tac } | [ "saturate" ] -> { Zify.saturate } END VERNAC COMMAND EXTEND ZifyPrint CLASSIFIED AS SIDEFF |[ "Show" "Zify" "InjTyp" ] -> { Zify.InjTable.print () } |[ "Show" "Zify" "BinOp" ] -> { Zify.BinOp.print () } |[ "Show" "Zify" "UnOp" ] -> { Zify.UnOp.print () } |[ "Show" "Zify" "CstOp"] -> { Zify.CstOp.print () } |[ "Show" "Zify" "BinRel"] -> { Zify.BinRel.print () } |[ "Show" "Zify" "Spec"] -> { Zify.Spec.print () } END coq-8.11.0/plugins/micromega/plugin_base.dune0000644000175000017500000000106013612664533021054 0ustar treinentreinen(library (name micromega_plugin) (public_name coq.plugins.micromega) ; be careful not to link the executable to the plugin! (modules (:standard \ csdpcert g_zify zify)) (synopsis "Coq's micromega plugin") (libraries num coq.plugins.ltac)) (executable (name csdpcert) (public_name csdpcert) (package coq) (modules csdpcert) (flags :standard -open Micromega_plugin) (libraries coq.plugins.micromega)) (library (name zify_plugin) (public_name coq.plugins.zify) (modules g_zify zify) (synopsis "Coq's zify plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/ltac/0000755000175000017500000000000013612664533014672 5ustar treinentreinencoq-8.11.0/plugins/ltac/tacarg.ml0000644000175000017500000000353013612664533016466 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (** Possible arguments of a tactic definition *) type 'a gen_tactic_arg = | TacGeneric of 'lev generic_argument | ConstrMayEval of ('trm,'cst,'pat) may_eval | Reference of 'ref | TacCall of ('ref * 'a gen_tactic_arg list) CAst.t | TacFreshId of string or_var list | Tacexp of 'tacexpr | TacPretype of 'trm | TacNumgoals constraint 'a = < term:'trm; dterm: 'dtrm; pattern:'pat; constant:'cst; reference:'ref; name:'nam; tacexpr:'tacexpr; level:'lev > (** Generic ltac expressions. 't : terms, 'p : patterns, 'c : constants, 'i : inductive, 'r : ltac refs, 'n : idents, 'l : levels *) and 'a gen_tactic_expr = | TacAtom of ('a gen_atomic_tactic_expr) CAst.t | TacThen of 'a gen_tactic_expr * 'a gen_tactic_expr | TacDispatch of 'a gen_tactic_expr list | TacExtendTac of 'a gen_tactic_expr array * 'a gen_tactic_expr * 'a gen_tactic_expr array | TacThens of 'a gen_tactic_expr * 'a gen_tactic_expr list | TacThens3parts of 'a gen_tactic_expr * 'a gen_tactic_expr array * 'a gen_tactic_expr * 'a gen_tactic_expr array | TacFirst of 'a gen_tactic_expr list | TacComplete of 'a gen_tactic_expr | TacSolve of 'a gen_tactic_expr list | TacTry of 'a gen_tactic_expr | TacOr of 'a gen_tactic_expr * 'a gen_tactic_expr | TacOnce of 'a gen_tactic_expr | TacExactlyOnce of 'a gen_tactic_expr | TacIfThenCatch of 'a gen_tactic_expr * 'a gen_tactic_expr * 'a gen_tactic_expr | TacOrelse of 'a gen_tactic_expr * 'a gen_tactic_expr | TacDo of int or_var * 'a gen_tactic_expr | TacTimeout of int or_var * 'a gen_tactic_expr | TacTime of string option * 'a gen_tactic_expr | TacRepeat of 'a gen_tactic_expr | TacProgress of 'a gen_tactic_expr | TacShowHyps of 'a gen_tactic_expr | TacAbstract of 'a gen_tactic_expr * Id.t option | TacId of 'n message_token list | TacFail of global_flag * int or_var * 'n message_token list | TacInfo of 'a gen_tactic_expr | TacLetIn of rec_flag * (lname * 'a gen_tactic_arg) list * 'a gen_tactic_expr | TacMatch of lazy_flag * 'a gen_tactic_expr * ('p,'a gen_tactic_expr) match_rule list | TacMatchGoal of lazy_flag * direction_flag * ('p,'a gen_tactic_expr) match_rule list | TacFun of 'a gen_tactic_fun_ast | TacArg of 'a gen_tactic_arg CAst.t | TacSelect of Goal_select.t * 'a gen_tactic_expr (* For ML extensions *) | TacML of (ml_tactic_entry * 'a gen_tactic_arg list) CAst.t (* For syntax extensions *) | TacAlias of (KerName.t * 'a gen_tactic_arg list) CAst.t constraint 'a = < term:'t; dterm: 'dtrm; pattern:'p; constant:'c; reference:'r; name:'n; tacexpr:'tacexpr; level:'l > and 'a gen_tactic_fun_ast = Name.t list * 'a gen_tactic_expr constraint 'a = < term:'t; dterm: 'dtrm; pattern:'p; constant:'c; reference:'r; name:'n; tacexpr:'te; level:'l > (** Globalized tactics *) type g_trm = Genintern.glob_constr_and_expr type g_pat = Genintern.glob_constr_pattern_and_expr type g_cst = evaluable_global_reference Genredexpr.and_short_name or_var type g_ref = ltac_constant located or_var type g_nam = lident type g_dispatch = < term:g_trm; dterm:g_trm; pattern:g_pat; constant:g_cst; reference:g_ref; name:g_nam; tacexpr:glob_tactic_expr; level:glevel > and glob_tactic_expr = g_dispatch gen_tactic_expr type glob_atomic_tactic_expr = g_dispatch gen_atomic_tactic_expr type glob_tactic_arg = g_dispatch gen_tactic_arg (** Raw tactics *) type r_ref = qualid type r_nam = lident type r_lev = rlevel type r_dispatch = < term:r_trm; dterm:r_trm; pattern:r_pat; constant:r_cst; reference:r_ref; name:r_nam; tacexpr:raw_tactic_expr; level:rlevel > and raw_tactic_expr = r_dispatch gen_tactic_expr type raw_atomic_tactic_expr = r_dispatch gen_atomic_tactic_expr type raw_tactic_arg = r_dispatch gen_tactic_arg (** Interpreted tactics *) type t_trm = EConstr.constr type t_pat = constr_pattern type t_cst = evaluable_global_reference type t_ref = ltac_constant located type t_nam = Id.t type t_dispatch = < term:t_trm; dterm:g_trm; pattern:t_pat; constant:t_cst; reference:t_ref; name:t_nam; tacexpr:unit; level:tlevel > type atomic_tactic_expr = t_dispatch gen_atomic_tactic_expr (** Misc *) type raw_red_expr = (r_trm, r_cst, r_pat) red_expr_gen type glob_red_expr = (g_trm, g_cst, g_pat) red_expr_gen (** Traces *) type ltac_call_kind = | LtacMLCall of glob_tactic_expr | LtacNotationCall of KerName.t | LtacNameCall of ltac_constant | LtacAtomCall of glob_atomic_tactic_expr | LtacVarCall of Id.t * glob_tactic_expr | LtacConstrInterp of Glob_term.glob_constr * Ltac_pretype.ltac_var_map type ltac_trace = ltac_call_kind Loc.located list type tacdef_body = | TacticDefinition of lident * raw_tactic_expr (* indicates that user employed ':=' in Ltac body *) | TacticRedefinition of qualid * raw_tactic_expr (* indicates that user employed '::=' in Ltac body *) coq-8.11.0/plugins/ltac/ltac_plugin.mlpack0000644000175000017500000000043613612664533020367 0ustar treinentreinenTacexpr Tacarg Tacsubst Tacenv Pptactic Pltac Taccoerce Tactic_debug Tacintern Profile_ltac Tactic_matching Tacinterp Tacentries Evar_tactics Tactic_option Extraargs G_obligations Coretactics Extratactics Profile_ltac_tactics G_auto G_class Rewrite G_rewrite G_eqdecide G_tactic G_ltac coq-8.11.0/plugins/ltac/tactic_option.ml0000644000175000017500000000415113612664533020064 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* obj = declare_object { (default_object name) with cache_function = cache; load_function = (fun _ -> load); open_function = (fun _ -> load); classify_function = (fun (local, tac) -> if local then Dispose else Substitute (local, tac)); subst_function = subst} in let put local tac = set_default_tactic local tac; Lib.add_anonymous_leaf (input (local, tac)) in let get () = !locality, Tacinterp.eval_tactic !default_tactic in let print () = Pptactic.pr_glob_tactic (Global.env ()) !default_tactic_expr ++ (if !locality then str" (locally defined)" else str" (globally defined)") in put, get, print coq-8.11.0/plugins/ltac/extraargs.mli0000644000175000017500000000502413612664533017376 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Pp.t val wit_rename : (Id.t * Id.t) Genarg.uniform_genarg_type val occurrences : (int list Locus.or_var) Pcoq.Entry.t val wit_occurrences : (int list Locus.or_var, int list Locus.or_var, int list) Genarg.genarg_type val pr_occurrences : int list Locus.or_var -> Pp.t val occurrences_of : int list -> Locus.occurrences val wit_natural : int Genarg.uniform_genarg_type val wit_glob : (constr_expr, glob_constr_and_expr, Tacinterp.interp_sign * glob_constr) Genarg.genarg_type val wit_lglob : (constr_expr, glob_constr_and_expr, Tacinterp.interp_sign * glob_constr) Genarg.genarg_type val wit_lconstr : (constr_expr, glob_constr_and_expr, EConstr.t) Genarg.genarg_type val wit_casted_constr : (constr_expr, glob_constr_and_expr, EConstr.t) Genarg.genarg_type val glob : constr_expr Pcoq.Entry.t val lglob : constr_expr Pcoq.Entry.t type 'id gen_place= ('id * Locus.hyp_location_flag,unit) location type loc_place = lident gen_place type place = Id.t gen_place val wit_hloc : (loc_place, loc_place, place) Genarg.genarg_type val hloc : loc_place Pcoq.Entry.t val pr_hloc : loc_place -> Pp.t val by_arg_tac : Tacexpr.raw_tactic_expr option Pcoq.Entry.t val wit_by_arg_tac : (raw_tactic_expr option, glob_tactic_expr option, Geninterp.Val.t option) Genarg.genarg_type val pr_by_arg_tac : Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> int * Notation_gram.parenRelation -> raw_tactic_expr -> Pp.t) -> raw_tactic_expr option -> Pp.t val test_lpar_id_colon : unit Pcoq.Entry.t val wit_test_lpar_id_colon : (unit, unit, unit) Genarg.genarg_type val wit_in_clause : (lident Locus.clause_expr, lident Locus.clause_expr, Id.t Locus.clause_expr) Genarg.genarg_type coq-8.11.0/plugins/ltac/g_auto.mlg0000644000175000017500000001641713612664533016662 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Eauto.e_assumption } END TACTIC EXTEND eexact | [ "eexact" constr(c) ] -> { Eauto.e_give_exact c } END { let pr_hintbases _prc _prlc _prt = Pptactic.pr_hintbases } ARGUMENT EXTEND hintbases TYPED AS preident list option PRINTED BY { pr_hintbases } | [ "with" "*" ] -> { None } | [ "with" ne_preident_list(l) ] -> { Some l } | [ ] -> { Some [] } END { let eval_uconstrs ist cs = let flags = { Pretyping.use_typeclasses = false; solve_unification_constraints = true; fail_evar = false; expand_evars = true; program_mode = false; polymorphic = false; } in let map c env sigma = c env sigma in List.map (fun c -> map (Tacinterp.type_uconstr ~flags ist c)) cs let pr_auto_using_raw env sigma _ _ _ = Pptactic.pr_auto_using @@ Ppconstr.pr_constr_expr env sigma let pr_auto_using_glob env sigma _ _ _ = Pptactic.pr_auto_using (fun (c,_) -> Printer.pr_glob_constr_env env c) let pr_auto_using env sigma _ _ _ = Pptactic.pr_auto_using @@ Printer.pr_closed_glob_env env sigma } ARGUMENT EXTEND auto_using TYPED AS uconstr list PRINTED BY { pr_auto_using env sigma } RAW_PRINTED BY { pr_auto_using_raw env sigma } GLOB_PRINTED BY { pr_auto_using_glob env sigma } | [ "using" ne_uconstr_list_sep(l, ",") ] -> { l } | [ ] -> { [] } END (** Auto *) TACTIC EXTEND trivial | [ "trivial" auto_using(lems) hintbases(db) ] -> { Auto.h_trivial (eval_uconstrs ist lems) db } END TACTIC EXTEND info_trivial | [ "info_trivial" auto_using(lems) hintbases(db) ] -> { Auto.h_trivial ~debug:Info (eval_uconstrs ist lems) db } END TACTIC EXTEND debug_trivial | [ "debug" "trivial" auto_using(lems) hintbases(db) ] -> { Auto.h_trivial ~debug:Debug (eval_uconstrs ist lems) db } END TACTIC EXTEND auto | [ "auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> { Auto.h_auto n (eval_uconstrs ist lems) db } END TACTIC EXTEND info_auto | [ "info_auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> { Auto.h_auto ~debug:Info n (eval_uconstrs ist lems) db } END TACTIC EXTEND debug_auto | [ "debug" "auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> { Auto.h_auto ~debug:Debug n (eval_uconstrs ist lems) db } END (** Eauto *) TACTIC EXTEND prolog | [ "prolog" "[" uconstr_list(l) "]" int_or_var(n) ] -> { Eauto.prolog_tac (eval_uconstrs ist l) n } END { let make_depth n = snd (Eauto.make_dimension n None) } TACTIC EXTEND eauto | [ "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) hintbases(db) ] -> { Eauto.gen_eauto (Eauto.make_dimension n p) (eval_uconstrs ist lems) db } END TACTIC EXTEND new_eauto | [ "new" "auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> { match db with | None -> Auto.new_full_auto (make_depth n) (eval_uconstrs ist lems) | Some l -> Auto.new_auto (make_depth n) (eval_uconstrs ist lems) l } END TACTIC EXTEND debug_eauto | [ "debug" "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) hintbases(db) ] -> { Eauto.gen_eauto ~debug:Debug (Eauto.make_dimension n p) (eval_uconstrs ist lems) db } END TACTIC EXTEND info_eauto | [ "info_eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) hintbases(db) ] -> { Eauto.gen_eauto ~debug:Info (Eauto.make_dimension n p) (eval_uconstrs ist lems) db } END TACTIC EXTEND dfs_eauto | [ "dfs" "eauto" int_or_var_opt(p) auto_using(lems) hintbases(db) ] -> { Eauto.gen_eauto (Eauto.make_dimension p None) (eval_uconstrs ist lems) db } END TACTIC EXTEND autounfold | [ "autounfold" hintbases(db) clause_dft_concl(cl) ] -> { Eauto.autounfold_tac db cl } END TACTIC EXTEND autounfold_one | [ "autounfold_one" hintbases(db) "in" hyp(id) ] -> { Eauto.autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) (Some (id, Locus.InHyp)) } | [ "autounfold_one" hintbases(db) ] -> { Eauto.autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) None } END TACTIC EXTEND unify | ["unify" constr(x) constr(y) ] -> { Tactics.unify x y } | ["unify" constr(x) constr(y) "with" preident(base) ] -> { let table = try Some (Hints.searchtable_map base) with Not_found -> None in match table with | None -> let msg = str "Hint table " ++ str base ++ str " not found" in Tacticals.New.tclZEROMSG msg | Some t -> let state = Hints.Hint_db.transparent_state t in Tactics.unify ~state x y } END { let deprecated_convert_concl_no_check = CWarnings.create ~name:"convert_concl_no_check" ~category:"deprecated" (fun () -> Pp.str "The syntax [convert_concl_no_check] is deprecated. Use [change_no_check] instead.") } TACTIC EXTEND convert_concl_no_check | ["convert_concl_no_check" constr(x) ] -> { deprecated_convert_concl_no_check (); Tactics.convert_concl ~check:false x DEFAULTcast } END { let pr_pre_hints_path_atom _ _ _ = Hints.pp_hints_path_atom Libnames.pr_qualid let pr_hints_path_atom _ _ _ = Hints.pp_hints_path_atom Printer.pr_global let glob_hints_path_atom ist = Hints.glob_hints_path_atom } ARGUMENT EXTEND hints_path_atom PRINTED BY { pr_hints_path_atom } GLOBALIZED BY { glob_hints_path_atom } RAW_PRINTED BY { pr_pre_hints_path_atom } GLOB_PRINTED BY { pr_hints_path_atom } | [ ne_global_list(g) ] -> { Hints.PathHints g } | [ "_" ] -> { Hints.PathAny } END { let pr_hints_path prc prx pry c = Hints.pp_hints_path c let pr_pre_hints_path prc prx pry c = Hints.pp_hints_path_gen Libnames.pr_qualid c let glob_hints_path ist = Hints.glob_hints_path } ARGUMENT EXTEND hints_path PRINTED BY { pr_hints_path } GLOBALIZED BY { glob_hints_path } RAW_PRINTED BY { pr_pre_hints_path } GLOB_PRINTED BY { pr_hints_path } | [ "(" hints_path(p) ")" ] -> { p } | [ hints_path(p) "*" ] -> { Hints.PathStar p } | [ "emp" ] -> { Hints.PathEmpty } | [ "eps" ] -> { Hints.PathEpsilon } | [ hints_path(p) "|" hints_path(q) ] -> { Hints.PathOr (p, q) } | [ hints_path_atom(a) ] -> { Hints.PathAtom a } | [ hints_path(p) hints_path(q) ] -> { Hints.PathSeq (p, q) } END ARGUMENT EXTEND opthints TYPED AS preident list option PRINTED BY { pr_hintbases } | [ ":" ne_preident_list(l) ] -> { Some l } | [ ] -> { None } END VERNAC COMMAND EXTEND HintCut CLASSIFIED AS SIDEFF | #[ locality = Attributes.locality; ] [ "Hint" "Cut" "[" hints_path(p) "]" opthints(dbnames) ] -> { let entry = Hints.HintsCutEntry (Hints.glob_hints_path p) in Hints.add_hints ~local:(Locality.make_section_locality locality) (match dbnames with None -> ["core"] | Some l -> l) entry; } END coq-8.11.0/plugins/ltac/tactic_debug.ml0000644000175000017500000003527413612664533017654 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Printer.pr_constr_pattern_env env sigma p) rl (* This module intends to be a beginning of debugger for tactic expressions. Currently, it is quite simple and we can hope to have, in the future, a more complete panel of commands dedicated to a proof assistant framework *) (* Debug information *) type debug_info = | DebugOn of int | DebugOff (* An exception handler *) let explain_logic_error e = CErrors.print e let explain_logic_error_no_anomaly e = CErrors.print_no_report e let msg_tac_debug s = Proofview.NonLogical.print_debug (s++fnl()) let msg_tac_notice s = Proofview.NonLogical.print_notice (s++fnl()) (* Prints the goal *) let db_pr_goal gl = let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in let penv = Termops.Internal.print_named_context env in let pc = Printer.pr_econstr_env env (Tacmach.New.project gl) concl in str" " ++ hv 0 (penv ++ fnl () ++ str "============================" ++ fnl () ++ str" " ++ pc) ++ fnl () let db_pr_goal = Proofview.Goal.enter begin fun gl -> let pg = db_pr_goal gl in Proofview.tclLIFT (msg_tac_notice (str "Goal:" ++ fnl () ++ pg)) end (* Prints the commands *) let help () = msg_tac_debug (str "Commands: = Continue" ++ fnl() ++ str " h/? = Help" ++ fnl() ++ str " r = Run times" ++ fnl() ++ str " r = Run up to next idtac " ++ fnl() ++ str " s = Skip" ++ fnl() ++ str " x = Exit") (* Prints the goal and the command to be executed *) let goal_com tac = Proofview.tclTHEN db_pr_goal (Proofview.tclLIFT (msg_tac_debug (str "Going to execute:" ++ fnl () ++ prtac tac))) (* [run (new_ref _)] gives us a ref shared among [NonLogical.t] expressions. It avoids parametrizing everything over a reference. *) let skipped = Proofview.NonLogical.run (Proofview.NonLogical.ref 0) let skip = Proofview.NonLogical.run (Proofview.NonLogical.ref 0) let breakpoint = Proofview.NonLogical.run (Proofview.NonLogical.ref None) let batch = ref false open Goptions let () = declare_bool_option { optdepr = false; optname = "Ltac batch debug"; optkey = ["Ltac";"Batch";"Debug"]; optread = (fun () -> !batch); optwrite = (fun x -> batch := x) } let rec drop_spaces inst i = if String.length inst > i && inst.[i] == ' ' then drop_spaces inst (i+1) else i let possibly_unquote s = if String.length s >= 2 && s.[0] == '"' && s.[String.length s - 1] == '"' then String.sub s 1 (String.length s - 2) else s (* (Re-)initialize debugger *) let db_initialize = let open Proofview.NonLogical in (skip:=0) >> (skipped:=0) >> (breakpoint:=None) let int_of_string s = try Proofview.NonLogical.return (int_of_string s) with e -> Proofview.NonLogical.raise e let string_get s i = try Proofview.NonLogical.return (String.get s i) with e -> Proofview.NonLogical.raise e let run_invalid_arg () = Proofview.NonLogical.raise (Invalid_argument "run_com") (* Gives the number of steps or next breakpoint of a run command *) let run_com inst = let open Proofview.NonLogical in string_get inst 0 >>= fun first_char -> if first_char ='r' then let i = drop_spaces inst 1 in if String.length inst > i then let s = String.sub inst i (String.length inst - i) in if inst.[0] >= '0' && inst.[0] <= '9' then int_of_string s >>= fun num -> (if num<0 then run_invalid_arg () else return ()) >> (skip:=num) >> (skipped:=0) else breakpoint:=Some (possibly_unquote s) else run_invalid_arg () else run_invalid_arg () (* Prints the run counter *) let run ini = let open Proofview.NonLogical in if not ini then begin Proofview.NonLogical.print_notice (str"\b\r\b\r") >> !skipped >>= fun skipped -> msg_tac_debug (str "Executed expressions: " ++ int skipped ++ fnl()) end >> !skipped >>= fun x -> skipped := x+1 else return () (* Prints the prompt *) let rec prompt level = (* spiwack: avoid overriding by the open below *) let runtrue = run true in begin let open Proofview.NonLogical in Proofview.NonLogical.print_notice (fnl () ++ str "TcDebug (" ++ int level ++ str ") > ") >> if Util.(!batch) then return (DebugOn (level+1)) else let exit = (skip:=0) >> (skipped:=0) >> raise Sys.Break in Proofview.NonLogical.catch Proofview.NonLogical.read_line begin function (e, info) -> match e with | End_of_file -> exit | e -> raise ~info e end >>= fun inst -> match inst with | "" -> return (DebugOn (level+1)) | "s" -> return (DebugOff) | "x" -> Proofview.NonLogical.print_char '\b' >> exit | "h"| "?" -> begin help () >> prompt level end | _ -> Proofview.NonLogical.catch (run_com inst >> runtrue >> return (DebugOn (level+1))) begin function (e, info) -> match e with | Failure _ | Invalid_argument _ -> prompt level | e -> raise ~info e end end (* Prints the state and waits for an instruction *) (* spiwack: the only reason why we need to take the continuation [f] as an argument rather than returning the new level directly seems to be that [f] is wrapped in with "explain_logic_error". I don't think it serves any purpose in the current design, so we could just drop that. *) let debug_prompt lev tac f = (* spiwack: avoid overriding by the open below *) let runfalse = run false in let open Proofview.NonLogical in let (>=) = Proofview.tclBIND in (* What to print and to do next *) let newlevel = Proofview.tclLIFT !skip >= fun initial_skip -> if Int.equal initial_skip 0 then Proofview.tclLIFT !breakpoint >= fun breakpoint -> if Option.is_empty breakpoint then Proofview.tclTHEN (goal_com tac) (Proofview.tclLIFT (prompt lev)) else Proofview.tclLIFT(runfalse >> return (DebugOn (lev+1))) else Proofview.tclLIFT begin (!skip >>= fun s -> skip:=s-1) >> runfalse >> !skip >>= fun new_skip -> (if Int.equal new_skip 0 then skipped:=0 else return ()) >> return (DebugOn (lev+1)) end in newlevel >= fun newlevel -> (* What to execute *) Proofview.tclOR (f newlevel) begin fun (reraise, info) -> Proofview.tclTHEN (Proofview.tclLIFT begin (skip:=0) >> (skipped:=0) >> if Logic.catchable_exception reraise then msg_tac_debug (str "Level " ++ int lev ++ str ": " ++ explain_logic_error reraise) else return () end) (Proofview.tclZERO ~info reraise) end let is_debug db = let open Proofview.NonLogical in !breakpoint >>= fun breakpoint -> match db, breakpoint with | DebugOff, _ -> return false | _, Some _ -> return false | _ -> !skip >>= fun skip -> return (Int.equal skip 0) (* Prints a constr *) let db_constr debug env sigma c = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then msg_tac_debug (str "Evaluated term: " ++ Printer.pr_econstr_env env sigma c) else return () (* Prints the pattern rule *) let db_pattern_rule debug env sigma num r = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then begin msg_tac_debug (str "Pattern rule " ++ int num ++ str ":" ++ fnl () ++ str "|" ++ spc () ++ prmatchrl env sigma r) end else return () (* Prints the hypothesis pattern identifier if it exists *) let hyp_bound = function | Anonymous -> str " (unbound)" | Name id -> str " (bound to " ++ Id.print id ++ str ")" (* Prints a matched hypothesis *) let db_matched_hyp debug env sigma (id,_,c) ido = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then msg_tac_debug (str "Hypothesis " ++ Id.print id ++ hyp_bound ido ++ str " has been matched: " ++ Printer.pr_econstr_env env sigma c) else return () (* Prints the matched conclusion *) let db_matched_concl debug env sigma c = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then msg_tac_debug (str "Conclusion has been matched: " ++ Printer.pr_econstr_env env sigma c) else return () (* Prints a success message when the goal has been matched *) let db_mc_pattern_success debug = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then msg_tac_debug (str "The goal has been successfully matched!" ++ fnl() ++ str "Let us execute the right-hand side part..." ++ fnl()) else return () (* Prints a failure message for an hypothesis pattern *) let db_hyp_pattern_failure debug env sigma (na,hyp) = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then msg_tac_debug (str "The pattern hypothesis" ++ hyp_bound na ++ str " cannot match: " ++ prmatchpatt env sigma hyp) else return () (* Prints a matching failure message for a rule *) let db_matching_failure debug = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then msg_tac_debug (str "This rule has failed due to matching errors!" ++ fnl() ++ str "Let us try the next one...") else return () (* Prints an evaluation failure message for a rule *) let db_eval_failure debug s = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then let s = str "message \"" ++ s ++ str "\"" in msg_tac_debug (str "This rule has failed due to \"Fail\" tactic (" ++ s ++ str ", level 0)!" ++ fnl() ++ str "Let us try the next one...") else return () (* Prints a logic failure message for a rule *) let db_logic_failure debug err = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then begin msg_tac_debug (explain_logic_error err) >> msg_tac_debug (str "This rule has failed due to a logic error!" ++ fnl() ++ str "Let us try the next one...") end else return () let is_breakpoint brkname s = match brkname, s with | Some s, MsgString s'::_ -> String.equal s s' | _ -> false let db_breakpoint debug s = let open Proofview.NonLogical in !breakpoint >>= fun opt_breakpoint -> match debug with | DebugOn lev when not (CList.is_empty s) && is_breakpoint opt_breakpoint s -> breakpoint:=None | _ -> return () (** Extrating traces *) let is_defined_ltac trace = let rec aux = function | (_, Tacexpr.LtacNameCall f) :: _ -> not (Tacenv.is_ltac_for_ml_tactic f) | (_, Tacexpr.LtacNotationCall f) :: _ -> true | (_, Tacexpr.LtacAtomCall _) :: _ -> false | _ :: tail -> aux tail | [] -> false in aux (List.rev trace) let explain_ltac_call_trace last trace loc = let calls = last :: List.rev_map snd trace in let pr_call ck = match ck with | Tacexpr.LtacNotationCall kn -> quote (Pptactic.pr_alias_key kn) | Tacexpr.LtacNameCall cst -> quote (Pptactic.pr_ltac_constant cst) | Tacexpr.LtacMLCall t -> quote (Pptactic.pr_glob_tactic (Global.env()) t) | Tacexpr.LtacVarCall (id,t) -> quote (Id.print id) ++ strbrk " (bound to " ++ Pptactic.pr_glob_tactic (Global.env()) t ++ str ")" | Tacexpr.LtacAtomCall te -> quote (Pptactic.pr_glob_tactic (Global.env()) (Tacexpr.TacAtom (CAst.make te))) | Tacexpr.LtacConstrInterp (c, { Ltac_pretype.ltac_constrs = vars }) -> quote (Printer.pr_glob_constr_env (Global.env()) c) ++ (if not (Id.Map.is_empty vars) then strbrk " (with " ++ prlist_with_sep pr_comma (fun (id,c) -> (* XXX: This hooks into the CErrors's additional error info API so it is tricky to provide the right env for now. *) let env = Global.env () in let sigma = Evd.from_env env in Id.print id ++ str ":=" ++ Printer.pr_lconstr_under_binders_env env sigma c) (List.rev (Id.Map.bindings vars)) ++ str ")" else mt()) in match calls with | [] -> mt () | [a] -> hov 0 (str "Ltac call to " ++ pr_call a ++ str " failed.") | _ -> let kind_of_last_call = match List.last calls with | Tacexpr.LtacConstrInterp _ -> ", last term evaluation failed." | _ -> ", last call failed." in hov 0 (str "In nested Ltac calls to " ++ pr_enum pr_call calls ++ strbrk kind_of_last_call) let skip_extensions trace = let rec aux = function | (_,Tacexpr.LtacNotationCall _ as tac) :: (_,Tacexpr.LtacMLCall _) :: tail -> (* Case of an ML defined tactic with entry of the form <<"foo" args>> *) (* see tacextend.mlp *) tac :: aux tail | t :: tail -> t :: aux tail | [] -> [] in List.rev (aux (List.rev trace)) let extract_ltac_trace ?loc trace = let trace = skip_extensions trace in let (tloc,c),tail = List.sep_last trace in if is_defined_ltac trace then (* We entered a user-defined tactic, we display the trace with location of the call *) let msg = hov 0 (explain_ltac_call_trace c tail loc ++ fnl()) in (if Loc.finer loc tloc then loc else tloc), Some msg else (* We entered a primitive tactic, we don't display trace but report on the finest location *) let best_loc = (* trace is with innermost call coming first *) let rec aux best_loc = function | (loc,_)::tail -> if Option.is_empty best_loc || not (Option.is_empty loc) && Loc.finer loc best_loc then aux loc tail else aux best_loc tail | [] -> best_loc in aux loc trace in best_loc, None let get_ltac_trace info = let ltac_trace = Exninfo.get info ltac_trace_info in let loc = Loc.get_loc info in match ltac_trace with | None -> None | Some trace -> Some (extract_ltac_trace ?loc trace) let () = CErrors.register_additional_error_info get_ltac_trace coq-8.11.0/plugins/ltac/g_tactic.mlg0000644000175000017500000007320413612664533017156 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* "; "<-" ; "by" ] let _ = List.iter CLexer.add_keyword tactic_kw let err () = raise Stream.Failure (* Hack to parse "(x:=t)" as an explicit argument without conflicts with the *) (* admissible notation "(x t)" *) let test_lpar_id_coloneq = Pcoq.Entry.of_parser "lpar_id_coloneq" (fun _ strm -> match stream_nth 0 strm with | KEYWORD "(" -> (match stream_nth 1 strm with | IDENT _ -> (match stream_nth 2 strm with | KEYWORD ":=" -> () | _ -> err ()) | _ -> err ()) | _ -> err ()) (* Hack to recognize "(x)" *) let test_lpar_id_rpar = Pcoq.Entry.of_parser "lpar_id_coloneq" (fun _ strm -> match stream_nth 0 strm with | KEYWORD "(" -> (match stream_nth 1 strm with | IDENT _ -> (match stream_nth 2 strm with | KEYWORD ")" -> () | _ -> err ()) | _ -> err ()) | _ -> err ()) (* idem for (x:=t) and (1:=t) *) let test_lpar_idnum_coloneq = Pcoq.Entry.of_parser "test_lpar_idnum_coloneq" (fun _ strm -> match stream_nth 0 strm with | KEYWORD "(" -> (match stream_nth 1 strm with | IDENT _ | NUMERAL _ -> (match stream_nth 2 strm with | KEYWORD ":=" -> () | _ -> err ()) | _ -> err ()) | _ -> err ()) (* idem for (x:t) *) open Extraargs (* idem for (x1..xn:t) [n^2 complexity but exceptional use] *) let check_for_coloneq = Pcoq.Entry.of_parser "lpar_id_colon" (fun _ strm -> let rec skip_to_rpar p n = match List.last (Stream.npeek n strm) with | KEYWORD "(" -> skip_to_rpar (p+1) (n+1) | KEYWORD ")" -> if Int.equal p 0 then n+1 else skip_to_rpar (p-1) (n+1) | KEYWORD "." -> err () | _ -> skip_to_rpar p (n+1) in let rec skip_names n = match List.last (Stream.npeek n strm) with | IDENT _ | KEYWORD "_" -> skip_names (n+1) | KEYWORD ":" -> skip_to_rpar 0 (n+1) (* skip a constr *) | _ -> err () in let rec skip_binders n = match List.last (Stream.npeek n strm) with | KEYWORD "(" -> skip_binders (skip_names (n+1)) | IDENT _ | KEYWORD "_" -> skip_binders (n+1) | KEYWORD ":=" -> () | _ -> err () in match stream_nth 0 strm with | KEYWORD "(" -> skip_binders 2 | _ -> err ()) let lookup_at_as_comma = Pcoq.Entry.of_parser "lookup_at_as_comma" (fun _ strm -> match stream_nth 0 strm with | KEYWORD (","|"at"|"as") -> () | _ -> err ()) open Constr open Prim open Pltac let mk_fix_tac (loc,id,bl,ann,ty) = let n = match bl,ann with [([_],_,_)], None -> 1 | _, Some x -> let ids = List.map (fun x -> x.CAst.v) (List.flatten (List.map (fun (nal,_,_) -> nal) bl)) in (try List.index Names.Name.equal x.CAst.v ids with Not_found -> user_err Pp.(str "No such fix variable.")) | _ -> user_err Pp.(str "Cannot guess decreasing argument of fix.") in let bl = List.map (fun (nal,bk,t) -> CLocalAssum (nal,bk,t)) bl in (id,n, CAst.make ~loc @@ CProdN(bl,ty)) let mk_cofix_tac (loc,id,bl,ann,ty) = let _ = Option.map (fun { CAst.loc = aloc } -> user_err ?loc:aloc ~hdr:"Constr:mk_cofix_tac" (Pp.str"Annotation forbidden in cofix expression.")) ann in let bl = List.map (fun (nal,bk,t) -> CLocalAssum (nal,bk,t)) bl in (id,if bl = [] then ty else CAst.make ~loc @@ CProdN(bl,ty)) (* Functions overloaded by quotifier *) let destruction_arg_of_constr (c,lbind as clbind) = match lbind with | NoBindings -> begin try ElimOnIdent (CAst.make ?loc:(Constrexpr_ops.constr_loc c) (Constrexpr_ops.coerce_to_id c).CAst.v) with e when CErrors.noncritical e -> ElimOnConstr clbind end | _ -> ElimOnConstr clbind let mkNumeral n = Numeral ((if 0<=n then SPlus else SMinus),NumTok.int (string_of_int (abs n))) let mkTacCase with_evar = function | [(clear,ElimOnConstr cl),(None,None),None],None -> TacCase (with_evar,(clear,cl)) (* Reinterpret numbers as a notation for terms *) | [(clear,ElimOnAnonHyp n),(None,None),None],None -> TacCase (with_evar, (clear,(CAst.make @@ CPrim (mkNumeral n), NoBindings))) (* Reinterpret ident as notations for variables in the context *) (* because we don't know if they are quantified or not *) | [(clear,ElimOnIdent id),(None,None),None],None -> TacCase (with_evar,(clear,(CAst.make @@ CRef (qualid_of_ident ?loc:id.CAst.loc id.CAst.v,None),NoBindings))) | ic -> if List.exists (function ((_, ElimOnAnonHyp _),_,_) -> true | _ -> false) (fst ic) then user_err Pp.(str "Use of numbers as direct arguments of 'case' is not supported."); TacInductionDestruct (false,with_evar,ic) let rec mkCLambdaN_simple_loc ?loc bll c = match bll with | ({CAst.loc = loc1}::_ as idl,bk,t) :: bll -> CAst.make ?loc @@ CLambdaN ([CLocalAssum (idl,bk,t)],mkCLambdaN_simple_loc ?loc:(Loc.merge_opt loc1 loc) bll c) | ([],_,_) :: bll -> mkCLambdaN_simple_loc ?loc bll c | [] -> c let mkCLambdaN_simple bl c = match bl with | [] -> c | h :: _ -> let loc = Loc.merge_opt (List.hd (pi1 h)).CAst.loc (Constrexpr_ops.constr_loc c) in mkCLambdaN_simple_loc ?loc bl c let loc_of_ne_list l = Loc.merge_opt (List.hd l).CAst.loc (List.last l).CAst.loc let all_concl_occs_clause = { onhyps=Some[]; concl_occs=AllOccurrences } let merge_occurrences loc cl = function | None -> if Locusops.clause_with_generic_occurrences cl then (None, cl) else user_err ~loc (str "Found an \"at\" clause without \"with\" clause.") | Some (occs, p) -> let ans = match occs with | AllOccurrences -> cl | _ -> begin match cl with | { onhyps = Some []; concl_occs = AllOccurrences } -> { onhyps = Some []; concl_occs = occs } | { onhyps = Some [(AllOccurrences, id), l]; concl_occs = NoOccurrences } -> { cl with onhyps = Some [(occs, id), l] } | _ -> if Locusops.clause_with_generic_occurrences cl then user_err ~loc (str "Unable to interpret the \"at\" clause; move it in the \"in\" clause.") else user_err ~loc (str "Cannot use clause \"at\" twice.") end in (Some p, ans) let warn_deprecated_eqn_syntax = CWarnings.create ~name:"deprecated-eqn-syntax" ~category:"deprecated" (fun arg -> strbrk (Printf.sprintf "Syntax \"_eqn:%s\" is deprecated. Please use \"eqn:%s\" instead." arg arg)) (* Auxiliary grammar rules *) open Pvernac.Vernac_ } GRAMMAR EXTEND Gram GLOBAL: simple_tactic constr_with_bindings quantified_hypothesis bindings red_expr int_or_var open_constr uconstr simple_intropattern in_clause clause_dft_concl hypident destruction_arg; int_or_var: [ [ n = integer -> { ArgArg n } | id = identref -> { ArgVar id } ] ] ; nat_or_var: [ [ n = natural -> { ArgArg n } | id = identref -> { ArgVar id } ] ] ; (* An identifier or a quotation meta-variable *) id_or_meta: [ [ id = identref -> { id } ] ] ; open_constr: [ [ c = constr -> { c } ] ] ; uconstr: [ [ c = constr -> { c } ] ] ; destruction_arg: [ [ n = natural -> { (None,ElimOnAnonHyp n) } | test_lpar_id_rpar; c = constr_with_bindings -> { (Some false,destruction_arg_of_constr c) } | c = constr_with_bindings_arg -> { on_snd destruction_arg_of_constr c } ] ] ; constr_with_bindings_arg: [ [ ">"; c = constr_with_bindings -> { (Some true,c) } | c = constr_with_bindings -> { (None,c) } ] ] ; quantified_hypothesis: [ [ id = ident -> { NamedHyp id } | n = natural -> { AnonHyp n } ] ] ; conversion: [ [ c = constr -> { (None, c) } | c1 = constr; "with"; c2 = constr -> { (Some (AllOccurrences,c1),c2) } | c1 = constr; "at"; occs = occs_nums; "with"; c2 = constr -> { (Some (occs,c1), c2) } ] ] ; occs_nums: [ [ nl = LIST1 nat_or_var -> { OnlyOccurrences nl } | "-"; n = nat_or_var; nl = LIST0 int_or_var -> (* have used int_or_var instead of nat_or_var for compatibility *) { AllOccurrencesBut (List.map (Locusops.or_var_map abs) (n::nl)) } ] ] ; occs: [ [ "at"; occs = occs_nums -> { occs } | -> { AllOccurrences } ] ] ; pattern_occ: [ [ c = constr; nl = occs -> { (nl,c) } ] ] ; ref_or_pattern_occ: (* If a string, it is interpreted as a ref (anyway a Coq string does not reduce) *) [ [ c = smart_global; nl = occs -> { nl,Inl c } | c = constr; nl = occs -> { nl,Inr c } ] ] ; unfold_occ: [ [ c = smart_global; nl = occs -> { (nl,c) } ] ] ; intropatterns: [ [ l = LIST0 intropattern -> { l } ] ] ; ne_intropatterns: [ [ l = LIST1 intropattern -> { l } ] ] ; or_and_intropattern: [ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> { IntroOrPattern tc } | "()" -> { IntroAndPattern [] } | "("; si = simple_intropattern; ")" -> { IntroAndPattern [si] } | "("; si = simple_intropattern; ","; tc = LIST1 simple_intropattern SEP "," ; ")" -> { IntroAndPattern (si::tc) } | "("; si = simple_intropattern; "&"; tc = LIST1 simple_intropattern SEP "&" ; ")" -> (* (A & B & C) is translated into (A,(B,C)) *) { let rec pairify = function | ([]|[_]|[_;_]) as l -> l | t::q -> [t; CAst.make ?loc:(loc_of_ne_list q) (IntroAction (IntroOrAndPattern (IntroAndPattern (pairify q))))] in IntroAndPattern (pairify (si::tc)) } ] ] ; equality_intropattern: [ [ "->" -> { IntroRewrite true } | "<-" -> { IntroRewrite false } | "[="; tc = intropatterns; "]" -> { IntroInjection tc } ] ] ; naming_intropattern: [ [ prefix = pattern_ident -> { IntroFresh prefix } | "?" -> { IntroAnonymous } | id = ident -> { IntroIdentifier id } ] ] ; intropattern: [ [ l = simple_intropattern -> { l } | "*" -> { CAst.make ~loc @@ IntroForthcoming true } | "**" -> { CAst.make ~loc @@ IntroForthcoming false } ] ] ; simple_intropattern: [ [ pat = simple_intropattern_closed; l = LIST0 ["%"; c = operconstr LEVEL "0" -> { c } ] -> { let {CAst.loc=loc0;v=pat} = pat in let f c pat = let loc1 = Constrexpr_ops.constr_loc c in let loc = Loc.merge_opt loc0 loc1 in IntroAction (IntroApplyOn (CAst.(make ?loc:loc1 c),CAst.(make ?loc pat))) in CAst.make ~loc @@ List.fold_right f l pat } ] ] ; simple_intropattern_closed: [ [ pat = or_and_intropattern -> { CAst.make ~loc @@ IntroAction (IntroOrAndPattern pat) } | pat = equality_intropattern -> { CAst.make ~loc @@ IntroAction pat } | "_" -> { CAst.make ~loc @@ IntroAction IntroWildcard } | pat = naming_intropattern -> { CAst.make ~loc @@ IntroNaming pat } ] ] ; simple_binding: [ [ "("; id = ident; ":="; c = lconstr; ")" -> { CAst.make ~loc (NamedHyp id, c) } | "("; n = natural; ":="; c = lconstr; ")" -> { CAst.make ~loc (AnonHyp n, c) } ] ] ; bindings: [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> { ExplicitBindings bl } | bl = LIST1 constr -> { ImplicitBindings bl } ] ] ; constr_with_bindings: [ [ c = constr; l = with_bindings -> { (c, l) } ] ] ; with_bindings: [ [ "with"; bl = bindings -> { bl } | -> { NoBindings } ] ] ; red_flags: [ [ IDENT "beta" -> { [FBeta] } | IDENT "iota" -> { [FMatch;FFix;FCofix] } | IDENT "match" -> { [FMatch] } | IDENT "fix" -> { [FFix] } | IDENT "cofix" -> { [FCofix] } | IDENT "zeta" -> { [FZeta] } | IDENT "delta"; d = delta_flag -> { [d] } ] ] ; delta_flag: [ [ "-"; "["; idl = LIST1 smart_global; "]" -> { FDeltaBut idl } | "["; idl = LIST1 smart_global; "]" -> { FConst idl } | -> { FDeltaBut [] } ] ] ; strategy_flag: [ [ s = LIST1 red_flags -> { Redops.make_red_flag (List.flatten s) } | d = delta_flag -> { all_with d } ] ] ; red_expr: [ [ IDENT "red" -> { Red false } | IDENT "hnf" -> { Hnf } | IDENT "simpl"; d = delta_flag; po = OPT ref_or_pattern_occ -> { Simpl (all_with d,po) } | IDENT "cbv"; s = strategy_flag -> { Cbv s } | IDENT "cbn"; s = strategy_flag -> { Cbn s } | IDENT "lazy"; s = strategy_flag -> { Lazy s } | IDENT "compute"; delta = delta_flag -> { Cbv (all_with delta) } | IDENT "vm_compute"; po = OPT ref_or_pattern_occ -> { CbvVm po } | IDENT "native_compute"; po = OPT ref_or_pattern_occ -> { CbvNative po } | IDENT "unfold"; ul = LIST1 unfold_occ SEP "," -> { Unfold ul } | IDENT "fold"; cl = LIST1 constr -> { Fold cl } | IDENT "pattern"; pl = LIST1 pattern_occ SEP"," -> { Pattern pl } | s = IDENT -> { ExtraRedExpr s } ] ] ; hypident: [ [ id = id_or_meta -> { let id : lident = id in id,InHyp } | "("; IDENT "type"; IDENT "of"; id = id_or_meta; ")" -> { let id : lident = id in id,InHypTypeOnly } | "("; IDENT "value"; IDENT "of"; id = id_or_meta; ")" -> { let id : lident = id in id,InHypValueOnly } ] ] ; hypident_occ: [ [ h=hypident; occs=occs -> { let (id,l) = h in let id : lident = id in ((occs,id),l) } ] ] ; in_clause: [ [ "*"; occs=occs -> { {onhyps=None; concl_occs=occs} } | "*"; "|-"; occs=concl_occ -> { {onhyps=None; concl_occs=occs} } | hl=LIST0 hypident_occ SEP","; "|-"; occs=concl_occ -> { {onhyps=Some hl; concl_occs=occs} } | hl=LIST0 hypident_occ SEP"," -> { {onhyps=Some hl; concl_occs=NoOccurrences} } ] ] ; clause_dft_concl: [ [ "in"; cl = in_clause -> { cl } | occs=occs -> { {onhyps=Some[]; concl_occs=occs} } | -> { all_concl_occs_clause } ] ] ; clause_dft_all: [ [ "in"; cl = in_clause -> { cl } | -> { {onhyps=None; concl_occs=AllOccurrences} } ] ] ; opt_clause: [ [ "in"; cl = in_clause -> { Some cl } | "at"; occs = occs_nums -> { Some {onhyps=Some[]; concl_occs=occs} } | -> { None } ] ] ; concl_occ: [ [ "*"; occs = occs -> { occs } | -> { NoOccurrences } ] ] ; in_hyp_list: [ [ "in"; idl = LIST1 id_or_meta -> { idl } | -> { [] } ] ] ; in_hyp_as: [ [ "in"; id = id_or_meta; ipat = as_ipat -> { Some (id,ipat) } | -> { None } ] ] ; orient: [ [ "->" -> { true } | "<-" -> { false } | -> { true } ] ] ; simple_binder: [ [ na=name -> { ([na],Default Glob_term.Explicit, CAst.make ~loc @@ CHole (Some (Evar_kinds.BinderType na.CAst.v), IntroAnonymous, None)) } | "("; nal=LIST1 name; ":"; c=lconstr; ")" -> { (nal,Default Glob_term.Explicit,c) } ] ] ; fixdecl: [ [ "("; id = ident; bl=LIST0 simple_binder; ann=fixannot; ":"; ty=lconstr; ")" -> { (loc, id, bl, ann, ty) } ] ] ; fixannot: [ [ "{"; IDENT "struct"; id=name; "}" -> { Some id } | -> { None } ] ] ; cofixdecl: [ [ "("; id = ident; bl=LIST0 simple_binder; ":"; ty=lconstr; ")" -> { (loc, id, bl, None, ty) } ] ] ; bindings_with_parameters: [ [ check_for_coloneq; "("; id = ident; bl = LIST0 simple_binder; ":="; c = lconstr; ")" -> { (id, mkCLambdaN_simple bl c) } ] ] ; eliminator: [ [ "using"; el = constr_with_bindings -> { el } ] ] ; as_ipat: [ [ "as"; ipat = simple_intropattern -> { Some ipat } | -> { None } ] ] ; or_and_intropattern_loc: [ [ ipat = or_and_intropattern -> { ArgArg (CAst.make ~loc ipat) } | locid = identref -> { ArgVar locid } ] ] ; as_or_and_ipat: [ [ "as"; ipat = or_and_intropattern_loc -> { Some ipat } | -> { None } ] ] ; eqn_ipat: [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> { Some (CAst.make ~loc pat) } | IDENT "_eqn"; ":"; pat = naming_intropattern -> { warn_deprecated_eqn_syntax ~loc "H"; Some (CAst.make ~loc pat) } | IDENT "_eqn" -> { warn_deprecated_eqn_syntax ~loc "?"; Some (CAst.make ~loc IntroAnonymous) } | -> { None } ] ] ; as_name: [ [ "as"; id = ident -> { Names.Name.Name id } | -> { Names.Name.Anonymous } ] ] ; by_tactic: [ [ "by"; tac = tactic_expr LEVEL "3" -> { Some tac } | -> { None } ] ] ; rewriter : [ [ "!"; c = constr_with_bindings_arg -> { (Equality.RepeatPlus,c) } | ["?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings_arg -> { (Equality.RepeatStar,c) } | n = natural; "!"; c = constr_with_bindings_arg -> { (Equality.Precisely n,c) } | n = natural; ["?" -> { () } | LEFTQMARK -> { () } ]; c = constr_with_bindings_arg -> { (Equality.UpTo n,c) } | n = natural; c = constr_with_bindings_arg -> { (Equality.Precisely n,c) } | c = constr_with_bindings_arg -> { (Equality.Precisely 1, c) } ] ] ; oriented_rewriter : [ [ b = orient; p = rewriter -> { let (m,c) = p in (b,m,c) } ] ] ; induction_clause: [ [ c = destruction_arg; pat = as_or_and_ipat; eq = eqn_ipat; cl = opt_clause -> { (c,(eq,pat),cl) } ] ] ; induction_clause_list: [ [ ic = LIST1 induction_clause SEP ","; el = OPT eliminator; cl_tolerance = opt_clause -> (* Condition for accepting "in" at the end by compatibility *) { match ic,el,cl_tolerance with | [c,pat,None],Some _,Some _ -> ([c,pat,cl_tolerance],el) | _,_,Some _ -> err () | _,_,None -> (ic,el) } ] ] ; simple_tactic: [ [ (* Basic tactics *) IDENT "intros"; pl = ne_intropatterns -> { TacAtom (CAst.make ~loc @@ TacIntroPattern (false,pl)) } | IDENT "intros" -> { TacAtom (CAst.make ~loc @@ TacIntroPattern (false,[CAst.make ~loc @@IntroForthcoming false])) } | IDENT "eintros"; pl = ne_intropatterns -> { TacAtom (CAst.make ~loc @@ TacIntroPattern (true,pl)) } | IDENT "eintros" -> { TacAtom (CAst.make ~loc @@ TacIntroPattern (true,[CAst.make ~loc @@IntroForthcoming false])) } | IDENT "apply"; cl = LIST1 constr_with_bindings_arg SEP ","; inhyp = in_hyp_as -> { TacAtom (CAst.make ~loc @@ TacApply (true,false,cl,inhyp)) } | IDENT "eapply"; cl = LIST1 constr_with_bindings_arg SEP ","; inhyp = in_hyp_as -> { TacAtom (CAst.make ~loc @@ TacApply (true,true,cl,inhyp)) } | IDENT "simple"; IDENT "apply"; cl = LIST1 constr_with_bindings_arg SEP ","; inhyp = in_hyp_as -> { TacAtom (CAst.make ~loc @@ TacApply (false,false,cl,inhyp)) } | IDENT "simple"; IDENT "eapply"; cl = LIST1 constr_with_bindings_arg SEP","; inhyp = in_hyp_as -> { TacAtom (CAst.make ~loc @@ TacApply (false,true,cl,inhyp)) } | IDENT "elim"; cl = constr_with_bindings_arg; el = OPT eliminator -> { TacAtom (CAst.make ~loc @@ TacElim (false,cl,el)) } | IDENT "eelim"; cl = constr_with_bindings_arg; el = OPT eliminator -> { TacAtom (CAst.make ~loc @@ TacElim (true,cl,el)) } | IDENT "case"; icl = induction_clause_list -> { TacAtom (CAst.make ~loc @@ mkTacCase false icl) } | IDENT "ecase"; icl = induction_clause_list -> { TacAtom (CAst.make ~loc @@ mkTacCase true icl) } | "fix"; id = ident; n = natural; "with"; fd = LIST1 fixdecl -> { TacAtom (CAst.make ~loc @@ TacMutualFix (id,n,List.map mk_fix_tac fd)) } | "cofix"; id = ident; "with"; fd = LIST1 cofixdecl -> { TacAtom (CAst.make ~loc @@ TacMutualCofix (id,List.map mk_cofix_tac fd)) } | IDENT "pose"; bl = bindings_with_parameters -> { let (id,b) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (false,Names.Name.Name id,b,Locusops.nowhere,true,None)) } | IDENT "pose"; b = constr; na = as_name -> { TacAtom (CAst.make ~loc @@ TacLetTac (false,na,b,Locusops.nowhere,true,None)) } | IDENT "epose"; bl = bindings_with_parameters -> { let (id,b) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (true,Names.Name id,b,Locusops.nowhere,true,None)) } | IDENT "epose"; b = constr; na = as_name -> { TacAtom (CAst.make ~loc @@ TacLetTac (true,na,b,Locusops.nowhere,true,None)) } | IDENT "set"; bl = bindings_with_parameters; p = clause_dft_concl -> { let (id,c) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (false,Names.Name.Name id,c,p,true,None)) } | IDENT "set"; c = constr; na = as_name; p = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacLetTac (false,na,c,p,true,None)) } | IDENT "eset"; bl = bindings_with_parameters; p = clause_dft_concl -> { let (id,c) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (true,Names.Name id,c,p,true,None)) } | IDENT "eset"; c = constr; na = as_name; p = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacLetTac (true,na,c,p,true,None)) } | IDENT "remember"; c = constr; na = as_name; e = eqn_ipat; p = clause_dft_all -> { TacAtom (CAst.make ~loc @@ TacLetTac (false,na,c,p,false,e)) } | IDENT "eremember"; c = constr; na = as_name; e = eqn_ipat; p = clause_dft_all -> { TacAtom (CAst.make ~loc @@ TacLetTac (true,na,c,p,false,e)) } (* Alternative syntax for "pose proof c as id" *) | IDENT "assert"; test_lpar_id_coloneq; "("; lid = identref; ":="; c = lconstr; ")" -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (false,true,None,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "eassert"; test_lpar_id_coloneq; "("; lid = identref; ":="; c = lconstr; ")" -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (true,true,None,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } (* Alternative syntax for "assert c as id by tac" *) | IDENT "assert"; test_lpar_id_colon; "("; lid = identref; ":"; c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (false,true,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "eassert"; test_lpar_id_colon; "("; lid = identref; ":"; c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (true,true,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } (* Alternative syntax for "enough c as id by tac" *) | IDENT "enough"; test_lpar_id_colon; "("; lid = identref; ":"; c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (false,false,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "eenough"; test_lpar_id_colon; "("; lid = identref; ":"; c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (true,false,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "assert"; c = constr; ipat = as_ipat; tac = by_tactic -> { TacAtom (CAst.make ~loc @@ TacAssert (false,true,Some tac,ipat,c)) } | IDENT "eassert"; c = constr; ipat = as_ipat; tac = by_tactic -> { TacAtom (CAst.make ~loc @@ TacAssert (true,true,Some tac,ipat,c)) } | IDENT "pose"; IDENT "proof"; c = lconstr; ipat = as_ipat -> { TacAtom (CAst.make ~loc @@ TacAssert (false,true,None,ipat,c)) } | IDENT "epose"; IDENT "proof"; c = lconstr; ipat = as_ipat -> { TacAtom (CAst.make ~loc @@ TacAssert (true,true,None,ipat,c)) } | IDENT "enough"; c = constr; ipat = as_ipat; tac = by_tactic -> { TacAtom (CAst.make ~loc @@ TacAssert (false,false,Some tac,ipat,c)) } | IDENT "eenough"; c = constr; ipat = as_ipat; tac = by_tactic -> { TacAtom (CAst.make ~loc @@ TacAssert (true,false,Some tac,ipat,c)) } | IDENT "generalize"; c = constr -> { TacAtom (CAst.make ~loc @@ TacGeneralize [((AllOccurrences,c),Names.Name.Anonymous)]) } | IDENT "generalize"; c = constr; l = LIST1 constr -> { let gen_everywhere c = ((AllOccurrences,c),Names.Name.Anonymous) in TacAtom (CAst.make ~loc @@ TacGeneralize (List.map gen_everywhere (c::l))) } | IDENT "generalize"; c = constr; lookup_at_as_comma; nl = occs; na = as_name; l = LIST0 [","; c = pattern_occ; na = as_name -> { (c,na) } ] -> { TacAtom (CAst.make ~loc @@ TacGeneralize (((nl,c),na)::l)) } (* Derived basic tactics *) | IDENT "induction"; ic = induction_clause_list -> { TacAtom (CAst.make ~loc @@ TacInductionDestruct (true,false,ic)) } | IDENT "einduction"; ic = induction_clause_list -> { TacAtom (CAst.make ~loc @@ TacInductionDestruct(true,true,ic)) } | IDENT "destruct"; icl = induction_clause_list -> { TacAtom (CAst.make ~loc @@ TacInductionDestruct(false,false,icl)) } | IDENT "edestruct"; icl = induction_clause_list -> { TacAtom (CAst.make ~loc @@ TacInductionDestruct(false,true,icl)) } (* Equality and inversion *) | IDENT "rewrite"; l = LIST1 oriented_rewriter SEP ","; cl = clause_dft_concl; t=by_tactic -> { TacAtom (CAst.make ~loc @@ TacRewrite (false,l,cl,t)) } | IDENT "erewrite"; l = LIST1 oriented_rewriter SEP ","; cl = clause_dft_concl; t=by_tactic -> { TacAtom (CAst.make ~loc @@ TacRewrite (true,l,cl,t)) } | IDENT "dependent"; k = [ IDENT "simple"; IDENT "inversion" -> { SimpleInversion } | IDENT "inversion" -> { FullInversion } | IDENT "inversion_clear" -> { FullInversionClear } ]; hyp = quantified_hypothesis; ids = as_or_and_ipat; co = OPT ["with"; c = constr -> { c } ] -> { TacAtom (CAst.make ~loc @@ TacInversion (DepInversion (k,co,ids),hyp)) } | IDENT "simple"; IDENT "inversion"; hyp = quantified_hypothesis; ids = as_or_and_ipat; cl = in_hyp_list -> { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp)) } | IDENT "inversion"; hyp = quantified_hypothesis; ids = as_or_and_ipat; cl = in_hyp_list -> { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (FullInversion, cl, ids), hyp)) } | IDENT "inversion_clear"; hyp = quantified_hypothesis; ids = as_or_and_ipat; cl = in_hyp_list -> { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp)) } | IDENT "inversion"; hyp = quantified_hypothesis; "using"; c = constr; cl = in_hyp_list -> { TacAtom (CAst.make ~loc @@ TacInversion (InversionUsing (c,cl), hyp)) } (* Conversion *) | IDENT "red"; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Red false, cl)) } | IDENT "hnf"; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Hnf, cl)) } | IDENT "simpl"; d = delta_flag; po = OPT ref_or_pattern_occ; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Simpl (all_with d, po), cl)) } | IDENT "cbv"; s = strategy_flag; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Cbv s, cl)) } | IDENT "cbn"; s = strategy_flag; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Cbn s, cl)) } | IDENT "lazy"; s = strategy_flag; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Lazy s, cl)) } | IDENT "compute"; delta = delta_flag; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Cbv (all_with delta), cl)) } | IDENT "vm_compute"; po = OPT ref_or_pattern_occ; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (CbvVm po, cl)) } | IDENT "native_compute"; po = OPT ref_or_pattern_occ; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (CbvNative po, cl)) } | IDENT "unfold"; ul = LIST1 unfold_occ SEP ","; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Unfold ul, cl)) } | IDENT "fold"; l = LIST1 constr; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Fold l, cl)) } | IDENT "pattern"; pl = LIST1 pattern_occ SEP","; cl = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacReduce (Pattern pl, cl)) } (* Change ne doit pas s'appliquer dans un Definition t := Eval ... *) | IDENT "change"; c = conversion; cl = clause_dft_concl -> { let (oc, c) = c in let p,cl = merge_occurrences loc cl oc in TacAtom (CAst.make ~loc @@ TacChange (true,p,c,cl)) } | IDENT "change_no_check"; c = conversion; cl = clause_dft_concl -> { let (oc, c) = c in let p,cl = merge_occurrences loc cl oc in TacAtom (CAst.make ~loc @@ TacChange (false,p,c,cl)) } ] ] ; END coq-8.11.0/plugins/ltac/tacsubst.ml0000644000175000017500000003057513612664533017066 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* subst_quantified_hypothesis subst b,subst_glob_constr subst c) let subst_bindings subst = function | NoBindings -> NoBindings | ImplicitBindings l -> ImplicitBindings (List.map (subst_glob_constr subst) l) | ExplicitBindings l -> ExplicitBindings (List.map (subst_binding subst) l) let subst_glob_with_bindings subst (c,bl) = (subst_glob_constr subst c, subst_bindings subst bl) let subst_glob_with_bindings_arg subst (clear,c) = (clear,subst_glob_with_bindings subst c) let rec subst_intro_pattern subst = CAst.map (function | IntroAction p -> IntroAction (subst_intro_pattern_action subst p) | IntroNaming _ | IntroForthcoming _ as x -> x) and subst_intro_pattern_action subst = let open CAst in function | IntroApplyOn ({loc;v=t},pat) -> IntroApplyOn (make ?loc @@ subst_glob_constr subst t,subst_intro_pattern subst pat) | IntroOrAndPattern l -> IntroOrAndPattern (subst_intro_or_and_pattern subst l) | IntroInjection l -> IntroInjection (List.map (subst_intro_pattern subst) l) | IntroWildcard | IntroRewrite _ as x -> x and subst_intro_or_and_pattern subst = function | IntroAndPattern l -> IntroAndPattern (List.map (subst_intro_pattern subst) l) | IntroOrPattern ll -> IntroOrPattern (List.map (List.map (subst_intro_pattern subst)) ll) let subst_destruction_arg subst = function | clear,ElimOnConstr c -> clear,ElimOnConstr (subst_glob_with_bindings subst c) | clear,ElimOnAnonHyp n as x -> x | clear,ElimOnIdent id as x -> x let subst_and_short_name f (c,n) = (* assert (n=None); *)(* since tacdef are strictly globalized *) (f c,None) let subst_located f = Loc.map f let subst_reference subst = Locusops.or_var_map (subst_located (subst_kn subst)) (*CSC: subst_global_reference is used "only" for RefArgType, that propagates to the syntactic non-terminals "global", used in commands such as Print. It is also used for non-evaluable references. *) let subst_global_reference subst = Locusops.or_var_map (subst_located (subst_global_reference subst)) let subst_evaluable subst = let subst_eval_ref = subst_evaluable_reference subst in Locusops.or_var_map (subst_and_short_name subst_eval_ref) let subst_constr_with_occurrences subst (l,c) = (l,subst_glob_constr subst c) let subst_glob_constr_or_pattern subst (bvars,c,p) = let env = Global.env () in let sigma = Evd.from_env env in (bvars,subst_glob_constr subst c,subst_pattern env sigma subst p) let subst_redexp subst = Redops.map_red_expr_gen (subst_glob_constr subst) (subst_evaluable subst) (subst_glob_constr_or_pattern subst) let subst_raw_may_eval subst = function | ConstrEval (r,c) -> ConstrEval (subst_redexp subst r,subst_glob_constr subst c) | ConstrContext (locid,c) -> ConstrContext (locid,subst_glob_constr subst c) | ConstrTypeOf c -> ConstrTypeOf (subst_glob_constr subst c) | ConstrTerm c -> ConstrTerm (subst_glob_constr subst c) let subst_match_pattern subst = function | Subterm (ido,pc) -> Subterm (ido,(subst_glob_constr_or_pattern subst pc)) | Term pc -> Term (subst_glob_constr_or_pattern subst pc) let rec subst_match_goal_hyps subst = function | Hyp (locs,mp) :: tl -> Hyp (locs,subst_match_pattern subst mp) :: subst_match_goal_hyps subst tl | Def (locs,mv,mp) :: tl -> Def (locs,subst_match_pattern subst mv, subst_match_pattern subst mp) :: subst_match_goal_hyps subst tl | [] -> [] let rec subst_atomic subst (t:glob_atomic_tactic_expr) = match t with (* Basic tactics *) | TacIntroPattern (ev,l) -> TacIntroPattern (ev,List.map (subst_intro_pattern subst) l) | TacApply (a,ev,cb,cl) -> TacApply (a,ev,List.map (subst_glob_with_bindings_arg subst) cb,cl) | TacElim (ev,cb,cbo) -> TacElim (ev,subst_glob_with_bindings_arg subst cb, Option.map (subst_glob_with_bindings subst) cbo) | TacCase (ev,cb) -> TacCase (ev,subst_glob_with_bindings_arg subst cb) | TacMutualFix (id,n,l) -> TacMutualFix(id,n,List.map (fun (id,n,c) -> (id,n,subst_glob_constr subst c)) l) | TacMutualCofix (id,l) -> TacMutualCofix (id, List.map (fun (id,c) -> (id,subst_glob_constr subst c)) l) | TacAssert (ev,b,otac,na,c) -> TacAssert (ev,b,Option.map (Option.map (subst_tactic subst)) otac,na, subst_glob_constr subst c) | TacGeneralize cl -> TacGeneralize (List.map (on_fst (subst_constr_with_occurrences subst))cl) | TacLetTac (ev,id,c,clp,b,eqpat) -> TacLetTac (ev,id,subst_glob_constr subst c,clp,b,eqpat) (* Derived basic tactics *) | TacInductionDestruct (isrec,ev,(l,el)) -> let l' = List.map (fun (c,ids,cls) -> subst_destruction_arg subst c, ids, cls) l in let el' = Option.map (subst_glob_with_bindings subst) el in TacInductionDestruct (isrec,ev,(l',el')) (* Conversion *) | TacReduce (r,cl) -> TacReduce (subst_redexp subst r, cl) | TacChange (check,op,c,cl) -> TacChange (check,Option.map (subst_glob_constr_or_pattern subst) op, subst_glob_constr subst c, cl) (* Equality and inversion *) | TacRewrite (ev,l,cl,by) -> TacRewrite (ev, List.map (fun (b,m,c) -> b,m,subst_glob_with_bindings_arg subst c) l, cl,Option.map (subst_tactic subst) by) | TacInversion (DepInversion (k,c,l),hyp) -> TacInversion (DepInversion (k,Option.map (subst_glob_constr subst) c,l),hyp) | TacInversion (NonDepInversion _,_) as x -> x | TacInversion (InversionUsing (c,cl),hyp) -> TacInversion (InversionUsing (subst_glob_constr subst c,cl),hyp) and subst_tactic subst (t:glob_tactic_expr) = match t with | TacAtom { CAst.v=t } -> TacAtom (CAst.make @@ subst_atomic subst t) | TacFun tacfun -> TacFun (subst_tactic_fun subst tacfun) | TacLetIn (r,l,u) -> let l = List.map (fun (n,b) -> (n,subst_tacarg subst b)) l in TacLetIn (r,l,subst_tactic subst u) | TacMatchGoal (lz,lr,lmr) -> TacMatchGoal(lz,lr, subst_match_rule subst lmr) | TacMatch (lz,c,lmr) -> TacMatch (lz,subst_tactic subst c,subst_match_rule subst lmr) | TacId _ | TacFail _ as x -> x | TacProgress tac -> TacProgress (subst_tactic subst tac:glob_tactic_expr) | TacShowHyps tac -> TacShowHyps (subst_tactic subst tac:glob_tactic_expr) | TacAbstract (tac,s) -> TacAbstract (subst_tactic subst tac,s) | TacThen (t1,t2) -> TacThen (subst_tactic subst t1, subst_tactic subst t2) | TacDispatch tl -> TacDispatch (List.map (subst_tactic subst) tl) | TacExtendTac (tf,t,tl) -> TacExtendTac (Array.map (subst_tactic subst) tf, subst_tactic subst t, Array.map (subst_tactic subst) tl) | TacThens (t,tl) -> TacThens (subst_tactic subst t, List.map (subst_tactic subst) tl) | TacThens3parts (t1,tf,t2,tl) -> TacThens3parts (subst_tactic subst t1,Array.map (subst_tactic subst) tf, subst_tactic subst t2,Array.map (subst_tactic subst) tl) | TacDo (n,tac) -> TacDo (n,subst_tactic subst tac) | TacTimeout (n,tac) -> TacTimeout (n,subst_tactic subst tac) | TacTime (s,tac) -> TacTime (s,subst_tactic subst tac) | TacTry tac -> TacTry (subst_tactic subst tac) | TacInfo tac -> TacInfo (subst_tactic subst tac) | TacRepeat tac -> TacRepeat (subst_tactic subst tac) | TacOr (tac1,tac2) -> TacOr (subst_tactic subst tac1,subst_tactic subst tac2) | TacOnce tac -> TacOnce (subst_tactic subst tac) | TacExactlyOnce tac -> TacExactlyOnce (subst_tactic subst tac) | TacIfThenCatch (tac,tact,tace) -> TacIfThenCatch ( subst_tactic subst tac, subst_tactic subst tact, subst_tactic subst tace) | TacOrelse (tac1,tac2) -> TacOrelse (subst_tactic subst tac1,subst_tactic subst tac2) | TacFirst l -> TacFirst (List.map (subst_tactic subst) l) | TacSolve l -> TacSolve (List.map (subst_tactic subst) l) | TacComplete tac -> TacComplete (subst_tactic subst tac) | TacArg { CAst.v=a } -> TacArg (CAst.make @@ subst_tacarg subst a) | TacSelect (s, tac) -> TacSelect (s, subst_tactic subst tac) (* For extensions *) | TacAlias { CAst.v=(s,l) } -> let s = subst_kn subst s in TacAlias (CAst.make (s,List.map (subst_tacarg subst) l)) | TacML { CAst.loc; v=(opn,l)} -> TacML CAst.(make ?loc (opn,List.map (subst_tacarg subst) l)) and subst_tactic_fun subst (var,body) = (var,subst_tactic subst body) and subst_tacarg subst = function | Reference r -> Reference (subst_reference subst r) | ConstrMayEval c -> ConstrMayEval (subst_raw_may_eval subst c) | TacCall { CAst.loc; v=(f,l) } -> TacCall CAst.(make ?loc (subst_reference subst f, List.map (subst_tacarg subst) l)) | TacFreshId _ as x -> x | TacPretype c -> TacPretype (subst_glob_constr subst c) | TacNumgoals -> TacNumgoals | Tacexp t -> Tacexp (subst_tactic subst t) | TacGeneric arg -> TacGeneric (subst_genarg subst arg) (* Reads the rules of a Match Context or a Match *) and subst_match_rule subst = function | (All tc)::tl -> (All (subst_tactic subst tc))::(subst_match_rule subst tl) | (Pat (rl,mp,tc))::tl -> let hyps = subst_match_goal_hyps subst rl in let pat = subst_match_pattern subst mp in Pat (hyps,pat,subst_tactic subst tc) ::(subst_match_rule subst tl) | [] -> [] and subst_genarg subst (GenArg (Glbwit wit, x)) = match wit with | ListArg wit -> let map x = let ans = subst_genarg subst (in_gen (glbwit wit) x) in out_gen (glbwit wit) ans in in_gen (glbwit (wit_list wit)) (List.map map x) | OptArg wit -> let ans = match x with | None -> in_gen (glbwit (wit_opt wit)) None | Some x -> let s = out_gen (glbwit wit) (subst_genarg subst (in_gen (glbwit wit) x)) in in_gen (glbwit (wit_opt wit)) (Some s) in ans | PairArg (wit1, wit2) -> let p, q = x in let p = out_gen (glbwit wit1) (subst_genarg subst (in_gen (glbwit wit1) p)) in let q = out_gen (glbwit wit2) (subst_genarg subst (in_gen (glbwit wit2) q)) in in_gen (glbwit (wit_pair wit1 wit2)) (p, q) | ExtraArg s -> Genintern.generic_substitute subst (in_gen (glbwit wit) x) (** Registering *) let () = Genintern.register_subst0 wit_int_or_var (fun _ v -> v); Genintern.register_subst0 wit_ref subst_global_reference; Genintern.register_subst0 wit_pre_ident (fun _ v -> v); Genintern.register_subst0 wit_ident (fun _ v -> v); Genintern.register_subst0 wit_var (fun _ v -> v); Genintern.register_subst0 wit_intropattern subst_intro_pattern [@warning "-3"]; Genintern.register_subst0 wit_simple_intropattern subst_intro_pattern; Genintern.register_subst0 wit_tactic subst_tactic; Genintern.register_subst0 wit_ltac subst_tactic; Genintern.register_subst0 wit_constr subst_glob_constr; Genintern.register_subst0 wit_clause_dft_concl (fun _ v -> v); Genintern.register_subst0 wit_uconstr (fun subst c -> subst_glob_constr subst c); Genintern.register_subst0 wit_open_constr (fun subst c -> subst_glob_constr subst c); Genintern.register_subst0 wit_red_expr subst_redexp; Genintern.register_subst0 wit_quant_hyp subst_declared_or_quantified_hypothesis; Genintern.register_subst0 wit_bindings subst_bindings; Genintern.register_subst0 wit_constr_with_bindings subst_glob_with_bindings; Genintern.register_subst0 wit_destruction_arg subst_destruction_arg; () coq-8.11.0/plugins/ltac/extratactics.mlg0000644000175000017500000010542313612664533020076 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* replace_in_clause_maybe_by c1 c2 cl (Option.map (Tacinterp.tactic_of_value ist) tac)) let replace_term ist dir_opt c cl = with_delayed_uconstr ist c (fun c -> replace_term dir_opt c cl) } TACTIC EXTEND replace | ["replace" uconstr(c1) "with" constr(c2) clause(cl) by_arg_tac(tac) ] -> { replace_in_clause_maybe_by ist c1 c2 cl tac } END TACTIC EXTEND replace_term_left | [ "replace" "->" uconstr(c) clause(cl) ] -> { replace_term ist (Some true) c cl } END TACTIC EXTEND replace_term_right | [ "replace" "<-" uconstr(c) clause(cl) ] -> { replace_term ist (Some false) c cl } END TACTIC EXTEND replace_term | [ "replace" uconstr(c) clause(cl) ] -> { replace_term ist None c cl } END { let induction_arg_of_quantified_hyp = function | AnonHyp n -> None,ElimOnAnonHyp n | NamedHyp id -> None,ElimOnIdent (CAst.make id) (* Versions *_main must come first!! so that "1" is interpreted as a ElimOnAnonHyp and not as a "constr", and "id" is interpreted as a ElimOnIdent and not as "constr" *) let mytclWithHoles tac with_evars c = Proofview.Goal.enter begin fun gl -> let env = Tacmach.New.pf_env gl in let sigma = Tacmach.New.project gl in let sigma',c = Tactics.force_destruction_arg with_evars env sigma c in Tacticals.New.tclWITHHOLES with_evars (tac with_evars (Some c)) sigma' end let elimOnConstrWithHoles tac with_evars c = Tacticals.New.tclDELAYEDWITHHOLES with_evars c (fun c -> tac with_evars (Some (None,ElimOnConstr c))) } TACTIC EXTEND simplify_eq | [ "simplify_eq" ] -> { dEq ~keep_proofs:None false None } | [ "simplify_eq" destruction_arg(c) ] -> { mytclWithHoles (dEq ~keep_proofs:None) false c } END TACTIC EXTEND esimplify_eq | [ "esimplify_eq" ] -> { dEq ~keep_proofs:None true None } | [ "esimplify_eq" destruction_arg(c) ] -> { mytclWithHoles (dEq ~keep_proofs:None) true c } END { let discr_main c = elimOnConstrWithHoles discr_tac false c } TACTIC EXTEND discriminate | [ "discriminate" ] -> { discr_tac false None } | [ "discriminate" destruction_arg(c) ] -> { mytclWithHoles discr_tac false c } END TACTIC EXTEND ediscriminate | [ "ediscriminate" ] -> { discr_tac true None } | [ "ediscriminate" destruction_arg(c) ] -> { mytclWithHoles discr_tac true c } END { let discrHyp id = Proofview.tclEVARMAP >>= fun sigma -> discr_main (fun env sigma -> (sigma, (EConstr.mkVar id, NoBindings))) let injection_main with_evars c = elimOnConstrWithHoles (injClause None None) with_evars c let isInjPat pat = match pat.CAst.v with IntroAction (IntroInjection _) -> Some pat.CAst.loc | _ -> None let decode_inj_ipat ?loc = function (* For the "as [= pat1 ... patn ]" syntax *) | [{ CAst.v = IntroAction (IntroInjection ipat) }] -> ipat (* For the "as pat1 ... patn" syntax *) | ([] | [_]) as ipat -> ipat | pat1::pat2::_ as ipat -> (* To be sure that there is no confusion of syntax, we check that no [= ...] occurs in the non-singleton list of patterns *) match isInjPat pat1 with | Some _ -> user_err ?loc:pat2.CAst.loc (str "Unexpected pattern.") | None -> match List.map_filter isInjPat ipat with | loc :: _ -> user_err ?loc (str "Unexpected injection pattern.") | [] -> ipat } TACTIC EXTEND injection | [ "injection" ] -> { injClause None None false None } | [ "injection" destruction_arg(c) ] -> { mytclWithHoles (injClause None None) false c } END TACTIC EXTEND einjection | [ "einjection" ] -> { injClause None None true None } | [ "einjection" destruction_arg(c) ] -> { mytclWithHoles (injClause None None) true c } END TACTIC EXTEND injection_as | [ "injection" "as" simple_intropattern_list(ipat)] -> { injClause None (Some (decode_inj_ipat ipat)) false None } | [ "injection" destruction_arg(c) "as" simple_intropattern_list(ipat)] -> { mytclWithHoles (injClause None (Some (decode_inj_ipat ipat))) false c } END TACTIC EXTEND einjection_as | [ "einjection" "as" simple_intropattern_list(ipat)] -> { injClause None (Some (decode_inj_ipat ipat)) true None } | [ "einjection" destruction_arg(c) "as" simple_intropattern_list(ipat)] -> { mytclWithHoles (injClause None (Some (decode_inj_ipat ipat))) true c } END TACTIC EXTEND simple_injection | [ "simple" "injection" ] -> { simpleInjClause None false None } | [ "simple" "injection" destruction_arg(c) ] -> { mytclWithHoles (simpleInjClause None) false c } END { let injHyp id = Proofview.tclEVARMAP >>= fun sigma -> injection_main false (fun env sigma -> (sigma, (EConstr.mkVar id, NoBindings))) } TACTIC EXTEND dependent_rewrite | [ "dependent" "rewrite" orient(b) constr(c) ] -> { rewriteInConcl b c } | [ "dependent" "rewrite" orient(b) constr(c) "in" hyp(id) ] -> { rewriteInHyp b c id } END TACTIC EXTEND cut_rewrite | [ "cutrewrite" orient(b) constr(eqn) ] -> { cutRewriteInConcl b eqn } | [ "cutrewrite" orient(b) constr(eqn) "in" hyp(id) ] -> { cutRewriteInHyp b eqn id } END (**********************************************************************) (* Decompose *) TACTIC EXTEND decompose_sum | [ "decompose" "sum" constr(c) ] -> { Elim.h_decompose_or c } END TACTIC EXTEND decompose_record | [ "decompose" "record" constr(c) ] -> { Elim.h_decompose_and c } END (**********************************************************************) (* Contradiction *) { open Contradiction } TACTIC EXTEND absurd | [ "absurd" constr(c) ] -> { absurd c } END { let onSomeWithHoles tac = function | None -> tac None | Some c -> Tacticals.New.tclDELAYEDWITHHOLES false c (fun c -> tac (Some c)) } TACTIC EXTEND contradiction | [ "contradiction" constr_with_bindings_opt(c) ] -> { onSomeWithHoles contradiction c } END (**********************************************************************) (* AutoRewrite *) { open Autorewrite } TACTIC EXTEND autorewrite | [ "autorewrite" "with" ne_preident_list(l) clause(cl) ] -> { auto_multi_rewrite l ( cl) } | [ "autorewrite" "with" ne_preident_list(l) clause(cl) "using" tactic(t) ] -> { auto_multi_rewrite_with (Tacinterp.tactic_of_value ist t) l cl } END TACTIC EXTEND autorewrite_star | [ "autorewrite" "*" "with" ne_preident_list(l) clause(cl) ] -> { auto_multi_rewrite ~conds:AllMatches l cl } | [ "autorewrite" "*" "with" ne_preident_list(l) clause(cl) "using" tactic(t) ] -> { auto_multi_rewrite_with ~conds:AllMatches (Tacinterp.tactic_of_value ist t) l cl } END (**********************************************************************) (* Rewrite star *) { let rewrite_star ist clause orient occs c (tac : Geninterp.Val.t option) = let tac' = Option.map (fun t -> Tacinterp.tactic_of_value ist t, FirstSolved) tac in with_delayed_uconstr ist c (fun c -> general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (c,NoBindings) true) } TACTIC EXTEND rewrite_star | [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) "at" occurrences(occ) by_arg_tac(tac) ] -> { rewrite_star ist (Some id) o (occurrences_of occ) c tac } | [ "rewrite" "*" orient(o) uconstr(c) "at" occurrences(occ) "in" hyp(id) by_arg_tac(tac) ] -> { rewrite_star ist (Some id) o (occurrences_of occ) c tac } | [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) by_arg_tac(tac) ] -> { rewrite_star ist (Some id) o Locus.AllOccurrences c tac } | [ "rewrite" "*" orient(o) uconstr(c) "at" occurrences(occ) by_arg_tac(tac) ] -> { rewrite_star ist None o (occurrences_of occ) c tac } | [ "rewrite" "*" orient(o) uconstr(c) by_arg_tac(tac) ] -> { rewrite_star ist None o Locus.AllOccurrences c tac } END (**********************************************************************) (* Hint Rewrite *) { let add_rewrite_hint ~poly bases ort t lcsr = let env = Global.env() in let sigma = Evd.from_env env in let f ce = let c, ctx = Constrintern.interp_constr env sigma ce in let c = EConstr.to_constr sigma c in let ctx = let ctx = UState.context_set ctx in if poly then ctx else (* This is a global universe context that shouldn't be refreshed at every use of the hint, declare it globally. *) (Declare.declare_universe_context ~poly:false ctx; Univ.ContextSet.empty) in CAst.make ?loc:(Constrexpr_ops.constr_loc ce) ((c, ctx), ort, Option.map (in_gen (rawwit wit_ltac)) t) in let eqs = List.map f lcsr in let add_hints base = add_rew_rules base eqs in List.iter add_hints bases let classify_hint _ = VtSideff ([], VtLater) } VERNAC COMMAND EXTEND HintRewrite CLASSIFIED BY { classify_hint } | #[ polymorphic; ] [ "Hint" "Rewrite" orient(o) ne_constr_list(l) ":" preident_list(bl) ] -> { add_rewrite_hint ~poly:polymorphic bl o None l } | #[ polymorphic; ] [ "Hint" "Rewrite" orient(o) ne_constr_list(l) "using" tactic(t) ":" preident_list(bl) ] -> { add_rewrite_hint ~poly:polymorphic bl o (Some t) l } | #[ polymorphic; ] [ "Hint" "Rewrite" orient(o) ne_constr_list(l) ] -> { add_rewrite_hint ~poly:polymorphic ["core"] o None l } | #[ polymorphic; ] [ "Hint" "Rewrite" orient(o) ne_constr_list(l) "using" tactic(t) ] -> { add_rewrite_hint ~poly:polymorphic ["core"] o (Some t) l } END (**********************************************************************) (* Refine *) { open EConstr open Vars let constr_flags () = { Pretyping.use_typeclasses = true; Pretyping.solve_unification_constraints = Pfedit.use_unification_heuristics (); Pretyping.fail_evar = false; Pretyping.expand_evars = true; Pretyping.program_mode = false; Pretyping.polymorphic = false; } let refine_tac ist simple with_classes c = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let flags = { (constr_flags ()) with Pretyping.use_typeclasses = with_classes } in let expected_type = Pretyping.OfType concl in let c = Tacinterp.type_uconstr ~flags ~expected_type ist c in let update = begin fun sigma -> c env sigma end in let refine = Refine.refine ~typecheck:false update in if simple then refine else refine <*> Tactics.New.reduce_after_refine <*> Proofview.shelve_unifiable end } TACTIC EXTEND refine | [ "refine" uconstr(c) ] -> { refine_tac ist false true c } END TACTIC EXTEND simple_refine | [ "simple" "refine" uconstr(c) ] -> { refine_tac ist true true c } END TACTIC EXTEND notcs_refine | [ "notypeclasses" "refine" uconstr(c) ] -> { refine_tac ist false false c } END TACTIC EXTEND notcs_simple_refine | [ "simple" "notypeclasses" "refine" uconstr(c) ] -> { refine_tac ist true false c } END (* Solve unification constraints using heuristics or fail if any remain *) TACTIC EXTEND solve_constraints | [ "solve_constraints" ] -> { Refine.solve_constraints } END (**********************************************************************) (* Inversion lemmas (Leminv) *) { open Inv open Leminv let seff id = VtSideff ([id], VtLater) } (*VERNAC ARGUMENT EXTEND sort_family | [ "Set" ] -> { InSet } | [ "Prop" ] -> { InProp } | [ "Type" ] -> { InType } END*) VERNAC COMMAND EXTEND DeriveInversionClear | #[ polymorphic; ] [ "Derive" "Inversion_clear" ident(na) "with" constr(c) "Sort" sort_family(s) ] => { seff na } -> { add_inversion_lemma_exn ~poly:polymorphic na c s false inv_clear_tac } | #[ polymorphic; ] [ "Derive" "Inversion_clear" ident(na) "with" constr(c) ] => { seff na } -> { add_inversion_lemma_exn ~poly:polymorphic na c Sorts.InProp false inv_clear_tac } END VERNAC COMMAND EXTEND DeriveInversion | #[ polymorphic; ] [ "Derive" "Inversion" ident(na) "with" constr(c) "Sort" sort_family(s) ] => { seff na } -> { add_inversion_lemma_exn ~poly:polymorphic na c s false inv_tac } | #[ polymorphic; ] [ "Derive" "Inversion" ident(na) "with" constr(c) ] => { seff na } -> { add_inversion_lemma_exn ~poly:polymorphic na c Sorts.InProp false inv_tac } END VERNAC COMMAND EXTEND DeriveDependentInversion | #[ polymorphic; ] [ "Derive" "Dependent" "Inversion" ident(na) "with" constr(c) "Sort" sort_family(s) ] => { seff na } -> { add_inversion_lemma_exn ~poly:polymorphic na c s true dinv_tac } END VERNAC COMMAND EXTEND DeriveDependentInversionClear | #[ polymorphic; ] [ "Derive" "Dependent" "Inversion_clear" ident(na) "with" constr(c) "Sort" sort_family(s) ] => { seff na } -> { add_inversion_lemma_exn ~poly:polymorphic na c s true dinv_clear_tac } END (**********************************************************************) (* Subst *) TACTIC EXTEND subst | [ "subst" ne_var_list(l) ] -> { subst l } | [ "subst" ] -> { subst_all () } END { let simple_subst_tactic_flags = { only_leibniz = true; rewrite_dependent_proof = false } } TACTIC EXTEND simple_subst | [ "simple" "subst" ] -> { subst_all ~flags:simple_subst_tactic_flags () } END { open Evar_tactics } (**********************************************************************) (* Evar creation *) (* TODO: add support for some test similar to g_constr.name_colon so that expressions like "evar (list A)" do not raise a syntax error *) TACTIC EXTEND evar | [ "evar" test_lpar_id_colon "(" ident(id) ":" lconstr(typ) ")" ] -> { let_evar (Name.Name id) typ } | [ "evar" constr(typ) ] -> { let_evar Name.Anonymous typ } END TACTIC EXTEND instantiate | [ "instantiate" "(" ident(id) ":=" lglob(c) ")" ] -> { Tacticals.New.tclTHEN (instantiate_tac_by_name id c) Proofview.V82.nf_evar_goals } | [ "instantiate" "(" integer(i) ":=" lglob(c) ")" hloc(hl) ] -> { Tacticals.New.tclTHEN (instantiate_tac i c hl) Proofview.V82.nf_evar_goals } | [ "instantiate" ] -> { Proofview.V82.nf_evar_goals } END (**********************************************************************) (** Nijmegen "step" tactic for setoid rewriting *) { open Tactics open Glob_term open Libobject open Lib (* Registered lemmas are expected to be of the form x R y -> y == z -> x R z (in the right table) x R y -> x == z -> z R y (in the left table) *) let transitivity_right_table = Summary.ref [] ~name:"transitivity-steps-r" let transitivity_left_table = Summary.ref [] ~name:"transitivity-steps-l" (* [step] tries to apply a rewriting lemma; then apply [tac] intended to complete to proof of the last hypothesis (assumed to state an equality) *) let step left x tac = let l = List.map (fun lem -> let lem = EConstr.of_constr lem in Tacticals.New.tclTHENLAST (apply_with_bindings (lem, ImplicitBindings [x])) tac) !(if left then transitivity_left_table else transitivity_right_table) in Tacticals.New.tclFIRST l (* Main function to push lemmas in persistent environment *) let cache_transitivity_lemma (_,(left,lem)) = if left then transitivity_left_table := lem :: !transitivity_left_table else transitivity_right_table := lem :: !transitivity_right_table let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref) let inTransitivity : bool * Constr.t -> obj = declare_object @@ global_object_nodischarge "TRANSITIVITY-STEPS" ~cache:cache_transitivity_lemma ~subst:(Some subst_transitivity_lemma) (* Main entry points *) let add_transitivity_lemma left lem = let env = Global.env () in let sigma = Evd.from_env env in let lem',ctx (*FIXME*) = Constrintern.interp_constr env sigma lem in let lem' = EConstr.to_constr sigma lem' in add_anonymous_leaf (inTransitivity (left,lem')) } (* Vernacular syntax *) TACTIC EXTEND stepl | ["stepl" constr(c) "by" tactic(tac) ] -> { step true c (Tacinterp.tactic_of_value ist tac) } | ["stepl" constr(c) ] -> { step true c (Proofview.tclUNIT ()) } END TACTIC EXTEND stepr | ["stepr" constr(c) "by" tactic(tac) ] -> { step false c (Tacinterp.tactic_of_value ist tac) } | ["stepr" constr(c) ] -> { step false c (Proofview.tclUNIT ()) } END VERNAC COMMAND EXTEND AddStepl CLASSIFIED AS SIDEFF | [ "Declare" "Left" "Step" constr(t) ] -> { add_transitivity_lemma true t } END VERNAC COMMAND EXTEND AddStepr CLASSIFIED AS SIDEFF | [ "Declare" "Right" "Step" constr(t) ] -> { add_transitivity_lemma false t } END (**********************************************************************) (* sozeau: abs/gen for induction on instantiated dependent inductives, using "Ford" induction as defined by Conor McBride *) TACTIC EXTEND generalize_eqs | ["generalize_eqs" hyp(id) ] -> { abstract_generalize ~generalize_vars:false id } END TACTIC EXTEND dep_generalize_eqs | ["dependent" "generalize_eqs" hyp(id) ] -> { abstract_generalize ~generalize_vars:false ~force_dep:true id } END TACTIC EXTEND generalize_eqs_vars | ["generalize_eqs_vars" hyp(id) ] -> { abstract_generalize ~generalize_vars:true id } END TACTIC EXTEND dep_generalize_eqs_vars | ["dependent" "generalize_eqs_vars" hyp(id) ] -> { abstract_generalize ~force_dep:true ~generalize_vars:true id } END (** Tactic to automatically simplify hypotheses of the form [Π Δ, x_i = t_i -> T] where [t_i] is closed w.r.t. Δ. Such hypotheses are automatically generated during dependent induction. For internal use. *) TACTIC EXTEND specialize_eqs | [ "specialize_eqs" hyp(id) ] -> { specialize_eqs id } END (**********************************************************************) (* A tactic that considers a given occurrence of [c] in [t] and *) (* abstract the minimal set of all the occurrences of [c] so that the *) (* abstraction [fun x -> t[x/c]] is well-typed *) (* *) (* Contributed by Chung-Kil Hur (Winter 2009) *) (**********************************************************************) { let subst_var_with_hole occ tid t = let occref = if occ > 0 then ref occ else Find_subterm.error_invalid_occurrence [occ] in let locref = ref 0 in let rec substrec x = match DAst.get x with | GVar id -> if Id.equal id tid then (decr occref; if Int.equal !occref 0 then x else (incr locref; DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ GHole (Evar_kinds.QuestionMark { Evar_kinds.qm_obligation=Evar_kinds.Define true; Evar_kinds.qm_name=Anonymous; Evar_kinds.qm_record_field=None; }, IntroAnonymous, None))) else x | _ -> map_glob_constr_left_to_right substrec x in let t' = substrec t in if !occref > 0 then Find_subterm.error_invalid_occurrence [occ] else t' let subst_hole_with_term occ tc t = let locref = ref 0 in let occref = ref occ in let rec substrec c = match DAst.get c with | GHole (Evar_kinds.QuestionMark { Evar_kinds.qm_obligation=Evar_kinds.Define true; Evar_kinds.qm_name=Anonymous; Evar_kinds.qm_record_field=None; }, IntroAnonymous, s) -> decr occref; if Int.equal !occref 0 then tc else (incr locref; DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ GHole (Evar_kinds.QuestionMark { Evar_kinds.qm_obligation=Evar_kinds.Define true; Evar_kinds.qm_name=Anonymous; Evar_kinds.qm_record_field=None; },IntroAnonymous,s)) | _ -> map_glob_constr_left_to_right substrec c in substrec t open Tacmach let hResolve id c occ t = Proofview.Goal.enter begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Termops.clear_named_body id (Proofview.Goal.env gl) in let concl = Proofview.Goal.concl gl in let env_ids = Termops.vars_of_env env in let c_raw = Detyping.detype Detyping.Now true env_ids env sigma c in let t_raw = Detyping.detype Detyping.Now true env_ids env sigma t in let rec resolve_hole t_hole = try Pretyping.understand env sigma t_hole with | Pretype_errors.PretypeError (_,_,Pretype_errors.UnsolvableImplicit _) as e -> let (e, info) = CErrors.push e in let loc_begin = Option.cata (fun l -> fst (Loc.unloc l)) 0 (Loc.get_loc info) in resolve_hole (subst_hole_with_term loc_begin c_raw t_hole) in let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in let sigma = Evd.merge_universe_context sigma ctx in let t_constr_type = Retyping.get_type_of env sigma t_constr in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (change_concl (mkLetIn (make_annot Name.Anonymous Sorts.Relevant,t_constr,t_constr_type,concl))) end let hResolve_auto id c t = let rec resolve_auto n = try hResolve id c n t with | UserError _ as e -> raise e | e when CErrors.noncritical e -> resolve_auto (n+1) in resolve_auto 1 } TACTIC EXTEND hresolve_core | [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "at" int_or_var(occ) "in" constr(t) ] -> { hResolve id c occ t } | [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "in" constr(t) ] -> { hResolve_auto id c t } END (** hget_evar *) TACTIC EXTEND hget_evar | [ "hget_evar" int_or_var(n) ] -> { Evar_tactics.hget_evar n } END (**********************************************************************) (**********************************************************************) (* A tactic that reduces one match t with ... by doing destruct t. *) (* if t is not a variable, the tactic does *) (* case_eq t;intros ... heq;rewrite heq in *|-. (but heq itself is *) (* preserved). *) (* Contributed by Julien Forest and Pierre Courtieu (july 2010) *) (**********************************************************************) { exception Found of unit Proofview.tactic let rewrite_except h = Proofview.Goal.enter begin fun gl -> let hyps = Tacmach.New.pf_ids_of_hyps gl in Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false)) hyps end let refl_equal () = Coqlib.lib_ref "core.eq.type" (* This is simply an implementation of the case_eq tactic. this code should be replaced by a call to the tactic but I don't know how to call it before it is defined. *) let mkCaseEq a : unit Proofview.tactic = Proofview.Goal.enter begin fun gl -> let type_of_a = Tacmach.New.pf_unsafe_type_of gl a in Tacticals.New.pf_constr_of_global (delayed_force refl_equal) >>= fun req -> Tacticals.New.tclTHENLIST [Tactics.generalize [(mkApp(req, [| type_of_a; a|]))]; Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in (* FIXME: this looks really wrong. Does anybody really use this tactic? *) let (_, c) = Tacred.pattern_occs [Locus.OnlyOccurrences [1], a] env (Evd.from_env env) concl in change_concl c end; simplest_case a] end let case_eq_intros_rewrite x = Proofview.Goal.enter begin fun gl -> let n = nb_prod (Tacmach.New.project gl) (Proofview.Goal.concl gl) in (* Pp.msgnl (Printer.pr_lconstr x); *) Tacticals.New.tclTHENLIST [ mkCaseEq x; Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let hyps = Tacmach.New.pf_ids_set_of_hyps gl in let n' = nb_prod (Tacmach.New.project gl) concl in let h = fresh_id_in_env hyps (Id.of_string "heq") (Proofview.Goal.env gl) in Tacticals.New.tclTHENLIST [ Tacticals.New.tclDO (n'-n-1) intro; introduction h; rewrite_except h] end ] end let rec find_a_destructable_match sigma t = let cl = induction_arg_of_quantified_hyp (NamedHyp (Id.of_string "x")) in let cl = [cl, (None, None), None], None in let dest = TacAtom (CAst.make @@ TacInductionDestruct(false, false, cl)) in match EConstr.kind sigma t with | Case (_,_,x,_) when closed0 sigma x -> if isVar sigma x then (* TODO check there is no rel n. *) raise (Found (Tacinterp.eval_tactic dest)) else (* let _ = Pp.msgnl (Printer.pr_lconstr x) in *) raise (Found (case_eq_intros_rewrite x)) | _ -> EConstr.iter sigma (fun c -> find_a_destructable_match sigma c) t let destauto t = Proofview.tclEVARMAP >>= fun sigma -> try find_a_destructable_match sigma t; Tacticals.New.tclZEROMSG (str "No destructable match found") with Found tac -> tac let destauto_in id = Proofview.Goal.enter begin fun gl -> let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in (* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) (* Pp.msgnl (Printer.pr_lconstr (ctype)); *) destauto ctype end } TACTIC EXTEND destauto | [ "destauto" ] -> { Proofview.Goal.enter begin fun gl -> destauto (Proofview.Goal.concl gl) end } | [ "destauto" "in" hyp(id) ] -> { destauto_in id } END (**********************************************************************) (**********************************************************************) (* A version of abstract constructing transparent terms *) (* Introduced by Jason Gross and Benjamin Delaware in June 2016 *) (**********************************************************************) TACTIC EXTEND transparent_abstract | [ "transparent_abstract" tactic3(t) ] -> { Proofview.Goal.enter begin fun gl -> Abstract.tclABSTRACT ~opaque:false None (Tacinterp.tactic_of_value ist t) end; } | [ "transparent_abstract" tactic3(t) "using" ident(id) ] -> { Proofview.Goal.enter begin fun gl -> Abstract.tclABSTRACT ~opaque:false (Some id) (Tacinterp.tactic_of_value ist t) end; } END (* ********************************************************************* *) TACTIC EXTEND constr_eq | [ "constr_eq" constr(x) constr(y) ] -> { Tactics.constr_eq ~strict:false x y } END TACTIC EXTEND constr_eq_strict | [ "constr_eq_strict" constr(x) constr(y) ] -> { Tactics.constr_eq ~strict:true x y } END TACTIC EXTEND constr_eq_nounivs | [ "constr_eq_nounivs" constr(x) constr(y) ] -> { Proofview.tclEVARMAP >>= fun sigma -> if eq_constr_nounivs sigma x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") } END TACTIC EXTEND is_evar | [ "is_evar" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | Evar _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar") } END TACTIC EXTEND has_evar | [ "has_evar" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> if Evarutil.has_undefined_evars sigma x then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars") } END TACTIC EXTEND is_hyp | [ "is_var" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | Var _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (str "Not a variable or hypothesis") } END TACTIC EXTEND is_fix | [ "is_fix" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | Fix _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a fix definition") } END TACTIC EXTEND is_cofix | [ "is_cofix" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | CoFix _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a cofix definition") } END TACTIC EXTEND is_ind | [ "is_ind" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | Ind _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not an (co)inductive datatype") } END TACTIC EXTEND is_constructor | [ "is_constructor" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | Construct _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constructor") } END TACTIC EXTEND is_proj | [ "is_proj" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | Proj _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a primitive projection") } END TACTIC EXTEND is_const | [ "is_const" constr(x) ] -> { Proofview.tclEVARMAP >>= fun sigma -> match EConstr.kind sigma x with | Const _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constant") } END (* Command to grab the evars left unresolved at the end of a proof. *) (* spiwack: I put it in extratactics because it is somewhat tied with the semantics of the LCF-style tactics, hence with the classic tactic mode. *) VERNAC COMMAND EXTEND GrabEvars STATE proof | [ "Grab" "Existential" "Variables" ] => { classify_as_proofstep } -> { fun ~pstate -> Proof_global.map_proof (fun p -> Proof.V82.grab_evars p) pstate } END (* Shelves all the goals under focus. *) TACTIC EXTEND shelve | [ "shelve" ] -> { Proofview.shelve } END (* Shelves the unifiable goals under focus, i.e. the goals which appear in other goals under focus (the unfocused goals are not considered). *) TACTIC EXTEND shelve_unifiable | [ "shelve_unifiable" ] -> { Proofview.shelve_unifiable } END (* Unshelves the goal shelved by the tactic. *) TACTIC EXTEND unshelve | [ "unshelve" tactic1(t) ] -> { Proofview.with_shelf (Tacinterp.tactic_of_value ist t) >>= fun (gls, ()) -> let gls = List.map Proofview.with_empty_state gls in Proofview.Unsafe.tclGETGOALS >>= fun ogls -> Proofview.Unsafe.tclSETGOALS (gls @ ogls) } END (* Command to add every unshelved variables to the focus *) VERNAC COMMAND EXTEND Unshelve STATE proof | [ "Unshelve" ] => { classify_as_proofstep } -> { fun ~pstate -> Proof_global.map_proof (fun p -> Proof.unshelve p) pstate } END (* Gives up on the goals under focus: the goals are considered solved, but the proof cannot be closed until the user goes back and solve these goals. *) TACTIC EXTEND give_up | [ "give_up" ] -> { Proofview.give_up } END (* cycles [n] goals *) TACTIC EXTEND cycle | [ "cycle" int_or_var(n) ] -> { Proofview.cycle n } END (* swaps goals number [i] and [j] *) TACTIC EXTEND swap | [ "swap" int_or_var(i) int_or_var(j) ] -> { Proofview.swap i j } END (* reverses the list of focused goals *) TACTIC EXTEND revgoals | [ "revgoals" ] -> { Proofview.revgoals } END { type cmp = | Eq | Lt | Le | Gt | Ge type 'i test = | Test of cmp * 'i * 'i let pr_cmp = function | Eq -> Pp.str"=" | Lt -> Pp.str"<" | Le -> Pp.str"<=" | Gt -> Pp.str">" | Ge -> Pp.str">=" let pr_cmp' _prc _prlc _prt = pr_cmp let pr_test_gen f (Test(c,x,y)) = Pp.(f x ++ pr_cmp c ++ f y) let pr_test = pr_test_gen (Pputils.pr_or_var Pp.int) let pr_test' _prc _prlc _prt = pr_test let pr_itest = pr_test_gen Pp.int let pr_itest' _prc _prlc _prt = pr_itest } ARGUMENT EXTEND comparison PRINTED BY { pr_cmp' } | [ "=" ] -> { Eq } | [ "<" ] -> { Lt } | [ "<=" ] -> { Le } | [ ">" ] -> { Gt } | [ ">=" ] -> { Ge } END { let interp_test ist gls = function | Test (c,x,y) -> project gls , Test(c,Tacinterp.interp_int_or_var ist x,Tacinterp.interp_int_or_var ist y) } ARGUMENT EXTEND test PRINTED BY { pr_itest' } INTERPRETED BY { interp_test } RAW_PRINTED BY { pr_test' } GLOB_PRINTED BY { pr_test' } | [ int_or_var(x) comparison(c) int_or_var(y) ] -> { Test(c,x,y) } END { let interp_cmp = function | Eq -> Int.equal | Lt -> ((<):int->int->bool) | Le -> ((<=):int->int->bool) | Gt -> ((>):int->int->bool) | Ge -> ((>=):int->int->bool) let run_test = function | Test(c,x,y) -> interp_cmp c x y let guard tst = if run_test tst then Proofview.tclUNIT () else let msg = Pp.(str"Condition not satisfied:"++ws 1++(pr_itest tst)) in Tacticals.New.tclZEROMSG msg } TACTIC EXTEND guard | [ "guard" test(tst) ] -> { guard tst } END { let decompose l c = Proofview.Goal.enter begin fun gl -> let sigma = Tacmach.New.project gl in let to_ind c = if isInd sigma c then fst (destInd sigma c) else user_err Pp.(str "not an inductive type") in let l = List.map to_ind l in Elim.h_decompose l c end } TACTIC EXTEND decompose | [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> { decompose l c } END (** library/keys *) VERNAC COMMAND EXTEND Declare_keys CLASSIFIED AS SIDEFF | [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> { let get_key c = let env = Global.env () in let evd = Evd.from_env env in let (evd, c) = Constrintern.interp_open_constr env evd c in let kind c = EConstr.kind evd c in Keys.constr_key kind c in let k1 = get_key c in let k2 = get_key c' in match k1, k2 with | Some k1, Some k2 -> Keys.declare_equiv_keys k1 k2 | _ -> () } END VERNAC COMMAND EXTEND Print_keys CLASSIFIED AS QUERY | [ "Print" "Equivalent" "Keys" ] -> { Feedback.msg_notice (Keys.pr_keys Printer.pr_global) } END VERNAC COMMAND EXTEND OptimizeProof | ![ proof ] [ "Optimize" "Proof" ] => { classify_as_proofstep } -> { fun ~pstate -> Proof_global.compact_the_proof pstate } | [ "Optimize" "Heap" ] => { classify_as_proofstep } -> { Gc.compact () } END (** tactic analogous to "OPTIMIZE HEAP" *) { let tclOPTIMIZE_HEAP = Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> Gc.compact ())) } TACTIC EXTEND optimize_heap | [ "optimize_heap" ] -> { tclOPTIMIZE_HEAP } END coq-8.11.0/plugins/ltac/tacentries.ml0000644000175000017500000006603513612664533017377 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* entry_name (** Quite ad-hoc *) let get_tacentry n m = let check_lvl n = Int.equal m n && not (Int.equal m 5) (* Because tactic5 is at binder_tactic *) && not (Int.equal m 0) (* Because tactic0 is at simple_tactic *) in if check_lvl n then EntryName (rawwit Tacarg.wit_tactic, Aself) else if check_lvl (n + 1) then EntryName (rawwit Tacarg.wit_tactic, Anext) else EntryName (rawwit Tacarg.wit_tactic, atactic n) let get_separator = function | None -> user_err Pp.(str "Missing separator.") | Some sep -> sep let check_separator ?loc = function | None -> () | Some _ -> user_err ?loc (str "Separator is only for arguments with suffix _list_sep.") let rec parse_user_entry ?loc s sep = let l = String.length s in if l > 8 && coincide s "ne_" 0 && coincide s "_list" (l - 5) then let entry = parse_user_entry ?loc (String.sub s 3 (l-8)) None in check_separator ?loc sep; Ulist1 entry else if l > 12 && coincide s "ne_" 0 && coincide s "_list_sep" (l-9) then let entry = parse_user_entry ?loc (String.sub s 3 (l-12)) None in Ulist1sep (entry, get_separator sep) else if l > 5 && coincide s "_list" (l-5) then let entry = parse_user_entry ?loc (String.sub s 0 (l-5)) None in check_separator ?loc sep; Ulist0 entry else if l > 9 && coincide s "_list_sep" (l-9) then let entry = parse_user_entry ?loc (String.sub s 0 (l-9)) None in Ulist0sep (entry, get_separator sep) else if l > 4 && coincide s "_opt" (l-4) then let entry = parse_user_entry ?loc (String.sub s 0 (l-4)) None in check_separator ?loc sep; Uopt entry else if Int.equal l 7 && coincide s "tactic" 0 && '5' >= s.[6] && s.[6] >= '0' then let n = Char.code s.[6] - 48 in check_separator ?loc sep; Uentryl ("tactic", n) else let _ = check_separator ?loc sep in Uentry s let interp_entry_name interp symb = let rec eval = function | Ulist1 e -> Ulist1 (eval e) | Ulist1sep (e, sep) -> Ulist1sep (eval e, sep) | Ulist0 e -> Ulist0 (eval e) | Ulist0sep (e, sep) -> Ulist0sep (eval e, sep) | Uopt e -> Uopt (eval e) | Uentry s -> Uentry (interp s None) | Uentryl (s, n) -> Uentryl (interp s (Some n), n) in eval symb (**********************************************************************) (** Grammar declaration for Tactic Notation (Coq level) *) let get_tactic_entry n = if Int.equal n 0 then Pltac.simple_tactic, None else if Int.equal n 5 then Pltac.binder_tactic, None else if 1<=n && n<5 then Pltac.tactic_expr, Some (Gramlib.Gramext.Level (string_of_int n)) else user_err Pp.(str ("Invalid Tactic Notation level: "^(string_of_int n)^".")) (**********************************************************************) (** State of the grammar extensions *) type tactic_grammar = { tacgram_level : int; tacgram_prods : Pptactic.grammar_terminals; } (* Declaration of the tactic grammar rule *) let head_is_ident tg = match tg.tacgram_prods with | TacTerm _ :: _ -> true | _ -> false let rec prod_item_of_symbol lev = function | Extend.Ulist1 s -> let EntryName (Rawwit typ, e) = prod_item_of_symbol lev s in EntryName (Rawwit (ListArg typ), Alist1 e) | Extend.Ulist0 s -> let EntryName (Rawwit typ, e) = prod_item_of_symbol lev s in EntryName (Rawwit (ListArg typ), Alist0 e) | Extend.Ulist1sep (s, sep) -> let EntryName (Rawwit typ, e) = prod_item_of_symbol lev s in EntryName (Rawwit (ListArg typ), Alist1sep (e, Atoken (CLexer.terminal sep))) | Extend.Ulist0sep (s, sep) -> let EntryName (Rawwit typ, e) = prod_item_of_symbol lev s in EntryName (Rawwit (ListArg typ), Alist0sep (e, Atoken (CLexer.terminal sep))) | Extend.Uopt s -> let EntryName (Rawwit typ, e) = prod_item_of_symbol lev s in EntryName (Rawwit (OptArg typ), Aopt e) | Extend.Uentry arg -> let ArgT.Any tag = arg in let wit = ExtraArg tag in EntryName (Rawwit wit, Extend.Aentry (genarg_grammar wit)) | Extend.Uentryl (s, n) -> let ArgT.Any tag = s in assert (coincide (ArgT.repr tag) "tactic" 0); get_tacentry n lev (** Tactic grammar extensions *) let add_tactic_entry (kn, ml, tg) state = let open Tacexpr in let entry, pos = get_tactic_entry tg.tacgram_level in let mkact loc l = let map arg = (* HACK to handle especially the tactic(...) entry *) let wit = Genarg.rawwit Tacarg.wit_tactic in if Genarg.has_type arg wit && not ml then Tacexp (Genarg.out_gen wit arg) else TacGeneric arg in let l = List.map map l in (TacAlias (CAst.make ~loc (kn,l)):raw_tactic_expr) in let () = if Int.equal tg.tacgram_level 0 && not (head_is_ident tg) then user_err Pp.(str "Notation for simple tactic must start with an identifier.") in let map = function | TacTerm s -> GramTerminal s | TacNonTerm (loc, (s, ido)) -> let EntryName (typ, e) = prod_item_of_symbol tg.tacgram_level s in GramNonTerminal (Loc.tag ?loc @@ (typ, e)) in let prods = List.map map tg.tacgram_prods in let rules = make_rule mkact prods in let r = ExtendRule (entry, None, (pos, [(None, None, [rules])])) in ([r], state) let tactic_grammar = create_grammar_command "TacticGrammar" add_tactic_entry let extend_tactic_grammar kn ml ntn = extend_grammar_command tactic_grammar (kn, ml, ntn) (**********************************************************************) (* Tactic Notation *) let entry_names = ref String.Map.empty let register_tactic_notation_entry name entry = let entry = match entry with | ExtraArg arg -> ArgT.Any arg | _ -> assert false in entry_names := String.Map.add name entry !entry_names let interp_prod_item = function | TacTerm s -> TacTerm s | TacNonTerm (loc, ((nt, sep), ido)) -> let symbol = parse_user_entry ?loc nt sep in let interp s = function | None -> if String.Map.mem s !entry_names then String.Map.find s !entry_names else begin match ArgT.name s with | None -> user_err Pp.(str ("Unknown entry "^s^".")) | Some arg -> arg end | Some n -> (* FIXME: do better someday *) assert (String.equal s "tactic"); begin match Tacarg.wit_tactic with | ExtraArg tag -> ArgT.Any tag end in let symbol = interp_entry_name interp symbol in TacNonTerm (loc, (symbol, ido)) let make_fresh_key = let id = Summary.ref ~name:"TACTIC-NOTATION-COUNTER" 0 in fun prods -> let cur = incr id; !id in let map = function | TacTerm s -> s | TacNonTerm _ -> "#" in let prods = String.concat "_" (List.map map prods) in (* We embed the hash of the kernel name in the label so that the identifier should be mostly unique. This ensures that including two modules together won't confuse the corresponding labels. *) let hash = (cur lxor (ModPath.hash (Lib.current_mp ()))) land 0x7FFFFFFF in let lbl = Id.of_string_soft (Printf.sprintf "%s_%08X" prods hash) in Lib.make_kn lbl type tactic_grammar_obj = { tacobj_key : KerName.t; tacobj_local : locality_flag; tacobj_tacgram : tactic_grammar; tacobj_body : Tacenv.alias_tactic; tacobj_forml : bool; } let pprule pa = { Pptactic.pptac_level = pa.tacgram_level; pptac_prods = pa.tacgram_prods; } let check_key key = if Tacenv.check_alias key then user_err Pp.(str "Conflicting tactic notations keys. This can happen when including \ twice the same module.") let cache_tactic_notation (_, tobj) = let key = tobj.tacobj_key in let () = check_key key in Tacenv.register_alias key tobj.tacobj_body; extend_tactic_grammar key tobj.tacobj_forml tobj.tacobj_tacgram; Pptactic.declare_notation_tactic_pprule key (pprule tobj.tacobj_tacgram) let open_tactic_notation i (_, tobj) = let key = tobj.tacobj_key in if Int.equal i 1 && not tobj.tacobj_local then extend_tactic_grammar key tobj.tacobj_forml tobj.tacobj_tacgram let load_tactic_notation i (_, tobj) = let key = tobj.tacobj_key in let () = check_key key in (* Only add the printing and interpretation rules. *) Tacenv.register_alias key tobj.tacobj_body; Pptactic.declare_notation_tactic_pprule key (pprule tobj.tacobj_tacgram); if Int.equal i 1 && not tobj.tacobj_local then extend_tactic_grammar key tobj.tacobj_forml tobj.tacobj_tacgram let subst_tactic_notation (subst, tobj) = let open Tacenv in let alias = tobj.tacobj_body in { tobj with tacobj_key = Mod_subst.subst_kn subst tobj.tacobj_key; tacobj_body = { alias with alias_body = Tacsubst.subst_tactic subst alias.alias_body }; } let classify_tactic_notation tacobj = Substitute tacobj let inTacticGrammar : tactic_grammar_obj -> obj = declare_object {(default_object "TacticGrammar") with open_function = open_tactic_notation; load_function = load_tactic_notation; cache_function = cache_tactic_notation; subst_function = subst_tactic_notation; classify_function = classify_tactic_notation} let cons_production_parameter = function | TacTerm _ -> None | TacNonTerm (_, (_, ido)) -> ido let add_glob_tactic_notation local ~level ?deprecation prods forml ids tac = let parule = { tacgram_level = level; tacgram_prods = prods; } in let open Tacenv in let tacobj = { tacobj_key = make_fresh_key prods; tacobj_local = local; tacobj_tacgram = parule; tacobj_body = { alias_args = ids; alias_body = tac; alias_deprecation = deprecation }; tacobj_forml = forml; } in Lib.add_anonymous_leaf (inTacticGrammar tacobj) let add_tactic_notation local n ?deprecation prods e = let ids = List.map_filter cons_production_parameter prods in let prods = List.map interp_prod_item prods in let tac = Tacintern.glob_tactic_env ids (Global.env()) e in add_glob_tactic_notation local ~level:n ?deprecation prods false ids tac (**********************************************************************) (* ML Tactic entries *) exception NonEmptyArgument (** ML tactic notations whose use can be restricted to an identifier are added as true Ltac entries. *) let extend_atomic_tactic name entries = let open Tacexpr in let map_prod prods = let (hd, rem) = match prods with | TacTerm s :: rem -> (s, rem) | _ -> assert false (* Not handled by the ML extension syntax *) in let empty_value = function | TacTerm s -> raise NonEmptyArgument | TacNonTerm (_, (symb, _)) -> let EntryName (typ, e) = prod_item_of_symbol 0 symb in let Genarg.Rawwit wit = typ in let inj x = TacArg (CAst.make @@ TacGeneric (Genarg.in_gen typ x)) in let default = epsilon_value inj e in match default with | None -> raise NonEmptyArgument | Some def -> Tacintern.intern_tactic_or_tacarg (Genintern.empty_glob_sign Environ.empty_env) def in try Some (hd, List.map empty_value rem) with NonEmptyArgument -> None in let entries = List.map map_prod entries in let add_atomic i args = match args with | None -> () | Some (id, args) -> let args = List.map (fun a -> Tacexp a) args in let entry = { mltac_name = name; mltac_index = i } in let body = TacML (CAst.make (entry, args)) in Tacenv.register_ltac false false (Names.Id.of_string id) body in List.iteri add_atomic entries let add_ml_tactic_notation name ~level ?deprecation prods = let len = List.length prods in let iter i prods = let open Tacexpr in let get_id = function | TacTerm s -> None | TacNonTerm (_, (_, ido)) -> ido in let ids = List.map_filter get_id prods in let entry = { mltac_name = name; mltac_index = len - i - 1 } in let map id = Reference (Locus.ArgVar (CAst.make id)) in let tac = TacML (CAst.make (entry, List.map map ids)) in add_glob_tactic_notation false ~level ?deprecation prods true ids tac in List.iteri iter (List.rev prods); (* We call [extend_atomic_tactic] only for "basic tactics" (the ones at tactic_expr level 0) *) if Int.equal level 0 then extend_atomic_tactic name prods (**********************************************************************) (** Ltac quotations *) let ltac_quotations = ref String.Set.empty let create_ltac_quotation name cast (e, l) = let () = if String.Set.mem name !ltac_quotations then failwith ("Ltac quotation " ^ name ^ " already registered") in let () = ltac_quotations := String.Set.add name !ltac_quotations in let entry = match l with | None -> Aentry e | Some l -> Aentryl (e, string_of_int l) in (* let level = Some "1" in *) let level = None in let assoc = None in let rule = Next (Next (Next (Next (Next (Stop, Atoken (CLexer.terminal name)), Atoken (CLexer.terminal ":")), Atoken (CLexer.terminal "(")), entry), Atoken (CLexer.terminal ")")) in let action _ v _ _ _ loc = cast (Some loc, v) in let gram = (level, assoc, [Rule (rule, action)]) in Pcoq.grammar_extend Pltac.tactic_arg None (None, [gram]) (** Command *) type tacdef_kind = | NewTac of Id.t | UpdateTac of Tacexpr.ltac_constant let is_defined_tac kn = try ignore (Tacenv.interp_ltac kn); true with Not_found -> false let warn_unusable_identifier = CWarnings.create ~name:"unusable-identifier" ~category:"parsing" (fun id -> strbrk "The Ltac name" ++ spc () ++ Id.print id ++ spc () ++ strbrk "may be unusable because of a conflict with a notation.") let register_ltac local ?deprecation tacl = let map tactic_body = match tactic_body with | Tacexpr.TacticDefinition ({CAst.loc;v=id}, body) -> let kn = Lib.make_kn id in let id_pp = Id.print id in let () = if is_defined_tac kn then CErrors.user_err ?loc (str "There is already an Ltac named " ++ id_pp ++ str".") in let is_shadowed = try match Pcoq.parse_string Pltac.tactic (Id.to_string id) with | Tacexpr.TacArg _ -> false | _ -> true (* most probably TacAtom, i.e. a primitive tactic ident *) with e when CErrors.noncritical e -> true (* prim tactics with args, e.g. "apply" *) in let () = if is_shadowed then warn_unusable_identifier id in NewTac id, body | Tacexpr.TacticRedefinition (qid, body) -> let kn = try Tacenv.locate_tactic qid with Not_found -> CErrors.user_err ?loc:qid.CAst.loc (str "There is no Ltac named " ++ pr_qualid qid ++ str ".") in UpdateTac kn, body in let rfun = List.map map tacl in let recvars = let fold accu (op, _) = match op with | UpdateTac _ -> accu | NewTac id -> (Lib.make_path id, Lib.make_kn id) :: accu in List.fold_left fold [] rfun in let ist = Tacintern.make_empty_glob_sign () in let map (name, body) = let body = Flags.with_option Tacintern.strict_check (Tacintern.intern_tactic_or_tacarg ist) body in (name, body) in let defs () = (* Register locally the tactic to handle recursivity. This function affects the whole environment, so that we transactify it afterwards. *) let iter_rec (sp, kn) = Tacenv.push_tactic (Nametab.Until 1) sp kn in let () = List.iter iter_rec recvars in List.map map rfun in (* STATE XXX: Review what is going on here. Why does this needs protection? Why is not the STM level protection enough? Fishy *) let defs = States.with_state_protection defs () in let iter (def, tac) = match def with | NewTac id -> Tacenv.register_ltac false local id tac ?deprecation; Flags.if_verbose Feedback.msg_info (Id.print id ++ str " is defined") | UpdateTac kn -> Tacenv.redefine_ltac local kn tac ?deprecation; let name = Tacenv.shortest_qualid_of_tactic kn in Flags.if_verbose Feedback.msg_info (Libnames.pr_qualid name ++ str " is redefined") in List.iter iter defs (** Queries *) let print_ltacs () = let entries = KNmap.bindings (Tacenv.ltac_entries ()) in let sort (kn1, _) (kn2, _) = KerName.compare kn1 kn2 in let entries = List.sort sort entries in let map (kn, entry) = let qid = try Some (Tacenv.shortest_qualid_of_tactic kn) with Not_found -> None in match qid with | None -> None | Some qid -> Some (qid, entry.Tacenv.tac_body) in let entries = List.map_filter map entries in let pr_entry (qid, body) = let (l, t) = match body with | Tacexpr.TacFun (l, t) -> (l, t) | _ -> ([], body) in let pr_ltac_fun_arg n = spc () ++ Name.print n in hov 2 (pr_qualid qid ++ prlist pr_ltac_fun_arg l) in Feedback.msg_notice (prlist_with_sep fnl pr_entry entries) let locatable_ltac = "Ltac" let () = let open Prettyp in let locate qid = try Some (Tacenv.locate_tactic qid) with Not_found -> None in let locate_all = Tacenv.locate_extended_all_tactic in let shortest_qualid = Tacenv.shortest_qualid_of_tactic in let name kn = str "Ltac" ++ spc () ++ pr_path (Tacenv.path_of_tactic kn) in let print kn = let qid = qualid_of_path (Tacenv.path_of_tactic kn) in Tacintern.print_ltac qid in let about = name in register_locatable locatable_ltac { locate; locate_all; shortest_qualid; name; print; about; } let print_located_tactic qid = Feedback.msg_notice (Prettyp.print_located_other locatable_ltac qid) (** Grammar *) let () = let entries = [ AnyEntry Pltac.tactic_expr; AnyEntry Pltac.binder_tactic; AnyEntry Pltac.simple_tactic; AnyEntry Pltac.tactic_arg; ] in register_grammars_by_name "tactic" entries let get_identifier i = (* Workaround for badly-designed generic arguments lacking a closure *) Names.Id.of_string_soft (Printf.sprintf "$%i" i) type _ ty_sig = | TyNil : (Geninterp.interp_sign -> unit Proofview.tactic) ty_sig | TyIdent : string * 'r ty_sig -> 'r ty_sig | TyArg : ('a, 'b, 'c) Extend.ty_user_symbol * 'r ty_sig -> ('c -> 'r) ty_sig type ty_ml = TyML : 'r ty_sig * 'r -> ty_ml let rec untype_user_symbol : type a b c. (a,b,c) ty_user_symbol -> Genarg.ArgT.any user_symbol = fun tu -> match tu with | TUlist1 l -> Ulist1(untype_user_symbol l) | TUlist1sep(l,s) -> Ulist1sep(untype_user_symbol l, s) | TUlist0 l -> Ulist0(untype_user_symbol l) | TUlist0sep(l,s) -> Ulist0sep(untype_user_symbol l, s) | TUopt(o) -> Uopt(untype_user_symbol o) | TUentry a -> Uentry (Genarg.ArgT.Any a) | TUentryl (a,i) -> Uentryl (Genarg.ArgT.Any a,i) let rec clause_of_sign : type a. int -> a ty_sig -> Genarg.ArgT.any Extend.user_symbol grammar_tactic_prod_item_expr list = fun i sign -> match sign with | TyNil -> [] | TyIdent (s, sig') -> TacTerm s :: clause_of_sign i sig' | TyArg (a, sig') -> let id = Some (get_identifier i) in TacNonTerm (None, (untype_user_symbol a, id)) :: clause_of_sign (i + 1) sig' let clause_of_ty_ml = function | TyML (t,_) -> clause_of_sign 1 t let rec eval_sign : type a. a ty_sig -> a -> Geninterp.Val.t list -> Geninterp.interp_sign -> unit Proofview.tactic = fun sign tac -> match sign with | TyNil -> begin fun vals ist -> match vals with | [] -> tac ist | _ :: _ -> assert false end | TyIdent (s, sig') -> eval_sign sig' tac | TyArg (a, sig') -> let f = eval_sign sig' in begin fun tac vals ist -> match vals with | [] -> assert false | v :: vals -> let v' = Taccoerce.Value.cast (topwit (Egramml.proj_symbol a)) v in f (tac v') vals ist end tac let eval : ty_ml -> Geninterp.Val.t list -> Geninterp.interp_sign -> unit Proofview.tactic = function | TyML (t,tac) -> eval_sign t tac let is_constr_entry = function | TUentry a -> Option.has_some @@ genarg_type_eq (ExtraArg a) Stdarg.wit_constr | _ -> false let rec only_constr : type a. a ty_sig -> bool = function | TyNil -> true | TyIdent(_,_) -> false | TyArg (u, s) -> if is_constr_entry u then only_constr s else false let rec mk_sign_vars : type a. int -> a ty_sig -> Name.t list = fun i tu -> match tu with | TyNil -> [] | TyIdent (_,s) -> mk_sign_vars i s | TyArg (_, s) -> Name (get_identifier i) :: mk_sign_vars (i + 1) s let dummy_id = Id.of_string "_" let lift_constr_tac_to_ml_tac vars tac = let tac _ ist = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let map = function | Anonymous -> None | Name id -> let c = Id.Map.find id ist.Geninterp.lfun in try Some (Taccoerce.Value.of_constr @@ Taccoerce.coerce_to_closed_constr env c) with Taccoerce.CannotCoerceTo ty -> Taccoerce.error_ltac_variable dummy_id (Some (env,sigma)) c ty in let args = List.map_filter map vars in tac args ist end in tac let tactic_extend plugin_name tacname ~level ?deprecation sign = let open Tacexpr in let ml_tactic_name = { mltac_tactic = tacname; mltac_plugin = plugin_name } in match sign with | [TyML (TyIdent (name, s),tac) as ml_tac] when only_constr s -> (* The extension is only made of a name followed by constr entries: we do not add any grammar nor printing rule and add it as a true Ltac definition. *) let vars = mk_sign_vars 1 s in let ml = { Tacexpr.mltac_name = ml_tactic_name; Tacexpr.mltac_index = 0 } in let tac = match s with | TyNil -> eval ml_tac (* Special handling of tactics without arguments: such tactics do not do a Proofview.Goal.nf_enter to compute their arguments. It matters for some whole-prof tactics like [shelve_unifiable]. *) | _ -> lift_constr_tac_to_ml_tac vars (eval ml_tac) in (* Arguments are not passed directly to the ML tactic in the TacML node, the ML tactic retrieves its arguments in the [ist] environment instead. This is the rôle of the [lift_constr_tac_to_ml_tac] function. *) let body = Tacexpr.TacFun (vars, Tacexpr.TacML (CAst.make (ml, [])))in let id = Names.Id.of_string name in let obj () = Tacenv.register_ltac true false id body ?deprecation in let () = Tacenv.register_ml_tactic ml_tactic_name [|tac|] in Mltop.declare_cache_obj obj plugin_name | _ -> let obj () = add_ml_tactic_notation ml_tactic_name ~level ?deprecation (List.map clause_of_ty_ml sign) in Tacenv.register_ml_tactic ml_tactic_name @@ Array.of_list (List.map eval sign); Mltop.declare_cache_obj obj plugin_name (** ARGUMENT EXTEND *) open Geninterp type ('a, 'b, 'c) argument_printer = 'a Pptactic.raw_extra_genarg_printer * 'b Pptactic.glob_extra_genarg_printer * 'c Pptactic.extra_genarg_printer type ('a, 'b) argument_intern = | ArgInternFun : ('a, 'b) Genintern.intern_fun -> ('a, 'b) argument_intern | ArgInternWit : ('a, 'b, 'c) Genarg.genarg_type -> ('a, 'b) argument_intern type 'b argument_subst = | ArgSubstFun : 'b Genintern.subst_fun -> 'b argument_subst | ArgSubstWit : ('a, 'b, 'c) Genarg.genarg_type -> 'b argument_subst type ('b, 'c) argument_interp = | ArgInterpRet : ('c, 'c) argument_interp | ArgInterpFun : ('b, Val.t) interp_fun -> ('b, 'c) argument_interp | ArgInterpWit : ('a, 'b, 'r) Genarg.genarg_type -> ('b, 'c) argument_interp | ArgInterpLegacy : (Geninterp.interp_sign -> Goal.goal Evd.sigma -> 'b -> Evd.evar_map * 'c) -> ('b, 'c) argument_interp type ('a, 'b, 'c) tactic_argument = { arg_parsing : 'a Vernacextend.argument_rule; arg_tag : 'c Val.tag option; arg_intern : ('a, 'b) argument_intern; arg_subst : 'b argument_subst; arg_interp : ('b, 'c) argument_interp; arg_printer : ('a, 'b, 'c) argument_printer; } let intern_fun (type a b c) name (arg : (a, b, c) tactic_argument) : (a, b) Genintern.intern_fun = match arg.arg_intern with | ArgInternFun f -> f | ArgInternWit wit -> fun ist v -> let ans = Genarg.out_gen (glbwit wit) (Tacintern.intern_genarg ist (Genarg.in_gen (rawwit wit) v)) in (ist, ans) let subst_fun (type a b c) (arg : (a, b, c) tactic_argument) : b Genintern.subst_fun = match arg.arg_subst with | ArgSubstFun f -> f | ArgSubstWit wit -> fun s v -> let ans = Genarg.out_gen (glbwit wit) (Tacsubst.subst_genarg s (Genarg.in_gen (glbwit wit) v)) in ans let interp_fun (type a b c) name (arg : (a, b, c) tactic_argument) (tag : c Val.tag) : (b, Val.t) interp_fun = match arg.arg_interp with | ArgInterpRet -> (fun ist v -> Ftactic.return (Geninterp.Val.inject tag v)) | ArgInterpFun f -> f | ArgInterpWit wit -> (fun ist x -> Tacinterp.interp_genarg ist (Genarg.in_gen (glbwit wit) x)) | ArgInterpLegacy f -> (fun ist v -> Ftactic.enter (fun gl -> let (sigma, v) = Tacmach.New.of_old (fun gl -> f ist gl v) gl in let v = Geninterp.Val.inject tag v in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Ftactic.return v) )) let argument_extend (type a b c) ~name (arg : (a, b, c) tactic_argument) = let wit = Genarg.create_arg name in let () = Genintern.register_intern0 wit (intern_fun name arg) in let () = Genintern.register_subst0 wit (subst_fun arg) in let tag = match arg.arg_tag with | None -> let () = register_val0 wit None in val_tag (topwit wit) | Some tag -> let () = register_val0 wit (Some tag) in tag in let () = register_interp0 wit (interp_fun name arg tag) in let entry = match arg.arg_parsing with | Vernacextend.Arg_alias e -> let () = Pcoq.register_grammar wit e in e | Vernacextend.Arg_rules rules -> let e = Pcoq.create_generic_entry Pcoq.utactic name (Genarg.rawwit wit) in let () = Pcoq.grammar_extend e None (None, [(None, None, rules)]) in e in let (rpr, gpr, tpr) = arg.arg_printer in let () = Pptactic.declare_extra_genarg_pprule wit rpr gpr tpr in let () = create_ltac_quotation name (fun (loc, v) -> Tacexpr.TacGeneric (Genarg.in_gen (Genarg.rawwit wit) v)) (entry, None) in (wit, entry) coq-8.11.0/plugins/ltac/tacinterp.ml0000644000175000017500000024150413612664533017223 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* a typed_abstract_argument_type -> bool = fun v wit -> let Val.Dyn (t, _) = v in let t' = match val_tag wit with | Val.Base t' -> t' | _ -> assert false (* not used in this module *) in match Val.eq t t' with | None -> false | Some Refl -> true let prj : type a. a Val.typ -> Val.t -> a option = fun t v -> let Val.Dyn (t', x) = v in match Val.eq t t' with | None -> None | Some Refl -> Some x let in_list tag v = let tag = match tag with Val.Base tag -> tag | _ -> assert false in Val.Dyn (Val.typ_list, List.map (fun x -> Val.Dyn (tag, x)) v) let in_gen wit v = let t = match val_tag wit with | Val.Base t -> t | _ -> assert false (* not used in this module *) in Val.Dyn (t, v) let out_gen wit v = let t = match val_tag wit with | Val.Base t -> t | _ -> assert false (* not used in this module *) in match prj t v with None -> assert false | Some x -> x let val_tag wit = val_tag (topwit wit) let pr_argument_type arg = let Val.Dyn (tag, _) = arg in Val.pr tag let safe_msgnl s = Proofview.NonLogical.catch (Proofview.NonLogical.print_debug (s++fnl())) (fun _ -> Proofview.NonLogical.print_warning (str "bug in the debugger: an exception is raised while printing debug information"++fnl())) type value = Val.t let push_appl appl args = match appl with | UnnamedAppl -> UnnamedAppl | GlbAppl l -> GlbAppl (List.map (fun (h,vs) -> (h,vs@args)) l) let pr_generic arg = let Val.Dyn (tag, _) = arg in str"<" ++ Val.pr tag ++ str ":(" ++ Pptactic.pr_value Pptactic.ltop arg ++ str ")>" let pr_appl h vs = Pptactic.pr_ltac_constant h ++ spc () ++ Pp.prlist_with_sep spc pr_generic vs let rec name_with_list appl t = match appl with | [] -> t | (h,vs)::l -> Proofview.Trace.name_tactic (fun _ _ -> pr_appl h vs) (name_with_list l t) let name_if_glob appl t = match appl with | UnnamedAppl -> t | GlbAppl l -> name_with_list l t let combine_appl appl1 appl2 = match appl1,appl2 with | UnnamedAppl,a | a,UnnamedAppl -> a | GlbAppl l1 , GlbAppl l2 -> GlbAppl (l2@l1) let of_tacvalue v = in_gen (topwit wit_tacvalue) v let to_tacvalue v = out_gen (topwit wit_tacvalue) v let log_trace = ref false let is_traced () = !log_trace || !Flags.profile_ltac (** More naming applications *) let name_vfun appl vle = if is_traced () && has_type vle (topwit wit_tacvalue) then match to_tacvalue vle with | VFun (appl0,trace,lfun,vars,t) -> of_tacvalue (VFun (combine_appl appl0 appl,trace,lfun,vars,t)) | _ -> vle else vle module TacStore = Geninterp.TacStore let f_avoid_ids : Id.Set.t TacStore.field = TacStore.field () (* ids inherited from the call context (needed to get fresh ids) *) let f_debug : debug_info TacStore.field = TacStore.field () let f_trace : ltac_trace TacStore.field = TacStore.field () (* Signature for interpretation: val_interp and interpretation functions *) type interp_sign = Geninterp.interp_sign = { lfun : value Id.Map.t ; poly : bool ; extra : TacStore.t } let extract_trace ist = if is_traced () then match TacStore.get ist.extra f_trace with | None -> [] | Some l -> l else [] let print_top_val env v = Pptactic.pr_value Pptactic.ltop v let catching_error call_trace fail (e, info) = let inner_trace = Option.default [] (Exninfo.get info ltac_trace_info) in if List.is_empty call_trace && List.is_empty inner_trace then fail (e, info) else begin assert (CErrors.noncritical e); (* preserved invariant *) let new_trace = inner_trace @ call_trace in let located_exc = (e, Exninfo.add info ltac_trace_info new_trace) in fail located_exc end let catch_error call_trace f x = try f x with e when CErrors.noncritical e -> let e = CErrors.push e in catching_error call_trace iraise e let wrap_error tac k = if is_traced () then Proofview.tclORELSE tac k else tac let catch_error_tac call_trace tac = wrap_error tac (catching_error call_trace (fun (e, info) -> Proofview.tclZERO ~info e)) let curr_debug ist = match TacStore.get ist.extra f_debug with | None -> DebugOff | Some level -> level let pr_closure env ist body = let pp_body = Pptactic.pr_glob_tactic env body in let pr_sep () = fnl () in let pr_iarg (id, arg) = let arg = pr_argument_type arg in hov 0 (Id.print id ++ spc () ++ str ":" ++ spc () ++ arg) in let pp_iargs = v 0 (prlist_with_sep pr_sep pr_iarg (Id.Map.bindings ist)) in pp_body ++ fnl() ++ str "in environment " ++ fnl() ++ pp_iargs let pr_inspect env expr result = let pp_expr = Pptactic.pr_glob_tactic env expr in let pp_result = if has_type result (topwit wit_tacvalue) then match to_tacvalue result with | VFun (_,_, ist, ul, b) -> let body = if List.is_empty ul then b else (TacFun (ul, b)) in str "a closure with body " ++ fnl() ++ pr_closure env ist body | VRec (ist, body) -> str "a recursive closure" ++ fnl () ++ pr_closure env !ist body else let pp_type = pr_argument_type result in str "an object of type" ++ spc () ++ pp_type in pp_expr ++ fnl() ++ str "this is " ++ pp_result (* Transforms an id into a constr if possible, or fails with Not_found *) let constr_of_id env id = EConstr.mkVar (let _ = Environ.lookup_named id env in id) (** Generic arguments : table of interpretation functions *) (* Some of the code further down depends on the fact that push_trace does not modify sigma (the evar map) *) let push_trace call ist = if is_traced () then match TacStore.get ist.extra f_trace with | None -> Proofview.tclUNIT [call] | Some trace -> Proofview.tclUNIT (call :: trace) else Proofview.tclUNIT [] let propagate_trace ist loc id v = if has_type v (topwit wit_tacvalue) then let tacv = to_tacvalue v in match tacv with | VFun (appl,_,lfun,it,b) -> let t = if List.is_empty it then b else TacFun (it,b) in push_trace(loc,LtacVarCall (id,t)) ist >>= fun trace -> let ans = VFun (appl,trace,lfun,it,b) in Proofview.tclUNIT (of_tacvalue ans) | _ -> Proofview.tclUNIT v else Proofview.tclUNIT v let append_trace trace v = if has_type v (topwit wit_tacvalue) then match to_tacvalue v with | VFun (appl,trace',lfun,it,b) -> of_tacvalue (VFun (appl,trace'@trace,lfun,it,b)) | _ -> v else v (* Dynamically check that an argument is a tactic *) let coerce_to_tactic loc id v = let fail () = user_err ?loc (str "Variable " ++ Id.print id ++ str " should be bound to a tactic.") in if has_type v (topwit wit_tacvalue) then let tacv = to_tacvalue v in match tacv with | VFun _ -> v | _ -> fail () else fail () let intro_pattern_of_ident id = make @@ IntroNaming (IntroIdentifier id) let value_of_ident id = in_gen (topwit wit_intro_pattern) (intro_pattern_of_ident id) let (+++) lfun1 lfun2 = Id.Map.fold Id.Map.add lfun1 lfun2 let extend_values_with_bindings (ln,lm) lfun = let of_cub c = match c with | [], c -> Value.of_constr c | _ -> in_gen (topwit wit_constr_under_binders) c in (* For compatibility, bound variables are visible only if no other binding of the same name exists *) let accu = Id.Map.map value_of_ident ln in let accu = lfun +++ accu in Id.Map.fold (fun id c accu -> Id.Map.add id (of_cub c) accu) lm accu (***************************************************************************) (* Evaluation/interpretation *) let is_variable env id = Id.List.mem id (ids_of_named_context (Environ.named_context env)) (* Debug reference *) let debug = ref DebugOff (* Sets the debugger mode *) let set_debug pos = debug := pos (* Gives the state of debug *) let get_debug () = !debug let debugging_step ist pp = match curr_debug ist with | DebugOn lev -> safe_msgnl (str "Level " ++ int lev ++ str": " ++ pp () ++ fnl()) | _ -> Proofview.NonLogical.return () let debugging_exception_step ist signal_anomaly e pp = let explain_exc = if signal_anomaly then explain_logic_error else explain_logic_error_no_anomaly in debugging_step ist (fun () -> pp() ++ spc() ++ str "raised the exception" ++ fnl() ++ explain_exc e) let ensure_freshness env = (* We anonymize declarations which we know will not be used *) (* This assumes that the original context had no rels *) process_rel_context (fun d e -> EConstr.push_rel (Context.Rel.Declaration.set_name Anonymous d) e) env (* Raise Not_found if not in interpretation sign *) let try_interp_ltac_var coerce ist env {loc;v=id} = let v = Id.Map.find id ist.lfun in try coerce v with CannotCoerceTo s -> Taccoerce.error_ltac_variable ?loc id env v s let interp_ltac_var coerce ist env locid = try try_interp_ltac_var coerce ist env locid with Not_found -> anomaly (str "Detected '" ++ Id.print locid.v ++ str "' as ltac var at interning time.") let interp_ident ist env sigma id = try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (make id) with Not_found -> id (* Interprets an optional identifier, bound or fresh *) let interp_name ist env sigma = function | Anonymous -> Anonymous | Name id -> Name (interp_ident ist env sigma id) let interp_intro_pattern_var loc ist env sigma id = try try_interp_ltac_var (coerce_to_intro_pattern sigma) ist (Some (env,sigma)) (make ?loc id) with Not_found -> IntroNaming (IntroIdentifier id) let interp_intro_pattern_naming_var loc ist env sigma id = try try_interp_ltac_var (coerce_to_intro_pattern_naming sigma) ist (Some (env,sigma)) (make ?loc id) with Not_found -> IntroIdentifier id let interp_int ist ({loc;v=id} as locid) = try try_interp_ltac_var coerce_to_int ist None locid with Not_found -> user_err ?loc ~hdr:"interp_int" (str "Unbound variable " ++ Id.print id ++ str".") let interp_int_or_var ist = function | ArgVar locid -> interp_int ist locid | ArgArg n -> n let interp_int_or_var_as_list ist = function | ArgVar ({v=id} as locid) -> (try coerce_to_int_or_var_list (Id.Map.find id ist.lfun) with Not_found | CannotCoerceTo _ -> [ArgArg (interp_int ist locid)]) | ArgArg n as x -> [x] let interp_int_or_var_list ist l = List.flatten (List.map (interp_int_or_var_as_list ist) l) (* Interprets a bound variable (especially an existing hypothesis) *) let interp_hyp ist env sigma ({loc;v=id} as locid) = (* Look first in lfun for a value coercible to a variable *) try try_interp_ltac_var (coerce_to_hyp env sigma) ist (Some (env,sigma)) locid with Not_found -> (* Then look if bound in the proof context at calling time *) if is_variable env id then id else Loc.raise ?loc (Logic.RefinerError (env, sigma, Logic.NoSuchHyp id)) let interp_hyp_list_as_list ist env sigma ({loc;v=id} as x) = try coerce_to_hyp_list env sigma (Id.Map.find id ist.lfun) with Not_found | CannotCoerceTo _ -> [interp_hyp ist env sigma x] let interp_hyp_list ist env sigma l = List.flatten (List.map (interp_hyp_list_as_list ist env sigma) l) let interp_reference ist env sigma = function | ArgArg (_,r) -> r | ArgVar {loc;v=id} -> try try_interp_ltac_var (coerce_to_reference sigma) ist (Some (env,sigma)) (make ?loc id) with Not_found -> try GlobRef.VarRef (get_id (Environ.lookup_named id env)) with Not_found -> Nametab.error_global_not_found (qualid_of_ident ?loc id) let try_interp_evaluable env (loc, id) = let v = Environ.lookup_named id env in match v with | LocalDef _ -> EvalVarRef id | _ -> error_not_evaluable (GlobRef.VarRef id) let interp_evaluable ist env sigma = function | ArgArg (r,Some {loc;v=id}) -> (* Maybe [id] has been introduced by Intro-like tactics *) begin try try_interp_evaluable env (loc, id) with Not_found -> match r with | EvalConstRef _ -> r | _ -> Nametab.error_global_not_found (qualid_of_ident ?loc id) end | ArgArg (r,None) -> r | ArgVar {loc;v=id} -> try try_interp_ltac_var (coerce_to_evaluable_ref env sigma) ist (Some (env,sigma)) (make ?loc id) with Not_found -> try try_interp_evaluable env (loc, id) with Not_found -> Nametab.error_global_not_found (qualid_of_ident ?loc id) (* Interprets an hypothesis name *) let interp_occurrences ist occs = Locusops.occurrences_map (interp_int_or_var_list ist) occs let interp_hyp_location ist env sigma ((occs,id),hl) = ((interp_occurrences ist occs,interp_hyp ist env sigma id),hl) let interp_hyp_location_list_as_list ist env sigma ((occs,id),hl as x) = match occs,hl with | AllOccurrences,InHyp -> List.map (fun id -> ((AllOccurrences,id),InHyp)) (interp_hyp_list_as_list ist env sigma id) | _,_ -> [interp_hyp_location ist env sigma x] let interp_hyp_location_list ist env sigma l = List.flatten (List.map (interp_hyp_location_list_as_list ist env sigma) l) let interp_clause ist env sigma { onhyps=ol; concl_occs=occs } : clause = { onhyps=Option.map (interp_hyp_location_list ist env sigma) ol; concl_occs=interp_occurrences ist occs } (* Interpretation of constructions *) (* Extract the constr list from lfun *) let extract_ltac_constr_values ist env = let fold id v accu = try let c = coerce_to_constr env v in Id.Map.add id c accu with CannotCoerceTo _ -> accu in Id.Map.fold fold ist.lfun Id.Map.empty (** ppedrot: I have changed the semantics here. Before this patch, closure was implemented as a list and a variable could be bound several times with different types, resulting in its possible appearance on both sides. This could barely be defined as a feature... *) (* Extract the identifier list from lfun: join all branches (what to do else?)*) let rec intropattern_ids accu {loc;v=pat} = match pat with | IntroNaming (IntroIdentifier id) -> Id.Set.add id accu | IntroAction (IntroOrAndPattern (IntroAndPattern l)) -> List.fold_left intropattern_ids accu l | IntroAction (IntroOrAndPattern (IntroOrPattern ll)) -> List.fold_left intropattern_ids accu (List.flatten ll) | IntroAction (IntroInjection l) -> List.fold_left intropattern_ids accu l | IntroAction (IntroApplyOn ({v=c},pat)) -> intropattern_ids accu pat | IntroNaming (IntroAnonymous | IntroFresh _) | IntroAction (IntroWildcard | IntroRewrite _) | IntroForthcoming _ -> accu let extract_ids ids lfun accu = let fold id v accu = if has_type v (topwit wit_intro_pattern) then let {v=ipat} = out_gen (topwit wit_intro_pattern) v in if Id.List.mem id ids then accu else intropattern_ids accu (make ipat) else accu in Id.Map.fold fold lfun accu let default_fresh_id = Id.of_string "H" let interp_fresh_id ist env sigma l = let extract_ident ist env sigma id = try try_interp_ltac_var (coerce_to_ident_not_fresh sigma) ist (Some (env,sigma)) (make id) with Not_found -> id in let ids = List.map_filter (function ArgVar {v=id} -> Some id | _ -> None) l in let avoid = match TacStore.get ist.extra f_avoid_ids with | None -> Id.Set.empty | Some l -> l in let avoid = extract_ids ids ist.lfun avoid in let id = if List.is_empty l then default_fresh_id else let s = String.concat "" (List.map (function | ArgArg s -> s | ArgVar {v=id} -> Id.to_string (extract_ident ist env sigma id)) l) in let s = if CLexer.is_keyword s then s^"0" else s in Id.of_string s in Tactics.fresh_id_in_env avoid id env (* Extract the uconstr list from lfun *) let extract_ltac_constr_context ist env sigma = let add_uconstr id v map = try Id.Map.add id (coerce_to_uconstr v) map with CannotCoerceTo _ -> map in let add_constr id v map = try Id.Map.add id (coerce_to_constr env v) map with CannotCoerceTo _ -> map in let add_ident id v map = try Id.Map.add id (coerce_var_to_ident false env sigma v) map with CannotCoerceTo _ -> map in let fold id v {idents;typed;untyped} = let idents = add_ident id v idents in let typed = add_constr id v typed in let untyped = add_uconstr id v untyped in { idents ; typed ; untyped } in let empty = { idents = Id.Map.empty ;typed = Id.Map.empty ; untyped = Id.Map.empty } in Id.Map.fold fold ist.lfun empty (** Significantly simpler than [interp_constr], to interpret an untyped constr, it suffices to adjoin a closure environment. *) let interp_glob_closure ist env sigma ?(kind=WithoutTypeConstraint) ?(pattern_mode=false) (term,term_expr_opt) = let closure = extract_ltac_constr_context ist env sigma in match term_expr_opt with | None -> { closure ; term } | Some term_expr -> (* If at toplevel (term_expr_opt<>None), the error can be due to an incorrect context at globalization time: we retype with the now known intros/lettac/inversion hypothesis names *) let constr_context = Id.Set.union (Id.Map.domain closure.typed) (Id.Map.domain closure.untyped) in let ltacvars = { ltac_vars = constr_context; ltac_bound = Id.Map.domain ist.lfun; ltac_extra = Genintern.Store.empty; } in { closure ; term = intern_gen kind ~pattern_mode ~ltacvars env sigma term_expr } let interp_uconstr ist env sigma c = interp_glob_closure ist env sigma c let interp_gen kind ist pattern_mode flags env sigma c = let kind_for_intern = match kind with OfType _ -> WithoutTypeConstraint | _ -> kind in let { closure = constrvars ; term } = interp_glob_closure ist env sigma ~kind:kind_for_intern ~pattern_mode c in let vars = { ltac_constrs = constrvars.typed; ltac_uconstrs = constrvars.untyped; ltac_idents = constrvars.idents; ltac_genargs = ist.lfun; } in (* Jason Gross: To avoid unnecessary modifications to tacinterp, as suggested by Arnaud Spiwack, we run push_trace immediately. We do this with the kludge of an empty proofview, and rely on the invariant that running the tactic returned by push_trace does not modify sigma. *) let (_, dummy_proofview) = Proofview.init sigma [] in (* Again this is called at times with no open proof! *) let name, poly = Id.of_string "tacinterp", ist.poly in let (trace,_,_,_) = Proofview.apply ~name ~poly env (push_trace (loc_of_glob_constr term,LtacConstrInterp (term,vars)) ist) dummy_proofview in let (evd,c) = catch_error trace (understand_ltac flags env sigma vars kind) term in (* spiwack: to avoid unnecessary modifications of tacinterp, as this function already use effect, I call [run] hoping it doesn't mess up with any assumption. *) Proofview.NonLogical.run (db_constr (curr_debug ist) env evd c); (evd,c) let constr_flags () = { use_typeclasses = true; solve_unification_constraints = true; fail_evar = true; expand_evars = true; program_mode = false; polymorphic = false; } (* Interprets a constr; expects evars to be solved *) let interp_constr_gen kind ist env sigma c = let flags = { (constr_flags ()) with polymorphic = ist.Geninterp.poly } in interp_gen kind ist false flags env sigma c let interp_constr = interp_constr_gen WithoutTypeConstraint let interp_type = interp_constr_gen IsType let open_constr_use_classes_flags () = { use_typeclasses = true; solve_unification_constraints = true; fail_evar = false; expand_evars = true; program_mode = false; polymorphic = false; } let open_constr_no_classes_flags () = { use_typeclasses = false; solve_unification_constraints = true; fail_evar = false; expand_evars = true; program_mode = false; polymorphic = false; } let pure_open_constr_flags = { use_typeclasses = false; solve_unification_constraints = true; fail_evar = false; expand_evars = false; program_mode = false; polymorphic = false; } (* Interprets an open constr *) let interp_open_constr ?(expected_type=WithoutTypeConstraint) ?(flags=open_constr_no_classes_flags ()) ist env sigma c = interp_gen expected_type ist false flags env sigma c let interp_open_constr_with_classes ?(expected_type=WithoutTypeConstraint) ist env sigma c = interp_gen expected_type ist false (open_constr_use_classes_flags ()) env sigma c let interp_pure_open_constr ist = interp_gen WithoutTypeConstraint ist false pure_open_constr_flags let interp_typed_pattern ist env sigma (_,c,_) = let sigma, c = interp_gen WithoutTypeConstraint ist true pure_open_constr_flags env sigma c in (* FIXME: it is necessary to be unsafe here because of the way we handle evars in the pretyper. Sometimes they get solved eagerly. *) pattern_of_constr env sigma (EConstr.Unsafe.to_constr c) (* Interprets a constr expression *) let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l = let try_expand_ltac_var sigma x = try match DAst.get (fst (dest_fun x)) with | GVar id -> let v = Id.Map.find id ist.lfun in sigma, List.map inj_fun (coerce_to_constr_list env v) | _ -> raise Not_found with CannotCoerceTo _ | Not_found -> (* dest_fun, List.assoc may raise Not_found *) let sigma, c = interp_fun ist env sigma x in sigma, [c] in let sigma, l = List.fold_left_map try_expand_ltac_var sigma l in sigma, List.flatten l let interp_constr_list ist env sigma c = interp_constr_in_compound_list (fun x -> x) (fun x -> x) interp_constr ist env sigma c let interp_open_constr_list = interp_constr_in_compound_list (fun x -> x) (fun x -> x) interp_open_constr (* Interprets a reduction expression *) let interp_unfold ist env sigma (occs,qid) = (interp_occurrences ist occs,interp_evaluable ist env sigma qid) let interp_flag ist env sigma red = { red with rConst = List.map (interp_evaluable ist env sigma) red.rConst } let interp_constr_with_occurrences ist env sigma (occs,c) = let (sigma,c_interp) = interp_constr ist env sigma c in sigma , (interp_occurrences ist occs, c_interp) let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = let p = match a with | Inl (ArgVar {loc;v=id}) -> (* This is the encoding of an ltac var supposed to be bound prioritary to an evaluable reference and otherwise to a constr (it is an encoding to satisfy the "union" type given to Simpl) *) let coerce_eval_ref_or_constr x = try Inl (coerce_to_evaluable_ref env sigma x) with CannotCoerceTo _ -> let c = coerce_to_closed_constr env x in Inr (pattern_of_constr env sigma (EConstr.to_constr sigma c)) in (try try_interp_ltac_var coerce_eval_ref_or_constr ist (Some (env,sigma)) (make ?loc id) with Not_found -> Nametab.error_global_not_found (qualid_of_ident ?loc id)) | Inl (ArgArg _ as b) -> Inl (interp_evaluable ist env sigma b) | Inr c -> Inr (interp_typed_pattern ist env sigma c) in interp_occurrences ist occs, p let interp_constr_with_occurrences_and_name_as_list = interp_constr_in_compound_list (fun c -> ((AllOccurrences,c),Anonymous)) (function ((occs,c),Anonymous) when occs == AllOccurrences -> c | _ -> raise Not_found) (fun ist env sigma (occ_c,na) -> let (sigma,c_interp) = interp_constr_with_occurrences ist env sigma occ_c in sigma, (c_interp, interp_name ist env sigma na)) let interp_red_expr ist env sigma = function | Unfold l -> sigma , Unfold (List.map (interp_unfold ist env sigma) l) | Fold l -> let (sigma,l_interp) = interp_constr_list ist env sigma l in sigma , Fold l_interp | Cbv f -> sigma , Cbv (interp_flag ist env sigma f) | Cbn f -> sigma , Cbn (interp_flag ist env sigma f) | Lazy f -> sigma , Lazy (interp_flag ist env sigma f) | Pattern l -> let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> interp_constr_with_occurrences ist env sigma c) l sigma in sigma , Pattern l_interp | Simpl (f,o) -> sigma , Simpl (interp_flag ist env sigma f, Option.map (interp_closed_typed_pattern_with_occurrences ist env sigma) o) | CbvVm o -> sigma , CbvVm (Option.map (interp_closed_typed_pattern_with_occurrences ist env sigma) o) | CbvNative o -> sigma , CbvNative (Option.map (interp_closed_typed_pattern_with_occurrences ist env sigma) o) | (Red _ | Hnf | ExtraRedExpr _ as r) -> sigma , r let interp_may_eval f ist env sigma = function | ConstrEval (r,c) -> let (sigma,redexp) = interp_red_expr ist env sigma r in let (sigma,c_interp) = f ist env sigma c in let (redfun, _) = Redexpr.reduction_of_red_expr env redexp in redfun env sigma c_interp | ConstrContext ({loc;v=s},c) -> (try let (sigma,ic) = f ist env sigma c in let ctxt = try_interp_ltac_var coerce_to_constr_context ist (Some (env, sigma)) (make ?loc s) in let ctxt = EConstr.Unsafe.to_constr ctxt in let ic = EConstr.Unsafe.to_constr ic in let c = subst_meta [Constr_matching.special_meta,ic] ctxt in Typing.solve_evars env sigma (EConstr.of_constr c) with | Not_found -> user_err ?loc ~hdr:"interp_may_eval" (str "Unbound context identifier" ++ Id.print s ++ str".")) | ConstrTypeOf c -> let (sigma,c_interp) = f ist env sigma c in let (sigma, t) = Typing.type_of ~refresh:true env sigma c_interp in (sigma, t) | ConstrTerm c -> try f ist env sigma c with reraise -> let reraise = CErrors.push reraise in (* spiwack: to avoid unnecessary modifications of tacinterp, as this function already use effect, I call [run] hoping it doesn't mess up with any assumption. *) Proofview.NonLogical.run (debugging_exception_step ist false (fst reraise) (fun () -> str"interpretation of term " ++ pr_glob_constr_env env (fst c))); iraise reraise (* Interprets a constr expression possibly to first evaluate *) let interp_constr_may_eval ist env sigma c = let (sigma,csr) = try interp_may_eval interp_constr ist env sigma c with reraise -> let reraise = CErrors.push reraise in (* spiwack: to avoid unnecessary modifications of tacinterp, as this function already use effect, I call [run] hoping it doesn't mess up with any assumption. *) Proofview.NonLogical.run (debugging_exception_step ist false (fst reraise) (fun () -> str"evaluation of term")); iraise reraise in begin (* spiwack: to avoid unnecessary modifications of tacinterp, as this function already use effect, I call [run] hoping it doesn't mess up with any assumption. *) Proofview.NonLogical.run (db_constr (curr_debug ist) env sigma csr); sigma , csr end (** TODO: should use dedicated printers *) let message_of_value v = let pr_with_env pr = Ftactic.enter begin fun gl -> Ftactic.return (pr (pf_env gl) (project gl)) end in let open Genprint in match generic_val_print v with | TopPrinterBasic pr -> Ftactic.return (pr ()) | TopPrinterNeedsContext pr -> pr_with_env pr | TopPrinterNeedsContextAndLevel { default_ensure_surrounded; printer } -> pr_with_env (fun env sigma -> printer env sigma default_ensure_surrounded) let interp_message_token ist = function | MsgString s -> Ftactic.return (str s) | MsgInt n -> Ftactic.return (int n) | MsgIdent {loc;v=id} -> let v = try Some (Id.Map.find id ist.lfun) with Not_found -> None in match v with | None -> Ftactic.lift (Tacticals.New.tclZEROMSG (Id.print id ++ str" not found.")) | Some v -> message_of_value v let interp_message ist l = let open Ftactic in Ftactic.List.map (interp_message_token ist) l >>= fun l -> Ftactic.return (prlist_with_sep spc (fun x -> x) l) let rec interp_intro_pattern ist env sigma = with_loc_val (fun ?loc -> function | IntroAction pat -> let (sigma,pat) = interp_intro_pattern_action ist env sigma pat in sigma, make ?loc @@ IntroAction pat | IntroNaming (IntroIdentifier id) -> sigma, make ?loc @@ interp_intro_pattern_var loc ist env sigma id | IntroNaming pat -> sigma, make ?loc @@ IntroNaming (interp_intro_pattern_naming loc ist env sigma pat) | IntroForthcoming _ as x -> sigma, make ?loc x) and interp_intro_pattern_naming loc ist env sigma = function | IntroFresh id -> IntroFresh (interp_ident ist env sigma id) | IntroIdentifier id -> interp_intro_pattern_naming_var loc ist env sigma id | IntroAnonymous as x -> x and interp_intro_pattern_action ist env sigma = function | IntroOrAndPattern l -> let (sigma,l) = interp_or_and_intro_pattern ist env sigma l in sigma, IntroOrAndPattern l | IntroInjection l -> let sigma,l = interp_intro_pattern_list_as_list ist env sigma l in sigma, IntroInjection l | IntroApplyOn ({loc;v=c},ipat) -> let c env sigma = interp_open_constr ist env sigma c in let sigma,ipat = interp_intro_pattern ist env sigma ipat in sigma, IntroApplyOn (make ?loc c,ipat) | IntroWildcard | IntroRewrite _ as x -> sigma, x and interp_or_and_intro_pattern ist env sigma = function | IntroAndPattern l -> let sigma, l = List.fold_left_map (interp_intro_pattern ist env) sigma l in sigma, IntroAndPattern l | IntroOrPattern ll -> let sigma, ll = List.fold_left_map (interp_intro_pattern_list_as_list ist env) sigma ll in sigma, IntroOrPattern ll and interp_intro_pattern_list_as_list ist env sigma = function | [{loc;v=IntroNaming (IntroIdentifier id)}] as l -> (try sigma, coerce_to_intro_pattern_list ?loc sigma (Id.Map.find id ist.lfun) with Not_found | CannotCoerceTo _ -> List.fold_left_map (interp_intro_pattern ist env) sigma l) | l -> List.fold_left_map (interp_intro_pattern ist env) sigma l let interp_intro_pattern_naming_option ist env sigma = function | None -> None | Some lpat -> Some (map_with_loc (fun ?loc pat -> interp_intro_pattern_naming loc ist env sigma pat) lpat) let interp_or_and_intro_pattern_option ist env sigma = function | None -> sigma, None | Some (ArgVar {loc;v=id}) -> (match interp_intro_pattern_var loc ist env sigma id with | IntroAction (IntroOrAndPattern l) -> sigma, Some (make ?loc l) | _ -> user_err ?loc (str "Cannot coerce to a disjunctive/conjunctive pattern.")) | Some (ArgArg {loc;v=l}) -> let sigma,l = interp_or_and_intro_pattern ist env sigma l in sigma, Some (make ?loc l) let interp_intro_pattern_option ist env sigma = function | None -> sigma, None | Some ipat -> let sigma, ipat = interp_intro_pattern ist env sigma ipat in sigma, Some ipat let interp_in_hyp_as ist env sigma (id,ipat) = let sigma, ipat = interp_intro_pattern_option ist env sigma ipat in sigma,(interp_hyp ist env sigma id,ipat) let interp_binding_name ist env sigma = function | AnonHyp n -> AnonHyp n | NamedHyp id -> (* If a name is bound, it has to be a quantified hypothesis *) (* user has to use other names for variables if these ones clash with *) (* a name intended to be used as a (non-variable) identifier *) try try_interp_ltac_var (coerce_to_quantified_hypothesis sigma) ist (Some (env,sigma)) (make id) with Not_found -> NamedHyp id let interp_declared_or_quantified_hypothesis ist env sigma = function | AnonHyp n -> AnonHyp n | NamedHyp id -> try try_interp_ltac_var (coerce_to_decl_or_quant_hyp sigma) ist (Some (env,sigma)) (make id) with Not_found -> NamedHyp id let interp_binding ist env sigma {loc;v=(b,c)} = let sigma, c = interp_open_constr ist env sigma c in sigma, (make ?loc (interp_binding_name ist env sigma b,c)) let interp_bindings ist env sigma = function | NoBindings -> sigma, NoBindings | ImplicitBindings l -> let sigma, l = interp_open_constr_list ist env sigma l in sigma, ImplicitBindings l | ExplicitBindings l -> let sigma, l = List.fold_left_map (interp_binding ist env) sigma l in sigma, ExplicitBindings l let interp_constr_with_bindings ist env sigma (c,bl) = let sigma, bl = interp_bindings ist env sigma bl in let sigma, c = interp_constr ist env sigma c in sigma, (c,bl) let interp_open_constr_with_bindings ist env sigma (c,bl) = let sigma, bl = interp_bindings ist env sigma bl in let sigma, c = interp_open_constr ist env sigma c in sigma, (c, bl) let loc_of_bindings = function | NoBindings -> None | ImplicitBindings l -> loc_of_glob_constr (fst (List.last l)) | ExplicitBindings l -> (List.last l).loc let interp_open_constr_with_bindings_loc ist ((c,_),bl as cb) = let loc1 = loc_of_glob_constr c in let loc2 = loc_of_bindings bl in let loc = Loc.merge_opt loc1 loc2 in let f env sigma = interp_open_constr_with_bindings ist env sigma cb in (loc,f) let interp_destruction_arg ist gl arg = match arg with | keep,ElimOnConstr c -> keep,ElimOnConstr begin fun env sigma -> interp_open_constr_with_bindings ist env sigma c end | keep,ElimOnAnonHyp n as x -> x | keep,ElimOnIdent {loc;v=id} -> let error () = user_err ?loc (strbrk "Cannot coerce " ++ Id.print id ++ strbrk " neither to a quantified hypothesis nor to a term.") in let try_cast_id id' = if Tactics.is_quantified_hypothesis id' gl then keep,ElimOnIdent (make ?loc id') else (keep, ElimOnConstr begin fun env sigma -> try (sigma, (constr_of_id env id', NoBindings)) with Not_found -> user_err ?loc ~hdr:"interp_destruction_arg" ( Id.print id ++ strbrk " binds to " ++ Id.print id' ++ strbrk " which is neither a declared nor a quantified hypothesis.") end) in try (* FIXME: should be moved to taccoerce *) let v = Id.Map.find id ist.lfun in if has_type v (topwit wit_intro_pattern) then let v = out_gen (topwit wit_intro_pattern) v in match v with | {v=IntroNaming (IntroIdentifier id)} -> try_cast_id id | _ -> error () else if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in try_cast_id id else if has_type v (topwit wit_int) then keep,ElimOnAnonHyp (out_gen (topwit wit_int) v) else match Value.to_constr v with | None -> error () | Some c -> keep,ElimOnConstr (fun env sigma -> (sigma, (c,NoBindings))) with Not_found -> (* We were in non strict (interactive) mode *) if Tactics.is_quantified_hypothesis id gl then keep,ElimOnIdent (make ?loc id) else let c = (DAst.make ?loc @@ GVar id,Some (make @@ CRef (qualid_of_ident ?loc id,None))) in let f env sigma = let (sigma,c) = interp_open_constr ist env sigma c in (sigma, (c,NoBindings)) in keep,ElimOnConstr f (* Associates variables with values and gives the remaining variables and values *) let head_with_value (lvar,lval) = let rec head_with_value_rec lacc = function | ([],[]) -> (lacc,[],[]) | (vr::tvr,ve::tve) -> (match vr with | Anonymous -> head_with_value_rec lacc (tvr,tve) | Name v -> head_with_value_rec ((v,ve)::lacc) (tvr,tve)) | (vr,[]) -> (lacc,vr,[]) | ([],ve) -> (lacc,[],ve) in head_with_value_rec [] (lvar,lval) (** [interp_context ctxt] interprets a context (as in {!Matching.matching_result}) into a context value of Ltac. *) let interp_context ctxt = in_gen (topwit wit_constr_context) ctxt (* Reads a pattern by substituting vars of lfun *) let use_types = false let eval_pattern lfun ist env sigma (bvars,(glob,_),pat as c) = if use_types then (bvars,interp_typed_pattern ist env sigma c) else (bvars,instantiate_pattern env sigma lfun pat) let read_pattern lfun ist env sigma = function | Subterm (ido,c) -> Subterm (ido,eval_pattern lfun ist env sigma c) | Term c -> Term (eval_pattern lfun ist env sigma c) (* Reads the hypotheses of a Match Context rule *) let cons_and_check_name id l = if Id.List.mem id l then user_err ~hdr:"read_match_goal_hyps" ( str "Hypothesis pattern-matching variable " ++ Id.print id ++ str " used twice in the same pattern.") else id::l let rec read_match_goal_hyps lfun ist env sigma lidh = function | (Hyp ({loc;v=na} as locna,mp))::tl -> let lidh' = Name.fold_right cons_and_check_name na lidh in Hyp (locna,read_pattern lfun ist env sigma mp):: (read_match_goal_hyps lfun ist env sigma lidh' tl) | (Def ({loc;v=na} as locna,mv,mp))::tl -> let lidh' = Name.fold_right cons_and_check_name na lidh in Def (locna,read_pattern lfun ist env sigma mv, read_pattern lfun ist env sigma mp):: (read_match_goal_hyps lfun ist env sigma lidh' tl) | [] -> [] (* Reads the rules of a Match Context or a Match *) let rec read_match_rule lfun ist env sigma = function | (All tc)::tl -> (All tc)::(read_match_rule lfun ist env sigma tl) | (Pat (rl,mp,tc))::tl -> Pat (read_match_goal_hyps lfun ist env sigma [] rl, read_pattern lfun ist env sigma mp,tc) :: read_match_rule lfun ist env sigma tl | [] -> [] (* Fully evaluate an untyped constr *) let type_uconstr ?(flags = (constr_flags ())) ?(expected_type = WithoutTypeConstraint) ist c = begin fun env sigma -> let { closure; term } = c in let vars = { ltac_constrs = closure.typed; ltac_uconstrs = closure.untyped; ltac_idents = closure.idents; ltac_genargs = Id.Map.empty; } in let flags = { flags with polymorphic = ist.Geninterp.poly } in understand_ltac flags env sigma vars expected_type term end let warn_deprecated_info = CWarnings.create ~name:"deprecated-info-tactical" ~category:"deprecated" (fun () -> strbrk "The general \"info\" tactic is currently not working." ++ spc()++ strbrk "There is an \"Info\" command to replace it." ++fnl () ++ strbrk "Some specific verbose tactics may also exist, such as info_eauto.") (* Interprets an l-tac expression into a value *) let rec val_interp ist ?(appl=UnnamedAppl) (tac:glob_tactic_expr) : Val.t Ftactic.t = (* The name [appl] of applied top-level Ltac names is ignored in [value_interp]. It is installed in the second step by a call to [name_vfun], because it gives more opportunities to detect a [VFun]. Otherwise a [Ltac t := let x := .. in tac] would never register its name since it is syntactically a let, not a function. *) let value_interp ist = match tac with | TacFun (it, body) -> Ftactic.return (of_tacvalue (VFun (UnnamedAppl,extract_trace ist, ist.lfun, it, body))) | TacLetIn (true,l,u) -> interp_letrec ist l u | TacLetIn (false,l,u) -> interp_letin ist l u | TacMatchGoal (lz,lr,lmr) -> interp_match_goal ist lz lr lmr | TacMatch (lz,c,lmr) -> interp_match ist lz c lmr | TacArg {loc;v} -> interp_tacarg ist v | t -> (* Delayed evaluation *) Ftactic.return (of_tacvalue (VFun (UnnamedAppl,extract_trace ist, ist.lfun, [], t))) in let open Ftactic in Control.check_for_interrupt (); match curr_debug ist with | DebugOn lev -> let eval v = let ist = { ist with extra = TacStore.set ist.extra f_debug v } in value_interp ist >>= fun v -> return (name_vfun appl v) in Tactic_debug.debug_prompt lev tac eval | _ -> value_interp ist >>= fun v -> return (name_vfun appl v) and eval_tactic ist tac : unit Proofview.tactic = match tac with | TacAtom {loc;v=t} -> let call = LtacAtomCall t in push_trace(loc,call) ist >>= fun trace -> Profile_ltac.do_profile "eval_tactic:2" trace (catch_error_tac trace (interp_atomic ist t)) | TacFun _ | TacLetIn _ | TacMatchGoal _ | TacMatch _ -> interp_tactic ist tac | TacId [] -> Proofview.tclLIFT (db_breakpoint (curr_debug ist) []) | TacId s -> let msgnl = let open Ftactic in interp_message ist s >>= fun msg -> return (hov 0 msg , hov 0 msg) in let print (_,msgnl) = Proofview.(tclLIFT (NonLogical.print_info msgnl)) in let log (msg,_) = Proofview.Trace.log (fun _ _ -> msg) in let break = Proofview.tclLIFT (db_breakpoint (curr_debug ist) s) in Ftactic.run msgnl begin fun msgnl -> print msgnl <*> log msgnl <*> break end | TacFail (g,n,s) -> let msg = interp_message ist s in let tac l = Tacticals.New.tclFAIL (interp_int_or_var ist n) l in let tac = match g with | TacLocal -> fun l -> Proofview.tclINDEPENDENT (tac l) | TacGlobal -> tac in Ftactic.run msg tac | TacProgress tac -> Tacticals.New.tclPROGRESS (interp_tactic ist tac) | TacShowHyps tac -> Proofview.V82.tactic begin tclSHOWHYPS (Proofview.V82.of_tactic (interp_tactic ist tac)) end | TacAbstract (t,ido) -> let call = LtacMLCall tac in push_trace(None,call) ist >>= fun trace -> Profile_ltac.do_profile "eval_tactic:TacAbstract" trace (catch_error_tac trace begin Proofview.Goal.enter begin fun gl -> Abstract.tclABSTRACT (Option.map (interp_ident ist (pf_env gl) (project gl)) ido) (interp_tactic ist t) end end) | TacThen (t1,t) -> Tacticals.New.tclTHEN (interp_tactic ist t1) (interp_tactic ist t) | TacDispatch tl -> Proofview.tclDISPATCH (List.map (interp_tactic ist) tl) | TacExtendTac (tf,t,tl) -> Proofview.tclEXTEND (Array.map_to_list (interp_tactic ist) tf) (interp_tactic ist t) (Array.map_to_list (interp_tactic ist) tl) | TacThens (t1,tl) -> Tacticals.New.tclTHENS (interp_tactic ist t1) (List.map (interp_tactic ist) tl) | TacThens3parts (t1,tf,t,tl) -> Tacticals.New.tclTHENS3PARTS (interp_tactic ist t1) (Array.map (interp_tactic ist) tf) (interp_tactic ist t) (Array.map (interp_tactic ist) tl) | TacDo (n,tac) -> Tacticals.New.tclDO (interp_int_or_var ist n) (interp_tactic ist tac) | TacTimeout (n,tac) -> Tacticals.New.tclTIMEOUT (interp_int_or_var ist n) (interp_tactic ist tac) | TacTime (s,tac) -> Tacticals.New.tclTIME s (interp_tactic ist tac) | TacTry tac -> Tacticals.New.tclTRY (interp_tactic ist tac) | TacRepeat tac -> Tacticals.New.tclREPEAT (interp_tactic ist tac) | TacOr (tac1,tac2) -> Tacticals.New.tclOR (interp_tactic ist tac1) (interp_tactic ist tac2) | TacOnce tac -> Tacticals.New.tclONCE (interp_tactic ist tac) | TacExactlyOnce tac -> Tacticals.New.tclEXACTLY_ONCE (interp_tactic ist tac) | TacIfThenCatch (t,tt,te) -> Tacticals.New.tclIFCATCH (interp_tactic ist t) (fun () -> interp_tactic ist tt) (fun () -> interp_tactic ist te) | TacOrelse (tac1,tac2) -> Tacticals.New.tclORELSE (interp_tactic ist tac1) (interp_tactic ist tac2) | TacFirst l -> Tacticals.New.tclFIRST (List.map (interp_tactic ist) l) | TacSolve l -> Tacticals.New.tclSOLVE (List.map (interp_tactic ist) l) | TacComplete tac -> Tacticals.New.tclCOMPLETE (interp_tactic ist tac) | TacArg a -> interp_tactic ist (TacArg a) | TacInfo tac -> warn_deprecated_info (); eval_tactic ist tac | TacSelect (sel, tac) -> Tacticals.New.tclSELECT sel (interp_tactic ist tac) (* For extensions *) | TacAlias {loc; v=(s,l)} -> let alias = Tacenv.interp_alias s in Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) -> let (>>=) = Ftactic.bind in let interp_vars = Ftactic.List.map (fun v -> interp_tacarg ist v) l in let tac l = let addvar x v accu = Id.Map.add x v accu in let lfun = List.fold_right2 addvar alias.Tacenv.alias_args l ist.lfun in Ftactic.lift (push_trace (loc,LtacNotationCall s) ist) >>= fun trace -> let ist = { lfun ; poly ; extra = TacStore.set ist.extra f_trace trace } in val_interp ist alias.Tacenv.alias_body >>= fun v -> Ftactic.lift (tactic_of_value ist v) in let tac = Ftactic.with_env interp_vars >>= fun (env, lr) -> let name _ _ = Pptactic.pr_alias (fun v -> print_top_val env v) 0 s lr in Proofview.Trace.name_tactic name (tac lr) (* spiwack: this use of name_tactic is not robust to a change of implementation of [Ftactic]. In such a situation, some more elaborate solution will have to be used. *) in let tac = let len1 = List.length alias.Tacenv.alias_args in let len2 = List.length l in if len1 = len2 then tac else Tacticals.New.tclZEROMSG (str "Arguments length mismatch: \ expected " ++ int len1 ++ str ", found " ++ int len2) in Ftactic.run tac (fun () -> Proofview.tclUNIT ()) | TacML {loc; v=(opn,l)} -> push_trace (Loc.tag ?loc @@ LtacMLCall tac) ist >>= fun trace -> let ist = { ist with extra = TacStore.set ist.extra f_trace trace; } in let tac = Tacenv.interp_ml_tactic opn in let args = Ftactic.List.map_right (fun a -> interp_tacarg ist a) l in let tac args = let name _ _ = Pptactic.pr_extend (fun v -> print_top_val () v) 0 opn args in Proofview.Trace.name_tactic name (catch_error_tac trace (tac args ist)) in Ftactic.run args tac and force_vrec ist v : Val.t Ftactic.t = if has_type v (topwit wit_tacvalue) then let v = to_tacvalue v in match v with | VRec (lfun,body) -> val_interp {ist with lfun = !lfun} body | v -> Ftactic.return (of_tacvalue v) else Ftactic.return v and interp_ltac_reference ?loc' mustbetac ist r : Val.t Ftactic.t = match r with | ArgVar {loc;v=id} -> let v = try Id.Map.find id ist.lfun with Not_found -> in_gen (topwit wit_var) id in let open Ftactic in force_vrec ist v >>= begin fun v -> Ftactic.lift (propagate_trace ist loc id v) >>= fun v -> if mustbetac then Ftactic.return (coerce_to_tactic loc id v) else Ftactic.return v end | ArgArg (loc,r) -> Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) -> let ids = extract_ids [] ist.lfun Id.Set.empty in let loc_info = (Option.default loc loc',LtacNameCall r) in let extra = TacStore.set ist.extra f_avoid_ids ids in push_trace loc_info ist >>= fun trace -> let extra = TacStore.set extra f_trace trace in let ist = { lfun = Id.Map.empty; poly; extra } in let appl = GlbAppl[r,[]] in Profile_ltac.do_profile "interp_ltac_reference" trace ~count_call:false (val_interp ~appl ist (Tacenv.interp_ltac r)) and interp_tacarg ist arg : Val.t Ftactic.t = match arg with | TacGeneric arg -> interp_genarg ist arg | Reference r -> interp_ltac_reference false ist r | ConstrMayEval c -> Ftactic.enter begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in let (sigma,c_interp) = interp_constr_may_eval ist env sigma c in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Ftactic.return (Value.of_constr c_interp)) end | TacCall { v=(r,[]) } -> interp_ltac_reference true ist r | TacCall { loc; v=(f,l) } -> let (>>=) = Ftactic.bind in interp_ltac_reference true ist f >>= fun fv -> Ftactic.List.map (fun a -> interp_tacarg ist a) l >>= fun largs -> interp_app loc ist fv largs | TacFreshId l -> Ftactic.enter begin fun gl -> let id = interp_fresh_id ist (pf_env gl) (project gl) l in Ftactic.return (in_gen (topwit wit_intro_pattern) (make @@ IntroNaming (IntroIdentifier id))) end | TacPretype c -> Ftactic.enter begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let c = interp_uconstr ist env sigma c in let (sigma, c) = type_uconstr ist c env sigma in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Ftactic.return (Value.of_constr c)) end | TacNumgoals -> Ftactic.lift begin let open Proofview.Notations in Proofview.numgoals >>= fun i -> Proofview.tclUNIT (Value.of_int i) end | Tacexp t -> val_interp ist t (* Interprets an application node *) and interp_app loc ist fv largs : Val.t Ftactic.t = Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) -> let (>>=) = Ftactic.bind in let fail = Tacticals.New.tclZEROMSG (str "Illegal tactic application.") in if has_type fv (topwit wit_tacvalue) then match to_tacvalue fv with (* if var=[] and body has been delayed by val_interp, then body is not a tactic that expects arguments. Otherwise Ltac goes into an infinite loop (val_interp puts a VFun back on body, and then interp_app is called again...) *) | (VFun(appl,trace,olfun,(_::_ as var),body) |VFun(appl,trace,olfun,([] as var), (TacFun _|TacLetIn _|TacMatchGoal _|TacMatch _| TacArg _ as body))) -> let (extfun,lvar,lval)=head_with_value (var,largs) in let fold accu (id, v) = Id.Map.add id v accu in let newlfun = List.fold_left fold olfun extfun in if List.is_empty lvar then begin wrap_error begin let ist = { lfun = newlfun ; poly ; extra = TacStore.set ist.extra f_trace [] } in Profile_ltac.do_profile "interp_app" trace ~count_call:false (catch_error_tac trace (val_interp ist body)) >>= fun v -> Ftactic.return (name_vfun (push_appl appl largs) v) end begin fun (e, info) -> Proofview.tclLIFT (debugging_exception_step ist false e (fun () -> str "evaluation")) <*> Proofview.tclZERO ~info e end end >>= fun v -> (* No errors happened, we propagate the trace *) let v = append_trace trace v in let call_debug env = Proofview.tclLIFT (debugging_step ist (fun () -> str"evaluation returns"++fnl()++pr_value env v)) in begin let open Genprint in match generic_val_print v with | TopPrinterBasic _ -> call_debug None | TopPrinterNeedsContext _ | TopPrinterNeedsContextAndLevel _ -> Proofview.Goal.enter (fun gl -> call_debug (Some (pf_env gl,project gl))) end <*> if List.is_empty lval then Ftactic.return v else interp_app loc ist v lval else Ftactic.return (of_tacvalue (VFun(push_appl appl largs,trace,newlfun,lvar,body))) | (VFun(appl,trace,olfun,[],body)) -> let extra_args = List.length largs in Tacticals.New.tclZEROMSG (str "Illegal tactic application: got " ++ str (string_of_int extra_args) ++ str " extra " ++ str (String.plural extra_args "argument") ++ str ".") | VRec(_,_) -> fail else fail (* Gives the tactic corresponding to the tactic value *) and tactic_of_value ist vle = if has_type vle (topwit wit_tacvalue) then match to_tacvalue vle with | VFun (appl,trace,lfun,[],t) -> Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) -> let ist = { lfun = lfun; poly; extra = TacStore.set ist.extra f_trace []; } in let tac = name_if_glob appl (eval_tactic ist t) in Profile_ltac.do_profile "tactic_of_value" trace (catch_error_tac trace tac) | VFun (appl,_,vmap,vars,_) -> let tactic_nm = match appl with UnnamedAppl -> "An unnamed user-defined tactic" | GlbAppl apps -> let nms = List.map (fun (kn,_) -> string_of_qualid (Tacenv.shortest_qualid_of_tactic kn)) apps in match nms with [] -> assert false | kn::_ -> "The user-defined tactic \"" ^ kn ^ "\"" (* TODO: when do we not have a singleton? *) in let numargs = List.length vars in let givenargs = List.map (fun (arg,_) -> Names.Id.to_string arg) (Names.Id.Map.bindings vmap) in let numgiven = List.length givenargs in Tacticals.New.tclZEROMSG (Pp.str tactic_nm ++ Pp.str " was not fully applied:" ++ spc() ++ (match numargs with 0 -> assert false | 1 -> Pp.str "There is a missing argument for variable " ++ (Name.print (List.hd vars)) | _ -> Pp.str "There are missing arguments for variables " ++ pr_enum Name.print vars) ++ Pp.pr_comma () ++ match numgiven with 0 -> Pp.str "no arguments at all were provided." | 1 -> Pp.str "an argument was provided for variable " ++ Pp.str (List.hd givenargs) ++ Pp.str "." | _ -> Pp.str "arguments were provided for variables " ++ pr_enum Pp.str givenargs ++ Pp.str ".") | VRec _ -> Tacticals.New.tclZEROMSG (str "A fully applied tactic is expected.") else if has_type vle (topwit wit_tactic) then let tac = out_gen (topwit wit_tactic) vle in tactic_of_value ist tac else Tacticals.New.tclZEROMSG (str "Expression does not evaluate to a tactic.") (* Interprets the clauses of a recursive LetIn *) and interp_letrec ist llc u = Proofview.tclUNIT () >>= fun () -> (* delay for the effects of [lref], just in case. *) let lref = ref ist.lfun in let fold accu ({v=na}, b) = let v = of_tacvalue (VRec (lref, TacArg (CAst.make b))) in Name.fold_right (fun id -> Id.Map.add id v) na accu in let lfun = List.fold_left fold ist.lfun llc in let () = lref := lfun in let ist = { ist with lfun } in val_interp ist u (* Interprets the clauses of a LetIn *) and interp_letin ist llc u = let rec fold lfun = function | [] -> let ist = { ist with lfun } in val_interp ist u | ({v=na}, body) :: defs -> Ftactic.bind (interp_tacarg ist body) (fun v -> fold (Name.fold_right (fun id -> Id.Map.add id v) na lfun) defs) in fold ist.lfun llc (** [interp_match_success lz ist succ] interprets a single matching success (of type {!Tactic_matching.t}). *) and interp_match_success ist { Tactic_matching.subst ; context ; terms ; lhs } = Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) -> let (>>=) = Ftactic.bind in let lctxt = Id.Map.map interp_context context in let hyp_subst = Id.Map.map Value.of_constr terms in let lfun = extend_values_with_bindings subst (lctxt +++ hyp_subst +++ ist.lfun) in let ist = { ist with lfun } in val_interp ist lhs >>= fun v -> if has_type v (topwit wit_tacvalue) then match to_tacvalue v with | VFun (appl,trace,lfun,[],t) -> let ist = { lfun = lfun ; poly ; extra = TacStore.set ist.extra f_trace trace } in let tac = eval_tactic ist t in let dummy = VFun (appl,extract_trace ist, Id.Map.empty, [], TacId []) in catch_error_tac trace (tac <*> Ftactic.return (of_tacvalue dummy)) | _ -> Ftactic.return v else Ftactic.return v (** [interp_match_successes lz ist s] interprets the stream of matching of successes [s]. If [lz] is set to true, then only the first success is considered, otherwise further successes are tried if the left-hand side fails. *) and interp_match_successes lz ist s = let general = let break (e, info) = match e with | FailError (0, _) -> None | FailError (n, s) -> Some (FailError (pred n, s), info) | _ -> None in Proofview.tclBREAK break s >>= fun ans -> interp_match_success ist ans in match lz with | General -> general | Select -> begin (* Only keep the first matching result, we don't backtrack on it *) let s = Proofview.tclONCE s in s >>= fun ans -> interp_match_success ist ans end | Once -> (* Once a tactic has succeeded, do not backtrack anymore *) Proofview.tclONCE general (* Interprets the Match expressions *) and interp_match ist lz constr lmr = let (>>=) = Ftactic.bind in begin wrap_error (interp_ltac_constr ist constr) begin function | (e, info) -> Proofview.tclLIFT (debugging_exception_step ist true e (fun () -> str "evaluation of the matched expression")) <*> Proofview.tclZERO ~info e end end >>= fun constr -> Ftactic.enter begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in let ilr = read_match_rule (extract_ltac_constr_values ist env) ist env sigma lmr in interp_match_successes lz ist (Tactic_matching.match_term env sigma constr ilr) end (* Interprets the Match Context expressions *) and interp_match_goal ist lz lr lmr = Ftactic.enter begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in let hyps = Proofview.Goal.hyps gl in let hyps = if lr then List.rev hyps else hyps in let concl = Proofview.Goal.concl gl in let ilr = read_match_rule (extract_ltac_constr_values ist env) ist env sigma lmr in interp_match_successes lz ist (Tactic_matching.match_goal env sigma hyps concl ilr) end (* Interprets extended tactic generic arguments *) and interp_genarg ist x : Val.t Ftactic.t = let open Ftactic.Notations in (* Ad-hoc handling of some types. *) let tag = genarg_tag x in if argument_type_eq tag (unquote (topwit (wit_list wit_var))) then interp_genarg_var_list ist x else if argument_type_eq tag (unquote (topwit (wit_list wit_constr))) then interp_genarg_constr_list ist x else let GenArg (Glbwit wit, x) = x in match wit with | ListArg wit -> let map x = interp_genarg ist (Genarg.in_gen (glbwit wit) x) in Ftactic.List.map map x >>= fun l -> Ftactic.return (Val.Dyn (Val.typ_list, l)) | OptArg wit -> begin match x with | None -> Ftactic.return (Val.Dyn (Val.typ_opt, None)) | Some x -> interp_genarg ist (Genarg.in_gen (glbwit wit) x) >>= fun x -> Ftactic.return (Val.Dyn (Val.typ_opt, Some x)) end | PairArg (wit1, wit2) -> let (p, q) = x in interp_genarg ist (Genarg.in_gen (glbwit wit1) p) >>= fun p -> interp_genarg ist (Genarg.in_gen (glbwit wit2) q) >>= fun q -> Ftactic.return (Val.Dyn (Val.typ_pair, (p, q))) | ExtraArg s -> Geninterp.interp wit ist x (** returns [true] for genargs which have the same meaning independently of goals. *) and interp_genarg_constr_list ist x = Ftactic.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let lc = Genarg.out_gen (glbwit (wit_list wit_constr)) x in let (sigma,lc) = interp_constr_list ist env sigma lc in let lc = in_list (val_tag wit_constr) lc in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Ftactic.return lc) end and interp_genarg_var_list ist x = Ftactic.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let lc = Genarg.out_gen (glbwit (wit_list wit_var)) x in let lc = interp_hyp_list ist env sigma lc in let lc = in_list (val_tag wit_var) lc in Ftactic.return lc end (* Interprets tactic expressions : returns a "constr" *) and interp_ltac_constr ist e : EConstr.t Ftactic.t = let (>>=) = Ftactic.bind in begin wrap_error (val_interp ist e) begin function (err, info) -> match err with | Not_found -> Ftactic.enter begin fun gl -> let env = Proofview.Goal.env gl in Proofview.tclLIFT begin debugging_step ist (fun () -> str "evaluation failed for" ++ fnl() ++ Pptactic.pr_glob_tactic env e) end <*> Proofview.tclZERO Not_found end | err -> Proofview.tclZERO ~info err end end >>= fun result -> Ftactic.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in try let cresult = coerce_to_closed_constr env result in Proofview.tclLIFT begin debugging_step ist (fun () -> Pptactic.pr_glob_tactic env e ++ fnl() ++ str " has value " ++ fnl() ++ pr_econstr_env env sigma cresult) end <*> Ftactic.return cresult with CannotCoerceTo _ -> let env = Proofview.Goal.env gl in Tacticals.New.tclZEROMSG (str "Must evaluate to a closed term" ++ fnl() ++ str "offending expression: " ++ fnl() ++ pr_inspect env e result) end (* Interprets tactic expressions : returns a "tactic" *) and interp_tactic ist tac : unit Proofview.tactic = Ftactic.run (val_interp ist tac) (fun v -> tactic_of_value ist v) (* Provides a "name" for the trace to atomic tactics *) and name_atomic ?env tacexpr tac : unit Proofview.tactic = begin match env with | Some e -> Proofview.tclUNIT e | None -> Proofview.tclENV end >>= fun env -> Proofview.tclEVARMAP >>= fun sigma -> let name _ _ = Pptactic.pr_atomic_tactic env sigma tacexpr in Proofview.Trace.name_tactic name tac (* Interprets a primitive tactic *) and interp_atomic ist tac : unit Proofview.tactic = match tac with (* Basic tactics *) | TacIntroPattern (ev,l) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let sigma,l' = interp_intro_pattern_list_as_list ist env sigma l in Tacticals.New.tclWITHHOLES ev (name_atomic ~env (TacIntroPattern (ev,l)) (* spiwack: print uninterpreted, not sure if it is the expected behaviour. *) (Tactics.intro_patterns ev l')) sigma end | TacApply (a,ev,cb,cl) -> (* spiwack: until the tactic is in the monad *) Proofview.Trace.name_tactic (fun _ _ -> Pp.str"") begin Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let l = List.map (fun (k,c) -> let loc, f = interp_open_constr_with_bindings_loc ist c in (k,(make ?loc f))) cb in let sigma,tac = match cl with | None -> sigma, Tactics.apply_with_delayed_bindings_gen a ev l | Some cl -> let sigma,(id,cl) = interp_in_hyp_as ist env sigma cl in sigma, Tactics.apply_delayed_in a ev id l cl in Tacticals.New.tclWITHHOLES ev tac sigma end end | TacElim (ev,(keep,cb),cbo) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let sigma, cb = interp_open_constr_with_bindings ist env sigma cb in let sigma, cbo = Option.fold_left_map (interp_open_constr_with_bindings ist env) sigma cbo in let named_tac = let tac = Tactics.elim ev keep cb cbo in name_atomic ~env (TacElim (ev,(keep,cb),cbo)) tac in Tacticals.New.tclWITHHOLES ev named_tac sigma end | TacCase (ev,(keep,cb)) -> Proofview.Goal.enter begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in let sigma, cb = interp_open_constr_with_bindings ist env sigma cb in let named_tac = let tac = Tactics.general_case_analysis ev keep cb in name_atomic ~env (TacCase(ev,(keep,cb))) tac in Tacticals.New.tclWITHHOLES ev named_tac sigma end | TacMutualFix (id,n,l) -> (* spiwack: until the tactic is in the monad *) Proofview.Trace.name_tactic (fun _ _ -> Pp.str"") begin Proofview.Goal.enter begin fun gl -> let env = pf_env gl in let f sigma (id,n,c) = let (sigma,c_interp) = interp_type ist env sigma c in sigma , (interp_ident ist env sigma id,n,c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Tactics.mutual_fix (interp_ident ist env sigma id) n l_interp 0) end end | TacMutualCofix (id,l) -> (* spiwack: until the tactic is in the monad *) Proofview.Trace.name_tactic (fun _ _ -> Pp.str"") begin Proofview.Goal.enter begin fun gl -> let env = pf_env gl in let f sigma (id,c) = let (sigma,c_interp) = interp_type ist env sigma c in sigma , (interp_ident ist env sigma id,c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Tactics.mutual_cofix (interp_ident ist env sigma id) l_interp 0) end end | TacAssert (ev,b,t,ipat,c) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let (sigma,c) = let expected_type = if Option.is_empty t then WithoutTypeConstraint else IsType in let flags = open_constr_use_classes_flags () in interp_open_constr ~expected_type ~flags ist env sigma c in let sigma, ipat' = interp_intro_pattern_option ist env sigma ipat in let tac = Option.map (Option.map (interp_tactic ist)) t in Tacticals.New.tclWITHHOLES ev (name_atomic ~env (TacAssert(ev,b,Option.map (Option.map ignore) t,ipat,c)) (Tactics.forward b tac ipat' c)) sigma end | TacGeneralize cl -> Proofview.Goal.enter begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in let sigma, cl = interp_constr_with_occurrences_and_name_as_list ist env sigma cl in Tacticals.New.tclWITHHOLES false (name_atomic ~env (TacGeneralize cl) (Tactics.generalize_gen cl)) sigma end | TacLetTac (ev,na,c,clp,b,eqpat) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let clp = interp_clause ist env sigma clp in let eqpat = interp_intro_pattern_naming_option ist env sigma eqpat in if Locusops.is_nowhere clp (* typically "pose" *) then (* We try to fully-typecheck the term *) let flags = open_constr_use_classes_flags () in let (sigma,c_interp) = interp_open_constr ~flags ist env sigma c in let na = interp_name ist env sigma na in let let_tac = if b then Tactics.pose_tac na c_interp else let id = Option.default (make IntroAnonymous) eqpat in let with_eq = Some (true, id) in Tactics.letin_tac with_eq na c_interp None Locusops.nowhere in Tacticals.New.tclWITHHOLES ev (name_atomic ~env (TacLetTac(ev,na,c_interp,clp,b,eqpat)) let_tac) sigma else (* We try to keep the pattern structure as much as possible *) let let_pat_tac b na c cl eqpat = let id = Option.default (make IntroAnonymous) eqpat in let with_eq = if b then None else Some (true,id) in Tactics.letin_pat_tac ev with_eq na c cl in let (sigma',c) = interp_pure_open_constr ist env sigma c in name_atomic ~env (TacLetTac(ev,na,c,clp,b,eqpat)) (Tacticals.New.tclWITHHOLES ev (let_pat_tac b (interp_name ist env sigma na) (sigma,c) clp eqpat) sigma') end (* Derived basic tactics *) | TacInductionDestruct (isrec,ev,(l,el)) -> (* spiwack: some unknown part of destruct needs the goal to be prenormalised. *) Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let sigma,l = List.fold_left_map begin fun sigma (c,(ipato,ipats),cls) -> (* TODO: move sigma as a side-effect *) (* spiwack: the [*p] variants are for printing *) let cp = c in let c = interp_destruction_arg ist gl c in let ipato = interp_intro_pattern_naming_option ist env sigma ipato in let ipatsp = ipats in let sigma,ipats = interp_or_and_intro_pattern_option ist env sigma ipats in let cls = Option.map (interp_clause ist env sigma) cls in sigma,((c,(ipato,ipats),cls),(cp,(ipato,ipatsp),cls)) end sigma l in let l,lp = List.split l in let sigma,el = Option.fold_left_map (interp_open_constr_with_bindings ist env) sigma el in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (name_atomic ~env (TacInductionDestruct(isrec,ev,(lp,el))) (Tactics.induction_destruct isrec ev (l,el))) end (* Conversion *) | TacReduce (r,cl) -> Proofview.Goal.enter begin fun gl -> let (sigma,r_interp) = interp_red_expr ist (pf_env gl) (project gl) r in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Tactics.reduce r_interp (interp_clause ist (pf_env gl) (project gl) cl)) end | TacChange (check,None,c,cl) -> (* spiwack: until the tactic is in the monad *) Proofview.Trace.name_tactic (fun _ _ -> Pp.str"") begin Proofview.Goal.enter begin fun gl -> let is_onhyps = match cl.onhyps with | None | Some [] -> true | _ -> false in let is_onconcl = match cl.concl_occs with | AtLeastOneOccurrence | AllOccurrences | NoOccurrences -> true | _ -> false in let c_interp patvars env sigma = let lfun' = Id.Map.fold (fun id c lfun -> Id.Map.add id (Value.of_constr c) lfun) patvars ist.lfun in let ist = { ist with lfun = lfun' } in if is_onhyps && is_onconcl then interp_type ist env sigma c else interp_constr ist env sigma c in Tactics.change ~check None c_interp (interp_clause ist (pf_env gl) (project gl) cl) end end | TacChange (check,Some op,c,cl) -> (* spiwack: until the tactic is in the monad *) Proofview.Trace.name_tactic (fun _ _ -> Pp.str"") begin Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let op = interp_typed_pattern ist env sigma op in let to_catch = function Not_found -> true | e -> CErrors.is_anomaly e in let c_interp patvars env sigma = let lfun' = Id.Map.fold (fun id c lfun -> Id.Map.add id (Value.of_constr c) lfun) patvars ist.lfun in let env = ensure_freshness env in let ist = { ist with lfun = lfun' } in try interp_constr ist env sigma c with e when to_catch e (* Hack *) -> user_err (strbrk "Failed to get enough information from the left-hand side to type the right-hand side.") in Tactics.change ~check (Some op) c_interp (interp_clause ist env sigma cl) end end (* Equality and inversion *) | TacRewrite (ev,l,cl,by) -> Proofview.Goal.enter begin fun gl -> let l' = List.map (fun (b,m,(keep,c)) -> let f env sigma = interp_open_constr_with_bindings ist env sigma c in (b,m,keep,f)) l in let env = Proofview.Goal.env gl in let sigma = project gl in let cl = interp_clause ist env sigma cl in name_atomic ~env (TacRewrite (ev,l,cl,Option.map ignore by)) (Equality.general_multi_rewrite ev l' cl (Option.map (fun by -> Tacticals.New.tclCOMPLETE (interp_tactic ist by), Equality.Naive) by)) end | TacInversion (DepInversion (k,c,ids),hyp) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let (sigma,c_interp) = match c with | None -> sigma , None | Some c -> let (sigma,c_interp) = interp_constr ist env sigma c in sigma , Some c_interp in let dqhyps = interp_declared_or_quantified_hypothesis ist env sigma hyp in let sigma,ids_interp = interp_or_and_intro_pattern_option ist env sigma ids in Tacticals.New.tclWITHHOLES false (name_atomic ~env (TacInversion(DepInversion(k,c_interp,ids),dqhyps)) (Inv.dinv k c_interp ids_interp dqhyps)) sigma end | TacInversion (NonDepInversion (k,idl,ids),hyp) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let hyps = interp_hyp_list ist env sigma idl in let dqhyps = interp_declared_or_quantified_hypothesis ist env sigma hyp in let sigma, ids_interp = interp_or_and_intro_pattern_option ist env sigma ids in Tacticals.New.tclWITHHOLES false (name_atomic ~env (TacInversion (NonDepInversion (k,hyps,ids),dqhyps)) (Inv.inv_clause k ids_interp hyps dqhyps)) sigma end | TacInversion (InversionUsing (c,idl),hyp) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let (sigma,c_interp) = interp_constr ist env sigma c in let dqhyps = interp_declared_or_quantified_hypothesis ist env sigma hyp in let hyps = interp_hyp_list ist env sigma idl in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (name_atomic ~env (TacInversion (InversionUsing (c_interp,hyps),dqhyps)) (Leminv.lemInv_clause dqhyps c_interp hyps)) end (* Initial call for interpretation *) let default_ist () = let extra = TacStore.set TacStore.empty f_debug (get_debug ()) in { lfun = Id.Map.empty; poly = false; extra = extra } let eval_tactic t = Proofview.tclUNIT () >>= fun () -> (* delay for [default_ist] *) Proofview.tclLIFT db_initialize <*> interp_tactic (default_ist ()) t let eval_tactic_ist ist t = Proofview.tclLIFT db_initialize <*> interp_tactic ist t (** FFI *) module Value = struct include Taccoerce.Value let of_closure ist tac = let closure = VFun (UnnamedAppl,extract_trace ist, ist.lfun, [], tac) in of_tacvalue closure (** Apply toplevel tactic values *) let apply (f : value) (args: value list) = let fold arg (i, vars, lfun) = let id = Id.of_string ("x" ^ string_of_int i) in let x = Reference (ArgVar CAst.(make id)) in (succ i, x :: vars, Id.Map.add id arg lfun) in let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in let lfun = Id.Map.add (Id.of_string "F") f lfun in let ist = { (default_ist ()) with lfun = lfun; } in let tac = TacArg(CAst.make @@ TacCall (CAst.make (ArgVar CAst.(make @@ Id.of_string "F"),args))) in eval_tactic_ist ist tac end (* globalization + interpretation *) let interp_tac_gen lfun avoid_ids debug t = Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let extra = TacStore.set TacStore.empty f_debug debug in let extra = TacStore.set extra f_avoid_ids avoid_ids in let ist = { lfun; poly; extra } in let ltacvars = Id.Map.domain lfun in interp_tactic ist (intern_pure_tactic { (Genintern.empty_glob_sign env) with ltacvars } t) end let interp t = interp_tac_gen Id.Map.empty Id.Set.empty (get_debug()) t (* Used to hide interpretation for pretty-print, now just launch tactics *) (* [global] means that [t] should be internalized outside of goals. *) let hide_interp global t ot = let hide_interp env = let ist = Genintern.empty_glob_sign env in let te = intern_pure_tactic ist t in let t = eval_tactic te in match ot with | None -> t | Some t' -> Tacticals.New.tclTHEN t t' in if global then Proofview.tclENV >>= fun env -> hide_interp env else Proofview.Goal.enter begin fun gl -> hide_interp (Proofview.Goal.env gl) end (***************************************************************************) (** Register standard arguments *) let register_interp0 wit f = let open Ftactic.Notations in let interp ist v = f ist v >>= fun v -> Ftactic.return (Val.inject (val_tag wit) v) in Geninterp.register_interp0 wit interp let def_intern ist x = (ist, x) let def_subst _ x = x let def_interp ist x = Ftactic.return x let declare_uniform t = Genintern.register_intern0 t def_intern; Genintern.register_subst0 t def_subst; register_interp0 t def_interp let () = declare_uniform wit_unit let () = declare_uniform wit_int let () = declare_uniform wit_bool let () = declare_uniform wit_string let lift f = (); fun ist x -> Ftactic.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in Ftactic.return (f ist env sigma x) end let lifts f = (); fun ist x -> Ftactic.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let (sigma, v) = f ist env sigma x in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (* FIXME once we don't need to catch side effects *) (Proofview.tclTHEN (Proofview.Unsafe.tclSETENV (Global.env())) (Ftactic.return v)) end let interp_bindings' ist bl = Ftactic.return begin fun env sigma -> interp_bindings ist env sigma bl end let interp_constr_with_bindings' ist c = Ftactic.return begin fun env sigma -> interp_constr_with_bindings ist env sigma c end let interp_open_constr_with_bindings' ist c = Ftactic.return begin fun env sigma -> interp_open_constr_with_bindings ist env sigma c end let interp_destruction_arg' ist c = Ftactic.enter begin fun gl -> Ftactic.return (interp_destruction_arg ist gl c) end let interp_pre_ident ist env sigma s = s |> Id.of_string |> interp_ident ist env sigma |> Id.to_string let () = register_interp0 wit_int_or_var (fun ist n -> Ftactic.return (interp_int_or_var ist n)); register_interp0 wit_ref (lift interp_reference); register_interp0 wit_pre_ident (lift interp_pre_ident); register_interp0 wit_ident (lift interp_ident); register_interp0 wit_var (lift interp_hyp); register_interp0 wit_intropattern (lifts interp_intro_pattern) [@warning "-3"]; register_interp0 wit_simple_intropattern (lifts interp_intro_pattern); register_interp0 wit_clause_dft_concl (lift interp_clause); register_interp0 wit_constr (lifts interp_constr); register_interp0 wit_tacvalue (fun ist v -> Ftactic.return v); register_interp0 wit_red_expr (lifts interp_red_expr); register_interp0 wit_quant_hyp (lift interp_declared_or_quantified_hypothesis); register_interp0 wit_open_constr (lifts interp_open_constr); register_interp0 wit_bindings interp_bindings'; register_interp0 wit_constr_with_bindings interp_constr_with_bindings'; register_interp0 wit_open_constr_with_bindings interp_open_constr_with_bindings'; register_interp0 wit_destruction_arg interp_destruction_arg'; () let () = let interp ist tac = Ftactic.return (Value.of_closure ist tac) in register_interp0 wit_tactic interp let () = let interp ist tac = interp_tactic ist tac >>= fun () -> Ftactic.return () in register_interp0 wit_ltac interp let () = register_interp0 wit_uconstr (fun ist c -> Ftactic.enter begin fun gl -> Ftactic.return (interp_uconstr ist (Proofview.Goal.env gl) (Tacmach.New.project gl) c) end) (***************************************************************************) (* Other entry points *) let val_interp ist tac k = Ftactic.run (val_interp ist tac) k let interp_ltac_constr ist c k = Ftactic.run (interp_ltac_constr ist c) k let interp_redexp env sigma r = let ist = default_ist () in let gist = Genintern.empty_glob_sign env in interp_red_expr ist env sigma (intern_red_expr gist r) (***************************************************************************) (* Backwarding recursive needs of tactic glob/interp/eval functions *) let _ = let eval lfun poly env sigma ty tac = let extra = TacStore.set TacStore.empty f_debug (get_debug ()) in let ist = { lfun; poly; extra; } in let tac = interp_tactic ist tac in (* EJGA: We should also pass the proof name if desired, for now poly seems like enough to get reasonable behavior in practice *) let name, poly = Id.of_string "ltac_gen", poly in let name, poly = Id.of_string "ltac_gen", poly in let (c, sigma) = Pfedit.refine_by_tactic ~name ~poly env sigma ty tac in (EConstr.of_constr c, sigma) in GlobEnv.register_constr_interp0 wit_tactic eval let vernac_debug b = set_debug (if b then Tactic_debug.DebugOn 0 else Tactic_debug.DebugOff) let () = let open Goptions in declare_bool_option { optdepr = false; optname = "Ltac debug"; optkey = ["Ltac";"Debug"]; optread = (fun () -> get_debug () != Tactic_debug.DebugOff); optwrite = vernac_debug } let () = let open Goptions in declare_bool_option { optdepr = false; optname = "Ltac Backtrace"; optkey = ["Ltac"; "Backtrace"]; optread = (fun () -> !log_trace); optwrite = (fun b -> log_trace := b) } let () = Hook.set Vernacentries.interp_redexp_hook interp_redexp coq-8.11.0/plugins/ltac/tactic_debug.mli0000644000175000017500000000632113612664533020014 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* glob_tactic_expr -> (debug_info -> 'a Proofview.tactic) -> 'a Proofview.tactic (** Initializes debugger *) val db_initialize : unit Proofview.NonLogical.t (** Prints a constr *) val db_constr : debug_info -> env -> evar_map -> constr -> unit Proofview.NonLogical.t (** Prints the pattern rule *) val db_pattern_rule : debug_info -> env -> evar_map -> int -> (Genintern.glob_constr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> unit Proofview.NonLogical.t (** Prints a matched hypothesis *) val db_matched_hyp : debug_info -> env -> evar_map -> Id.t * constr option * constr -> Name.t -> unit Proofview.NonLogical.t (** Prints the matched conclusion *) val db_matched_concl : debug_info -> env -> evar_map -> constr -> unit Proofview.NonLogical.t (** Prints a success message when the goal has been matched *) val db_mc_pattern_success : debug_info -> unit Proofview.NonLogical.t (** Prints a failure message for an hypothesis pattern *) val db_hyp_pattern_failure : debug_info -> env -> evar_map -> Name.t * constr_pattern match_pattern -> unit Proofview.NonLogical.t (** Prints a matching failure message for a rule *) val db_matching_failure : debug_info -> unit Proofview.NonLogical.t (** Prints an evaluation failure message for a rule *) val db_eval_failure : debug_info -> Pp.t -> unit Proofview.NonLogical.t (** An exception handler *) val explain_logic_error: exn -> Pp.t (** For use in the Ltac debugger: some exception that are usually consider anomalies are acceptable because they are caught later in the process that is being debugged. One should not require from users that they report these anomalies. *) val explain_logic_error_no_anomaly : exn -> Pp.t (** Prints a logic failure message for a rule *) val db_logic_failure : debug_info -> exn -> unit Proofview.NonLogical.t (** Prints a logic failure message for a rule *) val db_breakpoint : debug_info -> lident message_token list -> unit Proofview.NonLogical.t val extract_ltac_trace : ?loc:Loc.t -> Tacexpr.ltac_trace -> Pp.t option Loc.located coq-8.11.0/plugins/ltac/pptactic.mli0000644000175000017500000001352713612664533017214 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Evd.evar_map -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> raw_tactic_expr -> Pp.t) -> 'a -> Pp.t type 'a glob_extra_genarg_printer = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> glob_tactic_expr -> Pp.t) -> 'a -> Pp.t type 'a extra_genarg_printer = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> Val.t -> Pp.t) -> 'a -> Pp.t type 'a raw_extra_genarg_printer_with_level = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> raw_tactic_expr -> Pp.t) -> tolerability -> 'a -> Pp.t type 'a glob_extra_genarg_printer_with_level = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> glob_tactic_expr -> Pp.t) -> tolerability -> 'a -> Pp.t type 'a extra_genarg_printer_with_level = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> Val.t -> Pp.t) -> tolerability -> 'a -> Pp.t val declare_extra_genarg_pprule : ('a, 'b, 'c) genarg_type -> 'a raw_extra_genarg_printer -> 'b glob_extra_genarg_printer -> 'c extra_genarg_printer -> unit val declare_extra_genarg_pprule_with_level : ('a, 'b, 'c) genarg_type -> 'a raw_extra_genarg_printer_with_level -> 'b glob_extra_genarg_printer_with_level -> 'c extra_genarg_printer_with_level -> (* surroounded *) tolerability -> (* non-surroounded *) tolerability -> unit val declare_extra_vernac_genarg_pprule : ('a, 'b, 'c) genarg_type -> 'a raw_extra_genarg_printer -> unit type grammar_terminals = Genarg.ArgT.any Extend.user_symbol grammar_tactic_prod_item_expr list type pp_tactic = { pptac_level : int; pptac_prods : grammar_terminals; } val pr_goal_selector : toplevel:bool -> Goal_select.t -> Pp.t val declare_notation_tactic_pprule : KerName.t -> pp_tactic -> unit val pr_with_occurrences : ('a -> Pp.t) -> 'a Locus.with_occurrences -> Pp.t val pr_red_expr : env -> Evd.evar_map -> (env -> Evd.evar_map -> 'a -> Pp.t) * (env -> Evd.evar_map -> 'a -> Pp.t) * ('b -> Pp.t) * (env -> Evd.evar_map -> 'c -> Pp.t) -> ('a,'b,'c) Genredexpr.red_expr_gen -> Pp.t val pr_may_eval : env -> Evd.evar_map -> (env -> Evd.evar_map -> 'a -> Pp.t) -> (env -> Evd.evar_map -> 'a -> Pp.t) -> ('b -> Pp.t) -> (env -> Evd.evar_map -> 'c -> Pp.t) -> ('a,'b,'c) Genredexpr.may_eval -> Pp.t val pr_and_short_name : ('a -> Pp.t) -> 'a Genredexpr.and_short_name -> Pp.t val pr_evaluable_reference_env : env -> evaluable_global_reference -> Pp.t val pr_quantified_hypothesis : quantified_hypothesis -> Pp.t val pr_in_clause : ('a -> Pp.t) -> 'a Locus.clause_expr -> Pp.t val pr_clauses : (* default: *) bool option -> ('a -> Pp.t) -> 'a Locus.clause_expr -> Pp.t (* Some true = default is concl; Some false = default is all; None = no default *) val pr_raw_generic : env -> Evd.evar_map -> rlevel generic_argument -> Pp.t val pr_glb_generic : env -> Evd.evar_map -> glevel generic_argument -> Pp.t val pr_raw_extend: env -> Evd.evar_map -> int -> ml_tactic_entry -> raw_tactic_arg list -> Pp.t val pr_glob_extend: env -> Evd.evar_map -> int -> ml_tactic_entry -> glob_tactic_arg list -> Pp.t val pr_extend : (Val.t -> Pp.t) -> int -> ml_tactic_entry -> Val.t list -> Pp.t val pr_alias_key : Names.KerName.t -> Pp.t val pr_alias : (Val.t -> Pp.t) -> int -> Names.KerName.t -> Val.t list -> Pp.t val pr_ltac_constant : ltac_constant -> Pp.t val pr_raw_tactic : env -> Evd.evar_map -> raw_tactic_expr -> Pp.t val pr_raw_tactic_level : env -> Evd.evar_map -> tolerability -> raw_tactic_expr -> Pp.t val pr_glob_tactic : env -> glob_tactic_expr -> Pp.t val pr_atomic_tactic : env -> Evd.evar_map -> atomic_tactic_expr -> Pp.t val pr_hintbases : string list option -> Pp.t val pr_auto_using : ('constr -> Pp.t) -> 'constr list -> Pp.t val pr_match_pattern : ('a -> Pp.t) -> 'a match_pattern -> Pp.t val pr_match_rule : bool -> ('a -> Pp.t) -> ('b -> Pp.t) -> ('b, 'a) match_rule -> Pp.t val pr_value : tolerability -> Val.t -> Pp.t val ltop : tolerability val make_constr_printer : (env -> Evd.evar_map -> tolerability -> 'a -> Pp.t) -> 'a Genprint.top_printer coq-8.11.0/plugins/ltac/tacinterp.mli0000644000175000017500000001205413612664533017370 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t val to_constr : t -> constr option val of_int : int -> t val to_int : t -> int option val to_list : t -> t list option val of_closure : Geninterp.interp_sign -> glob_tactic_expr -> t val cast : 'a typed_abstract_argument_type -> Geninterp.Val.t -> 'a val apply : t -> t list -> unit Proofview.tactic end (** Values for interpretation *) type value = Value.t module TacStore : Store.S with type t = Geninterp.TacStore.t and type 'a field = 'a Geninterp.TacStore.field (** Signature for interpretation: val\_interp and interpretation functions *) type interp_sign = Geninterp.interp_sign = { lfun : value Id.Map.t ; poly : bool ; extra : TacStore.t } open Genintern val f_avoid_ids : Id.Set.t TacStore.field val f_debug : debug_info TacStore.field val extract_ltac_constr_values : interp_sign -> Environ.env -> Ltac_pretype.constr_under_binders Id.Map.t (** Given an interpretation signature, extract all values which are coercible to a [constr]. *) (** Sets the debugger mode *) val set_debug : debug_info -> unit (** Gives the state of debug *) val get_debug : unit -> debug_info val type_uconstr : ?flags:Pretyping.inference_flags -> ?expected_type:Pretyping.typing_constraint -> Geninterp.interp_sign -> Ltac_pretype.closed_glob_constr -> constr Tactypes.delayed_open (** Adds an interpretation function for extra generic arguments *) val interp_genarg : interp_sign -> glob_generic_argument -> Value.t Ftactic.t (** Interprets any expression *) val val_interp : interp_sign -> glob_tactic_expr -> (value -> unit Proofview.tactic) -> unit Proofview.tactic (** Interprets an expression that evaluates to a constr *) val interp_ltac_constr : interp_sign -> glob_tactic_expr -> (constr -> unit Proofview.tactic) -> unit Proofview.tactic (** Interprets redexp arguments *) val interp_redexp : Environ.env -> Evd.evar_map -> raw_red_expr -> Evd.evar_map * red_expr (** Interprets tactic expressions *) val interp_hyp : interp_sign -> Environ.env -> Evd.evar_map -> lident -> Id.t val interp_glob_closure : interp_sign -> Environ.env -> Evd.evar_map -> ?kind:Pretyping.typing_constraint -> ?pattern_mode:bool -> glob_constr_and_expr -> Ltac_pretype.closed_glob_constr val interp_uconstr : interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Ltac_pretype.closed_glob_constr val interp_constr_gen : Pretyping.typing_constraint -> interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Evd.evar_map * constr val interp_bindings : interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr bindings -> Evd.evar_map * constr bindings val interp_open_constr : ?expected_type:Pretyping.typing_constraint -> ?flags:Pretyping.inference_flags -> interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Evd.evar_map * EConstr.constr val interp_open_constr_with_classes : ?expected_type:Pretyping.typing_constraint -> interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Evd.evar_map * EConstr.constr val interp_open_constr_with_bindings : interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr with_bindings -> Evd.evar_map * EConstr.constr with_bindings (** Initial call for interpretation *) val eval_tactic : glob_tactic_expr -> unit Proofview.tactic val eval_tactic_ist : interp_sign -> glob_tactic_expr -> unit Proofview.tactic (** Same as [eval_tactic], but with the provided [interp_sign]. *) val tactic_of_value : interp_sign -> Value.t -> unit Proofview.tactic (** Globalization + interpretation *) val interp_tac_gen : value Id.Map.t -> Id.Set.t -> debug_info -> raw_tactic_expr -> unit Proofview.tactic val interp : raw_tactic_expr -> unit Proofview.tactic (** Hides interpretation for pretty-print *) val hide_interp : bool -> raw_tactic_expr -> unit Proofview.tactic option -> unit Proofview.tactic (** Internals that can be useful for syntax extensions. *) val interp_ltac_var : (value -> 'a) -> interp_sign -> (Environ.env * Evd.evar_map) option -> lident -> 'a val interp_int : interp_sign -> lident -> int val interp_int_or_var : interp_sign -> int Locus.or_var -> int val default_ist : unit -> Geninterp.interp_sign (** Empty ist with debug set on the current value. *) coq-8.11.0/plugins/ltac/tauto_plugin.mlpack0000644000175000017500000000000613612664533020571 0ustar treinentreinenTauto coq-8.11.0/plugins/ltac/profile_ltac.ml0000644000175000017500000003704213612664533017675 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* strbrk "Ltac Profiler cannot yet handle backtracking \ into multi-success tactics; profiling results may be wildly inaccurate.") let warn_encountered_multi_success_backtracking () = if !encountered_multi_success_backtracking then warn_profile_backtracking () let encounter_multi_success_backtracking () = if not !encountered_multi_success_backtracking then begin encountered_multi_success_backtracking := true; warn_encountered_multi_success_backtracking () end (* *************** tree data structure for profiling ****************** *) type treenode = { name : M.key; total : float; local : float; ncalls : int; max_total : float; children : treenode M.t } let empty_treenode name = { name; total = 0.0; local = 0.0; ncalls = 0; max_total = 0.0; children = M.empty; } let root = "root" module Local = Summary.Local let stack = Local.ref ~name:"LtacProf-stack" [empty_treenode root] let reset_profile_tmp () = Local.(stack := [empty_treenode root]); encountered_multi_success_backtracking := false (* ************** XML Serialization ********************* *) let rec of_ltacprof_tactic (name, t) = assert (String.equal name t.name); let open Xml_datatype in let total = string_of_float t.total in let local = string_of_float t.local in let ncalls = string_of_int t.ncalls in let max_total = string_of_float t.max_total in let children = List.map of_ltacprof_tactic (M.bindings t.children) in Element ("ltacprof_tactic", [ ("name", name); ("total",total); ("local",local); ("ncalls",ncalls); ("max_total",max_total)], children) let of_ltacprof_results t = let open Xml_datatype in assert(String.equal t.name root); let children = List.map of_ltacprof_tactic (M.bindings t.children) in Element ("ltacprof", [("total_time", string_of_float t.total)], children) let rec to_ltacprof_tactic m xml = let open Xml_datatype in match xml with | Element ("ltacprof_tactic", [("name", name); ("total",total); ("local",local); ("ncalls",ncalls); ("max_total",max_total)], xs) -> let node = { name; total = float_of_string total; local = float_of_string local; ncalls = int_of_string ncalls; max_total = float_of_string max_total; children = List.fold_left to_ltacprof_tactic M.empty xs; } in M.add name node m | _ -> CErrors.anomaly Pp.(str "Malformed ltacprof_tactic XML.") let to_ltacprof_results xml = let open Xml_datatype in match xml with | Element ("ltacprof", [("total_time", t)], xs) -> { name = root; total = float_of_string t; ncalls = 0; max_total = 0.0; local = 0.0; children = List.fold_left to_ltacprof_tactic M.empty xs } | _ -> CErrors.anomaly Pp.(str "Malformed ltacprof XML.") let feedback_results results = Feedback.(feedback (Custom (None, "ltacprof_results", of_ltacprof_results results))) (* ************** pretty printing ************************************* *) let format_sec x = (Printf.sprintf "%.3fs" x) let format_ratio x = (Printf.sprintf "%.1f%%" (100. *. x)) let padl n s = ws (max 0 (n - utf8_length s)) ++ str s let padr_with c n s = let ulength = utf8_length s in str (utf8_sub s 0 n) ++ str (String.make (max 0 (n - ulength)) c) let rec list_iter_is_last f = function | [] -> [] | [x] -> [f true x] | x :: xs -> f false x :: list_iter_is_last f xs let header = str " tactic local total calls max " ++ fnl () ++ str "────────────────────────────────────────┴──────┴──────┴───────┴─────────┘" ++ fnl () let rec print_node ~filter all_total indent prefix (s, e) = h 0 ( padr_with '-' 40 (prefix ^ s ^ " ") ++ padl 7 (format_ratio (e.local /. all_total)) ++ padl 7 (format_ratio (e.total /. all_total)) ++ padl 8 (string_of_int e.ncalls) ++ padl 10 (format_sec (e.max_total)) ) ++ fnl () ++ print_table ~filter all_total indent false e.children and print_table ~filter all_total indent first_level table = let fold _ n l = let s, total = n.name, n.total in if filter s total then (s, n) :: l else l in let ls = M.fold fold table [] in match ls with | [s, n] when not first_level -> v 0 (print_node ~filter all_total indent (indent ^ "└") (s, n)) | _ -> let ls = List.sort (fun (_, { total = s1 }) (_, { total = s2}) -> compare s2 s1) ls in let iter is_last = let sep0 = if first_level then "" else if is_last then " " else " │" in let sep1 = if first_level then "─" else if is_last then " └─" else " ├─" in print_node ~filter all_total (indent ^ sep0) (indent ^ sep1) in prlist (fun pr -> pr) (list_iter_is_last iter ls) let to_string ~filter ?(cutoff=0.0) node = let tree = node.children in let all_total = M.fold (fun _ { total } a -> total +. a) node.children 0.0 in let flat_tree = let global = ref M.empty in let find_tactic tname l = try M.find tname !global with Not_found -> let e = empty_treenode tname in global := M.add tname e !global; e in let add_tactic tname stats = global := M.add tname stats !global in let sum_stats add_total { name; total = t1; local = l1; ncalls = n1; max_total = m1 } { total = t2; local = l2; ncalls = n2; max_total = m2 } = { name; total = if add_total then t1 +. t2 else t1; local = l1 +. l2; ncalls = n1 + n2; max_total = if add_total then max m1 m2 else m1; children = M.empty; } in let rec cumulate table = let iter _ ({ name; children } as statistics) = if filter name then begin let stats' = find_tactic name global in add_tactic name (sum_stats true stats' statistics); end; cumulate children in M.iter iter table in cumulate tree; !global in warn_encountered_multi_success_backtracking (); let filter s n = filter s && (all_total <= 0.0 || n /. all_total >= cutoff /. 100.0) in let msg = h 0 (str "total time: " ++ padl 11 (format_sec (all_total))) ++ fnl () ++ fnl () ++ header ++ print_table ~filter all_total "" true flat_tree ++ fnl () ++ header ++ print_table ~filter all_total "" true tree in msg (* ******************** profiling code ************************************** *) let get_child name node = try M.find name node.children with Not_found -> empty_treenode name let time () = let times = Unix.times () in times.Unix.tms_utime +. times.Unix.tms_stime let string_of_call ck = let s = string_of_ppcmds (match ck with | Tacexpr.LtacNotationCall s -> Pptactic.pr_alias_key s | Tacexpr.LtacNameCall cst -> Pptactic.pr_ltac_constant cst | Tacexpr.LtacVarCall (id, t) -> Names.Id.print id | Tacexpr.LtacAtomCall te -> (Pptactic.pr_glob_tactic (Global.env ()) (Tacexpr.TacAtom (CAst.make te))) | Tacexpr.LtacConstrInterp (c, _) -> pr_glob_constr_env (Global.env ()) c | Tacexpr.LtacMLCall te -> (Pptactic.pr_glob_tactic (Global.env ()) te) ) in let s = String.map (fun c -> if c = '\n' then ' ' else c) s in let s = try String.sub s 0 (CString.string_index_from s 0 "(*") with Not_found -> s in String.trim s let rec merge_sub_tree name tree acc = try let t = M.find name acc in let t = { name; total = t.total +. tree.total; ncalls = t.ncalls + tree.ncalls; local = t.local +. tree.local; max_total = max t.max_total tree.max_total; children = M.fold merge_sub_tree tree.children t.children; } in M.add name t acc with Not_found -> M.add name tree acc let merge_roots ?(disjoint=true) t1 t2 = assert(String.equal t1.name t2.name); { name = t1.name; ncalls = t1.ncalls + t2.ncalls; local = if disjoint then t1.local +. t2.local else t1.local; total = if disjoint then t1.total +. t2.total else t1.total; max_total = if disjoint then max t1.max_total t2.max_total else t1.max_total; children = M.fold merge_sub_tree t2.children t1.children } let rec find_in_stack what acc = function | [] -> None | { name } as x :: rest when String.equal name what -> Some(acc, x, rest) | { name } as x :: rest -> find_in_stack what (x :: acc) rest let exit_tactic ~count_call start_time c = let diff = time () -. start_time in match Local.(!stack) with | [] | [_] -> (* oops, our stack is invalid *) encounter_multi_success_backtracking (); reset_profile_tmp () | node :: (parent :: rest as full_stack) -> let name = string_of_call c in if not (String.equal name node.name) then (* oops, our stack is invalid *) encounter_multi_success_backtracking (); let node = { node with total = node.total +. diff; local = node.local +. diff; ncalls = node.ncalls + (if count_call then 1 else 0); max_total = max node.max_total diff; } in (* updating the stack *) let parent = match find_in_stack node.name [] full_stack with | None -> (* no rec-call, we graft the subtree *) let parent = { parent with local = parent.local -. diff; children = M.add node.name node parent.children } in Local.(stack := parent :: rest); parent | Some(to_update, self, rest) -> (* we coalesce the rec-call and update the lower stack *) let self = merge_roots ~disjoint:false self node in let updated_stack = List.fold_left (fun s x -> (try M.find x.name (List.hd s).children with Not_found -> x) :: s) (self :: rest) to_update in Local.(stack := updated_stack); List.hd Local.(!stack) in (* Calls are over, we reset the stack and send back data *) if rest == [] && get_profiling () then begin assert(String.equal root parent.name); reset_profile_tmp (); feedback_results parent end let tclFINALLY tac (finally : unit Proofview.tactic) = let open Proofview.Notations in Proofview.tclIFCATCH tac (fun v -> finally <*> Proofview.tclUNIT v) (fun (exn, info) -> finally <*> Proofview.tclZERO ~info exn) let do_profile s call_trace ?(count_call=true) tac = let open Proofview.Notations in Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> if !is_profiling then match call_trace, Local.(!stack) with | (_, c) :: _, parent :: rest -> let name = string_of_call c in let node = get_child name parent in Local.(stack := node :: parent :: rest); Some (time ()) | _ :: _, [] -> assert false | _ -> None else None)) >>= function | Some start_time -> tclFINALLY tac (Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> (match call_trace with | (_, c) :: _ -> exit_tactic ~count_call start_time c | [] -> ())))) | None -> tac (* ************** Accumulation of data from workers ************************* *) let get_local_profiling_results () = List.hd Local.(!stack) (* We maintain our own cache of document data, given that the semantics of the STM implies that synchronized state for opaque proofs will be lost on QED. This provides some complications later on as we will have to simulate going back on the document on our own. *) module DData = struct type t = Feedback.doc_id * Stateid.t let compare x y = compare x y end module SM = Map.Make(DData) let data = ref SM.empty let _ = Feedback.(add_feeder (function | { doc_id = d; span_id = s; contents = Custom (_, "ltacprof_results", xml) } -> let results = to_ltacprof_results xml in let other_results = (* Multi success can cause this *) try SM.find (d,s) !data with Not_found -> empty_treenode root in data := SM.add (d,s) (merge_roots results other_results) !data | _ -> ())) let reset_profile () = reset_profile_tmp (); data := SM.empty (* ****************************** Named timers ****************************** *) let timer_data = ref M.empty let timer_name = function | Some v -> v | None -> "" let restart_timer name = timer_data := M.add (timer_name name) (System.get_time ()) !timer_data let get_timer name = try M.find (timer_name name) !timer_data with Not_found -> System.get_time () let finish_timing ~prefix name = let tend = System.get_time () in let tstart = get_timer name in Feedback.msg_notice(str prefix ++ pr_opt str name ++ str " ran for " ++ System.fmt_time_difference tstart tend) (* ******************** *) let print_results_filter ~cutoff ~filter = (* The STM doesn't provide yet a proper document query and traversal API, thus we need to re-check if some states are current anymore (due to backtracking) using the `state_of_id` API. *) let valid (did,id) _ = Stm.(state_of_id ~doc:(get_doc did) id) <> `Expired in data := SM.filter valid !data; let results = SM.fold (fun _ -> merge_roots ~disjoint:true) !data (empty_treenode root) in let results = merge_roots results Local.(CList.last !stack) in Feedback.msg_notice (to_string ~cutoff ~filter results) ;; let print_results ~cutoff = print_results_filter ~cutoff ~filter:(fun _ -> true) let print_results_tactic tactic = print_results_filter ~cutoff:!Flags.profile_ltac_cutoff ~filter:(fun s -> String.(equal tactic (sub (s ^ ".") 0 (min (1+length s) (length tactic))))) let do_print_results_at_close () = if get_profiling () then print_results ~cutoff:!Flags.profile_ltac_cutoff let _ = Declaremods.append_end_library_hook do_print_results_at_close let () = let open Goptions in declare_bool_option { optdepr = false; optname = "Ltac Profiling"; optkey = ["Ltac"; "Profiling"]; optread = get_profiling; optwrite = set_profiling } coq-8.11.0/plugins/ltac/pltac.ml0000644000175000017500000000534113612664533016332 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* v | e -> Tacexp (e:raw_tactic_expr) let genarg_of_unit () = in_gen (rawwit Stdarg.wit_unit) () let genarg_of_int n = in_gen (rawwit Stdarg.wit_int) n let genarg_of_ipattern pat = in_gen (rawwit Tacarg.wit_simple_intropattern) pat let genarg_of_uconstr c = in_gen (rawwit Stdarg.wit_uconstr) c let in_tac tac = in_gen (rawwit Tacarg.wit_ltac) tac let reference_to_id qid = if Libnames.qualid_is_ident qid then CAst.make ?loc:qid.CAst.loc @@ Libnames.qualid_basename qid else CErrors.user_err ?loc:qid.CAst.loc (str "This expression should be a simple identifier.") let tactic_mode = Entry.create "vernac:tactic_command" let new_entry name = let e = Entry.create name in e let toplevel_selector = new_entry "vernac:toplevel_selector" let tacdef_body = new_entry "tactic:tacdef_body" (* Registers [tactic_mode] as a parser for proof editing *) let classic_proof_mode = Pvernac.register_proof_mode "Classic" tactic_mode (* Hack to parse "[ id" without dropping [ *) let test_bracket_ident = Pcoq.Entry.of_parser "test_bracket_ident" (fun _ strm -> match stream_nth 0 strm with | KEYWORD "[" -> (match stream_nth 1 strm with | IDENT _ -> () | _ -> raise Stream.Failure) | _ -> raise Stream.Failure) (* Tactics grammar rules *) let hint = G_proofs.hint } GRAMMAR EXTEND Gram GLOBAL: tactic tacdef_body tactic_expr binder_tactic tactic_arg command hint tactic_mode constr_may_eval constr_eval toplevel_selector operconstr; tactic_then_last: [ [ "|"; lta = LIST0 (OPT tactic_expr) SEP "|" -> { Array.map (function None -> TacId [] | Some t -> t) (Array.of_list lta) } | -> { [||] } ] ] ; tactic_then_gen: [ [ ta = tactic_expr; "|"; tg = tactic_then_gen -> { let (first,last) = tg in (ta::first, last) } | ta = tactic_expr; ".."; l = tactic_then_last -> { ([], Some (ta, l)) } | ".."; l = tactic_then_last -> { ([], Some (TacId [], l)) } | ta = tactic_expr -> { ([ta], None) } | "|"; tg = tactic_then_gen -> { let (first,last) = tg in (TacId [] :: first, last) } | -> { ([TacId []], None) } ] ] ; tactic_then_locality: (* [true] for the local variant [TacThens] and [false] for [TacExtend] *) [ [ "[" ; l = OPT">" -> { if Option.is_empty l then true else false } ] ] ; tactic_expr: [ "5" RIGHTA [ te = binder_tactic -> { te } ] | "4" LEFTA [ ta0 = tactic_expr; ";"; ta1 = binder_tactic -> { TacThen (ta0, ta1) } | ta0 = tactic_expr; ";"; ta1 = tactic_expr -> { TacThen (ta0,ta1) } | ta0 = tactic_expr; ";"; l = tactic_then_locality; tg = tactic_then_gen; "]" -> { let (first,tail) = tg in match l , tail with | false , Some (t,last) -> TacThen (ta0,TacExtendTac (Array.of_list first, t, last)) | true , Some (t,last) -> TacThens3parts (ta0, Array.of_list first, t, last) | false , None -> TacThen (ta0,TacDispatch first) | true , None -> TacThens (ta0,first) } ] | "3" RIGHTA [ IDENT "try"; ta = tactic_expr -> { TacTry ta } | IDENT "do"; n = int_or_var; ta = tactic_expr -> { TacDo (n,ta) } | IDENT "timeout"; n = int_or_var; ta = tactic_expr -> { TacTimeout (n,ta) } | IDENT "time"; s = OPT string; ta = tactic_expr -> { TacTime (s,ta) } | IDENT "repeat"; ta = tactic_expr -> { TacRepeat ta } | IDENT "progress"; ta = tactic_expr -> { TacProgress ta } | IDENT "once"; ta = tactic_expr -> { TacOnce ta } | IDENT "exactly_once"; ta = tactic_expr -> { TacExactlyOnce ta } | IDENT "infoH"; ta = tactic_expr -> { TacShowHyps ta } (*To do: put Abstract in Refiner*) | IDENT "abstract"; tc = NEXT -> { TacAbstract (tc,None) } | IDENT "abstract"; tc = NEXT; "using"; s = ident -> { TacAbstract (tc,Some s) } | sel = selector; ta = tactic_expr -> { TacSelect (sel, ta) } ] (*End of To do*) | "2" RIGHTA [ ta0 = tactic_expr; "+"; ta1 = binder_tactic -> { TacOr (ta0,ta1) } | ta0 = tactic_expr; "+"; ta1 = tactic_expr -> { TacOr (ta0,ta1) } | IDENT "tryif" ; ta = tactic_expr ; "then" ; tat = tactic_expr ; "else" ; tae = tactic_expr -> { TacIfThenCatch(ta,tat,tae) } | ta0 = tactic_expr; "||"; ta1 = binder_tactic -> { TacOrelse (ta0,ta1) } | ta0 = tactic_expr; "||"; ta1 = tactic_expr -> { TacOrelse (ta0,ta1) } ] | "1" RIGHTA [ b = match_key; IDENT "goal"; "with"; mrl = match_context_list; "end" -> { TacMatchGoal (b,false,mrl) } | b = match_key; IDENT "reverse"; IDENT "goal"; "with"; mrl = match_context_list; "end" -> { TacMatchGoal (b,true,mrl) } | b = match_key; c = tactic_expr; "with"; mrl = match_list; "end" -> { TacMatch (b,c,mrl) } | IDENT "first" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> { TacFirst l } | IDENT "solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> { TacSolve l } | IDENT "idtac"; l = LIST0 message_token -> { TacId l } | g=failkw; n = [ n = int_or_var -> { n } | -> { fail_default_value } ]; l = LIST0 message_token -> { TacFail (g,n,l) } | st = simple_tactic -> { st } | a = tactic_arg -> { TacArg(CAst.make ~loc a) } | r = reference; la = LIST0 tactic_arg_compat -> { TacArg(CAst.make ~loc @@ TacCall (CAst.make ~loc (r,la))) } ] | "0" [ "("; a = tactic_expr; ")" -> { a } | "["; ">"; tg = tactic_then_gen; "]" -> { let (tf,tail) = tg in begin match tail with | Some (t,tl) -> TacExtendTac(Array.of_list tf,t,tl) | None -> TacDispatch tf end } | a = tactic_atom -> { TacArg (CAst.make ~loc a) } ] ] ; failkw: [ [ IDENT "fail" -> { TacLocal } | IDENT "gfail" -> { TacGlobal } ] ] ; (* binder_tactic: level 5 of tactic_expr *) binder_tactic: [ RIGHTA [ "fun"; it = LIST1 input_fun ; "=>"; body = tactic_expr LEVEL "5" -> { TacFun (it,body) } | "let"; isrec = [IDENT "rec" -> { true } | -> { false } ]; llc = LIST1 let_clause SEP "with"; "in"; body = tactic_expr LEVEL "5" -> { TacLetIn (isrec,llc,body) } | IDENT "info"; tc = tactic_expr LEVEL "5" -> { TacInfo tc } ] ] ; (* Tactic arguments to the right of an application *) tactic_arg_compat: [ [ a = tactic_arg -> { a } | c = Constr.constr -> { (match c with { CAst.v = CRef (r,None) } -> Reference r | c -> ConstrMayEval (ConstrTerm c)) } (* Unambiguous entries: tolerated w/o "ltac:" modifier *) | "()" -> { TacGeneric (genarg_of_unit ()) } ] ] ; (* Can be used as argument and at toplevel in tactic expressions. *) tactic_arg: [ [ c = constr_eval -> { ConstrMayEval c } | IDENT "fresh"; l = LIST0 fresh_id -> { TacFreshId l } | IDENT "type_term"; c=uconstr -> { TacPretype c } | IDENT "numgoals" -> { TacNumgoals } ] ] ; (* If a qualid is given, use its short name. TODO: have the shortest non ambiguous name where dots are replaced by "_"? Probably too verbose most of the time. *) fresh_id: [ [ s = STRING -> { Locus.ArgArg s (*| id = ident -> Locus.ArgVar (!@loc,id)*) } | qid = qualid -> { Locus.ArgVar (CAst.make ~loc @@ Libnames.qualid_basename qid) } ] ] ; constr_eval: [ [ IDENT "eval"; rtc = red_expr; "in"; c = Constr.constr -> { ConstrEval (rtc,c) } | IDENT "context"; id = identref; "["; c = Constr.lconstr; "]" -> { ConstrContext (id,c) } | IDENT "type"; IDENT "of"; c = Constr.constr -> { ConstrTypeOf c } ] ] ; constr_may_eval: (* For extensions *) [ [ c = constr_eval -> { c } | c = Constr.constr -> { ConstrTerm c } ] ] ; tactic_atom: [ [ n = integer -> { TacGeneric (genarg_of_int n) } | r = reference -> { TacCall (CAst.make ~loc (r,[])) } | "()" -> { TacGeneric (genarg_of_unit ()) } ] ] ; match_key: [ [ "match" -> { Once } | "lazymatch" -> { Select } | "multimatch" -> { General } ] ] ; input_fun: [ [ "_" -> { Name.Anonymous } | l = ident -> { Name.Name l } ] ] ; let_clause: [ [ idr = identref; ":="; te = tactic_expr -> { (CAst.map (fun id -> Name id) idr, arg_of_expr te) } | na = ["_" -> { CAst.make ~loc Anonymous } ]; ":="; te = tactic_expr -> { (na, arg_of_expr te) } | idr = identref; args = LIST1 input_fun; ":="; te = tactic_expr -> { (CAst.map (fun id -> Name id) idr, arg_of_expr (TacFun(args,te))) } ] ] ; match_pattern: [ [ IDENT "context"; oid = OPT Constr.ident; "["; pc = Constr.lconstr_pattern; "]" -> { Subterm (oid, pc) } | pc = Constr.lconstr_pattern -> { Term pc } ] ] ; match_hyps: [ [ na = name; ":"; mp = match_pattern -> { Hyp (na, mp) } | na = name; ":="; "["; mpv = match_pattern; "]"; ":"; mpt = match_pattern -> { Def (na, mpv, mpt) } | na = name; ":="; mpv = match_pattern -> { let t, ty = match mpv with | Term t -> (match t with | { CAst.v = CCast (t, (CastConv ty | CastVM ty | CastNative ty)) } -> Term t, Some (Term ty) | _ -> mpv, None) | _ -> mpv, None in Def (na, t, Option.default (Term (CAst.make @@ CHole (None, IntroAnonymous, None))) ty) } ] ] ; match_context_rule: [ [ largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; "=>"; te = tactic_expr -> { Pat (largs, mp, te) } | "["; largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; "]"; "=>"; te = tactic_expr -> { Pat (largs, mp, te) } | "_"; "=>"; te = tactic_expr -> { All te } ] ] ; match_context_list: [ [ mrl = LIST1 match_context_rule SEP "|" -> { mrl } | "|"; mrl = LIST1 match_context_rule SEP "|" -> { mrl } ] ] ; match_rule: [ [ mp = match_pattern; "=>"; te = tactic_expr -> { Pat ([],mp,te) } | "_"; "=>"; te = tactic_expr -> { All te } ] ] ; match_list: [ [ mrl = LIST1 match_rule SEP "|" -> { mrl } | "|"; mrl = LIST1 match_rule SEP "|" -> { mrl } ] ] ; message_token: [ [ id = identref -> { MsgIdent id } | s = STRING -> { MsgString s } | n = integer -> { MsgInt n } ] ] ; ltac_def_kind: [ [ ":=" -> { false } | "::=" -> { true } ] ] ; (* Definitions for tactics *) tacdef_body: [ [ name = Constr.global; it=LIST1 input_fun; redef = ltac_def_kind; body = tactic_expr -> { if redef then Tacexpr.TacticRedefinition (name, TacFun (it, body)) else let id = reference_to_id name in Tacexpr.TacticDefinition (id, TacFun (it, body)) } | name = Constr.global; redef = ltac_def_kind; body = tactic_expr -> { if redef then Tacexpr.TacticRedefinition (name, body) else let id = reference_to_id name in Tacexpr.TacticDefinition (id, body) } ] ] ; tactic: [ [ tac = tactic_expr -> { tac } ] ] ; range_selector: [ [ n = natural ; "-" ; m = natural -> { (n, m) } | n = natural -> { (n, n) } ] ] ; (* We unfold a range selectors list once so that we can make a special case * for a unique SelectNth selector. *) range_selector_or_nth: [ [ n = natural ; "-" ; m = natural; l = OPT [","; l = LIST1 range_selector SEP "," -> { l } ] -> { Goal_select.SelectList ((n, m) :: Option.default [] l) } | n = natural; l = OPT [","; l = LIST1 range_selector SEP "," -> { l } ] -> { let open Goal_select in Option.cata (fun l -> SelectList ((n, n) :: l)) (SelectNth n) l } ] ] ; selector_body: [ [ l = range_selector_or_nth -> { l } | test_bracket_ident; "["; id = ident; "]" -> { Goal_select.SelectId id } ] ] ; selector: [ [ IDENT "only"; sel = selector_body; ":" -> { sel } ] ] ; toplevel_selector: [ [ sel = selector_body; ":" -> { sel } | "!"; ":" -> { Goal_select.SelectAlreadyFocused } | IDENT "all"; ":" -> { Goal_select.SelectAll } ] ] ; tactic_mode: [ [ g = OPT toplevel_selector; tac = G_vernac.query_command -> { tac g } | g = OPT toplevel_selector; "{" -> { Vernacexpr.VernacSubproof g } ] ] ; command: [ [ IDENT "Proof"; "with"; ta = Pltac.tactic; l = OPT [ "using"; l = G_vernac.section_subset_expr -> { l } ] -> { Vernacexpr.VernacProof (Some (in_tac ta), l) } | IDENT "Proof"; "using"; l = G_vernac.section_subset_expr; ta = OPT [ "with"; ta = Pltac.tactic -> { in_tac ta } ] -> { Vernacexpr.VernacProof (ta,Some l) } ] ] ; hint: [ [ IDENT "Extern"; n = natural; c = OPT Constr.constr_pattern ; "=>"; tac = Pltac.tactic -> { Hints.HintsExtern (n,c, in_tac tac) } ] ] ; operconstr: LEVEL "0" [ [ IDENT "ltac"; ":"; "("; tac = Pltac.tactic_expr; ")" -> { let arg = Genarg.in_gen (Genarg.rawwit Tacarg.wit_tactic) tac in CAst.make ~loc @@ CHole (None, IntroAnonymous, Some arg) } ] ] ; END { open Stdarg open Tacarg open Vernacextend open Goptions open Libnames let print_info_trace = ref None let () = declare_int_option { optdepr = false; optname = "print info trace"; optkey = ["Info" ; "Level"]; optread = (fun () -> !print_info_trace); optwrite = fun n -> print_info_trace := n; } let vernac_solve ~pstate n info tcom b = let open Goal_select in let pstate, status = Proof_global.map_fold_proof_endline (fun etac p -> let with_end_tac = if b then Some etac else None in let global = match n with SelectAll | SelectList _ -> true | _ -> false in let info = Option.append info !print_info_trace in let (p,status) = Pfedit.solve n info (Tacinterp.hide_interp global tcom None) ?with_end_tac p in (* in case a strict subtree was completed, go back to the top of the prooftree *) let p = Proof.maximal_unfocus Vernacentries.command_focus p in p,status) pstate in if not status then Feedback.feedback Feedback.AddedAxiom; pstate let pr_ltac_selector s = Pptactic.pr_goal_selector ~toplevel:true s } VERNAC ARGUMENT EXTEND ltac_selector PRINTED BY { pr_ltac_selector } | [ toplevel_selector(s) ] -> { s } END { let pr_ltac_info n = str "Info" ++ spc () ++ int n } VERNAC ARGUMENT EXTEND ltac_info PRINTED BY { pr_ltac_info } | [ "Info" natural(n) ] -> { n } END { let pr_ltac_use_default b = if b then (* Bug: a space is inserted before "..." *) str ".." else mt () } VERNAC ARGUMENT EXTEND ltac_use_default PRINTED BY { pr_ltac_use_default } | [ "." ] -> { false } | [ "..." ] -> { true } END { let is_anonymous_abstract = function | TacAbstract (_,None) -> true | TacSolve [TacAbstract (_,None)] -> true | _ -> false let rm_abstract = function | TacAbstract (t,_) -> t | TacSolve [TacAbstract (t,_)] -> TacSolve [t] | x -> x let is_explicit_terminator = function TacSolve _ -> true | _ -> false } VERNAC { tactic_mode } EXTEND VernacSolve STATE proof | [ ltac_selector_opt(g) ltac_info_opt(n) tactic(t) ltac_use_default(def) ] => { classify_as_proofstep } -> { let g = Option.default (Goal_select.get_default_goal_selector ()) g in vernac_solve g n t def } | [ "par" ":" ltac_info_opt(n) tactic(t) ltac_use_default(def) ] => { let anon_abstracting_tac = is_anonymous_abstract t in let solving_tac = is_explicit_terminator t in let parallel = `Yes (solving_tac,anon_abstracting_tac) in let pbr = if solving_tac then Some "par" else None in VtProofStep{ parallel = parallel; proof_block_detection = pbr } } -> { let t = rm_abstract t in vernac_solve Goal_select.SelectAll n t def } END { let pr_ltac_tactic_level n = str "(at level " ++ int n ++ str ")" } VERNAC ARGUMENT EXTEND ltac_tactic_level PRINTED BY { pr_ltac_tactic_level } | [ "(" "at" "level" natural(n) ")" ] -> { n } END VERNAC ARGUMENT EXTEND ltac_production_sep | [ "," string(sep) ] -> { sep } END { let pr_ltac_production_item = function | Tacentries.TacTerm s -> quote (str s) | Tacentries.TacNonTerm (_, ((arg, None), None)) -> str arg | Tacentries.TacNonTerm (_, ((arg, Some _), None)) -> assert false | Tacentries.TacNonTerm (_, ((arg, sep), Some id)) -> let sep = match sep with | None -> mt () | Some sep -> str "," ++ spc () ++ quote (str sep) in str arg ++ str "(" ++ Id.print id ++ sep ++ str ")" } VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY { pr_ltac_production_item } | [ string(s) ] -> { Tacentries.TacTerm s } | [ ident(nt) "(" ident(p) ltac_production_sep_opt(sep) ")" ] -> { Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, sep), Some p)) } | [ ident(nt) ] -> { Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, None), None)) } END VERNAC COMMAND EXTEND VernacTacticNotation | #[ deprecation; locality; ] [ "Tactic" "Notation" ltac_tactic_level_opt(n) ne_ltac_production_item_list(r) ":=" tactic(e) ] => { VtSideff ([], VtNow) } -> { let n = Option.default 0 n in Tacentries.add_tactic_notation (Locality.make_module_locality locality) n ?deprecation r e; } END VERNAC COMMAND EXTEND VernacPrintLtac CLASSIFIED AS QUERY | [ "Print" "Ltac" reference(r) ] -> { Feedback.msg_notice (Tacintern.print_ltac r) } END VERNAC COMMAND EXTEND VernacLocateLtac CLASSIFIED AS QUERY | [ "Locate" "Ltac" reference(r) ] -> { Tacentries.print_located_tactic r } END { let pr_ltac_ref = Libnames.pr_qualid let pr_tacdef_body env sigma tacdef_body = let id, redef, body = match tacdef_body with | TacticDefinition ({CAst.v=id}, body) -> Id.print id, false, body | TacticRedefinition (id, body) -> pr_ltac_ref id, true, body in let idl, body = match body with | Tacexpr.TacFun (idl,b) -> idl,b | _ -> [], body in id ++ prlist (function Name.Anonymous -> str " _" | Name.Name id -> spc () ++ Id.print id) idl ++ (if redef then str" ::=" else str" :=") ++ brk(1,1) ++ Pptactic.pr_raw_tactic env sigma body } VERNAC ARGUMENT EXTEND ltac_tacdef_body PRINTED BY { pr_tacdef_body env sigma } | [ tacdef_body(t) ] -> { t } END VERNAC COMMAND EXTEND VernacDeclareTacticDefinition | #[ deprecation; locality; ] [ "Ltac" ne_ltac_tacdef_body_list_sep(l, "with") ] => { VtSideff (List.map (function | TacticDefinition ({CAst.v=r},_) -> r | TacticRedefinition (qid,_) -> qualid_basename qid) l, VtLater) } -> { Tacentries.register_ltac (Locality.make_module_locality locality) ?deprecation l; } END VERNAC COMMAND EXTEND VernacPrintLtacs CLASSIFIED AS QUERY | [ "Print" "Ltac" "Signatures" ] -> { Tacentries.print_ltacs () } END coq-8.11.0/plugins/ltac/tacenv.ml0000644000175000017500000001341413612664533016507 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* CErrors.anomaly (str "Unknown tactic alias: " ++ KerName.print key ++ str ".") let check_alias key = KNmap.mem key !alias_map (** ML tactic extensions (TacML) *) type ml_tactic = Geninterp.Val.t list -> Geninterp.interp_sign -> unit Proofview.tactic module MLName = struct type t = ml_tactic_name let compare tac1 tac2 = let c = String.compare tac1.mltac_tactic tac2.mltac_tactic in if c = 0 then String.compare tac1.mltac_plugin tac2.mltac_plugin else c end module MLTacMap = Map.Make(MLName) let pr_tacname t = str t.mltac_plugin ++ str "::" ++ str t.mltac_tactic let tac_tab = ref MLTacMap.empty let register_ml_tactic ?(overwrite = false) s (t : ml_tactic array) = let () = if MLTacMap.mem s !tac_tab then if overwrite then tac_tab := MLTacMap.remove s !tac_tab else CErrors.anomaly (str "Cannot redeclare tactic " ++ pr_tacname s ++ str ".") in tac_tab := MLTacMap.add s t !tac_tab let interp_ml_tactic { mltac_name = s; mltac_index = i } = try let tacs = MLTacMap.find s !tac_tab in let () = if Array.length tacs <= i then raise Not_found in tacs.(i) with Not_found -> CErrors.user_err (str "The tactic " ++ pr_tacname s ++ str " is not installed.") (***************************************************************************) (* Tactic registration *) (* Summary and Object declaration *) open Libobject type ltac_entry = { tac_for_ml : bool; tac_body : glob_tactic_expr; tac_redef : ModPath.t list; tac_deprecation : Deprecation.t option } let mactab = Summary.ref (KNmap.empty : ltac_entry KNmap.t) ~name:"tactic-definition" let ltac_entries () = !mactab let interp_ltac r = (KNmap.find r !mactab).tac_body let is_ltac_for_ml_tactic r = (KNmap.find r !mactab).tac_for_ml let add ~deprecation kn b t = let entry = { tac_for_ml = b; tac_body = t; tac_redef = []; tac_deprecation = deprecation; } in mactab := KNmap.add kn entry !mactab let replace kn path t = let path = KerName.modpath path in let entry _ e = { e with tac_body = t; tac_redef = path :: e.tac_redef } in mactab := KNmap.modify kn entry !mactab let tac_deprecation kn = try (KNmap.find kn !mactab).tac_deprecation with Not_found -> None let load_md i ((sp, kn), (local, id, b, t, deprecation)) = match id with | None -> let () = if not local then push_tactic (Nametab.Until i) sp kn in add ~deprecation kn b t | Some kn0 -> replace kn0 kn t let open_md i ((sp, kn), (local, id, b, t, deprecation)) = match id with | None -> let () = if not local then push_tactic (Nametab.Exactly i) sp kn in add ~deprecation kn b t | Some kn0 -> replace kn0 kn t let cache_md ((sp, kn), (local, id ,b, t, deprecation)) = match id with | None -> let () = push_tactic (Nametab.Until 1) sp kn in add ~deprecation kn b t | Some kn0 -> replace kn0 kn t let subst_kind subst id = match id with | None -> None | Some kn -> Some (Mod_subst.subst_kn subst kn) let subst_md (subst, (local, id, b, t, deprecation)) = (local, subst_kind subst id, b, Tacsubst.subst_tactic subst t, deprecation) let classify_md (local, _, _, _, _ as o) = Substitute o let inMD : bool * ltac_constant option * bool * glob_tactic_expr * Deprecation.t option -> obj = declare_object {(default_object "TAC-DEFINITION") with cache_function = cache_md; load_function = load_md; open_function = open_md; subst_function = subst_md; classify_function = classify_md} let register_ltac for_ml local ?deprecation id tac = ignore (Lib.add_leaf id (inMD (local, None, for_ml, tac, deprecation))) let redefine_ltac local ?deprecation kn tac = Lib.add_anonymous_leaf (inMD (local, Some kn, false, tac, deprecation)) coq-8.11.0/plugins/ltac/tacarg.mli0000644000175000017500000000500513612664533016636 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ev :: acc | _ -> EConstr.fold sigma evrec acc c in evrec [] c let instantiate_tac n c ido = Proofview.V82.tactic begin fun gl -> let sigma = gl.sigma in let evl = match ido with ConclLocation () -> evar_list sigma (pf_concl gl) | HypLocation (id,hloc) -> let decl = Environ.lookup_named id (pf_env gl) in match hloc with InHyp -> (match decl with | LocalAssum (_,typ) -> evar_list sigma (EConstr.of_constr typ) | _ -> user_err Pp.(str "Please be more specific: in type or value?")) | InHypTypeOnly -> evar_list sigma (EConstr.of_constr (NamedDecl.get_type decl)) | InHypValueOnly -> (match decl with | LocalDef (_,body,_) -> evar_list sigma (EConstr.of_constr body) | _ -> user_err Pp.(str "Not a defined hypothesis.")) in if List.length evl < n then user_err Pp.(str "Not enough uninstantiated existential variables."); if n <= 0 then user_err Pp.(str "Incorrect existential variable index."); let evk,_ = List.nth evl (n-1) in instantiate_evar evk c sigma gl end let instantiate_tac_by_name id c = Proofview.V82.tactic begin fun gl -> let sigma = gl.sigma in let evk = try Evd.evar_key id sigma with Not_found -> user_err Pp.(str "Unknown existential variable.") in instantiate_evar evk c sigma gl end let let_evar name typ = let src = (Loc.tag Evar_kinds.GoalEvar) in Proofview.Goal.enter begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let sigma, _ = Typing.sort_of env sigma typ in let id = match name with | Name.Anonymous -> let id = Namegen.id_of_name_using_hdchar env sigma typ name in Namegen.next_ident_away_in_goal id (Termops.vars_of_env env) | Name.Name id -> id in let (sigma, evar) = Evarutil.new_evar env sigma ~src ~naming:(Namegen.IntroFresh id) typ in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Tactics.pose_tac (Name.Name id) evar) end let hget_evar n = let open EConstr in Proofview.Goal.enter begin fun gl -> let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in let evl = evar_list sigma concl in if List.length evl < n then user_err Pp.(str "Not enough uninstantiated existential variables."); if n <= 0 then user_err Pp.(str "Incorrect existential variable index."); let ev = List.nth evl (n-1) in let ev_type = EConstr.existential_type sigma ev in let r = Retyping.relevance_of_type (Proofview.Goal.env gl) sigma ev_type in Tactics.change_concl (mkLetIn (make_annot Name.Anonymous r,mkEvar ev,ev_type,concl)) end coq-8.11.0/plugins/ltac/pltac.mli0000644000175000017500000000343113612664533016501 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ('a * Tacexpr.ltac_call_kind) list -> ?count_call:bool -> 'b Proofview.tactic -> 'b Proofview.tactic val set_profiling : bool -> unit (* Cut off results < than specified cutoff *) val print_results : cutoff:float -> unit val print_results_tactic : string -> unit val reset_profile : unit -> unit val restart_timer : string option -> unit val finish_timing : prefix:string -> string option -> unit val do_print_results_at_close : unit -> unit (* The collected statistics for a tactic. The timing data is collected over all * instances of a given tactic from its parent. E.g. if tactic 'aaa' calls * 'foo' twice, then 'aaa' will contain just one entry for 'foo' with the * statistics of the two invocations combined, and also combined over all * invocations of 'aaa'. * total: time spent running this tactic and its subtactics (seconds) * local: time spent running this tactic, minus its subtactics (seconds) * ncalls: the number of invocations of this tactic that have been made * max_total: the greatest running time of a single invocation (seconds) *) type treenode = { name : CString.Map.key; total : float; local : float; ncalls : int; max_total : float; children : treenode CString.Map.t } (* Returns the profiling results known by the current process *) val get_local_profiling_results : unit -> treenode val feedback_results : treenode -> unit coq-8.11.0/plugins/ltac/tactic_matching.ml0000644000175000017500000003560713612664533020360 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constr_matching.bound_ident_map * Ltac_pretype.extended_patvar_map = fun (l, lc) -> (l, Id.Map.map (fun c -> [], c) lc) (** Adds a binding to a {!Id.Map.t} if the identifier is [Some id] *) let id_map_try_add id x m = match id with | Some id -> Id.Map.add id (Lazy.force x) m | None -> m (** Adds a binding to a {!Id.Map.t} if the name is [Name id] *) let id_map_try_add_name id x m = match id with | Name id -> Id.Map.add id x m | Anonymous -> m (** Takes the union of two {!Id.Map.t}. If there is conflict, the binding of the right-hand argument shadows that of the left-hand argument. *) let id_map_right_biased_union m1 m2 = if Id.Map.is_empty m1 then m2 (* Don't reconstruct the whole map *) else Id.Map.fold Id.Map.add m2 m1 (** Tests whether the substitution [s] is empty. *) let is_empty_subst (ln,lm) = Id.Map.(is_empty ln && is_empty lm) (** {6 Non-linear patterns} *) (** The patterns of Ltac are not necessarily linear. Non-linear pattern are partially handled by the {!Matching} module, however goal patterns are not primitive to {!Matching}, hence we must deal with non-linearity between hypotheses and conclusion. Subterms are considered equal up to the equality implemented in [equal_instances]. *) (* spiwack: it doesn't seem to be quite the same rule for non-linear term patterns and non-linearity between hypotheses and/or conclusion. Indeed, in [Matching], matching is made modulo syntactic equality, and here we merge modulo conversion. It may be a good idea to have an entry point of [Matching] with a partial substitution as argument instead of merging substitution here. That would ensure consistency. *) let equal_instances env sigma (ctx',c') (ctx,c) = (* How to compare instances? Do we want the terms to be convertible? unifiable? Do we want the universe levels to be relevant? (historically, conv_x is used) *) CList.equal Id.equal ctx ctx' && Reductionops.is_conv env sigma c' c (** Merges two substitutions. Raises [Not_coherent_metas] when encountering two instances of the same metavariable which are not equal according to {!equal_instances}. *) exception Not_coherent_metas let verify_metas_coherence env sigma (ln1,lcm) (ln,lm) = let merge id oc1 oc2 = match oc1, oc2 with | None, None -> None | None, Some c | Some c, None -> Some c | Some c1, Some c2 -> if equal_instances env sigma c1 c2 then Some c1 else raise Not_coherent_metas in let (+++) lfun1 lfun2 = Id.Map.fold Id.Map.add lfun1 lfun2 in (* ppedrot: Is that even correct? *) let merged = ln +++ ln1 in (merged, Id.Map.merge merge lcm lm) let matching_error = CErrors.UserError (Some "tactic matching" , Pp.str "No matching clauses for match.") let imatching_error = (matching_error, Exninfo.null) (** A functor is introduced to share the environment and the evar_map. They do not change and it would be a pity to introduce closures everywhere just for the occasional calls to {!equal_instances}. *) module type StaticEnvironment = sig val env : Environ.env val sigma : Evd.evar_map end module PatternMatching (E:StaticEnvironment) = struct (** {6 The pattern-matching monad } *) (** To focus on the algorithmic portion of pattern-matching, the bookkeeping is relegated to a monad: the composition of the backtracking monad of {!IStream.t} with a "writer" effect. *) (* spiwack: as we don't benefit from the various stream optimisations of Haskell, it may be costly to give the monad in direct style such as here. We may want to use some continuation passing style. *) type 'a tac = 'a Proofview.tactic type 'a m = { stream : 'r. ('a -> unit t -> 'r tac) -> unit t -> 'r tac } (** The empty substitution. *) let empty_subst = Id.Map.empty , Id.Map.empty (** Composes two substitutions using {!verify_metas_coherence}. It must be a monoid with neutral element {!empty_subst}. Raises [Not_coherent_metas] when composition cannot be achieved. *) let subst_prod s1 s2 = if is_empty_subst s1 then s2 else if is_empty_subst s2 then s1 else verify_metas_coherence E.env E.sigma s1 s2 (** The empty context substitution. *) let empty_context_subst = Id.Map.empty (** Compose two context substitutions, in case of conflict the right hand substitution shadows the left hand one. *) let context_subst_prod = id_map_right_biased_union (** The empty term substitution. *) let empty_term_subst = Id.Map.empty (** Compose two terms substitutions, in case of conflict the right hand substitution shadows the left hand one. *) let term_subst_prod = id_map_right_biased_union (** Merge two writers (and ignore the first value component). *) let merge m1 m2 = try Some { subst = subst_prod m1.subst m2.subst; context = context_subst_prod m1.context m2.context; terms = term_subst_prod m1.terms m2.terms; lhs = m2.lhs; } with Not_coherent_metas -> None (** Monadic [return]: returns a single success with empty substitutions. *) let return (type a) (lhs:a) : a m = { stream = fun k ctx -> k lhs ctx } (** Monadic bind: each success of [x] is replaced by the successes of [f x]. The substitutions of [x] and [f x] are composed, dropping the apparent successes when the substitutions are not coherent. *) let (>>=) (type a) (type b) (m:a m) (f:a -> b m) : b m = { stream = fun k ctx -> m.stream (fun x ctx -> (f x).stream k ctx) ctx } (** A variant of [(>>=)] when the first argument returns [unit]. *) let (<*>) (type a) (m:unit m) (y:a m) : a m = { stream = fun k ctx -> m.stream (fun () ctx -> y.stream k ctx) ctx } (** Failure of the pattern-matching monad: no success. *) let fail (type a) : a m = { stream = fun _ _ -> Proofview.tclZERO matching_error } let run (m : 'a m) = let ctx = { subst = empty_subst ; context = empty_context_subst ; terms = empty_term_subst ; lhs = (); } in let eval lhs ctx = Proofview.tclUNIT { ctx with lhs } in m.stream eval ctx (** Chooses in a list, in the same order as the list *) let rec pick (l:'a list) (e, info) : 'a m = match l with | [] -> { stream = fun _ _ -> Proofview.tclZERO ~info e } | x :: l -> { stream = fun k ctx -> Proofview.tclOR (k x ctx) (fun e -> (pick l e).stream k ctx) } let pick l = pick l imatching_error (** Declares a substitution, a context substitution and a term substitution. *) let put subst context terms : unit m = let s = { subst ; context ; terms ; lhs = () } in { stream = fun k ctx -> match merge s ctx with None -> Proofview.tclZERO matching_error | Some s -> k () s } (** Declares a substitution. *) let put_subst subst : unit m = put subst empty_context_subst empty_term_subst (** Declares a term substitution. *) let put_terms terms : unit m = put empty_subst empty_context_subst terms (** {6 Pattern-matching} *) (** [wildcard_match_term lhs] matches a term against a wildcard pattern ([_ => lhs]). It has a single success with an empty substitution. *) let wildcard_match_term = return (** [pattern_match_term refresh pat term lhs] returns the possible matchings of [term] with the pattern [pat => lhs]. If refresh is true, refreshes the universes of [term]. *) let pattern_match_term refresh pat term lhs = (* let term = if refresh then Termops.refresh_universes_strict term else term in *) match pat with | Term p -> begin try put_subst (Constr_matching.extended_matches E.env E.sigma p term) <*> return lhs with Constr_matching.PatternMatchingFailure -> fail end | Subterm (id_ctxt,p) -> let rec map s (e, info) = { stream = fun k ctx -> match IStream.peek s with | IStream.Nil -> Proofview.tclZERO ~info e | IStream.Cons ({ Constr_matching.m_sub ; m_ctx }, s) -> let subst = adjust m_sub in let context = id_map_try_add id_ctxt m_ctx Id.Map.empty in let terms = empty_term_subst in let nctx = { subst ; context ; terms ; lhs = () } in match merge ctx nctx with | None -> (map s (e, info)).stream k ctx | Some nctx -> Proofview.tclOR (k lhs nctx) (fun e -> (map s e).stream k ctx) } in map (Constr_matching.match_subterm E.env E.sigma p term) imatching_error (** [rule_match_term term rule] matches the term [term] with the matching rule [rule]. *) let rule_match_term term = function | All lhs -> wildcard_match_term lhs | Pat ([],pat,lhs) -> pattern_match_term false pat term lhs | Pat _ -> (* Rules with hypotheses, only work in match goal. *) fail (** [match_term term rules] matches the term [term] with the set of matching rules [rules].*) let rec match_term (e, info) term rules = match rules with | [] -> { stream = fun _ _ -> Proofview.tclZERO ~info e } | r :: rules -> { stream = fun k ctx -> let head = rule_match_term term r in let tail e = match_term e term rules in Proofview.tclOR (head.stream k ctx) (fun e -> (tail e).stream k ctx) } (** [hyp_match_type hypname pat hyps] matches a single hypothesis pattern [hypname:pat] against the hypotheses in [hyps]. Tries the hypotheses in order. For each success returns the name of the matched hypothesis. *) let hyp_match_type hypname pat hyps = pick hyps >>= fun decl -> let id = NamedDecl.get_id decl in let refresh = is_local_def decl in pattern_match_term refresh pat (NamedDecl.get_type decl) () <*> put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*> return id (** [hyp_match_type hypname bodypat typepat hyps] matches a single hypothesis pattern [hypname := bodypat : typepat] against the hypotheses in [hyps].Tries the hypotheses in order. For each success returns the name of the matched hypothesis. *) let hyp_match_body_and_type hypname bodypat typepat hyps = pick hyps >>= function | LocalDef (id,body,hyp) -> pattern_match_term false bodypat body () <*> pattern_match_term true typepat hyp () <*> put_terms (id_map_try_add_name hypname (EConstr.mkVar id.binder_name) empty_term_subst) <*> return id.binder_name | LocalAssum (id,hyp) -> fail (** [hyp_match pat hyps] dispatches to {!hyp_match_type} or {!hyp_match_body_and_type} depending on whether [pat] is [Hyp _] or [Def _]. *) let hyp_match pat hyps = match pat with | Hyp ({CAst.v=hypname},typepat) -> hyp_match_type hypname typepat hyps | Def ({CAst.v=hypname},bodypat,typepat) -> hyp_match_body_and_type hypname bodypat typepat hyps (** [hyp_pattern_list_match pats hyps lhs], matches the list of patterns [pats] against the hypotheses in [hyps], and eventually returns [lhs]. *) let rec hyp_pattern_list_match pats hyps lhs = match pats with | pat::pats -> hyp_match pat hyps >>= fun matched_hyp -> (* spiwack: alternatively it is possible to return the list with the matched hypothesis removed directly in [hyp_match]. *) let select_matched_hyp decl = Id.equal (NamedDecl.get_id decl) matched_hyp in let hyps = CList.remove_first select_matched_hyp hyps in hyp_pattern_list_match pats hyps lhs | [] -> return lhs (** [rule_match_goal hyps concl rule] matches the rule [rule] against the goal [hyps|-concl]. *) let rule_match_goal hyps concl = function | All lhs -> wildcard_match_term lhs | Pat (hyppats,conclpat,lhs) -> (* the rules are applied from the topmost one (in the concrete syntax) to the bottommost. *) let hyppats = List.rev hyppats in pattern_match_term false conclpat concl () <*> hyp_pattern_list_match hyppats hyps lhs (** [match_goal hyps concl rules] matches the goal [hyps|-concl] with the set of matching rules [rules]. *) let rec match_goal (e, info) hyps concl rules = match rules with | [] -> { stream = fun _ _ -> Proofview.tclZERO ~info e } | r :: rules -> { stream = fun k ctx -> let head = rule_match_goal hyps concl r in let tail e = match_goal e hyps concl rules in Proofview.tclOR (head.stream k ctx) (fun e -> (tail e).stream k ctx) } end (** [match_term env sigma term rules] matches the term [term] with the set of matching rules [rules]. The environment [env] and the evar_map [sigma] are not currently used, but avoid code duplication. *) let match_term env sigma term rules = let module E = struct let env = env let sigma = sigma end in let module M = PatternMatching(E) in M.run (M.match_term imatching_error term rules) (** [match_goal env sigma hyps concl rules] matches the goal [hyps|-concl] with the set of matching rules [rules]. The environment [env] and the evar_map [sigma] are used to check convertibility for pattern variables shared between hypothesis patterns or the conclusion pattern. *) let match_goal env sigma hyps concl rules = let module E = struct let env = env let sigma = sigma end in let module M = PatternMatching(E) in M.run (M.match_goal imatching_error hyps concl rules) coq-8.11.0/plugins/ltac/rewrite.mli0000644000175000017500000000755313612664533017070 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* strategy val map_strategy : ('a -> 'b) -> ('c -> 'd) -> ('a, 'c) strategy_ast -> ('b, 'd) strategy_ast val pr_strategy : ('a -> Pp.t) -> ('b -> Pp.t) -> ('a, 'b) strategy_ast -> Pp.t (** Entry point for user-level "rewrite_strat" *) val cl_rewrite_clause_strat : strategy -> Id.t option -> unit Proofview.tactic (** Entry point for user-level "setoid_rewrite" *) val cl_rewrite_clause : interp_sign * (glob_constr_and_expr * glob_constr_and_expr bindings) -> bool -> Locus.occurrences -> Id.t option -> unit Proofview.tactic val is_applied_rewrite_relation : env -> evar_map -> rel_context -> constr -> types option val declare_relation : rewrite_attributes -> ?binders:local_binder_expr list -> constr_expr -> constr_expr -> Id.t -> constr_expr option -> constr_expr option -> constr_expr option -> unit val add_setoid : rewrite_attributes -> local_binder_expr list -> constr_expr -> constr_expr -> constr_expr -> Id.t -> unit val add_morphism_interactive : rewrite_attributes -> constr_expr -> Id.t -> Lemmas.t val add_morphism_as_parameter : rewrite_attributes -> constr_expr -> Id.t -> unit val add_morphism : rewrite_attributes -> local_binder_expr list -> constr_expr -> constr_expr -> Id.t -> Lemmas.t val get_reflexive_proof : env -> evar_map -> constr -> constr -> evar_map * constr val get_symmetric_proof : env -> evar_map -> constr -> constr -> evar_map * constr val get_transitive_proof : env -> evar_map -> constr -> constr -> evar_map * constr val default_morphism : (types * constr option) option list * (types * types option) option -> constr -> constr * constr val setoid_symmetry : unit Proofview.tactic val setoid_symmetry_in : Id.t -> unit Proofview.tactic val setoid_reflexivity : unit Proofview.tactic val setoid_transitivity : constr option -> unit Proofview.tactic val apply_strategy : strategy -> Environ.env -> Names.Id.Set.t -> constr -> bool * constr -> evars -> rewrite_result coq-8.11.0/plugins/ltac/tacintern.ml0000644000175000017500000010147413612664533017222 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Anonymous | Name id -> Name (intern_ident l ist id) let strict_check = ref false let adjust_loc loc = if !strict_check then None else loc (* Globalize a name which must be bound -- actually just check it is bound *) let intern_hyp ist ({loc;v=id} as locid) = if not !strict_check then locid else if find_ident id ist then make id else CErrors.user_err ?loc Pp.(str "Hypothesis" ++ spc () ++ Id.print id ++ spc() ++ str "was not found in the current environment.") let intern_or_var f ist = function | ArgVar locid -> ArgVar (intern_hyp ist locid) | ArgArg x -> ArgArg (f x) let intern_int_or_var = intern_or_var (fun (n : int) -> n) let intern_string_or_var = intern_or_var (fun (s : string) -> s) let intern_global_reference ist qid = if qualid_is_ident qid && find_var (qualid_basename qid) ist then ArgVar (make ?loc:qid.CAst.loc @@ qualid_basename qid) else try ArgArg (qid.CAst.loc,locate_global_with_alias qid) with Not_found -> Nametab.error_global_not_found qid let intern_ltac_variable ist qid = if qualid_is_ident qid && find_var (qualid_basename qid) ist then (* A local variable of any type *) ArgVar (make ?loc:qid.CAst.loc @@ qualid_basename qid) else raise Not_found let intern_constr_reference strict ist qid = let id = qualid_basename qid in if qualid_is_ident qid && not strict && find_hyp (qualid_basename qid) ist then (DAst.make @@ GVar id), Some (make @@ CRef (qid,None)) else if qualid_is_ident qid && find_var (qualid_basename qid) ist then (DAst.make @@ GVar id), if strict then None else Some (make @@ CRef (qid,None)) else DAst.make @@ GRef (locate_global_with_alias qid,None), if strict then None else Some (make @@ CRef (qid,None)) (* Internalize an isolated reference in position of tactic *) let warn_deprecated_tactic = Deprecation.create_warning ~object_name:"Tactic" ~warning_name:"deprecated-tactic" pr_qualid let warn_deprecated_alias = Deprecation.create_warning ~object_name:"Tactic Notation" ~warning_name:"deprecated-tactic-notation" Pptactic.pr_alias_key let intern_isolated_global_tactic_reference qid = let loc = qid.CAst.loc in let kn = Tacenv.locate_tactic qid in Option.iter (fun depr -> warn_deprecated_tactic ?loc (qid,depr)) @@ Tacenv.tac_deprecation kn; TacCall (CAst.make ?loc (ArgArg (loc,kn),[])) let intern_isolated_tactic_reference strict ist qid = (* An ltac reference *) try Reference (intern_ltac_variable ist qid) with Not_found -> (* A global tactic *) try intern_isolated_global_tactic_reference qid with Not_found -> (* Tolerance for compatibility, allow not to use "constr:" *) try ConstrMayEval (ConstrTerm (intern_constr_reference strict ist qid)) with Not_found -> (* Reference not found *) Nametab.error_global_not_found qid (* Internalize an applied tactic reference *) let intern_applied_global_tactic_reference qid = let loc = qid.CAst.loc in let kn = Tacenv.locate_tactic qid in Option.iter (fun depr -> warn_deprecated_tactic ?loc (qid,depr)) @@ Tacenv.tac_deprecation kn; ArgArg (loc,kn) let intern_applied_tactic_reference ist qid = (* An ltac reference *) try intern_ltac_variable ist qid with Not_found -> (* A global tactic *) try intern_applied_global_tactic_reference qid with Not_found -> (* Reference not found *) Nametab.error_global_not_found qid (* Intern a reference parsed in a non-tactic entry *) let intern_non_tactic_reference strict ist qid = (* An ltac reference *) try Reference (intern_ltac_variable ist qid) with Not_found -> (* A constr reference *) try ConstrMayEval (ConstrTerm (intern_constr_reference strict ist qid)) with Not_found -> (* Tolerance for compatibility, allow not to use "ltac:" *) try intern_isolated_global_tactic_reference qid with Not_found -> (* By convention, use IntroIdentifier for unbound ident, when not in a def *) if qualid_is_ident qid && not strict then let id = qualid_basename qid in let ipat = in_gen (glbwit wit_intro_pattern) (make ?loc:qid.CAst.loc @@ IntroNaming (IntroIdentifier id)) in TacGeneric ipat else (* Reference not found *) Nametab.error_global_not_found qid let intern_message_token ist = function | (MsgString _ | MsgInt _ as x) -> x | MsgIdent id -> MsgIdent (intern_hyp ist id) let intern_message ist = List.map (intern_message_token ist) let intern_quantified_hypothesis ist = function | AnonHyp n -> AnonHyp n | NamedHyp id -> (* Uncomment to disallow "intros until n" in ltac when n is not bound *) NamedHyp ((*snd (intern_hyp ist (dloc,*)id(* ))*)) let intern_binding_name ist x = (* We use identifier both for variables and binding names *) (* Todo: consider the body of the lemma to which the binding refer and if a term w/o ltac vars, check the name is indeed quantified *) x let intern_constr_gen pattern_mode isarity {ltacvars=lfun; genv=env; extra; intern_sign} c = let warn = if !strict_check then fun x -> x else Constrintern.for_grammar in let scope = if isarity then Pretyping.IsType else Pretyping.WithoutTypeConstraint in let ltacvars = { Constrintern.ltac_vars = lfun; ltac_bound = Id.Set.empty; ltac_extra = extra; } in let c' = warn (Constrintern.intern_core scope ~pattern_mode ~ltacvars env Evd.(from_env env) intern_sign) c in (c',if !strict_check then None else Some c) let intern_constr = intern_constr_gen false false let intern_type = intern_constr_gen false true (* Globalize bindings *) let intern_binding ist = map (fun (b,c) -> intern_binding_name ist b,intern_constr ist c) let intern_bindings ist = function | NoBindings -> NoBindings | ImplicitBindings l -> ImplicitBindings (List.map (intern_constr ist) l) | ExplicitBindings l -> ExplicitBindings (List.map (intern_binding ist) l) let intern_constr_with_bindings ist (c,bl) = (intern_constr ist c, intern_bindings ist bl) let intern_constr_with_bindings_arg ist (clear,c) = (clear,intern_constr_with_bindings ist c) let rec intern_intro_pattern lf ist = map (function | IntroNaming pat -> IntroNaming (intern_intro_pattern_naming lf ist pat) | IntroAction pat -> IntroAction (intern_intro_pattern_action lf ist pat) | IntroForthcoming _ as x -> x) and intern_intro_pattern_naming lf ist = function | IntroIdentifier id -> IntroIdentifier (intern_ident lf ist id) | IntroFresh id -> IntroFresh (intern_ident lf ist id) | IntroAnonymous as x -> x and intern_intro_pattern_action lf ist = function | IntroOrAndPattern l -> IntroOrAndPattern (intern_or_and_intro_pattern lf ist l) | IntroInjection l -> IntroInjection (List.map (intern_intro_pattern lf ist) l) | IntroWildcard | IntroRewrite _ as x -> x | IntroApplyOn ({loc;v=c},pat) -> IntroApplyOn (make ?loc @@ intern_constr ist c, intern_intro_pattern lf ist pat) and intern_or_and_intro_pattern lf ist = function | IntroAndPattern l -> IntroAndPattern (List.map (intern_intro_pattern lf ist) l) | IntroOrPattern ll -> IntroOrPattern (List.map (List.map (intern_intro_pattern lf ist)) ll) let intern_or_and_intro_pattern_loc lf ist = function | ArgVar {v=id} as x -> if find_var id ist then x else user_err Pp.(str "Disjunctive/conjunctive introduction pattern expected.") | ArgArg ll -> ArgArg (map (fun l -> intern_or_and_intro_pattern lf ist l) ll) let intern_intro_pattern_naming_loc lf ist = map (fun pat -> intern_intro_pattern_naming lf ist pat) (* TODO: catch ltac vars *) let intern_destruction_arg ist = function | clear,ElimOnConstr c -> clear,ElimOnConstr (intern_constr_with_bindings ist c) | clear,ElimOnAnonHyp n as x -> x | clear,ElimOnIdent {loc;v=id} -> if !strict_check then (* If in a defined tactic, no intros-until *) let c, p = intern_constr ist (make @@ CRef (qualid_of_ident id, None)) in match DAst.get c with | GVar id -> clear,ElimOnIdent (make ?loc:c.loc id) | _ -> clear,ElimOnConstr ((c, p), NoBindings) else clear,ElimOnIdent (make ?loc id) let short_name = function | {v=AN qid} when qualid_is_ident qid && not !strict_check -> Some (make ?loc:qid.CAst.loc @@ qualid_basename qid) | _ -> None let intern_evaluable_global_reference ist qid = try evaluable_of_global_reference ist.genv (locate_global_with_alias ~head:true qid) with Not_found -> if qualid_is_ident qid && not !strict_check then EvalVarRef (qualid_basename qid) else Nametab.error_global_not_found qid let intern_evaluable_reference_or_by_notation ist = function | {v=AN r} -> intern_evaluable_global_reference ist r | {v=ByNotation (ntn,sc);loc} -> evaluable_of_global_reference ist.genv (Notation.interp_notation_as_global_reference ?loc GlobRef.(function ConstRef _ | VarRef _ -> true | _ -> false) ntn sc) (* Globalize a reduction expression *) let intern_evaluable ist r = let f ist r = let e = intern_evaluable_reference_or_by_notation ist r in let na = short_name r in ArgArg (e,na) in match r with | {v=AN qid} when qualid_is_ident qid && find_var (qualid_basename qid) ist -> ArgVar (make ?loc:qid.CAst.loc @@ qualid_basename qid) | {v=AN qid} when qualid_is_ident qid && not !strict_check && find_hyp (qualid_basename qid) ist -> let id = qualid_basename qid in ArgArg (EvalVarRef id, Some (make ?loc:qid.CAst.loc id)) | _ -> f ist r let intern_unfold ist (l,qid) = (l,intern_evaluable ist qid) let intern_flag ist red = { red with rConst = List.map (intern_evaluable ist) red.rConst } let intern_constr_with_occurrences ist (l,c) = (l,intern_constr ist c) let intern_constr_pattern ist ~as_type ~ltacvars pc = let ltacvars = { Constrintern.ltac_vars = ltacvars; ltac_bound = Id.Set.empty; ltac_extra = ist.extra; } in let metas,pat = Constrintern.intern_constr_pattern ist.genv Evd.(from_env ist.genv) ~as_type ~ltacvars pc in let (glob,_ as c) = intern_constr_gen true false ist pc in let bound_names = Glob_ops.bound_glob_vars glob in metas,(bound_names,c,pat) let dummy_pat = PRel 0 let intern_typed_pattern ist ~as_type ~ltacvars p = (* we cannot ensure in non strict mode that the pattern is closed *) (* keeping a constr_expr copy is too complicated and we want anyway to *) (* type it, so we remember the pattern as a glob_constr only *) let metas,pat = if !strict_check then let ltacvars = { Constrintern.ltac_vars = ltacvars; ltac_bound = Id.Set.empty; ltac_extra = ist.extra; } in Constrintern.intern_constr_pattern ist.genv Evd.(from_env ist.genv) ~as_type ~ltacvars p else [], dummy_pat in let (glob,_ as c) = intern_constr_gen true false ist p in let bound_names = Glob_ops.bound_glob_vars glob in metas,(bound_names,c,pat) let intern_typed_pattern_or_ref_with_occurrences ist (l,p) = let interp_ref r = try Inl (intern_evaluable ist r) with e when Logic.catchable_exception e -> (* Compatibility. In practice, this means that the code above is useless. Still the idea of having either an evaluable ref or a pattern seems interesting, with "head" reduction in case of an evaluable ref, and "strong" reduction in the subterm matched when a pattern *) let r = match r with | {v=AN r} -> r | {loc} -> (qualid_of_path ?loc (Nametab.path_of_global (smart_global r))) in let sign = { Constrintern.ltac_vars = ist.ltacvars; ltac_bound = Id.Set.empty; ltac_extra = ist.extra; } in let c = Constrintern.interp_reference sign r in match DAst.get c with | GRef (r,None) -> Inl (ArgArg (evaluable_of_global_reference ist.genv r,None)) | GVar id -> let r = evaluable_of_global_reference ist.genv (GlobRef.VarRef id) in Inl (ArgArg (r,None)) | _ -> let bound_names = Glob_ops.bound_glob_vars c in Inr (bound_names,(c,None),dummy_pat) in (l, match p with | Inl r -> interp_ref r | Inr { v = CAppExpl((None,r,None),[]) } -> (* We interpret similarly @ref and ref *) interp_ref (make @@ AN r) | Inr c -> Inr (snd (intern_typed_pattern ist ~as_type:false ~ltacvars:ist.ltacvars c))) (* This seems fairly hacky, but it's the first way I've found to get proper globalization of [unfold]. --adamc *) let dump_glob_red_expr = function | Unfold occs -> List.iter (fun (_, r) -> try Dumpglob.add_glob ?loc:r.loc (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) occs | Cbv grf | Lazy grf -> List.iter (fun r -> try Dumpglob.add_glob ?loc:r.loc (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) grf.rConst | _ -> () let intern_red_expr ist = function | Unfold l -> Unfold (List.map (intern_unfold ist) l) | Fold l -> Fold (List.map (intern_constr ist) l) | Cbv f -> Cbv (intern_flag ist f) | Cbn f -> Cbn (intern_flag ist f) | Lazy f -> Lazy (intern_flag ist f) | Pattern l -> Pattern (List.map (intern_constr_with_occurrences ist) l) | Simpl (f,o) -> Simpl (intern_flag ist f, Option.map (intern_typed_pattern_or_ref_with_occurrences ist) o) | CbvVm o -> CbvVm (Option.map (intern_typed_pattern_or_ref_with_occurrences ist) o) | CbvNative o -> CbvNative (Option.map (intern_typed_pattern_or_ref_with_occurrences ist) o) | (Red _ | Hnf | ExtraRedExpr _ as r ) -> r let intern_in_hyp_as ist lf (id,ipat) = (intern_hyp ist id, Option.map (intern_intro_pattern lf ist) ipat) let intern_hyp_list ist = List.map (intern_hyp ist) let intern_inversion_strength lf ist = function | NonDepInversion (k,idl,ids) -> NonDepInversion (k,intern_hyp_list ist idl, Option.map (intern_or_and_intro_pattern_loc lf ist) ids) | DepInversion (k,copt,ids) -> DepInversion (k, Option.map (intern_constr ist) copt, Option.map (intern_or_and_intro_pattern_loc lf ist) ids) | InversionUsing (c,idl) -> InversionUsing (intern_constr ist c, intern_hyp_list ist idl) (* Interprets an hypothesis name *) let intern_hyp_location ist ((occs,id),hl) = ((Locusops.occurrences_map (List.map (intern_int_or_var ist)) occs, intern_hyp ist id), hl) (* Reads a pattern *) let intern_pattern ist ?(as_type=false) ltacvars = function | Subterm (ido,pc) -> let (metas,pc) = intern_constr_pattern ist ~as_type:false ~ltacvars pc in ido, metas, Subterm (ido,pc) | Term pc -> let (metas,pc) = intern_constr_pattern ist ~as_type ~ltacvars pc in None, metas, Term pc let intern_constr_may_eval ist = function | ConstrEval (r,c) -> ConstrEval (intern_red_expr ist r,intern_constr ist c) | ConstrContext (locid,c) -> ConstrContext (intern_hyp ist locid,intern_constr ist c) | ConstrTypeOf c -> ConstrTypeOf (intern_constr ist c) | ConstrTerm c -> ConstrTerm (intern_constr ist c) let name_cons accu = function | Anonymous -> accu | Name id -> Id.Set.add id accu let opt_cons accu = function | None -> accu | Some id -> Id.Set.add id accu (* Reads the hypotheses of a "match goal" rule *) let rec intern_match_goal_hyps ist ?(as_type=false) lfun = function | (Hyp ({v=na} as locna,mp))::tl -> let ido, metas1, pat = intern_pattern ist ~as_type:true lfun mp in let lfun, metas2, hyps = intern_match_goal_hyps ist lfun tl in let lfun' = name_cons (opt_cons lfun ido) na in lfun', metas1@metas2, Hyp (locna,pat)::hyps | (Def ({v=na} as locna,mv,mp))::tl -> let ido, metas1, patv = intern_pattern ist ~as_type:false lfun mv in let ido', metas2, patt = intern_pattern ist ~as_type:true lfun mp in let lfun, metas3, hyps = intern_match_goal_hyps ist ~as_type lfun tl in let lfun' = name_cons (opt_cons (opt_cons lfun ido) ido') na in lfun', metas1@metas2@metas3, Def (locna,patv,patt)::hyps | [] -> lfun, [], [] (* Utilities *) let extract_let_names lrc = let fold accu ({loc;v=name}, _) = Nameops.Name.fold_right (fun id accu -> if Id.Set.mem id accu then user_err ?loc ~hdr:"glob_tactic" (str "This variable is bound several times.") else Id.Set.add id accu) name accu in List.fold_left fold Id.Set.empty lrc let clause_app f = function { onhyps=None; concl_occs=nl } -> { onhyps=None; concl_occs=nl } | { onhyps=Some l; concl_occs=nl } -> { onhyps=Some(List.map f l); concl_occs=nl} (* Globalizes tactics : raw_tactic_expr -> glob_tactic_expr *) let rec intern_atomic lf ist x = match (x:raw_atomic_tactic_expr) with (* Basic tactics *) | TacIntroPattern (ev,l) -> TacIntroPattern (ev,List.map (intern_intro_pattern lf ist) l) | TacApply (a,ev,cb,inhyp) -> TacApply (a,ev,List.map (intern_constr_with_bindings_arg ist) cb, Option.map (intern_in_hyp_as ist lf) inhyp) | TacElim (ev,cb,cbo) -> TacElim (ev,intern_constr_with_bindings_arg ist cb, Option.map (intern_constr_with_bindings ist) cbo) | TacCase (ev,cb) -> TacCase (ev,intern_constr_with_bindings_arg ist cb) | TacMutualFix (id,n,l) -> let f (id,n,c) = (intern_ident lf ist id,n,intern_type ist c) in TacMutualFix (intern_ident lf ist id, n, List.map f l) | TacMutualCofix (id,l) -> let f (id,c) = (intern_ident lf ist id,intern_type ist c) in TacMutualCofix (intern_ident lf ist id, List.map f l) | TacAssert (ev,b,otac,ipat,c) -> TacAssert (ev,b,Option.map (Option.map (intern_pure_tactic ist)) otac, Option.map (intern_intro_pattern lf ist) ipat, intern_constr_gen false (not (Option.is_empty otac)) ist c) | TacGeneralize cl -> TacGeneralize (List.map (fun (c,na) -> intern_constr_with_occurrences ist c, intern_name lf ist na) cl) | TacLetTac (ev,na,c,cls,b,eqpat) -> let na = intern_name lf ist na in TacLetTac (ev,na,intern_constr ist c, (clause_app (intern_hyp_location ist) cls),b, (Option.map (intern_intro_pattern_naming_loc lf ist) eqpat)) (* Derived basic tactics *) | TacInductionDestruct (ev,isrec,(l,el)) -> TacInductionDestruct (ev,isrec,(List.map (fun (c,(ipato,ipats),cls) -> (intern_destruction_arg ist c, (Option.map (intern_intro_pattern_naming_loc lf ist) ipato, Option.map (intern_or_and_intro_pattern_loc lf ist) ipats), Option.map (clause_app (intern_hyp_location ist)) cls)) l, Option.map (intern_constr_with_bindings ist) el)) (* Conversion *) | TacReduce (r,cl) -> dump_glob_red_expr r; TacReduce (intern_red_expr ist r, clause_app (intern_hyp_location ist) cl) | TacChange (check,None,c,cl) -> let is_onhyps = match cl.onhyps with | None | Some [] -> true | _ -> false in let is_onconcl = match cl.concl_occs with | AtLeastOneOccurrence | AllOccurrences | NoOccurrences -> true | _ -> false in TacChange (check,None, (if is_onhyps && is_onconcl then intern_type ist c else intern_constr ist c), clause_app (intern_hyp_location ist) cl) | TacChange (check,Some p,c,cl) -> let { ltacvars } = ist in let metas,pat = intern_typed_pattern ist ~as_type:false ~ltacvars p in let fold accu x = Id.Set.add x accu in let ltacvars = List.fold_left fold ltacvars metas in let ist' = { ist with ltacvars } in TacChange (check,Some pat,intern_constr ist' c, clause_app (intern_hyp_location ist) cl) (* Equality and inversion *) | TacRewrite (ev,l,cl,by) -> TacRewrite (ev, List.map (fun (b,m,c) -> (b,m,intern_constr_with_bindings_arg ist c)) l, clause_app (intern_hyp_location ist) cl, Option.map (intern_pure_tactic ist) by) | TacInversion (inv,hyp) -> TacInversion (intern_inversion_strength lf ist inv, intern_quantified_hypothesis ist hyp) and intern_tactic onlytac ist tac = snd (intern_tactic_seq onlytac ist tac) and intern_tactic_seq onlytac ist = function | TacAtom { loc; v=t } -> let lf = ref ist.ltacvars in let t = intern_atomic lf ist t in !lf, TacAtom (CAst.make ?loc:(adjust_loc loc) t) | TacFun tacfun -> ist.ltacvars, TacFun (intern_tactic_fun ist tacfun) | TacLetIn (isrec,l,u) -> let ltacvars = Id.Set.union (extract_let_names l) ist.ltacvars in let ist' = { ist with ltacvars } in let l = List.map (fun (n,b) -> (n,intern_tacarg !strict_check false (if isrec then ist' else ist) b)) l in ist.ltacvars, TacLetIn (isrec,l,intern_tactic onlytac ist' u) | TacMatchGoal (lz,lr,lmr) -> ist.ltacvars, TacMatchGoal(lz,lr, intern_match_rule onlytac ist ~as_type:true lmr) | TacMatch (lz,c,lmr) -> ist.ltacvars, TacMatch (lz,intern_tactic_or_tacarg ist c,intern_match_rule onlytac ist lmr) | TacId l -> ist.ltacvars, TacId (intern_message ist l) | TacFail (g,n,l) -> ist.ltacvars, TacFail (g,intern_int_or_var ist n,intern_message ist l) | TacProgress tac -> ist.ltacvars, TacProgress (intern_pure_tactic ist tac) | TacShowHyps tac -> ist.ltacvars, TacShowHyps (intern_pure_tactic ist tac) | TacAbstract (tac,s) -> ist.ltacvars, TacAbstract (intern_pure_tactic ist tac,s) | TacThen (t1,t2) -> let lfun', t1 = intern_tactic_seq onlytac ist t1 in let lfun'', t2 = intern_tactic_seq onlytac { ist with ltacvars = lfun' } t2 in lfun'', TacThen (t1,t2) | TacDispatch tl -> ist.ltacvars , TacDispatch (List.map (intern_pure_tactic ist) tl) | TacExtendTac (tf,t,tl) -> ist.ltacvars , TacExtendTac (Array.map (intern_pure_tactic ist) tf, intern_pure_tactic ist t, Array.map (intern_pure_tactic ist) tl) | TacThens3parts (t1,tf,t2,tl) -> let lfun', t1 = intern_tactic_seq onlytac ist t1 in let ist' = { ist with ltacvars = lfun' } in (* Que faire en cas de (tac complexe avec Match et Thens; tac2) ?? *) lfun', TacThens3parts (t1,Array.map (intern_pure_tactic ist') tf,intern_pure_tactic ist' t2, Array.map (intern_pure_tactic ist') tl) | TacThens (t,tl) -> let lfun', t = intern_tactic_seq true ist t in let ist' = { ist with ltacvars = lfun' } in (* Que faire en cas de (tac complexe avec Match et Thens; tac2) ?? *) lfun', TacThens (t, List.map (intern_pure_tactic ist') tl) | TacDo (n,tac) -> ist.ltacvars, TacDo (intern_int_or_var ist n,intern_pure_tactic ist tac) | TacTry tac -> ist.ltacvars, TacTry (intern_pure_tactic ist tac) | TacInfo tac -> ist.ltacvars, TacInfo (intern_pure_tactic ist tac) | TacRepeat tac -> ist.ltacvars, TacRepeat (intern_pure_tactic ist tac) | TacTimeout (n,tac) -> ist.ltacvars, TacTimeout (intern_int_or_var ist n,intern_tactic onlytac ist tac) | TacTime (s,tac) -> ist.ltacvars, TacTime (s,intern_tactic onlytac ist tac) | TacOr (tac1,tac2) -> ist.ltacvars, TacOr (intern_pure_tactic ist tac1,intern_pure_tactic ist tac2) | TacOnce tac -> ist.ltacvars, TacOnce (intern_pure_tactic ist tac) | TacExactlyOnce tac -> ist.ltacvars, TacExactlyOnce (intern_pure_tactic ist tac) | TacIfThenCatch (tac,tact,tace) -> ist.ltacvars, TacIfThenCatch ( intern_pure_tactic ist tac, intern_pure_tactic ist tact, intern_pure_tactic ist tace) | TacOrelse (tac1,tac2) -> ist.ltacvars, TacOrelse (intern_pure_tactic ist tac1,intern_pure_tactic ist tac2) | TacFirst l -> ist.ltacvars, TacFirst (List.map (intern_pure_tactic ist) l) | TacSolve l -> ist.ltacvars, TacSolve (List.map (intern_pure_tactic ist) l) | TacComplete tac -> ist.ltacvars, TacComplete (intern_pure_tactic ist tac) | TacArg { loc; v=a } -> ist.ltacvars, intern_tactic_as_arg loc onlytac ist a | TacSelect (sel, tac) -> ist.ltacvars, TacSelect (sel, intern_pure_tactic ist tac) (* For extensions *) | TacAlias { loc; v=(s,l) } -> let alias = Tacenv.interp_alias s in Option.iter (fun o -> warn_deprecated_alias ?loc (s,o)) @@ alias.Tacenv.alias_deprecation; let l = List.map (intern_tacarg !strict_check false ist) l in ist.ltacvars, TacAlias (CAst.make ?loc (s,l)) | TacML { loc; v=(opn,l) } -> let _ignore = Tacenv.interp_ml_tactic opn in ist.ltacvars, TacML CAst.(make ?loc (opn,List.map (intern_tacarg !strict_check false ist) l)) and intern_tactic_as_arg loc onlytac ist a = match intern_tacarg !strict_check onlytac ist a with | TacCall _ | Reference _ | TacGeneric _ as a -> TacArg CAst.(make ?loc a) | Tacexp a -> a | ConstrMayEval _ | TacFreshId _ | TacPretype _ | TacNumgoals as a -> if onlytac then error_tactic_expected ?loc else TacArg CAst.(make ?loc a) and intern_tactic_or_tacarg ist = intern_tactic false ist and intern_pure_tactic ist = intern_tactic true ist and intern_tactic_fun ist (var,body) = let lfun = List.fold_left name_cons ist.ltacvars var in (var,intern_tactic_or_tacarg { ist with ltacvars = lfun } body) and intern_tacarg strict onlytac ist = function | Reference r -> intern_non_tactic_reference strict ist r | ConstrMayEval c -> ConstrMayEval (intern_constr_may_eval ist c) | TacCall { loc; v=(f,[]) } -> intern_isolated_tactic_reference strict ist f | TacCall { loc; v=(f,l) } -> TacCall (CAst.make ?loc ( intern_applied_tactic_reference ist f, List.map (intern_tacarg !strict_check false ist) l)) | TacFreshId x -> TacFreshId (List.map (intern_string_or_var ist) x) | TacPretype c -> TacPretype (intern_constr ist c) | TacNumgoals -> TacNumgoals | Tacexp t -> Tacexp (intern_tactic onlytac ist t) | TacGeneric arg -> let arg = intern_genarg ist arg in TacGeneric arg (* Reads the rules of a Match Context or a Match *) and intern_match_rule onlytac ist ?(as_type=false) = function | (All tc)::tl -> All (intern_tactic onlytac ist tc) :: (intern_match_rule onlytac ist ~as_type tl) | (Pat (rl,mp,tc))::tl -> let {ltacvars=lfun; genv=env} = ist in let lfun',metas1,hyps = intern_match_goal_hyps ist ~as_type lfun rl in let ido,metas2,pat = intern_pattern ist ~as_type lfun mp in let fold accu x = Id.Set.add x accu in let ltacvars = List.fold_left fold (opt_cons lfun' ido) metas1 in let ltacvars = List.fold_left fold ltacvars metas2 in let ist' = { ist with ltacvars } in Pat (hyps,pat,intern_tactic onlytac ist' tc) :: (intern_match_rule onlytac ist ~as_type tl) | [] -> [] and intern_genarg ist (GenArg (Rawwit wit, x)) = match wit with | ListArg wit -> let map x = let ans = intern_genarg ist (in_gen (rawwit wit) x) in out_gen (glbwit wit) ans in in_gen (glbwit (wit_list wit)) (List.map map x) | OptArg wit -> let ans = match x with | None -> in_gen (glbwit (wit_opt wit)) None | Some x -> let s = out_gen (glbwit wit) (intern_genarg ist (in_gen (rawwit wit) x)) in in_gen (glbwit (wit_opt wit)) (Some s) in ans | PairArg (wit1, wit2) -> let p, q = x in let p = out_gen (glbwit wit1) (intern_genarg ist (in_gen (rawwit wit1) p)) in let q = out_gen (glbwit wit2) (intern_genarg ist (in_gen (rawwit wit2) q)) in in_gen (glbwit (wit_pair wit1 wit2)) (p, q) | ExtraArg s -> snd (Genintern.generic_intern ist (in_gen (rawwit wit) x)) (** Other entry points *) let glob_tactic x = Flags.with_option strict_check (intern_pure_tactic (make_empty_glob_sign ())) x let glob_tactic_env l env x = let ltacvars = List.fold_left (fun accu x -> Id.Set.add x accu) Id.Set.empty l in Flags.with_option strict_check (intern_pure_tactic { (Genintern.empty_glob_sign env) with ltacvars }) x let split_ltac_fun = function | TacFun (l,t) -> (l,t) | t -> ([],t) let pr_ltac_fun_arg n = spc () ++ Name.print n let print_ltac id = try let kn = Tacenv.locate_tactic id in let entries = Tacenv.ltac_entries () in let tac = KNmap.find kn entries in let filter mp = try Some (Nametab.shortest_qualid_of_module mp) with Not_found -> None in let mods = List.map_filter filter tac.Tacenv.tac_redef in let redefined = match mods with | [] -> mt () | mods -> let redef = prlist_with_sep fnl pr_qualid mods in fnl () ++ str "Redefined by:" ++ fnl () ++ redef in let l,t = split_ltac_fun tac.Tacenv.tac_body in hv 2 ( hov 2 (str "Ltac" ++ spc() ++ pr_qualid id ++ prlist pr_ltac_fun_arg l ++ spc () ++ str ":=") ++ spc() ++ Pptactic.pr_glob_tactic (Global.env ()) t) ++ redefined with Not_found -> user_err ~hdr:"print_ltac" (pr_qualid id ++ spc() ++ str "is not a user defined tactic.") (** Registering *) let lift intern = (); fun ist x -> (ist, intern ist x) let () = let intern_intro_pattern ist pat = let lf = ref Id.Set.empty in let ans = intern_intro_pattern lf ist pat in let ist = { ist with ltacvars = !lf } in (ist, ans) in Genintern.register_intern0 wit_intropattern intern_intro_pattern [@warning "-3"]; Genintern.register_intern0 wit_simple_intropattern intern_intro_pattern let () = let intern_clause ist cl = let ans = clause_app (intern_hyp_location ist) cl in (ist, ans) in Genintern.register_intern0 wit_clause_dft_concl intern_clause let intern_ident' ist id = let lf = ref Id.Set.empty in (ist, intern_ident lf ist id) let intern_ltac ist tac = Flags.with_option strict_check (fun () -> intern_pure_tactic ist tac) () let () = Genintern.register_intern0 wit_int_or_var (lift intern_int_or_var); Genintern.register_intern0 wit_ref (lift intern_global_reference); Genintern.register_intern0 wit_pre_ident (fun ist c -> (ist,c)); Genintern.register_intern0 wit_ident intern_ident'; Genintern.register_intern0 wit_var (lift intern_hyp); Genintern.register_intern0 wit_tactic (lift intern_tactic_or_tacarg); Genintern.register_intern0 wit_ltac (lift intern_ltac); Genintern.register_intern0 wit_quant_hyp (lift intern_quantified_hypothesis); Genintern.register_intern0 wit_constr (fun ist c -> (ist,intern_constr ist c)); Genintern.register_intern0 wit_uconstr (fun ist c -> (ist,intern_constr ist c)); Genintern.register_intern0 wit_open_constr (fun ist c -> (ist,intern_constr ist c)); Genintern.register_intern0 wit_red_expr (lift intern_red_expr); Genintern.register_intern0 wit_bindings (lift intern_bindings); Genintern.register_intern0 wit_constr_with_bindings (lift intern_constr_with_bindings); Genintern.register_intern0 wit_destruction_arg (lift intern_destruction_arg); () (** Substitution for notations containing tactic-in-terms *) let notation_subst bindings tac = let fold id c accu = let loc = Glob_ops.loc_of_glob_constr (fst c) in let c = ConstrMayEval (ConstrTerm c) in (make ?loc @@ Name id, c) :: accu in let bindings = Id.Map.fold fold bindings [] in (* This is theoretically not correct due to potential variable capture, but Ltac has no true variables so one cannot simply substitute *) TacLetIn (false, bindings, tac) let () = Genintern.register_ntn_subst0 wit_tactic notation_subst coq-8.11.0/plugins/ltac/g_class.mlg0000644000175000017500000001041013612664533017002 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let gr = Smartlocate.global_with_alias r in let ev = Tacred.evaluable_of_global_reference (Global.env ()) gr in Classes.set_typeclass_transparency ev (Locality.make_section_locality None) b) cl } VERNAC COMMAND EXTEND Typeclasses_Unfold_Settings CLASSIFIED AS SIDEFF | [ "Typeclasses" "Transparent" reference_list(cl) ] -> { set_transparency cl true } END VERNAC COMMAND EXTEND Typeclasses_Rigid_Settings CLASSIFIED AS SIDEFF | [ "Typeclasses" "Opaque" reference_list(cl) ] -> { set_transparency cl false } END { let pr_debug _prc _prlc _prt b = if b then Pp.str "debug" else Pp.mt() } ARGUMENT EXTEND debug TYPED AS bool PRINTED BY { pr_debug } | [ "debug" ] -> { true } | [ ] -> { false } END { let pr_search_strategy _prc _prlc _prt = function | Some Dfs -> Pp.str "dfs" | Some Bfs -> Pp.str "bfs" | None -> Pp.mt () } ARGUMENT EXTEND eauto_search_strategy PRINTED BY { pr_search_strategy } | [ "(bfs)" ] -> { Some Bfs } | [ "(dfs)" ] -> { Some Dfs } | [ ] -> { None } END (* true = All transparent, false = Opaque if possible *) VERNAC COMMAND EXTEND Typeclasses_Settings CLASSIFIED AS SIDEFF | [ "Typeclasses" "eauto" ":=" debug(d) eauto_search_strategy(s) int_opt(depth) ] -> { set_typeclasses_debug d; Option.iter set_typeclasses_strategy s; set_typeclasses_depth depth } END (** Compatibility: typeclasses eauto has 8.5 and 8.6 modes *) TACTIC EXTEND typeclasses_eauto | [ "typeclasses" "eauto" "bfs" int_or_var_opt(d) "with" ne_preident_list(l) ] -> { typeclasses_eauto ~strategy:Bfs ~depth:d l } | [ "typeclasses" "eauto" int_or_var_opt(d) "with" ne_preident_list(l) ] -> { typeclasses_eauto ~depth:d l } | [ "typeclasses" "eauto" int_or_var_opt(d) ] -> { typeclasses_eauto ~only_classes:true ~depth:d [Class_tactics.typeclasses_db] } END TACTIC EXTEND head_of_constr | [ "head_of_constr" ident(h) constr(c) ] -> { head_of_constr h c } END TACTIC EXTEND not_evar | [ "not_evar" constr(ty) ] -> { not_evar ty } END TACTIC EXTEND is_ground | [ "is_ground" constr(ty) ] -> { is_ground ty } END { let deprecated_autoapply_using = CWarnings.create ~name:"autoapply-using" ~category:"deprecated" (fun () -> Pp.str "The syntax [autoapply ... using] is deprecated. Use [autoapply ... with] instead.") } TACTIC EXTEND autoapply | [ "autoapply" constr(c) "using" preident(i) ] -> { deprecated_autoapply_using (); autoapply c i } | [ "autoapply" constr(c) "with" preident(i) ] -> { autoapply c i } END { (** TODO: DEPRECATE *) (* A progress test that allows to see if the evars have changed *) open Constr open Proofview.Notations let rec eq_constr_mod_evars sigma x y = let open EConstr in match EConstr.kind sigma x, EConstr.kind sigma y with | Evar (e1, l1), Evar (e2, l2) when not (Evar.equal e1 e2) -> true | _, _ -> compare_constr sigma (fun x y -> eq_constr_mod_evars sigma x y) x y let progress_evars t = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let check = Proofview.Goal.enter begin fun gl' -> let sigma = Tacmach.New.project gl' in let newconcl = Proofview.Goal.concl gl' in if eq_constr_mod_evars sigma concl newconcl then Tacticals.New.tclFAIL 0 (Pp.str"No progress made (modulo evars)") else Proofview.tclUNIT () end in t <*> check end } TACTIC EXTEND progress_evars | [ "progress_evars" tactic(t) ] -> { progress_evars (Tacinterp.tactic_of_value ist t) } END coq-8.11.0/plugins/ltac/tactic_option.mli0000644000175000017500000000164213612664533020237 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* string -> (* put *) (locality_flag -> glob_tactic_expr -> unit) * (* get *) (unit -> locality_flag * unit Proofview.tactic) * (* print *) (unit -> Pp.t) coq-8.11.0/plugins/ltac/tauto.mli0000644000175000017500000000000013612664533016517 0ustar treinentreinencoq-8.11.0/plugins/ltac/g_obligations.mlg0000644000175000017500000001310013612664533020206 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* snd (get_default_tactic ()) end in Obligations.default_tactic := tac let with_tac f tac = let env = Genintern.empty_glob_sign (Global.env ()) in let tac = match tac with | None -> None | Some tac -> let tac = Genarg.in_gen (Genarg.rawwit wit_ltac) tac in let _, tac = Genintern.generic_intern env tac in Some tac in f tac (* We define new entries for programs, with the use of this module * Subtac. These entries are named Subtac. *) module Tactic = Pltac open Pcoq let sigref loc = mkRefC (Libnames.qualid_of_string ~loc "Coq.Init.Specif.sig") type 'a withtac_argtype = (Tacexpr.raw_tactic_expr option, 'a) Genarg.abstract_argument_type let wit_withtac : Tacexpr.raw_tactic_expr option Genarg.uniform_genarg_type = Genarg.create_arg "withtac" let withtac = Pcoq.create_generic_entry Pcoq.utactic "withtac" (Genarg.rawwit wit_withtac) } GRAMMAR EXTEND Gram GLOBAL: withtac; withtac: [ [ "with"; t = Tactic.tactic -> { Some t } | -> { None } ] ] ; Constr.closed_binder: [[ "("; id=Prim.name; ":"; t=Constr.lconstr; "|"; c=Constr.lconstr; ")" -> { let typ = mkAppC (sigref loc, [mkLambdaC ([id], default_binder_kind, t, c)]) in [CLocalAssum ([id], default_binder_kind, typ)] } ] ]; END { open Obligations let obligation obl tac = with_tac (fun t -> Obligations.obligation obl t) tac let next_obligation obl tac = with_tac (fun t -> Obligations.next_obligation obl t) tac let classify_obbl _ = Vernacextend.(VtStartProof (Doesn'tGuaranteeOpacity,[])) } VERNAC COMMAND EXTEND Obligations CLASSIFIED BY { classify_obbl } STATE open_proof | [ "Obligation" integer(num) "of" ident(name) ":" lglob(t) withtac(tac) ] -> { obligation (num, Some name, Some t) tac } | [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] -> { obligation (num, Some name, None) tac } | [ "Obligation" integer(num) ":" lglob(t) withtac(tac) ] -> { obligation (num, None, Some t) tac } | [ "Obligation" integer(num) withtac(tac) ] -> { obligation (num, None, None) tac } | [ "Next" "Obligation" "of" ident(name) withtac(tac) ] -> { next_obligation (Some name) tac } | [ "Next" "Obligation" withtac(tac) ] -> { next_obligation None tac } END VERNAC COMMAND EXTEND Solve_Obligation CLASSIFIED AS SIDEFF | [ "Solve" "Obligation" integer(num) "of" ident(name) "with" tactic(t) ] -> { try_solve_obligation num (Some name) (Some (Tacinterp.interp t)) } | [ "Solve" "Obligation" integer(num) "with" tactic(t) ] -> { try_solve_obligation num None (Some (Tacinterp.interp t)) } END VERNAC COMMAND EXTEND Solve_Obligations CLASSIFIED AS SIDEFF | [ "Solve" "Obligations" "of" ident(name) "with" tactic(t) ] -> { try_solve_obligations (Some name) (Some (Tacinterp.interp t)) } | [ "Solve" "Obligations" "with" tactic(t) ] -> { try_solve_obligations None (Some (Tacinterp.interp t)) } | [ "Solve" "Obligations" ] -> { try_solve_obligations None None } END VERNAC COMMAND EXTEND Solve_All_Obligations CLASSIFIED AS SIDEFF | [ "Solve" "All" "Obligations" "with" tactic(t) ] -> { solve_all_obligations (Some (Tacinterp.interp t)) } | [ "Solve" "All" "Obligations" ] -> { solve_all_obligations None } END VERNAC COMMAND EXTEND Admit_Obligations CLASSIFIED AS SIDEFF | [ "Admit" "Obligations" "of" ident(name) ] -> { admit_obligations (Some name) } | [ "Admit" "Obligations" ] -> { admit_obligations None } END VERNAC COMMAND EXTEND Set_Solver CLASSIFIED AS SIDEFF | #[ locality = Attributes.locality; ] [ "Obligation" "Tactic" ":=" tactic(t) ] -> { set_default_tactic (Locality.make_section_locality locality) (Tacintern.glob_tactic t); } END { open Pp } VERNAC COMMAND EXTEND Show_Solver CLASSIFIED AS QUERY | [ "Show" "Obligation" "Tactic" ] -> { Feedback.msg_notice (str"Program obligation tactic is " ++ print_default_tactic ()) } END VERNAC COMMAND EXTEND Show_Obligations CLASSIFIED AS QUERY | [ "Obligations" "of" ident(name) ] -> { show_obligations (Some name) } | [ "Obligations" ] -> { show_obligations None } END VERNAC COMMAND EXTEND Show_Preterm CLASSIFIED AS QUERY | [ "Preterm" "of" ident(name) ] -> { Feedback.msg_notice (show_term (Some name)) } | [ "Preterm" ] -> { Feedback.msg_notice (show_term None) } END { (* Declare a printer for the content of Program tactics *) let () = let printer env sigma _ _ _ = function | None -> mt () | Some tac -> str "with" ++ spc () ++ Pptactic.pr_raw_tactic env sigma tac in Pptactic.declare_extra_vernac_genarg_pprule wit_withtac printer } coq-8.11.0/plugins/ltac/tacentries.mli0000644000175000017500000001300213612664533017532 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ?deprecation:Deprecation.t -> Tacexpr.tacdef_body list -> unit (** Adds new Ltac definitions to the environment. *) (** {5 Tactic Notations} *) type 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr = | TacTerm of string | TacNonTerm of ('a * Names.Id.t option) Loc.located type raw_argument = string * string option (** An argument type as provided in Tactic notations, i.e. a string like "ne_foo_list_opt" together with a separator that only makes sense in the "_sep" cases. *) type argument = Genarg.ArgT.any Extend.user_symbol (** A fully resolved argument type given as an AST with generic arguments on the leaves. *) val add_tactic_notation : locality_flag -> int -> ?deprecation:Deprecation.t -> raw_argument grammar_tactic_prod_item_expr list -> raw_tactic_expr -> unit (** [add_tactic_notation local level prods expr] adds a tactic notation in the environment at level [level] with locality [local] made of the grammar productions [prods] and returning the body [expr] *) val register_tactic_notation_entry : string -> ('a, 'b, 'c) Genarg.genarg_type -> unit (** Register an argument under a given entry name for tactic notations. When translating [raw_argument] into [argument], atomic names will be first looked up according to names registered through this function and fallback to finding an argument by name (as in {!Genarg}) if there is none matching. *) val add_ml_tactic_notation : ml_tactic_name -> level:int -> ?deprecation:Deprecation.t -> argument grammar_tactic_prod_item_expr list list -> unit (** A low-level variant of {!add_tactic_notation} used by the TACTIC EXTEND ML-side macro. *) (** {5 Tactic Quotations} *) val create_ltac_quotation : string -> ('grm Loc.located -> raw_tactic_arg) -> ('grm Pcoq.Entry.t * int option) -> unit (** [create_ltac_quotation name f e] adds a quotation rule to Ltac, that is, Ltac grammar now accepts arguments of the form ["name" ":" "(" ")"], and generates an argument using [f] on the entry parsed by [e]. *) (** {5 Queries} *) val print_ltacs : unit -> unit (** Display the list of ltac definitions currently available. *) val print_located_tactic : Libnames.qualid -> unit (** Display the absolute name of a tactic. *) (** {5 TACTIC EXTEND} *) type _ ty_sig = | TyNil : (Geninterp.interp_sign -> unit Proofview.tactic) ty_sig | TyIdent : string * 'r ty_sig -> 'r ty_sig | TyArg : ('a, 'b, 'c) Extend.ty_user_symbol * 'r ty_sig -> ('c -> 'r) ty_sig type ty_ml = TyML : 'r ty_sig * 'r -> ty_ml val tactic_extend : string -> string -> level:Int.t -> ?deprecation:Deprecation.t -> ty_ml list -> unit (** {5 ARGUMENT EXTEND} *) (** This is the main entry point for the ARGUMENT EXTEND macro that allows to easily create user-made Ltac arguments. Each argument has three type parameters. See {!Genarg} for more details. There are two kinds of Ltac arguments, uniform and non-uniform. The former have the same type at each level (raw, glob, top) while the latter may vary. When declaring an argument one must provide the following data: - Internalization : raw -> glob - Substitution : glob -> glob - Interpretation : glob -> Ltac dynamic value - Printing for every level - An optional toplevel tag of type top (with the proviso that the interpretation function only produces values with this tag) This data can be either given explicitly with the [Fun] constructors, or it can be inherited from another argument with the [Wit] constructors. *) type ('a, 'b, 'c) argument_printer = 'a Pptactic.raw_extra_genarg_printer * 'b Pptactic.glob_extra_genarg_printer * 'c Pptactic.extra_genarg_printer type ('a, 'b) argument_intern = | ArgInternFun : ('a, 'b) Genintern.intern_fun -> ('a, 'b) argument_intern | ArgInternWit : ('a, 'b, 'c) Genarg.genarg_type -> ('a, 'b) argument_intern type 'b argument_subst = | ArgSubstFun : 'b Genintern.subst_fun -> 'b argument_subst | ArgSubstWit : ('a, 'b, 'c) Genarg.genarg_type -> 'b argument_subst type ('b, 'c) argument_interp = | ArgInterpRet : ('c, 'c) argument_interp | ArgInterpFun : ('b, Geninterp.Val.t) Geninterp.interp_fun -> ('b, 'c) argument_interp | ArgInterpWit : ('a, 'b, 'r) Genarg.genarg_type -> ('b, 'c) argument_interp | ArgInterpLegacy : (Geninterp.interp_sign -> Goal.goal Evd.sigma -> 'b -> Evd.evar_map * 'c) -> ('b, 'c) argument_interp type ('a, 'b, 'c) tactic_argument = { arg_parsing : 'a Vernacextend.argument_rule; arg_tag : 'c Geninterp.Val.tag option; arg_intern : ('a, 'b) argument_intern; arg_subst : 'b argument_subst; arg_interp : ('b, 'c) argument_interp; arg_printer : ('a, 'b, 'c) argument_printer; } val argument_extend : name:string -> ('a, 'b, 'c) tactic_argument -> ('a, 'b, 'c) Genarg.genarg_type * 'a Pcoq.Entry.t coq-8.11.0/plugins/ltac/g_eqdecide.mlg0000644000175000017500000000240413612664533017444 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { decideEqualityGoal } END TACTIC EXTEND compare | [ "compare" constr(c1) constr(c2) ] -> { compare c1 c2 } END coq-8.11.0/plugins/ltac/tactic_matching.mli0000644000175000017500000000452413612664533020523 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Evd.evar_map -> EConstr.constr -> (Constr_matching.binding_bound_vars * Pattern.constr_pattern, Tacexpr.glob_tactic_expr) Tacexpr.match_rule list -> Tacexpr.glob_tactic_expr t Proofview.tactic (** [match_goal env sigma hyps concl rules] matches the goal [hyps|-concl] with the set of matching rules [rules]. The environment [env] and the evar_map [sigma] are used to check convertibility for pattern variables shared between hypothesis patterns or the conclusion pattern. *) val match_goal: Environ.env -> Evd.evar_map -> EConstr.named_context -> EConstr.constr -> (Constr_matching.binding_bound_vars * Pattern.constr_pattern, Tacexpr.glob_tactic_expr) Tacexpr.match_rule list -> Tacexpr.glob_tactic_expr t Proofview.tactic coq-8.11.0/plugins/ltac/taccoerce.ml0000644000175000017500000003432413612664533017162 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t | _ -> CErrors.anomaly (Pp.str "Not a base val.") let (wit_constr_context : (Empty.t, Empty.t, EConstr.constr) Genarg.genarg_type) = let wit = Genarg.create_arg "constr_context" in let () = register_val0 wit None in let () = Genprint.register_val_print0 (base_val_typ wit) (Pptactic.make_constr_printer Printer.pr_econstr_n_env) in wit (* includes idents known to be bound and references *) let (wit_constr_under_binders : (Empty.t, Empty.t, Ltac_pretype.constr_under_binders) Genarg.genarg_type) = let wit = Genarg.create_arg "constr_under_binders" in let () = register_val0 wit None in let () = Genprint.register_val_print0 (base_val_typ wit) (fun c -> Genprint.TopPrinterNeedsContext (fun env sigma -> Printer.pr_constr_under_binders_env env sigma c)) in wit (** All the types considered here are base types *) let val_tag wit = match val_tag wit with | Val.Base t -> t | _ -> assert false let has_type : type a. Val.t -> a typed_abstract_argument_type -> bool = fun v wit -> let Val.Dyn (t, _) = v in match Val.eq t (val_tag wit) with | None -> false | Some Refl -> true let prj : type a. a Val.typ -> Val.t -> a option = fun t v -> let Val.Dyn (t', x) = v in match Val.eq t t' with | None -> None | Some Refl -> Some x let in_gen wit v = Val.Dyn (val_tag wit, v) let out_gen wit v = match prj (val_tag wit) v with None -> assert false | Some x -> x module Value = struct type t = Val.t let of_constr c = in_gen (topwit wit_constr) c let to_constr v = if has_type v (topwit wit_constr) then let c = out_gen (topwit wit_constr) v in Some c else if has_type v (topwit wit_constr_under_binders) then let vars, c = out_gen (topwit wit_constr_under_binders) v in match vars with [] -> Some c | _ -> None else None let of_uconstr c = in_gen (topwit wit_uconstr) c let to_uconstr v = if has_type v (topwit wit_uconstr) then Some (out_gen (topwit wit_uconstr) v) else None let of_int i = in_gen (topwit wit_int) i let to_int v = if has_type v (topwit wit_int) then Some (out_gen (topwit wit_int) v) else None let to_list v = prj Val.typ_list v let to_option v = prj Val.typ_opt v let to_pair v = prj Val.typ_pair v let cast_error wit v = let pr_v = Pptactic.pr_value Pptactic.ltop v in let Val.Dyn (tag, _) = v in let tag = Val.pr tag in CErrors.user_err (str "Type error: value " ++ pr_v ++ str " is a " ++ tag ++ str " while type " ++ Val.pr wit ++ str " was expected.") let unbox wit v ans = match ans with | None -> cast_error wit v | Some x -> x let rec prj : type a. a Val.tag -> Val.t -> a = fun tag v -> match tag with | Val.List tag -> List.map (fun v -> prj tag v) (unbox Val.typ_list v (to_list v)) | Val.Opt tag -> Option.map (fun v -> prj tag v) (unbox Val.typ_opt v (to_option v)) | Val.Pair (tag1, tag2) -> let (x, y) = unbox Val.typ_pair v (to_pair v) in (prj tag1 x, prj tag2 y) | Val.Base t -> let Val.Dyn (t', x) = v in match Val.eq t t' with | None -> cast_error t v | Some Refl -> x let rec tag_of_arg : type a b c. (a, b, c) genarg_type -> c Val.tag = fun wit -> match wit with | ExtraArg _ -> Geninterp.val_tag (topwit wit) | ListArg t -> Val.List (tag_of_arg t) | OptArg t -> Val.Opt (tag_of_arg t) | PairArg (t1, t2) -> Val.Pair (tag_of_arg t1, tag_of_arg t2) let val_cast arg v = prj (tag_of_arg arg) v let cast (Topwit wit) v = val_cast wit v end let is_variable env id = Id.List.mem id (Termops.ids_of_named_context (Environ.named_context env)) (* Transforms an id into a constr if possible, or fails with Not_found *) let constr_of_id env id = EConstr.mkVar (let _ = Environ.lookup_named id env in id) (* Gives the constr corresponding to a Constr_context tactic_arg *) let coerce_to_constr_context v = if has_type v (topwit wit_constr_context) then out_gen (topwit wit_constr_context) v else raise (CannotCoerceTo "a term context") let is_intro_pattern v = if has_type v (topwit wit_intro_pattern) then Some (out_gen (topwit wit_intro_pattern) v).CAst.v else None (* Interprets an identifier which must be fresh *) let coerce_var_to_ident fresh env sigma v = let fail () = raise (CannotCoerceTo "a fresh identifier") in match is_intro_pattern v with | Some (IntroNaming (IntroIdentifier id)) -> id | Some _ -> fail () | None -> if has_type v (topwit wit_intro_pattern) then match out_gen (topwit wit_intro_pattern) v with | { CAst.v=IntroNaming (IntroIdentifier id)} -> id | _ -> fail () else if has_type v (topwit wit_var) then out_gen (topwit wit_var) v else match Value.to_constr v with | None -> fail () | Some c -> (* We need it fresh for intro e.g. in "Tac H = clear H; intro H" *) if isVar sigma c && not (fresh && is_variable env (destVar sigma c)) then destVar sigma c else fail () (* Interprets, if possible, a constr to an identifier which may not be fresh but suitable to be given to the fresh tactic. Works for vars, constants, inductive, constructors and sorts. *) let coerce_to_ident_not_fresh sigma v = let id_of_name = function | Name.Anonymous -> Id.of_string "x" | Name.Name x -> x in let fail () = raise (CannotCoerceTo "an identifier") in match is_intro_pattern v with | Some (IntroNaming (IntroIdentifier id)) -> id | Some _ -> fail () | None -> if has_type v (topwit wit_var) then out_gen (topwit wit_var) v else match Value.to_constr v with | None -> fail () | Some c -> match EConstr.kind sigma c with | Var id -> id | Meta m -> id_of_name (Evd.meta_name sigma m) | Evar (kn,_) -> begin match Evd.evar_ident kn sigma with | None -> fail () | Some id -> id end | Const (cst,_) -> Label.to_id (Constant.label cst) | Construct (cstr,_) -> let ref = GlobRef.ConstructRef cstr in let basename = Nametab.basename_of_global ref in basename | Ind (ind,_) -> let ref = GlobRef.IndRef ind in let basename = Nametab.basename_of_global ref in basename | Sort s -> begin match ESorts.kind sigma s with | Sorts.SProp -> Label.to_id (Label.make "SProp") | Sorts.Prop -> Label.to_id (Label.make "Prop") | Sorts.Set -> Label.to_id (Label.make "Set") | Sorts.Type _ -> Label.to_id (Label.make "Type") end | _ -> fail() let coerce_to_intro_pattern sigma v = match is_intro_pattern v with | Some pat -> pat | None -> if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in IntroNaming (IntroIdentifier id) else match Value.to_constr v with | Some c when isVar sigma c -> (* This happens e.g. in definitions like "Tac H = clear H; intro H" *) (* but also in "destruct H as (H,H')" *) IntroNaming (IntroIdentifier (destVar sigma c)) | _ -> raise (CannotCoerceTo "an introduction pattern") let coerce_to_intro_pattern_naming sigma v = match coerce_to_intro_pattern sigma v with | IntroNaming pat -> pat | _ -> raise (CannotCoerceTo "a naming introduction pattern") let coerce_to_hint_base v = match is_intro_pattern v with | Some (IntroNaming (IntroIdentifier id)) -> Id.to_string id | Some _ | None -> raise (CannotCoerceTo "a hint base name") let coerce_to_int v = if has_type v (topwit wit_int) then out_gen (topwit wit_int) v else raise (CannotCoerceTo "an integer") let coerce_to_constr env v = let fail () = raise (CannotCoerceTo "a term") in match is_intro_pattern v with | Some (IntroNaming (IntroIdentifier id)) -> (try ([], constr_of_id env id) with Not_found -> fail ()) | Some _ -> fail () | None -> if has_type v (topwit wit_constr) then let c = out_gen (topwit wit_constr) v in ([], c) else if has_type v (topwit wit_constr_under_binders) then out_gen (topwit wit_constr_under_binders) v else if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in (try [], constr_of_id env id with Not_found -> fail ()) else fail () let coerce_to_uconstr v = if has_type v (topwit wit_uconstr) then out_gen (topwit wit_uconstr) v else raise (CannotCoerceTo "an untyped term") let coerce_to_closed_constr env v = let ids,c = coerce_to_constr env v in let () = if not (List.is_empty ids) then raise (CannotCoerceTo "a term") in c let coerce_to_evaluable_ref env sigma v = let fail () = raise (CannotCoerceTo "an evaluable reference") in let ev = match is_intro_pattern v with | Some (IntroNaming (IntroIdentifier id)) when is_variable env id -> EvalVarRef id | Some _ -> fail () | None -> if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in if Id.List.mem id (Termops.ids_of_context env) then EvalVarRef id else fail () else if has_type v (topwit wit_ref) then let open GlobRef in let r = out_gen (topwit wit_ref) v in match r with | VarRef var -> EvalVarRef var | ConstRef c -> EvalConstRef c | IndRef _ | ConstructRef _ -> fail () else match Value.to_constr v with | Some c when isConst sigma c -> EvalConstRef (fst (destConst sigma c)) | Some c when isVar sigma c -> EvalVarRef (destVar sigma c) | _ -> fail () in if Tacred.is_evaluable env ev then ev else fail () let coerce_to_constr_list env v = let v = Value.to_list v in match v with | Some l -> let map v = coerce_to_closed_constr env v in List.map map l | None -> raise (CannotCoerceTo "a term list") let coerce_to_intro_pattern_list ?loc sigma v = match Value.to_list v with | None -> raise (CannotCoerceTo "an intro pattern list") | Some l -> let map v = CAst.make ?loc @@ coerce_to_intro_pattern sigma v in List.map map l let coerce_to_hyp env sigma v = let fail () = raise (CannotCoerceTo "a variable") in match is_intro_pattern v with | Some (IntroNaming (IntroIdentifier id)) when is_variable env id -> id | Some _ -> fail () | None -> if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in if is_variable env id then id else fail () else match Value.to_constr v with | Some c when isVar sigma c -> destVar sigma c | _ -> fail () let coerce_to_hyp_list env sigma v = let v = Value.to_list v in match v with | Some l -> let map n = coerce_to_hyp env sigma n in List.map map l | None -> raise (CannotCoerceTo "a variable list") (* Interprets a qualified name *) let coerce_to_reference sigma v = match Value.to_constr v with | Some c -> begin try fst (Termops.global_of_constr sigma c) with Not_found -> raise (CannotCoerceTo "a reference") end | None -> raise (CannotCoerceTo "a reference") (* Quantified named or numbered hypothesis or hypothesis in context *) (* (as in Inversion) *) let coerce_to_quantified_hypothesis sigma v = match is_intro_pattern v with | Some (IntroNaming (IntroIdentifier id)) -> NamedHyp id | Some _ -> raise (CannotCoerceTo "a quantified hypothesis") | None -> if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in NamedHyp id else if has_type v (topwit wit_int) then AnonHyp (out_gen (topwit wit_int) v) else match Value.to_constr v with | Some c when isVar sigma c -> NamedHyp (destVar sigma c) | _ -> raise (CannotCoerceTo "a quantified hypothesis") (* Quantified named or numbered hypothesis or hypothesis in context *) (* (as in Inversion) *) let coerce_to_decl_or_quant_hyp sigma v = if has_type v (topwit wit_int) then AnonHyp (out_gen (topwit wit_int) v) else try coerce_to_quantified_hypothesis sigma v with CannotCoerceTo _ -> raise (CannotCoerceTo "a declared or quantified hypothesis") let coerce_to_int_or_var_list v = match Value.to_list v with | None -> raise (CannotCoerceTo "an int list") | Some l -> let map n = Locus.ArgArg (coerce_to_int n) in List.map map l (** Abstract application, to print ltac functions *) type appl = | UnnamedAppl (** For generic applications: nothing is printed *) | GlbAppl of (Names.KerName.t * Val.t list) list (** For calls to global constants, some may alias other. *) (* Values for interpretation *) type tacvalue = | VFun of appl*Tacexpr.ltac_trace * Val.t Id.Map.t * Name.t list * Tacexpr.glob_tactic_expr | VRec of Val.t Id.Map.t ref * Tacexpr.glob_tactic_expr let (wit_tacvalue : (Empty.t, tacvalue, tacvalue) Genarg.genarg_type) = let wit = Genarg.create_arg "tacvalue" in let () = register_val0 wit None in let () = Genprint.register_val_print0 (base_val_typ wit) (fun _ -> Genprint.TopPrinterBasic (fun () -> str "")) in wit let pr_argument_type arg = let Val.Dyn (tag, _) = arg in Val.pr tag (** TODO: unify printing of generic Ltac values in case of coercion failure. *) (* Displays a value *) let pr_value env v = let pr_with_env pr = match env with | Some (env,sigma) -> pr env sigma | None -> str "a value of type" ++ spc () ++ pr_argument_type v in let open Genprint in match generic_val_print v with | TopPrinterBasic pr -> pr () | TopPrinterNeedsContext pr -> pr_with_env pr | TopPrinterNeedsContextAndLevel { default_already_surrounded; printer } -> pr_with_env (fun env sigma -> printer env sigma default_already_surrounded) let error_ltac_variable ?loc id env v s = CErrors.user_err ?loc (str "Ltac variable " ++ Id.print id ++ strbrk " is bound to" ++ spc () ++ pr_value env v ++ spc () ++ strbrk "which cannot be coerced to " ++ str s ++ str".") coq-8.11.0/plugins/ltac/tacenv.mli0000644000175000017500000000706513612664533016665 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* full_path -> ltac_constant -> unit val locate_tactic : qualid -> ltac_constant val locate_extended_all_tactic : qualid -> ltac_constant list val exists_tactic : full_path -> bool val path_of_tactic : ltac_constant -> full_path val shortest_qualid_of_tactic : ltac_constant -> qualid (** {5 Tactic notations} *) type alias = KerName.t (** Type of tactic alias, used in the [TacAlias] node. *) type alias_tactic = { alias_args: Id.t list; alias_body: glob_tactic_expr; alias_deprecation: Deprecation.t option; } (** Contents of a tactic notation *) val register_alias : alias -> alias_tactic -> unit (** Register a tactic alias. *) val interp_alias : alias -> alias_tactic (** Recover the body of an alias. Raises an anomaly if it does not exist. *) val check_alias : alias -> bool (** Returns [true] if an alias is defined, false otherwise. *) (** {5 Coq tactic definitions} *) val register_ltac : bool -> bool -> ?deprecation:Deprecation.t -> Id.t -> glob_tactic_expr -> unit (** Register a new Ltac with the given name and body. The first boolean indicates whether this is done from ML side, rather than Coq side. If the second boolean flag is set to true, then this is a local definition. It also puts the Ltac name in the nametab, so that it can be used unqualified. *) val redefine_ltac : bool -> ?deprecation:Deprecation.t -> KerName.t -> glob_tactic_expr -> unit (** Replace a Ltac with the given name and body. If the boolean flag is set to true, then this is a local redefinition. *) val interp_ltac : KerName.t -> glob_tactic_expr (** Find a user-defined tactic by name. Raise [Not_found] if it is absent. *) val is_ltac_for_ml_tactic : KerName.t -> bool (** Whether the tactic is defined from ML-side *) val tac_deprecation : KerName.t -> Deprecation.t option (** The tactic deprecation notice, if any *) type ltac_entry = { tac_for_ml : bool; (** Whether the tactic is defined from ML-side *) tac_body : glob_tactic_expr; (** The current body of the tactic *) tac_redef : ModPath.t list; (** List of modules redefining the tactic in reverse chronological order *) tac_deprecation : Deprecation.t option; (** Deprecation notice to be printed when the tactic is used *) } val ltac_entries : unit -> ltac_entry KNmap.t (** Low-level access to all Ltac entries currently defined. *) (** {5 ML tactic extensions} *) type ml_tactic = Val.t list -> Geninterp.interp_sign -> unit Proofview.tactic (** Type of external tactics, used by [TacML]. *) val register_ml_tactic : ?overwrite:bool -> ml_tactic_name -> ml_tactic array -> unit (** Register an external tactic. *) val interp_ml_tactic : ml_tactic_entry -> ml_tactic (** Get the named tactic. Raises a user error if it does not exist. *) coq-8.11.0/plugins/ltac/tacexpr.mli0000644000175000017500000002550313612664533017050 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (** Possible arguments of a tactic definition *) type 'a gen_tactic_arg = | TacGeneric of 'lev generic_argument | ConstrMayEval of ('trm,'cst,'pat) may_eval | Reference of 'ref | TacCall of ('ref * 'a gen_tactic_arg list) CAst.t | TacFreshId of string or_var list | Tacexp of 'tacexpr | TacPretype of 'trm | TacNumgoals constraint 'a = < term:'trm; dterm: 'dtrm; pattern:'pat; constant:'cst; reference:'ref; name:'nam; tacexpr:'tacexpr; level:'lev > (** Generic ltac expressions. 't : terms, 'p : patterns, 'c : constants, 'i : inductive, 'r : ltac refs, 'n : idents, 'l : levels *) and 'a gen_tactic_expr = | TacAtom of ('a gen_atomic_tactic_expr) CAst.t | TacThen of 'a gen_tactic_expr * 'a gen_tactic_expr | TacDispatch of 'a gen_tactic_expr list | TacExtendTac of 'a gen_tactic_expr array * 'a gen_tactic_expr * 'a gen_tactic_expr array | TacThens of 'a gen_tactic_expr * 'a gen_tactic_expr list | TacThens3parts of 'a gen_tactic_expr * 'a gen_tactic_expr array * 'a gen_tactic_expr * 'a gen_tactic_expr array | TacFirst of 'a gen_tactic_expr list | TacComplete of 'a gen_tactic_expr | TacSolve of 'a gen_tactic_expr list | TacTry of 'a gen_tactic_expr | TacOr of 'a gen_tactic_expr * 'a gen_tactic_expr | TacOnce of 'a gen_tactic_expr | TacExactlyOnce of 'a gen_tactic_expr | TacIfThenCatch of 'a gen_tactic_expr * 'a gen_tactic_expr * 'a gen_tactic_expr | TacOrelse of 'a gen_tactic_expr * 'a gen_tactic_expr | TacDo of int or_var * 'a gen_tactic_expr | TacTimeout of int or_var * 'a gen_tactic_expr | TacTime of string option * 'a gen_tactic_expr | TacRepeat of 'a gen_tactic_expr | TacProgress of 'a gen_tactic_expr | TacShowHyps of 'a gen_tactic_expr | TacAbstract of 'a gen_tactic_expr * Id.t option | TacId of 'n message_token list | TacFail of global_flag * int or_var * 'n message_token list | TacInfo of 'a gen_tactic_expr | TacLetIn of rec_flag * (lname * 'a gen_tactic_arg) list * 'a gen_tactic_expr | TacMatch of lazy_flag * 'a gen_tactic_expr * ('p,'a gen_tactic_expr) match_rule list | TacMatchGoal of lazy_flag * direction_flag * ('p,'a gen_tactic_expr) match_rule list | TacFun of 'a gen_tactic_fun_ast | TacArg of 'a gen_tactic_arg CAst.t | TacSelect of Goal_select.t * 'a gen_tactic_expr (* For ML extensions *) | TacML of (ml_tactic_entry * 'a gen_tactic_arg list) CAst.t (* For syntax extensions *) | TacAlias of (KerName.t * 'a gen_tactic_arg list) CAst.t constraint 'a = < term:'t; dterm: 'dtrm; pattern:'p; constant:'c; reference:'r; name:'n; tacexpr:'tacexpr; level:'l > and 'a gen_tactic_fun_ast = Name.t list * 'a gen_tactic_expr constraint 'a = < term:'t; dterm: 'dtrm; pattern:'p; constant:'c; reference:'r; name:'n; tacexpr:'te; level:'l > (** Globalized tactics *) type g_trm = Genintern.glob_constr_and_expr type g_pat = Genintern.glob_constr_pattern_and_expr type g_cst = evaluable_global_reference Genredexpr.and_short_name or_var type g_ref = ltac_constant located or_var type g_nam = lident type g_dispatch = < term:g_trm; dterm:g_trm; pattern:g_pat; constant:g_cst; reference:g_ref; name:g_nam; tacexpr:glob_tactic_expr; level:glevel > and glob_tactic_expr = g_dispatch gen_tactic_expr type glob_atomic_tactic_expr = g_dispatch gen_atomic_tactic_expr type glob_tactic_arg = g_dispatch gen_tactic_arg (** Raw tactics *) type r_ref = qualid type r_nam = lident type r_lev = rlevel type r_dispatch = < term:r_trm; dterm:r_trm; pattern:r_pat; constant:r_cst; reference:r_ref; name:r_nam; tacexpr:raw_tactic_expr; level:rlevel > and raw_tactic_expr = r_dispatch gen_tactic_expr type raw_atomic_tactic_expr = r_dispatch gen_atomic_tactic_expr type raw_tactic_arg = r_dispatch gen_tactic_arg (** Interpreted tactics *) type t_trm = EConstr.constr type t_pat = constr_pattern type t_cst = evaluable_global_reference type t_ref = ltac_constant located type t_nam = Id.t type t_dispatch = < term:t_trm; dterm:g_trm; pattern:t_pat; constant:t_cst; reference:t_ref; name:t_nam; tacexpr:unit; level:tlevel > type atomic_tactic_expr = t_dispatch gen_atomic_tactic_expr (** Misc *) type raw_red_expr = (r_trm, r_cst, r_pat) red_expr_gen type glob_red_expr = (g_trm, g_cst, g_pat) red_expr_gen (** Traces *) type ltac_call_kind = | LtacMLCall of glob_tactic_expr | LtacNotationCall of KerName.t | LtacNameCall of ltac_constant | LtacAtomCall of glob_atomic_tactic_expr | LtacVarCall of Id.t * glob_tactic_expr | LtacConstrInterp of Glob_term.glob_constr * Ltac_pretype.ltac_var_map type ltac_trace = ltac_call_kind Loc.located list type tacdef_body = | TacticDefinition of lident * raw_tactic_expr (* indicates that user employed ':=' in Ltac body *) | TacticRedefinition of qualid * raw_tactic_expr (* indicates that user employed '::=' in Ltac body *) coq-8.11.0/plugins/ltac/pptactic.ml0000644000175000017500000015413613612664533017045 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Evd.evar_map -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> raw_tactic_expr -> Pp.t) -> 'a -> Pp.t type 'a glob_extra_genarg_printer = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> glob_tactic_expr -> Pp.t) -> 'a -> Pp.t type 'a extra_genarg_printer = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> Val.t -> Pp.t) -> 'a -> Pp.t type 'a raw_extra_genarg_printer_with_level = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> constr_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> raw_tactic_expr -> Pp.t) -> tolerability -> 'a -> Pp.t type 'a glob_extra_genarg_printer_with_level = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> glob_constr_and_expr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> glob_tactic_expr -> Pp.t) -> tolerability -> 'a -> Pp.t type 'a extra_genarg_printer_with_level = Environ.env -> Evd.evar_map -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.t) -> (Environ.env -> Evd.evar_map -> tolerability -> Val.t -> Pp.t) -> tolerability -> 'a -> Pp.t let string_of_genarg_arg (ArgumentType arg) = let rec aux : type a b c. (a, b, c) genarg_type -> string = function | ListArg t -> aux t ^ "_list" | OptArg t -> aux t ^ "_opt" | PairArg (t1, t2) -> assert false (* No parsing/printing rule for it *) | ExtraArg s -> ArgT.repr s in aux arg let keyword x = tag_keyword (str x) let primitive x = tag_primitive (str x) let has_type (Val.Dyn (tag, _)) t = match Val.eq tag t with | None -> false | Some _ -> true let unbox : type a. Val.t -> a Val.typ -> a= fun (Val.Dyn (tag, x)) t -> match Val.eq tag t with | None -> assert false | Some Refl -> x let rec pr_value lev v : Pp.t = if has_type v Val.typ_list then pr_sequence (fun x -> pr_value lev x) (unbox v Val.typ_list) else if has_type v Val.typ_opt then pr_opt_no_spc (fun x -> pr_value lev x) (unbox v Val.typ_opt) else if has_type v Val.typ_pair then let (v1, v2) = unbox v Val.typ_pair in str "(" ++ pr_value lev v1 ++ str ", " ++ pr_value lev v2 ++ str ")" else let Val.Dyn (tag, x) = v in let name = Val.repr tag in let default = str "<" ++ str name ++ str ">" in match ArgT.name name with | None -> default | Some (ArgT.Any arg) -> let wit = ExtraArg arg in match val_tag (Topwit wit) with | Val.Base t -> begin match Val.eq t tag with | None -> default | Some Refl -> let open Genprint in match generic_top_print (in_gen (Topwit wit) x) with | TopPrinterBasic pr -> pr () | TopPrinterNeedsContext pr -> let env = Global.env() in pr env (Evd.from_env env) | TopPrinterNeedsContextAndLevel { default_ensure_surrounded; printer } -> let env = Global.env() in printer env (Evd.from_env env) default_ensure_surrounded end | _ -> default let pr_with_occurrences pr c = Ppred.pr_with_occurrences pr keyword c let pr_red_expr env sigma pr c = Ppred.pr_red_expr_env env sigma pr keyword c let pr_may_eval env sigma test prc prlc pr2 pr3 = function | ConstrEval (r,c) -> hov 0 (keyword "eval" ++ brk (1,1) ++ pr_red_expr env sigma (prc,prlc,pr2,pr3) r ++ spc () ++ keyword "in" ++ spc() ++ prc env sigma c) | ConstrContext ({CAst.v=id},c) -> hov 0 (keyword "context" ++ spc () ++ pr_id id ++ spc () ++ str "[ " ++ prlc env sigma c ++ str " ]") | ConstrTypeOf c -> hov 1 (keyword "type of" ++ spc() ++ prc env sigma c) | ConstrTerm c when test c -> h 0 (str "(" ++ prc env sigma c ++ str ")") | ConstrTerm c -> prc env sigma c let pr_may_eval env sigma a = pr_may_eval env sigma (fun _ -> false) a let pr_arg pr x = spc () ++ pr x let pr_and_short_name pr (c,_) = pr c let pr_evaluable_reference = function | EvalVarRef id -> pr_id id | EvalConstRef sp -> pr_global (GlobRef.ConstRef sp) let pr_quantified_hypothesis = function | AnonHyp n -> int n | NamedHyp id -> pr_id id let pr_clear_flag clear_flag pp x = match clear_flag with | Some false -> surround (pp x) | Some true -> str ">" ++ pp x | None -> pp x let pr_with_bindings prc prlc (c,bl) = prc c ++ Miscprint.pr_bindings prc prlc bl let pr_with_bindings_arg prc prlc (clear_flag,c) = pr_clear_flag clear_flag (pr_with_bindings prc prlc) c let pr_with_constr prc = function | None -> mt () | Some c -> spc () ++ hov 1 (keyword "with" ++ spc () ++ prc c) let pr_message_token prid = function | MsgString s -> tag_string (qs s) | MsgInt n -> int n | MsgIdent id -> prid id let pr_fresh_ids = prlist (fun s -> spc() ++ pr_or_var (fun s -> tag_string (qs s)) s) let with_evars ev s = if ev then "e" ^ s else s let rec tacarg_using_rule_token pr_gen = function | [] -> [] | TacTerm s :: l -> keyword s :: tacarg_using_rule_token pr_gen l | TacNonTerm (_, ((symb, arg), _)) :: l -> pr_gen symb arg :: tacarg_using_rule_token pr_gen l let pr_tacarg_using_rule pr_gen l = let l = match l with | TacTerm s :: l -> (* First terminal token should be considered as the name of the tactic, so we tag it differently than the other terminal tokens. *) primitive s :: tacarg_using_rule_token pr_gen l | _ -> tacarg_using_rule_token pr_gen l in pr_sequence (fun x -> x) l let pr_extend_gen pr_gen _ { mltac_name = s; mltac_index = i } l = let name = str s.mltac_plugin ++ str "::" ++ str s.mltac_tactic ++ str "@" ++ int i in let args = match l with | [] -> mt () | _ -> spc() ++ pr_sequence pr_gen l in str "<" ++ name ++ str ">" ++ args let rec pr_user_symbol = function | Extend.Ulist1 tkn -> "ne_" ^ pr_user_symbol tkn ^ "_list" | Extend.Ulist1sep (tkn, _) -> "ne_" ^ pr_user_symbol tkn ^ "_list" | Extend.Ulist0 tkn -> pr_user_symbol tkn ^ "_list" | Extend.Ulist0sep (tkn, _) -> pr_user_symbol tkn ^ "_list" | Extend.Uopt tkn -> pr_user_symbol tkn ^ "_opt" | Extend.Uentry tag -> let ArgT.Any tag = tag in ArgT.repr tag | Extend.Uentryl (_, lvl) -> "tactic" ^ string_of_int lvl let pr_alias_key key = try let prods = (KNmap.find key !prnotation_tab).pptac_prods in let pr = function | TacTerm s -> primitive s | TacNonTerm (_, (symb, _)) -> str (Printf.sprintf "(%s)" (pr_user_symbol symb)) in pr_sequence pr prods with Not_found -> (* FIXME: This key, moreover printed with a low-level printer, has no meaning user-side *) KerName.print key let pr_alias_gen pr_gen lev key l = try let pp = KNmap.find key !prnotation_tab in let rec pack prods args = match prods, args with | [], [] -> [] | TacTerm s :: prods, args -> TacTerm s :: pack prods args | TacNonTerm (_, (_, None)) :: prods, args -> pack prods args | TacNonTerm (loc, (symb, (Some _ as ido))) :: prods, arg :: args -> TacNonTerm (loc, ((symb, arg), ido)) :: pack prods args | _ -> raise Not_found in let prods = pack pp.pptac_prods l in let p = pr_tacarg_using_rule pr_gen prods in if pp.pptac_level > lev then surround p else p with Not_found -> let pr _ = str "_" in KerName.print key ++ spc() ++ pr_sequence pr l ++ str" (* Generic printer *)" let pr_farg prtac arg = prtac (1, Any) (TacArg (CAst.make arg)) let is_genarg tag wit = let ArgT.Any tag = tag in argument_type_eq (ArgumentType (ExtraArg tag)) wit let get_list : type l. l generic_argument -> l generic_argument list option = function (GenArg (wit, arg)) -> match wit with | Rawwit (ListArg wit) -> Some (List.map (in_gen (rawwit wit)) arg) | Glbwit (ListArg wit) -> Some (List.map (in_gen (glbwit wit)) arg) | _ -> None let get_opt : type l. l generic_argument -> l generic_argument option option = function (GenArg (wit, arg)) -> match wit with | Rawwit (OptArg wit) -> Some (Option.map (in_gen (rawwit wit)) arg) | Glbwit (OptArg wit) -> Some (Option.map (in_gen (glbwit wit)) arg) | _ -> None let rec pr_any_arg : type l. (_ -> l generic_argument -> Pp.t) -> _ -> l generic_argument -> Pp.t = fun prtac symb arg -> match symb with | Extend.Uentry tag when is_genarg tag (genarg_tag arg) -> prtac (1, Any) arg | Extend.Ulist1 s | Extend.Ulist0 s -> begin match get_list arg with | None -> str "ltac:(" ++ prtac (1, Any) arg ++ str ")" | Some l -> pr_sequence (pr_any_arg prtac s) l end | Extend.Ulist1sep (s, sep) | Extend.Ulist0sep (s, sep) -> begin match get_list arg with | None -> str "ltac:(" ++ prtac (1, Any) arg ++ str ")" | Some l -> prlist_with_sep (fun () -> str sep) (pr_any_arg prtac s) l end | Extend.Uopt s -> begin match get_opt arg with | None -> str "ltac:(" ++ prtac (1, Any) arg ++ str ")" | Some l -> pr_opt (pr_any_arg prtac s) l end | Extend.Uentry _ | Extend.Uentryl _ -> str "ltac:(" ++ prtac (1, Any) arg ++ str ")" let pr_targ prtac symb arg = match symb with | Extend.Uentry tag when is_genarg tag (ArgumentType wit_tactic) -> prtac (1, Any) arg | Extend.Uentryl (_, l) -> prtac (l, Any) arg | _ -> match arg with | TacGeneric arg -> let pr l arg = prtac l (TacGeneric arg) in pr_any_arg pr symb arg | _ -> str "ltac:(" ++ prtac (1, Any) arg ++ str ")" let pr_raw_extend_rec prtac = pr_extend_gen (pr_farg prtac) let pr_glob_extend_rec prtac = pr_extend_gen (pr_farg prtac) let pr_raw_alias prtac lev key args = pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (CAst.make a)))) lev key args let pr_glob_alias prtac lev key args = pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (CAst.make a)))) lev key args (**********************************************************************) (* The tactic printer *) let strip_prod_binders_expr n ty = let rec strip_ty acc n ty = match ty.CAst.v with Constrexpr.CProdN(bll,a) -> let bll = List.map (function | CLocalAssum (nal,_,t) -> nal,t | _ -> user_err Pp.(str "Cannot translate fix tactic: not only products")) bll in let nb = List.fold_left (fun i (nal,t) -> i + List.length nal) 0 bll in if nb >= n then (List.rev (bll@acc)), a else strip_ty (bll@acc) (n-nb) a | _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in strip_ty [] n ty let pr_ltac_or_var pr = function | ArgArg x -> pr x | ArgVar {CAst.loc;v=id} -> pr_with_comments ?loc (pr_id id) let pr_ltac_constant kn = if !Flags.in_debugger then KerName.print kn else try pr_qualid (Tacenv.shortest_qualid_of_tactic kn) with Not_found -> (* local tactic not accessible anymore *) str "<" ++ KerName.print kn ++ str ">" let pr_evaluable_reference_env env = function | EvalVarRef id -> pr_id id | EvalConstRef sp -> Nametab.pr_global_env (Termops.vars_of_env env) (GlobRef.ConstRef sp) let pr_as_disjunctive_ipat prc ipatl = keyword "as" ++ spc () ++ pr_or_var (fun {CAst.loc;v=p} -> Miscprint.pr_or_and_intro_pattern prc p) ipatl let pr_eqn_ipat {CAst.v=ipat} = keyword "eqn:" ++ Miscprint.pr_intro_pattern_naming ipat let pr_with_induction_names prc = function | None, None -> mt () | Some eqpat, None -> hov 1 (pr_eqn_ipat eqpat) | None, Some ipat -> hov 1 (pr_as_disjunctive_ipat prc ipat) | Some eqpat, Some ipat -> hov 1 (pr_as_disjunctive_ipat prc ipat ++ spc () ++ pr_eqn_ipat eqpat) let pr_as_intro_pattern prc ipat = spc () ++ hov 1 (keyword "as" ++ spc () ++ Miscprint.pr_intro_pattern prc ipat) let pr_with_inversion_names prc = function | None -> mt () | Some ipat -> pr_as_disjunctive_ipat prc ipat let pr_as_ipat prc = function | None -> mt () | Some ipat -> pr_as_intro_pattern prc ipat let pr_as_name = function | Anonymous -> mt () | Name id -> spc () ++ keyword "as" ++ spc () ++ pr_lident (CAst.make id) let pr_pose_as_style prc na c = spc() ++ prc c ++ pr_as_name na let pr_pose prc prlc na c = match na with | Anonymous -> spc() ++ prc c | Name id -> spc() ++ surround (pr_id id ++ str " :=" ++ spc() ++ prlc c) let pr_assertion prc prdc _prlc ipat c = match ipat with (* Use this "optimisation" or use only the general case ? | IntroIdentifier id -> spc() ++ surround (pr_intro_pattern ipat ++ str " :" ++ spc() ++ prlc c) *) | ipat -> spc() ++ prc c ++ pr_as_ipat prdc ipat let pr_assumption prc prdc prlc ipat c = match ipat with (* Use this "optimisation" or use only the general case ?*) (* it seems that this "optimisation" is somehow more natural *) | Some {CAst.v=IntroNaming (IntroIdentifier id)} -> spc() ++ surround (pr_id id ++ str " :" ++ spc() ++ prlc c) | ipat -> spc() ++ prc c ++ pr_as_ipat prdc ipat let pr_by_tactic prt = function | Some tac -> keyword "by" ++ spc () ++ prt tac | None -> mt() let pr_hyp_location pr_id = function | occs, InHyp -> pr_with_occurrences pr_id occs | occs, InHypTypeOnly -> pr_with_occurrences (fun id -> str "(" ++ keyword "type of" ++ spc () ++ pr_id id ++ str ")" ) occs | occs, InHypValueOnly -> pr_with_occurrences (fun id -> str "(" ++ keyword "value of" ++ spc () ++ pr_id id ++ str ")" ) occs let pr_in pp = hov 0 (keyword "in" ++ pp) let pr_simple_hyp_clause pr_id = function | [] -> mt () | l -> pr_in (spc () ++ prlist_with_sep spc pr_id l) let pr_in_hyp_as prc pr_id = function | None -> mt () | Some (id,ipat) -> pr_in (spc () ++ pr_id id) ++ pr_as_ipat prc ipat let pr_in_clause pr_id = function | { onhyps=None; concl_occs=NoOccurrences } -> (str "* |-") | { onhyps=None; concl_occs=occs } -> (pr_with_occurrences (fun () -> str "*") (occs,())) | { onhyps=Some l; concl_occs=NoOccurrences } -> prlist_with_sep (fun () -> str ", ") (pr_hyp_location pr_id) l | { onhyps=Some l; concl_occs=occs } -> let pr_occs = pr_with_occurrences (fun () -> str" |- *") (occs,()) in (prlist_with_sep (fun () -> str", ") (pr_hyp_location pr_id) l ++ pr_occs) (* Some true = default is concl; Some false = default is all; None = no default *) let pr_clauses has_default pr_id = function | { onhyps=Some []; concl_occs=occs } when (match has_default with Some true -> true | _ -> false) -> pr_with_occurrences mt (occs,()) | { onhyps=None; concl_occs=AllOccurrences } when (match has_default with Some false -> true | _ -> false) -> mt () | { onhyps=None; concl_occs=NoOccurrences } -> pr_in (str " * |-") | { onhyps=None; concl_occs=occs } -> pr_in (pr_with_occurrences (fun () -> str " *") (occs,())) | { onhyps=Some l; concl_occs=occs } -> let pr_occs = match occs with | NoOccurrences -> mt () | _ -> pr_with_occurrences (fun () -> str" |- *") (occs,()) in pr_in (prlist_with_sep (fun () -> str",") (fun id -> spc () ++ pr_hyp_location pr_id id) l ++ pr_occs) let pr_orient b = if b then mt () else str "<- " let pr_multi = let open Equality in function | Precisely 1 -> mt () | Precisely n -> int n ++ str "!" | UpTo n -> int n ++ str "?" | RepeatStar -> str "?" | RepeatPlus -> str "!" let pr_core_destruction_arg prc prlc = function | ElimOnConstr c -> pr_with_bindings prc prlc c | ElimOnIdent {CAst.loc;v=id} -> pr_with_comments ?loc (pr_id id) | ElimOnAnonHyp n -> int n let pr_destruction_arg prc prlc (clear_flag,h) = pr_clear_flag clear_flag (pr_core_destruction_arg prc prlc) h let pr_inversion_kind = let open Inv in function | SimpleInversion -> primitive "simple inversion" | FullInversion -> primitive "inversion" | FullInversionClear -> primitive "inversion_clear" let pr_range_selector (i, j) = if Int.equal i j then int i else int i ++ str "-" ++ int j let pr_goal_selector toplevel = let open Goal_select in function | SelectAlreadyFocused -> str "!:" | SelectNth i -> int i ++ str ":" | SelectList l -> prlist_with_sep (fun () -> str ", ") pr_range_selector l ++ str ":" | SelectId id -> str "[" ++ Id.print id ++ str "]:" | SelectAll -> assert toplevel; str "all:" let pr_goal_selector ~toplevel s = (if toplevel then mt () else str "only ") ++ pr_goal_selector toplevel s let pr_lazy = function | General -> keyword "multi" | Select -> keyword "lazy" | Once -> mt () let pr_match_pattern pr_pat = function | Term a -> pr_pat a | Subterm (None,a) -> keyword "context" ++ str" [ " ++ pr_pat a ++ str " ]" | Subterm (Some id,a) -> keyword "context" ++ spc () ++ pr_id id ++ str "[ " ++ pr_pat a ++ str " ]" let pr_match_hyps pr_pat = function | Hyp (nal,mp) -> pr_lname nal ++ str ":" ++ pr_match_pattern pr_pat mp | Def (nal,mv,mp) -> pr_lname nal ++ str ":=" ++ pr_match_pattern pr_pat mv ++ str ":" ++ pr_match_pattern pr_pat mp let pr_match_rule m pr pr_pat = function | Pat ([],mp,t) when m -> pr_match_pattern pr_pat mp ++ spc () ++ str "=>" ++ brk (1,4) ++ pr t (* | Pat (rl,mp,t) -> hv 0 (prlist_with_sep pr_comma (pr_match_hyps pr_pat) rl ++ (if rl <> [] then spc () else mt ()) ++ hov 0 (str "|-" ++ spc () ++ pr_match_pattern pr_pat mp ++ spc () ++ str "=>" ++ brk (1,4) ++ pr t)) *) | Pat (rl,mp,t) -> hov 0 ( hv 0 (prlist_with_sep pr_comma (pr_match_hyps pr_pat) rl) ++ (if not (List.is_empty rl) then spc () else mt ()) ++ hov 0 ( str "|-" ++ spc () ++ pr_match_pattern pr_pat mp ++ spc () ++ str "=>" ++ brk (1,4) ++ pr t)) | All t -> str "_" ++ spc () ++ str "=>" ++ brk (1,4) ++ pr t let pr_funvar n = spc () ++ Name.print n let pr_let_clause k pr_gen pr_arg (na,(bl,t)) = let pr = function | TacGeneric arg -> let name = string_of_genarg_arg (genarg_tag arg) in if name = "unit" || name = "int" then (* Hard-wired parsing rules *) pr_gen arg else str name ++ str ":" ++ surround (pr_gen arg) | _ -> pr_arg (TacArg (CAst.make t)) in hov 0 (keyword k ++ spc () ++ pr_lname na ++ prlist pr_funvar bl ++ str " :=" ++ brk (1,1) ++ pr t) let pr_let_clauses recflag pr_gen pr = function | hd::tl -> hv 0 (pr_let_clause (if recflag then "let rec" else "let") pr_gen pr hd ++ prlist (fun t -> spc () ++ pr_let_clause "with" pr_gen pr t) tl) | [] -> anomaly (Pp.str "LetIn must declare at least one binding.") let pr_seq_body pr tl = hv 0 (str "[ " ++ prlist_with_sep (fun () -> spc () ++ str "| ") pr tl ++ str " ]") let pr_dispatch pr tl = hv 0 (str "[>" ++ prlist_with_sep (fun () -> spc () ++ str "| ") pr tl ++ str " ]") let pr_opt_tactic pr = function | TacId [] -> mt () | t -> pr t let pr_tac_extend_gen pr tf tm tl = prvect_with_sep mt (fun t -> pr t ++ spc () ++ str "| ") tf ++ pr_opt_tactic pr tm ++ str ".." ++ prvect_with_sep mt (fun t -> spc () ++ str "| " ++ pr t) tl let pr_then_gen pr tf tm tl = hv 0 (str "[ " ++ pr_tac_extend_gen pr tf tm tl ++ str " ]") let pr_tac_extend pr tf tm tl = hv 0 (str "[>" ++ pr_tac_extend_gen pr tf tm tl ++ str " ]") let pr_hintbases = function | None -> keyword "with" ++ str" *" | Some [] -> mt () | Some l -> hov 2 (keyword "with" ++ prlist (fun s -> spc () ++ str s) l) let pr_auto_using prc = function | [] -> mt () | l -> hov 2 (keyword "using" ++ spc () ++ prlist_with_sep pr_comma prc l) let pr_then () = str ";" let ltop = (5,E) let lseq = 4 let ltactical = 3 let lorelse = 2 let llet = 5 let lfun = 5 let lcomplete = 1 let labstract = 3 let lmatch = 1 let latom = 0 let lcall = 1 let leval = 1 let ltatom = 1 let linfo = 5 let level_of (n,p) = match p with E -> n | L -> n-1 | Prec n -> n | Any -> lseq (** A printer for tactics that polymorphically works on the three "raw", "glob" and "typed" levels *) type 'a printer = { pr_tactic : tolerability -> 'tacexpr -> Pp.t; pr_constr : Environ.env -> Evd.evar_map -> 'trm -> Pp.t; pr_lconstr : Environ.env -> Evd.evar_map -> 'trm -> Pp.t; pr_dconstr : Environ.env -> Evd.evar_map -> 'dtrm -> Pp.t; pr_pattern : Environ.env -> Evd.evar_map -> 'pat -> Pp.t; pr_lpattern : Environ.env -> Evd.evar_map -> 'pat -> Pp.t; pr_constant : 'cst -> Pp.t; pr_reference : 'ref -> Pp.t; pr_name : 'nam -> Pp.t; pr_generic : Environ.env -> Evd.evar_map -> 'lev generic_argument -> Pp.t; pr_extend : int -> ml_tactic_entry -> 'a gen_tactic_arg list -> Pp.t; pr_alias : int -> KerName.t -> 'a gen_tactic_arg list -> Pp.t; } constraint 'a = < term :'trm; dterm :'dtrm; pattern :'pat; constant :'cst; reference :'ref; name :'nam; tacexpr :'tacexpr; level :'lev > let pr_atom env sigma pr strip_prod_binders tag_atom = let pr_with_bindings = pr_with_bindings (pr.pr_constr env sigma) (pr.pr_lconstr env sigma) in let pr_with_bindings_arg_full = pr_with_bindings_arg in let pr_with_bindings_arg = pr_with_bindings_arg (pr.pr_constr env sigma) (pr.pr_lconstr env sigma) in let pr_red_expr = pr_red_expr env sigma (pr.pr_constr,pr.pr_lconstr,pr.pr_constant,pr.pr_pattern) in let _pr_constrarg c = spc () ++ pr.pr_constr env sigma c in let pr_lconstrarg c = spc () ++ pr.pr_lconstr env sigma c in let pr_intarg n = spc () ++ int n in (* Some printing combinators *) let pr_eliminator cb = keyword "using" ++ pr_arg pr_with_bindings cb in let pr_binder_fix (nal,t) = (* match t with | CHole _ -> spc() ++ prlist_with_sep spc (pr_lname) nal | _ ->*) let s = prlist_with_sep spc pr_lname nal ++ str ":" ++ pr.pr_lconstr env sigma t in spc() ++ hov 1 (str"(" ++ s ++ str")") in let pr_fix_tac (id,n,c) = let rec set_nth_name avoid n = function (nal,ty)::bll -> if n <= List.length nal then match List.chop (n-1) nal with _, {CAst.v=Name id} :: _ -> id, (nal,ty)::bll | bef, {CAst.loc;v=Anonymous} :: aft -> let id = next_ident_away (Id.of_string"y") avoid in id, ((bef@(CAst.make ?loc @@ Name id)::aft, ty)::bll) | _ -> assert false else let (id,bll') = set_nth_name avoid (n-List.length nal) bll in (id,(nal,ty)::bll') | [] -> assert false in let (bll,ty) = strip_prod_binders n c in let names = List.fold_left (fun ln (nal,_) -> List.fold_left (fun ln na -> match na with { CAst.v=Name id } -> Id.Set.add id ln | _ -> ln) ln nal) Id.Set.empty bll in let idarg,bll = set_nth_name names n bll in let annot = if Int.equal (Id.Set.cardinal names) 1 then mt () else spc() ++ str"{" ++ keyword "struct" ++ spc () ++ pr_id idarg ++ str"}" in hov 1 (str"(" ++ pr_id id ++ prlist pr_binder_fix bll ++ annot ++ str" :" ++ (pr_lconstrarg ty) ++ str")") in (* spc() ++ hov 0 (pr_id id ++ pr_intarg n ++ str":" ++ _pr_constrarg c) *) let pr_cofix_tac (id,c) = hov 1 (str"(" ++ pr_id id ++ str" :" ++ pr_lconstrarg c ++ str")") in (* Printing tactics as arguments *) let rec pr_atom0 a = tag_atom a (match a with | TacIntroPattern (false,[]) -> primitive "intros" | TacIntroPattern (true,[]) -> primitive "eintros" | t -> str "(" ++ pr_atom1 t ++ str ")" ) (* Main tactic printer *) and pr_atom1 a = tag_atom a (match a with (* Basic tactics *) | TacIntroPattern (_,[]) as t -> pr_atom0 t | TacIntroPattern (ev,(_::_ as p)) -> hov 1 (primitive (if ev then "eintros" else "intros") ++ (match p with | [{CAst.v=IntroForthcoming false}] -> mt () | _ -> spc () ++ prlist_with_sep spc (Miscprint.pr_intro_pattern @@ pr.pr_dconstr env sigma) p)) | TacApply (a,ev,cb,inhyp) -> hov 1 ( (if a then mt() else primitive "simple ") ++ primitive (with_evars ev "apply") ++ spc () ++ prlist_with_sep pr_comma pr_with_bindings_arg cb ++ pr_non_empty_arg (pr_in_hyp_as (pr.pr_dconstr env sigma) pr.pr_name) inhyp ) | TacElim (ev,cb,cbo) -> hov 1 ( primitive (with_evars ev "elim") ++ pr_arg pr_with_bindings_arg cb ++ pr_opt pr_eliminator cbo) | TacCase (ev,cb) -> hov 1 (primitive (with_evars ev "case") ++ spc () ++ pr_with_bindings_arg cb) | TacMutualFix (id,n,l) -> hov 1 ( primitive "fix" ++ spc () ++ pr_id id ++ pr_intarg n ++ spc() ++ keyword "with" ++ spc () ++ prlist_with_sep spc pr_fix_tac l) | TacMutualCofix (id,l) -> hov 1 ( primitive "cofix" ++ spc () ++ pr_id id ++ spc() ++ keyword "with" ++ spc () ++ prlist_with_sep spc pr_cofix_tac l ) | TacAssert (ev,b,Some tac,ipat,c) -> hov 1 ( primitive (if b then if ev then "eassert" else "assert" else if ev then "eenough" else "enough") ++ pr_assumption (pr.pr_constr env sigma) (pr.pr_dconstr env sigma) (pr.pr_lconstr env sigma) ipat c ++ pr_non_empty_arg (pr_by_tactic (pr.pr_tactic (ltactical,E))) tac ) | TacAssert (ev,_,None,ipat,c) -> hov 1 ( primitive (if ev then "epose proof" else "pose proof") ++ pr_assertion (pr.pr_constr env sigma) (pr.pr_dconstr env sigma) (pr.pr_lconstr env sigma) ipat c ) | TacGeneralize l -> hov 1 ( primitive "generalize" ++ spc () ++ prlist_with_sep pr_comma (fun (cl,na) -> pr_with_occurrences (pr.pr_constr env sigma) cl ++ pr_as_name na) l ) | TacLetTac (ev,na,c,cl,true,_) when Locusops.is_nowhere cl -> hov 1 (primitive (if ev then "epose" else "pose") ++ pr_pose (pr.pr_constr env sigma) (pr.pr_lconstr env sigma) na c) | TacLetTac (ev,na,c,cl,b,e) -> hov 1 ( primitive (if b then if ev then "eset" else "set" else if ev then "eremember" else "remember") ++ (if b then pr_pose (pr.pr_constr env sigma) (pr.pr_lconstr env sigma) na c else pr_pose_as_style (pr.pr_constr env sigma) na c) ++ pr_opt (fun p -> pr_eqn_ipat p ++ spc ()) e ++ pr_non_empty_arg (pr_clauses (Some b) pr.pr_name) cl) (* | TacInstantiate (n,c,ConclLocation ()) -> hov 1 (str "instantiate" ++ spc() ++ hov 1 (str"(" ++ pr_arg int n ++ str" :=" ++ pr_lconstrarg c ++ str ")" )) | TacInstantiate (n,c,HypLocation (id,hloc)) -> hov 1 (str "instantiate" ++ spc() ++ hov 1 (str"(" ++ pr_arg int n ++ str" :=" ++ pr_lconstrarg c ++ str ")" ) ++ str "in" ++ pr_hyp_location pr.pr_name (id,[],(hloc,ref None))) *) (* Derived basic tactics *) | TacInductionDestruct (isrec,ev,(l,el)) -> hov 1 ( primitive (with_evars ev (if isrec then "induction" else "destruct")) ++ spc () ++ prlist_with_sep pr_comma (fun (h,ids,cl) -> pr_destruction_arg (pr.pr_dconstr env sigma) (pr.pr_dconstr env sigma) h ++ pr_non_empty_arg (pr_with_induction_names (pr.pr_dconstr env sigma)) ids ++ pr_opt (pr_clauses None pr.pr_name) cl) l ++ pr_opt pr_eliminator el ) (* Conversion *) | TacReduce (r,h) -> hov 1 ( pr_red_expr r ++ pr_non_empty_arg (pr_clauses (Some true) pr.pr_name) h ) | TacChange (check,op,c,h) -> let name = if check then "change_no_check" else "change" in hov 1 ( primitive name ++ brk (1,1) ++ ( match op with None -> mt () | Some p -> pr.pr_pattern env sigma p ++ spc () ++ keyword "with" ++ spc () ) ++ pr.pr_dconstr env sigma c ++ pr_non_empty_arg (pr_clauses (Some true) pr.pr_name) h ) (* Equality and inversion *) | TacRewrite (ev,l,cl,tac) -> hov 1 ( primitive (with_evars ev "rewrite") ++ spc () ++ prlist_with_sep (fun () -> str ","++spc()) (fun (b,m,c) -> pr_orient b ++ pr_multi m ++ pr_with_bindings_arg_full (pr.pr_dconstr env sigma) (pr.pr_dconstr env sigma) c) l ++ pr_non_empty_arg (pr_clauses (Some true) pr.pr_name) cl ++ pr_non_empty_arg (pr_by_tactic (pr.pr_tactic (ltactical,E))) tac ) | TacInversion (DepInversion (k,c,ids),hyp) -> hov 1 ( primitive "dependent " ++ pr_inversion_kind k ++ spc () ++ pr_quantified_hypothesis hyp ++ pr_with_inversion_names (pr.pr_dconstr env sigma) ids ++ pr_with_constr (pr.pr_constr env sigma) c ) | TacInversion (NonDepInversion (k,cl,ids),hyp) -> hov 1 ( pr_inversion_kind k ++ spc () ++ pr_quantified_hypothesis hyp ++ pr_non_empty_arg (pr_with_inversion_names @@ pr.pr_dconstr env sigma) ids ++ pr_non_empty_arg (pr_simple_hyp_clause pr.pr_name) cl ) | TacInversion (InversionUsing (c,cl),hyp) -> hov 1 ( primitive "inversion" ++ spc() ++ pr_quantified_hypothesis hyp ++ spc () ++ keyword "using" ++ spc () ++ pr.pr_constr env sigma c ++ pr_non_empty_arg (pr_simple_hyp_clause pr.pr_name) cl ) ) in pr_atom1 let make_pr_tac env sigma pr strip_prod_binders tag_atom tag = let extract_binders = function | Tacexp (TacFun (lvar,body)) -> (lvar,Tacexp body) | body -> ([],body) in let rec pr_tac inherited tac = let return (doc, l) = (tag tac doc, l) in let (strm, prec) = return (match tac with | TacAbstract (t,None) -> keyword "abstract " ++ pr_tac (labstract,L) t, labstract | TacAbstract (t,Some s) -> hov 0 ( keyword "abstract" ++ str" (" ++ pr_tac (labstract,L) t ++ str")" ++ spc () ++ keyword "using" ++ spc () ++ pr_id s), labstract | TacLetIn (recflag,llc,u) -> let llc = List.map (fun (id,t) -> (id,extract_binders t)) llc in v 0 (hv 0 ( pr_let_clauses recflag (pr.pr_generic env sigma) (pr_tac ltop) llc ++ spc () ++ keyword "in" ) ++ fnl () ++ pr_tac (llet,E) u), llet | TacMatch (lz,t,lrul) -> hov 0 ( pr_lazy lz ++ keyword "match" ++ spc () ++ pr_tac ltop t ++ spc () ++ keyword "with" ++ prlist (fun r -> fnl () ++ str "| " ++ pr_match_rule true (pr_tac ltop) (pr.pr_lpattern env sigma) r ) lrul ++ fnl() ++ keyword "end"), lmatch | TacMatchGoal (lz,lr,lrul) -> hov 0 ( pr_lazy lz ++ keyword (if lr then "match reverse goal with" else "match goal with") ++ prlist (fun r -> fnl () ++ str "| " ++ pr_match_rule false (pr_tac ltop) (pr.pr_lpattern env sigma) r ) lrul ++ fnl() ++ keyword "end"), lmatch | TacFun (lvar,body) -> hov 2 ( keyword "fun" ++ prlist pr_funvar lvar ++ str " =>" ++ spc () ++ pr_tac (lfun,E) body), lfun | TacThens (t,tl) -> hov 1 ( pr_tac (lseq,E) t ++ pr_then () ++ spc () ++ pr_seq_body (pr_opt_tactic (pr_tac ltop)) tl), lseq | TacThen (t1,t2) -> hov 1 ( pr_tac (lseq,E) t1 ++ pr_then () ++ spc () ++ pr_tac (lseq,L) t2), lseq | TacDispatch tl -> pr_dispatch (pr_tac ltop) tl, lseq | TacExtendTac (tf,t,tr) -> pr_tac_extend (pr_tac ltop) tf t tr , lseq | TacThens3parts (t1,tf,t2,tl) -> hov 1 ( pr_tac (lseq,E) t1 ++ pr_then () ++ spc () ++ pr_then_gen (pr_tac ltop) tf t2 tl), lseq | TacTry t -> hov 1 ( keyword "try" ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacDo (n,t) -> hov 1 ( str "do" ++ spc () ++ pr_or_var int n ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacTimeout (n,t) -> hov 1 ( keyword "timeout " ++ pr_or_var int n ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacTime (s,t) -> hov 1 ( keyword "time" ++ pr_opt qstring s ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacRepeat t -> hov 1 ( keyword "repeat" ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacProgress t -> hov 1 ( keyword "progress" ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacShowHyps t -> hov 1 ( keyword "infoH" ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacInfo t -> hov 1 ( keyword "info" ++ spc () ++ pr_tac (ltactical,E) t), linfo | TacOr (t1,t2) -> hov 1 ( pr_tac (lorelse,L) t1 ++ spc () ++ str "+" ++ brk (1,1) ++ pr_tac (lorelse,E) t2), lorelse | TacOnce t -> hov 1 ( keyword "once" ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacExactlyOnce t -> hov 1 ( keyword "exactly_once" ++ spc () ++ pr_tac (ltactical,E) t), ltactical | TacIfThenCatch (t,tt,te) -> hov 1 ( str"tryif" ++ spc() ++ pr_tac (ltactical,E) t ++ brk(1,1) ++ str"then" ++ spc() ++ pr_tac (ltactical,E) tt ++ brk(1,1) ++ str"else" ++ spc() ++ pr_tac (ltactical,E) te ++ brk(1,1)), ltactical | TacOrelse (t1,t2) -> hov 1 ( pr_tac (lorelse,L) t1 ++ spc () ++ str "||" ++ brk (1,1) ++ pr_tac (lorelse,E) t2), lorelse | TacFail (g,n,l) -> let arg = match n with | ArgArg 0 -> mt () | _ -> pr_arg (pr_or_var int) n in let name = match g with | TacGlobal -> keyword "gfail" | TacLocal -> keyword "fail" in hov 1 ( name ++ arg ++ prlist (pr_arg (pr_message_token pr.pr_name)) l), latom | TacFirst tl -> keyword "first" ++ spc () ++ pr_seq_body (pr_tac ltop) tl, llet | TacSolve tl -> keyword "solve" ++ spc () ++ pr_seq_body (pr_tac ltop) tl, llet | TacComplete t -> pr_tac (lcomplete,E) t, lcomplete | TacSelect (s, tac) -> pr_goal_selector ~toplevel:false s ++ spc () ++ pr_tac ltop tac, latom | TacId l -> keyword "idtac" ++ prlist (pr_arg (pr_message_token pr.pr_name)) l, latom | TacAtom { CAst.loc; v=t } -> pr_with_comments ?loc (hov 1 (pr_atom env sigma pr strip_prod_binders tag_atom t)), ltatom | TacArg { CAst.v=Tacexp e } -> pr_tac inherited e, latom | TacArg { CAst.v=ConstrMayEval (ConstrTerm c) } -> keyword "constr:" ++ pr.pr_constr env sigma c, latom | TacArg { CAst.v=ConstrMayEval c } -> pr_may_eval env sigma pr.pr_constr pr.pr_lconstr pr.pr_constant pr.pr_pattern c, leval | TacArg { CAst.v=TacFreshId l } -> primitive "fresh" ++ pr_fresh_ids l, latom | TacArg { CAst.v=TacGeneric arg } -> pr.pr_generic env sigma arg, latom | TacArg { CAst.v=TacCall {CAst.v=(f,[])} } -> pr.pr_reference f, latom | TacArg { CAst.v=TacCall {CAst.loc; v=(f,l)} } -> pr_with_comments ?loc (hov 1 ( pr.pr_reference f ++ spc () ++ prlist_with_sep spc pr_tacarg l)), lcall | TacArg { CAst.v=a } -> pr_tacarg a, latom | TacML { CAst.loc; v=(s,l) } -> pr_with_comments ?loc (pr.pr_extend 1 s l), lcall | TacAlias { CAst.loc; v=(kn,l) } -> pr_with_comments ?loc (pr.pr_alias (level_of inherited) kn l), latom ) in if prec_less prec inherited then strm else str"(" ++ strm ++ str")" and pr_tacarg = function | Reference r -> pr.pr_reference r | ConstrMayEval c -> pr_may_eval env sigma pr.pr_constr pr.pr_lconstr pr.pr_constant pr.pr_pattern c | TacFreshId l -> keyword "fresh" ++ pr_fresh_ids l | TacPretype c -> keyword "type_term" ++ pr.pr_constr env sigma c | TacNumgoals -> keyword "numgoals" | (TacCall _|Tacexp _ | TacGeneric _) as a -> hov 0 (keyword "ltac:" ++ surround (pr_tac ltop (TacArg (CAst.make a)))) in pr_tac let strip_prod_binders_glob_constr n (ty,_) = let rec strip_ty acc n ty = if Int.equal n 0 then (List.rev acc, (ty,None)) else match DAst.get ty with Glob_term.GProd(na,Glob_term.Explicit,a,b) -> strip_ty (([CAst.make na],(a,None))::acc) (n-1) b | _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in strip_ty [] n ty let raw_printers = (strip_prod_binders_expr) let rec pr_raw_tactic_level env sigma n (t:raw_tactic_expr) = let pr = { pr_tactic = pr_raw_tactic_level env sigma; pr_constr = pr_constr_expr; pr_dconstr = pr_constr_expr; pr_lconstr = pr_lconstr_expr; pr_pattern = pr_constr_pattern_expr; pr_lpattern = pr_lconstr_pattern_expr; pr_constant = pr_or_by_notation pr_qualid; pr_reference = pr_qualid; pr_name = pr_lident; pr_generic = Pputils.pr_raw_generic; pr_extend = pr_raw_extend_rec @@ pr_raw_tactic_level env sigma; pr_alias = pr_raw_alias @@ pr_raw_tactic_level env sigma; } in make_pr_tac env sigma pr raw_printers tag_raw_atomic_tactic_expr tag_raw_tactic_expr n t let pr_raw_tactic env sigma = pr_raw_tactic_level env sigma ltop let pr_and_constr_expr pr (c,_) = pr c let pr_pat_and_constr_expr pr (_,(c,_),_) = pr c let pr_glob_tactic_level env n t = let glob_printers = (strip_prod_binders_glob_constr) in let rec prtac n (t:glob_tactic_expr) = let pr = { pr_tactic = prtac; pr_constr = (fun env sigma -> pr_and_constr_expr (pr_glob_constr_env env)); pr_dconstr = (fun env sigma -> pr_and_constr_expr (pr_glob_constr_env env)); pr_lconstr = (fun env sigma -> pr_and_constr_expr (pr_lglob_constr_env env)); pr_pattern = (fun env sigma -> pr_pat_and_constr_expr (pr_glob_constr_env env)); pr_lpattern = (fun env sigma -> pr_pat_and_constr_expr (pr_lglob_constr_env env)); pr_constant = pr_or_var (pr_and_short_name (pr_evaluable_reference_env env)); pr_reference = pr_ltac_or_var (pr_located pr_ltac_constant); pr_name = pr_lident; pr_generic = Pputils.pr_glb_generic; pr_extend = pr_glob_extend_rec prtac; pr_alias = pr_glob_alias prtac; } in make_pr_tac env (Evd.from_env env) pr glob_printers tag_glob_atomic_tactic_expr tag_glob_tactic_expr n t in prtac n t let pr_glob_tactic env = pr_glob_tactic_level env ltop let strip_prod_binders_constr n ty = let ty = EConstr.Unsafe.to_constr ty in let rec strip_ty acc n ty = if n=0 then (List.rev acc, EConstr.of_constr ty) else match Constr.kind ty with | Constr.Prod(na,a,b) -> strip_ty (([CAst.make na.Context.binder_name],EConstr.of_constr a)::acc) (n-1) b | _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in strip_ty [] n ty let pr_atomic_tactic_level env sigma t = let prtac (t:atomic_tactic_expr) = let pr = { pr_tactic = (fun _ _ -> str ""); pr_constr = pr_econstr_env; pr_dconstr = (fun env sigma -> pr_and_constr_expr (pr_glob_constr_env env)); pr_lconstr = pr_leconstr_env; pr_pattern = pr_constr_pattern_env; pr_lpattern = pr_lconstr_pattern_env; pr_constant = pr_evaluable_reference_env env; pr_reference = pr_located pr_ltac_constant; pr_name = pr_id; (* Those are not used by the atomic printer *) pr_generic = (fun _ -> assert false); pr_extend = (fun _ _ _ -> assert false); pr_alias = (fun _ _ _ -> assert false); } in pr_atom env sigma pr strip_prod_binders_constr tag_atomic_tactic_expr t in prtac t let pr_raw_generic = Pputils.pr_raw_generic let pr_glb_generic = Pputils.pr_glb_generic let pr_raw_extend env sigma = pr_raw_extend_rec @@ pr_raw_tactic_level env sigma let pr_glob_extend env sigma = pr_glob_extend_rec (pr_glob_tactic_level env) let pr_alias pr lev key args = pr_alias_gen (fun _ arg -> pr arg) lev key args let pr_extend pr lev ml args = pr_extend_gen pr lev ml args let pr_atomic_tactic env sigma c = pr_atomic_tactic_level env sigma c let declare_extra_genarg_pprule wit (f : 'a raw_extra_genarg_printer) (g : 'b glob_extra_genarg_printer) (h : 'c extra_genarg_printer) = begin match wit with | ExtraArg _ -> () | _ -> user_err Pp.(str "Can declare a pretty-printing rule only for extra argument types.") end; let f x = Genprint.PrinterBasic (fun env sigma -> f env sigma pr_constr_expr pr_lconstr_expr pr_raw_tactic_level x) in let g x = Genprint.PrinterBasic (fun env sigma -> g env sigma (fun env sigma -> pr_and_constr_expr (pr_glob_constr_env env)) (fun env sigma -> pr_and_constr_expr (pr_lglob_constr_env env)) (fun env sigma -> pr_glob_tactic_level env) x) in let h x = Genprint.TopPrinterNeedsContext (fun env sigma -> h env sigma pr_econstr_env pr_leconstr_env (fun _env _sigma _ _ -> str "") x) in Genprint.register_print0 wit f g h let declare_extra_genarg_pprule_with_level wit (f : 'a raw_extra_genarg_printer_with_level) (g : 'b glob_extra_genarg_printer_with_level) (h : 'c extra_genarg_printer_with_level) default_surrounded default_non_surrounded = begin match wit with | ExtraArg s -> () | _ -> user_err Pp.(str "Can declare a pretty-printing rule only for extra argument types.") end; let open Genprint in let f x = PrinterNeedsLevel { default_already_surrounded = default_surrounded; default_ensure_surrounded = default_non_surrounded; printer = (fun env sigma n -> f env sigma pr_constr_expr pr_lconstr_expr pr_raw_tactic_level n x) } in let g x = PrinterNeedsLevel { default_already_surrounded = default_surrounded; default_ensure_surrounded = default_non_surrounded; printer = (fun env sigma n -> g env sigma (fun env sigma -> pr_and_constr_expr (pr_glob_constr_env env)) (fun env sigma -> pr_and_constr_expr (pr_lglob_constr_env env)) (fun env sigma -> pr_glob_tactic_level env) n x) } in let h x = TopPrinterNeedsContextAndLevel { default_already_surrounded = default_surrounded; default_ensure_surrounded = default_non_surrounded; printer = (fun env sigma n -> h env sigma pr_econstr_env pr_leconstr_env (fun _env _sigma _ _ -> str "") n x) } in Genprint.register_print0 wit f g h let declare_extra_vernac_genarg_pprule wit f = let f x = Genprint.PrinterBasic (fun env sigma -> f env sigma pr_constr_expr pr_lconstr_expr pr_raw_tactic_level x) in Genprint.register_vernac_print0 wit f (** Registering *) let pr_intro_pattern_env p = Genprint.TopPrinterNeedsContext (fun env sigma -> let print_constr c = let (sigma, c) = c env sigma in pr_econstr_env env sigma c in Miscprint.pr_intro_pattern print_constr p) let pr_red_expr_env r = Genprint.TopPrinterNeedsContext (fun env sigma -> pr_red_expr env sigma (pr_econstr_env, pr_leconstr_env, pr_evaluable_reference_env env, pr_constr_pattern_env) r) let pr_bindings_env bl = Genprint.TopPrinterNeedsContext (fun env sigma -> let sigma, bl = bl env sigma in Miscprint.pr_bindings (pr_econstr_env env sigma) (pr_leconstr_env env sigma) bl) let pr_with_bindings_env bl = Genprint.TopPrinterNeedsContext (fun env sigma -> let sigma, bl = bl env sigma in pr_with_bindings (pr_econstr_env env sigma) (pr_leconstr_env env sigma) bl) let pr_destruction_arg_env c = Genprint.TopPrinterNeedsContext (fun env sigma -> let sigma, c = match c with | clear_flag,ElimOnConstr g -> let sigma,c = g env sigma in sigma,(clear_flag,ElimOnConstr c) | clear_flag,ElimOnAnonHyp n as x -> sigma, x | clear_flag,ElimOnIdent id as x -> sigma, x in pr_destruction_arg (pr_econstr_env env sigma) (pr_leconstr_env env sigma) c) let make_constr_printer f c = Genprint.TopPrinterNeedsContextAndLevel { Genprint.default_already_surrounded = Ppconstr.ltop; Genprint.default_ensure_surrounded = Ppconstr.lsimpleconstr; Genprint.printer = (fun env sigma n -> f env sigma n c)} let lift f a = Genprint.PrinterBasic (fun env sigma -> f a) let lift_env f a = Genprint.PrinterBasic (fun env sigma -> f env sigma a) let lift_top f a = Genprint.TopPrinterBasic (fun () -> f a) let register_basic_print0 wit f g h = Genprint.register_print0 wit (lift f) (lift g) (lift_top h) let pr_glob_constr_pptac env sigma c = pr_glob_constr_env env c let pr_lglob_constr_pptac env sigma c = pr_lglob_constr_env env c let pr_raw_intro_pattern = lift_env (fun env sigma -> Miscprint.pr_intro_pattern @@ pr_constr_expr env sigma) let pr_glob_intro_pattern = lift_env (fun env sigma -> Miscprint.pr_intro_pattern (fun (c,_) -> pr_glob_constr_pptac env sigma c)) let () = let pr_bool b = if b then str "true" else str "false" in let pr_unit _ = str "()" in let open Genprint in register_basic_print0 wit_int_or_var (pr_or_var int) (pr_or_var int) int; register_basic_print0 wit_ref pr_qualid (pr_or_var (pr_located pr_global)) pr_global; register_basic_print0 wit_ident pr_id pr_id pr_id; register_basic_print0 wit_var pr_lident pr_lident pr_id; register_print0 wit_intropattern pr_raw_intro_pattern pr_glob_intro_pattern pr_intro_pattern_env [@warning "-3"]; register_print0 wit_simple_intropattern pr_raw_intro_pattern pr_glob_intro_pattern pr_intro_pattern_env; Genprint.register_print0 wit_clause_dft_concl (lift (pr_clauses (Some true) pr_lident)) (lift (pr_clauses (Some true) pr_lident)) (fun c -> Genprint.TopPrinterBasic (fun () -> pr_clauses (Some true) (fun id -> pr_lident (CAst.make id)) c)) ; Genprint.register_print0 wit_constr (lift_env Ppconstr.pr_lconstr_expr) (lift_env (fun env sigma (c, _) -> pr_lglob_constr_pptac env sigma c)) (make_constr_printer Printer.pr_econstr_n_env) ; Genprint.register_print0 wit_uconstr (lift_env Ppconstr.pr_constr_expr) (lift_env (fun env sigma (c,_) -> pr_glob_constr_pptac env sigma c)) (make_constr_printer Printer.pr_closed_glob_n_env) ; Genprint.register_print0 wit_open_constr (lift_env Ppconstr.pr_constr_expr) (lift_env (fun env sigma (c, _) -> pr_glob_constr_pptac env sigma c)) (make_constr_printer Printer.pr_econstr_n_env) ; Genprint.register_print0 wit_red_expr (lift_env (fun env sigma -> pr_red_expr env sigma (pr_constr_expr, pr_lconstr_expr, pr_or_by_notation pr_qualid, pr_constr_pattern_expr))) (lift_env (fun env sigma -> pr_red_expr env sigma ((fun env sigma -> pr_and_constr_expr @@ pr_glob_constr_pptac env sigma), (fun env sigma -> pr_and_constr_expr @@ pr_lglob_constr_pptac env sigma), pr_or_var (pr_and_short_name pr_evaluable_reference), (fun env sigma -> pr_pat_and_constr_expr @@ pr_glob_constr_pptac env sigma)))) pr_red_expr_env ; register_basic_print0 wit_quant_hyp pr_quantified_hypothesis pr_quantified_hypothesis pr_quantified_hypothesis; register_print0 wit_bindings (lift_env (fun env sigma -> Miscprint.pr_bindings_no_with (pr_constr_expr env sigma) (pr_lconstr_expr env sigma))) (lift_env (fun env sigma -> Miscprint.pr_bindings_no_with (pr_and_constr_expr @@ pr_glob_constr_pptac env sigma) (pr_and_constr_expr @@ pr_lglob_constr_pptac env sigma))) pr_bindings_env ; register_print0 wit_constr_with_bindings (lift_env (fun env sigma -> pr_with_bindings (pr_constr_expr env sigma) (pr_lconstr_expr env sigma))) (lift_env (fun env sigma -> pr_with_bindings (pr_and_constr_expr @@ pr_glob_constr_pptac env sigma) (pr_and_constr_expr @@ pr_lglob_constr_pptac env sigma))) pr_with_bindings_env ; register_print0 wit_open_constr_with_bindings (lift_env (fun env sigma -> pr_with_bindings (pr_constr_expr env sigma) (pr_lconstr_expr env sigma))) (lift_env (fun env sigma -> pr_with_bindings (pr_and_constr_expr @@ pr_glob_constr_pptac env sigma) (pr_and_constr_expr @@ pr_lglob_constr_pptac env sigma))) pr_with_bindings_env ; register_print0 Tacarg.wit_destruction_arg (lift_env (fun env sigma -> pr_destruction_arg (pr_constr_expr env sigma) (pr_lconstr_expr env sigma))) (lift_env (fun env sigma -> pr_destruction_arg (pr_and_constr_expr @@ pr_glob_constr_pptac env sigma) (pr_and_constr_expr @@ pr_lglob_constr_pptac env sigma))) pr_destruction_arg_env ; register_basic_print0 Stdarg.wit_int int int int; register_basic_print0 Stdarg.wit_bool pr_bool pr_bool pr_bool; register_basic_print0 Stdarg.wit_unit pr_unit pr_unit pr_unit; register_basic_print0 Stdarg.wit_pre_ident str str str; register_basic_print0 Stdarg.wit_string qstring qstring qstring let () = let printer env sigma _ _ prtac = prtac env sigma in declare_extra_genarg_pprule_with_level wit_tactic printer printer printer ltop (0,E) let () = let pr_unit _env _sigma _ _ _ _ () = str "()" in let printer env sigma _ _ prtac = prtac env sigma in declare_extra_genarg_pprule_with_level wit_ltac printer printer pr_unit ltop (0,E) coq-8.11.0/plugins/ltac/evar_tactics.mli0000644000175000017500000000207213612664533020045 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Tacinterp.interp_sign * Glob_term.glob_constr -> (Id.t * hyp_location_flag, unit) location -> unit Proofview.tactic val instantiate_tac_by_name : Id.t -> Tacinterp.interp_sign * Glob_term.glob_constr -> unit Proofview.tactic val let_evar : Name.t -> EConstr.types -> unit Proofview.tactic val hget_evar : int -> unit Proofview.tactic coq-8.11.0/plugins/ltac/profile_ltac_tactics.mlg0000644000175000017500000000543113612664533021553 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* set_profiling b)) let tclRESET_PROFILE = Proofview.tclLIFT (Proofview.NonLogical.make reset_profile) let tclSHOW_PROFILE ~cutoff = Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> print_results ~cutoff)) let tclSHOW_PROFILE_TACTIC s = Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> print_results_tactic s)) let tclRESTART_TIMER s = Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> restart_timer s)) let tclFINISH_TIMING ?(prefix="Timer") (s : string option) = Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> finish_timing ~prefix s)) } TACTIC EXTEND start_ltac_profiling | [ "start" "ltac" "profiling" ] -> { tclSET_PROFILING true } END TACTIC EXTEND stop_ltac_profiling | [ "stop" "ltac" "profiling" ] -> { tclSET_PROFILING false } END TACTIC EXTEND reset_ltac_profile | [ "reset" "ltac" "profile" ] -> { tclRESET_PROFILE } END TACTIC EXTEND show_ltac_profile | [ "show" "ltac" "profile" ] -> { tclSHOW_PROFILE ~cutoff:!Flags.profile_ltac_cutoff } | [ "show" "ltac" "profile" "cutoff" int(n) ] -> { tclSHOW_PROFILE ~cutoff:(float_of_int n) } | [ "show" "ltac" "profile" string(s) ] -> { tclSHOW_PROFILE_TACTIC s } END TACTIC EXTEND restart_timer | [ "restart_timer" string_opt(s) ] -> { tclRESTART_TIMER s } END TACTIC EXTEND finish_timing | [ "finish_timing" string_opt(s) ] -> { tclFINISH_TIMING ~prefix:"Timer" s } | [ "finish_timing" "(" string(prefix) ")" string_opt(s) ] -> { tclFINISH_TIMING ~prefix s } END VERNAC COMMAND EXTEND ResetLtacProfiling CLASSIFIED AS SIDEFF | [ "Reset" "Ltac" "Profile" ] -> { reset_profile () } END VERNAC COMMAND EXTEND ShowLtacProfile CLASSIFIED AS QUERY | [ "Show" "Ltac" "Profile" ] -> { print_results ~cutoff:!Flags.profile_ltac_cutoff } | [ "Show" "Ltac" "Profile" "CutOff" int(n) ] -> { print_results ~cutoff:(float_of_int n) } END VERNAC COMMAND EXTEND ShowLtacProfileTactic CLASSIFIED AS QUERY | [ "Show" "Ltac" "Profile" string(s) ] -> { print_results_tactic s } END coq-8.11.0/plugins/ltac/g_rewrite.mlg0000644000175000017500000003262213612664533017367 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { bl } END { type raw_strategy = (constr_expr, Tacexpr.raw_red_expr) strategy_ast type glob_strategy = (glob_constr_and_expr, Tacexpr.raw_red_expr) strategy_ast let interp_strategy ist gl s = let sigma = project gl in sigma, strategy_of_ast s let glob_strategy ist s = map_strategy (Tacintern.intern_constr ist) (fun c -> c) s let subst_strategy s str = str let pr_strategy _ _ _ (s : strategy) = Pp.str "" let pr_raw_strategy env sigma prc prlc _ (s : raw_strategy) = let prr = Pptactic.pr_red_expr env sigma (prc, prlc, Pputils.pr_or_by_notation Libnames.pr_qualid, prc) in Rewrite.pr_strategy (prc env sigma) prr s let pr_glob_strategy env sigma prc prlc _ (s : glob_strategy) = let prr = Pptactic.pr_red_expr env sigma (Ppconstr.pr_constr_expr, Ppconstr.pr_lconstr_expr, Pputils.pr_or_by_notation Libnames.pr_qualid, Ppconstr.pr_constr_expr) in Rewrite.pr_strategy (prc env sigma) prr s } ARGUMENT EXTEND rewstrategy PRINTED BY { pr_strategy } INTERPRETED BY { interp_strategy } GLOBALIZED BY { glob_strategy } SUBSTITUTED BY { subst_strategy } RAW_PRINTED BY { pr_raw_strategy env sigma } GLOB_PRINTED BY { pr_glob_strategy env sigma } | [ glob(c) ] -> { StratConstr (c, true) } | [ "<-" constr(c) ] -> { StratConstr (c, false) } | [ "subterms" rewstrategy(h) ] -> { StratUnary (Subterms, h) } | [ "subterm" rewstrategy(h) ] -> { StratUnary (Subterm, h) } | [ "innermost" rewstrategy(h) ] -> { StratUnary(Innermost, h) } | [ "outermost" rewstrategy(h) ] -> { StratUnary(Outermost, h) } | [ "bottomup" rewstrategy(h) ] -> { StratUnary(Bottomup, h) } | [ "topdown" rewstrategy(h) ] -> { StratUnary(Topdown, h) } | [ "id" ] -> { StratId } | [ "fail" ] -> { StratFail } | [ "refl" ] -> { StratRefl } | [ "progress" rewstrategy(h) ] -> { StratUnary (Progress, h) } | [ "try" rewstrategy(h) ] -> { StratUnary (Try, h) } | [ "any" rewstrategy(h) ] -> { StratUnary (Any, h) } | [ "repeat" rewstrategy(h) ] -> { StratUnary (Repeat, h) } | [ rewstrategy(h) ";" rewstrategy(h') ] -> { StratBinary (Compose, h, h') } | [ "(" rewstrategy(h) ")" ] -> { h } | [ "choice" rewstrategy(h) rewstrategy(h') ] -> { StratBinary (Choice, h, h') } | [ "old_hints" preident(h) ] -> { StratHints (true, h) } | [ "hints" preident(h) ] -> { StratHints (false, h) } | [ "terms" constr_list(h) ] -> { StratTerms h } | [ "eval" red_expr(r) ] -> { StratEval r } | [ "fold" constr(c) ] -> { StratFold c } END (* By default the strategy for "rewrite_db" is top-down *) { let db_strat db = StratUnary (Topdown, StratHints (false, db)) let cl_rewrite_clause_db db = cl_rewrite_clause_strat (strategy_of_ast (db_strat db)) } TACTIC EXTEND rewrite_strat | [ "rewrite_strat" rewstrategy(s) "in" hyp(id) ] -> { cl_rewrite_clause_strat s (Some id) } | [ "rewrite_strat" rewstrategy(s) ] -> { cl_rewrite_clause_strat s None } | [ "rewrite_db" preident(db) "in" hyp(id) ] -> { cl_rewrite_clause_db db (Some id) } | [ "rewrite_db" preident(db) ] -> { cl_rewrite_clause_db db None } END { let clsubstitute o c = Proofview.Goal.enter begin fun gl -> let is_tac id = match DAst.get (fst (fst (snd c))) with GVar id' when Id.equal id' id -> true | _ -> false in let hyps = Tacmach.New.pf_ids_of_hyps gl in Tacticals.New.tclMAP (fun cl -> match cl with | Some id when is_tac id -> Tacticals.New.tclIDTAC | _ -> cl_rewrite_clause c o AllOccurrences cl) (None :: List.map (fun id -> Some id) hyps) end } TACTIC EXTEND substitute | [ "substitute" orient(o) glob_constr_with_bindings(c) ] -> { clsubstitute o c } END (* Compatibility with old Setoids *) TACTIC EXTEND setoid_rewrite | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) ] -> { cl_rewrite_clause c o AllOccurrences None } | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) ] -> { cl_rewrite_clause c o AllOccurrences (Some id) } | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) ] -> { cl_rewrite_clause c o (occurrences_of occ) None } | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) "in" hyp(id)] -> { cl_rewrite_clause c o (occurrences_of occ) (Some id) } | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) "at" occurrences(occ)] -> { cl_rewrite_clause c o (occurrences_of occ) (Some id) } END VERNAC COMMAND EXTEND AddRelation CLASSIFIED AS SIDEFF | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> { declare_relation atts a aeq n (Some lemma1) (Some lemma2) None } | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "as" ident(n) ] -> { declare_relation atts a aeq n (Some lemma1) None None } | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "as" ident(n) ] -> { declare_relation atts a aeq n None None None } END VERNAC COMMAND EXTEND AddRelation2 CLASSIFIED AS SIDEFF | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> { declare_relation atts a aeq n None (Some lemma2) None } | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts a aeq n None (Some lemma2) (Some lemma3) } END VERNAC COMMAND EXTEND AddRelation3 CLASSIFIED AS SIDEFF | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts a aeq n (Some lemma1) None (Some lemma3) } | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts a aeq n (Some lemma1) (Some lemma2) (Some lemma3) } | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts a aeq n None None (Some lemma3) } END { type binders_argtype = local_binder_expr list let wit_binders = (Genarg.create_arg "binders" : binders_argtype Genarg.uniform_genarg_type) let binders = Pcoq.create_generic_entry Pcoq.utactic "binders" (Genarg.rawwit wit_binders) let () = let raw_printer env sigma _ _ _ l = Pp.pr_non_empty_arg (Ppconstr.pr_binders env sigma) l in Pptactic.declare_extra_vernac_genarg_pprule wit_binders raw_printer } GRAMMAR EXTEND Gram GLOBAL: binders; binders: [ [ b = Pcoq.Constr.binders -> { b } ] ]; END VERNAC COMMAND EXTEND AddParametricRelation CLASSIFIED AS SIDEFF | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n (Some lemma1) (Some lemma2) None } | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n (Some lemma1) None None } | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n None None None } END VERNAC COMMAND EXTEND AddParametricRelation2 CLASSIFIED AS SIDEFF | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n None (Some lemma2) None } | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n None (Some lemma2) (Some lemma3) } END VERNAC COMMAND EXTEND AddParametricRelation3 CLASSIFIED AS SIDEFF | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n (Some lemma1) None (Some lemma3) } | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n (Some lemma1) (Some lemma2) (Some lemma3) } | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> { declare_relation atts ~binders:b a aeq n None None (Some lemma3) } END VERNAC COMMAND EXTEND AddSetoid1 CLASSIFIED AS SIDEFF | #[ atts = rewrite_attributes; ] [ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] -> { add_setoid atts [] a aeq t n } | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] -> { add_setoid atts binders a aeq t n } | #[ atts = rewrite_attributes; ] ![ open_proof ] [ "Add" "Morphism" constr(m) ":" ident(n) ] => { VtStartProof(GuaranteesOpacity, [n]) } -> { if Lib.is_modtype () then CErrors.user_err Pp.(str "Add Morphism cannot be used in a module type. Use Parameter Morphism instead."); add_morphism_interactive atts m n } | #[ atts = rewrite_attributes; ] [ "Declare" "Morphism" constr(m) ":" ident(n) ] => { VtSideff([n], VtLater) } -> { add_morphism_as_parameter atts m n } | #[ atts = rewrite_attributes; ] ![ open_proof ] [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ] => { VtStartProof(GuaranteesOpacity,[n]) } -> { add_morphism atts [] m s n } | #[ atts = rewrite_attributes; ] ![ open_proof ] [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m) "with" "signature" lconstr(s) "as" ident(n) ] => { VtStartProof(GuaranteesOpacity,[n]) } -> { add_morphism atts binders m s n } END TACTIC EXTEND setoid_symmetry | [ "setoid_symmetry" ] -> { setoid_symmetry } | [ "setoid_symmetry" "in" hyp(n) ] -> { setoid_symmetry_in n } END TACTIC EXTEND setoid_reflexivity | [ "setoid_reflexivity" ] -> { setoid_reflexivity } END TACTIC EXTEND setoid_transitivity | [ "setoid_transitivity" constr(t) ] -> { setoid_transitivity (Some t) } | [ "setoid_etransitivity" ] -> { setoid_transitivity None } END VERNAC COMMAND EXTEND PrintRewriteHintDb CLASSIFIED AS QUERY | [ "Print" "Rewrite" "HintDb" preident(s) ] -> { Feedback.msg_notice (Autorewrite.print_rewrite_hintdb s) } END coq-8.11.0/plugins/ltac/tauto.ml0000644000175000017500000002277013612664533016370 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* c | None -> failwith "tauto: anomaly" (** Parametrization of tauto *) type tauto_flags = { (* Whether conjunction and disjunction are restricted to binary connectives *) binary_mode : bool; (* Whether compatibility for buggy detection of binary connective is on *) binary_mode_bugged_detection : bool; (* Whether conjunction and disjunction are restricted to the connectives *) (* having the structure of "and" and "or" (up to the choice of sorts) in *) (* contravariant position in an hypothesis *) strict_in_contravariant_hyp : bool; (* Whether conjunction and disjunction are restricted to the connectives *) (* having the structure of "and" and "or" (up to the choice of sorts) in *) (* an hypothesis and in the conclusion *) strict_in_hyp_and_ccl : bool; (* Whether unit type includes equality types *) strict_unit : bool; } let tag_tauto_flags : tauto_flags Val.typ = Val.create "tauto_flags" let assoc_flags ist : tauto_flags = let Val.Dyn (tag, v) = Id.Map.find (Names.Id.of_string "tauto_flags") ist.lfun in match Val.eq tag tag_tauto_flags with | None -> assert false | Some Refl -> v (* Whether inner not are unfolded *) let negation_unfolding = ref true open Goptions let () = declare_bool_option { optdepr = false; optname = "unfolding of not in intuition"; optkey = ["Intuition";"Negation";"Unfolding"]; optread = (fun () -> !negation_unfolding); optwrite = (:=) negation_unfolding } (** Base tactics *) let idtac = Proofview.tclUNIT () let fail = Proofview.tclINDEPENDENT (tclFAIL 0 (Pp.mt ())) let intro = Tactics.intro let assert_ ?by c = let tac = match by with | None -> None | Some tac -> Some (Some tac) in Proofview.tclINDEPENDENT (Tactics.forward true tac None c) let apply c = Tactics.apply c let clear id = Tactics.clear [id] let assumption = Tactics.assumption let split = Tactics.split_with_bindings false [Tactypes.NoBindings] (** Test *) let is_empty _ ist = Proofview.tclENV >>= fun genv -> Proofview.tclEVARMAP >>= fun sigma -> if is_empty_type genv sigma (assoc_var "X1" ist) then idtac else fail (* Strictly speaking, this exceeds the propositional fragment as it matches also equality types (and solves them if a reflexivity) *) let is_unit_or_eq _ ist = Proofview.tclENV >>= fun genv -> Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let test = if flags.strict_unit then is_unit_type else is_unit_or_eq_type in if test genv sigma (assoc_var "X1" ist) then idtac else fail let bugged_is_binary sigma t = isApp sigma t && let (hdapp,args) = decompose_app sigma t in match EConstr.kind sigma hdapp with | Ind (ind,u) -> let (mib,mip) = Global.lookup_inductive ind in Int.equal mib.Declarations.mind_nparams 2 | _ -> false (** Dealing with conjunction *) let is_conj _ ist = Proofview.tclENV >>= fun genv -> Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let ind = assoc_var "X1" ist in if (not flags.binary_mode_bugged_detection || bugged_is_binary sigma ind) && is_conjunction genv sigma ~strict:flags.strict_in_hyp_and_ccl ~onlybinary:flags.binary_mode ind then idtac else fail let flatten_contravariant_conj _ ist = Proofview.tclENV >>= fun genv -> Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in let c = assoc_var "X2" ist in let hyp = assoc_var "id" ist in match match_with_conjunction genv sigma ~strict:flags.strict_in_contravariant_hyp ~onlybinary:flags.binary_mode typ with | Some (_,args) -> let newtyp = List.fold_right (fun a b -> mkArrow a Sorts.Relevant b) args c in let intros = tclMAP (fun _ -> intro) args in let by = tclTHENLIST [intros; apply hyp; split; assumption] in tclTHENLIST [assert_ ~by newtyp; clear (destVar sigma hyp)] | _ -> fail (** Dealing with disjunction *) let is_disj _ ist = Proofview.tclENV >>= fun genv -> Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let t = assoc_var "X1" ist in if (not flags.binary_mode_bugged_detection || bugged_is_binary sigma t) && is_disjunction genv sigma ~strict:flags.strict_in_hyp_and_ccl ~onlybinary:flags.binary_mode t then idtac else fail let flatten_contravariant_disj _ ist = Proofview.tclENV >>= fun genv -> Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in let c = assoc_var "X2" ist in let hyp = assoc_var "id" ist in match match_with_disjunction genv sigma ~strict:flags.strict_in_contravariant_hyp ~onlybinary:flags.binary_mode typ with | Some (_,args) -> let map i arg = let typ = mkArrow arg Sorts.Relevant c in let ci = Tactics.constructor_tac false None (succ i) Tactypes.NoBindings in let by = tclTHENLIST [intro; apply hyp; ci; assumption] in assert_ ~by typ in let tacs = List.mapi map args in let tac0 = clear (destVar sigma hyp) in tclTHEN (tclTHENLIST tacs) tac0 | _ -> fail let evalglobref_of_globref = function | GlobRef.VarRef v -> EvalVarRef v | GlobRef.ConstRef c -> EvalConstRef c | GlobRef.IndRef _ | GlobRef.ConstructRef _ -> assert false let make_unfold name = let const = evalglobref_of_globref (Coqlib.lib_ref name) in Locus.(AllOccurrences, ArgArg (const, None)) let reduction_not_iff _ ist = let make_reduce c = TacAtom (CAst.make @@ TacReduce (Genredexpr.Unfold c, Locusops.allHypsAndConcl)) in let tac = match !negation_unfolding with | true -> make_reduce [make_unfold "core.not.type"] | false -> TacId [] in eval_tactic_ist ist tac let apply_nnpp _ ist = let nnpp = "core.nnpp.type" in Proofview.tclBIND (Proofview.tclUNIT ()) begin fun () -> if Coqlib.has_ref nnpp then Tacticals.New.pf_constr_of_global (Coqlib.lib_ref nnpp) >>= apply else tclFAIL 0 (Pp.mt ()) end (* This is the uniform mode dealing with ->, not, iff and types isomorphic to /\ and *, \/ and +, False and Empty_set, True and unit, _and_ eq-like types. For the moment not and iff are still always unfolded. *) let tauto_uniform_unit_flags = { binary_mode = true; binary_mode_bugged_detection = false; strict_in_contravariant_hyp = true; strict_in_hyp_and_ccl = true; strict_unit = false } (* This is the compatibility mode (not used) *) let _tauto_legacy_flags = { binary_mode = true; binary_mode_bugged_detection = true; strict_in_contravariant_hyp = true; strict_in_hyp_and_ccl = false; strict_unit = false } (* This is the improved mode *) let tauto_power_flags = { binary_mode = false; (* support n-ary connectives *) binary_mode_bugged_detection = false; strict_in_contravariant_hyp = false; (* supports non-regular connectives *) strict_in_hyp_and_ccl = false; strict_unit = false } let with_flags flags _ ist = let f = CAst.make @@ Id.of_string "f" in let x = CAst.make @@ Id.of_string "x" in let arg = Val.Dyn (tag_tauto_flags, flags) in let ist = { ist with lfun = Id.Map.add x.CAst.v arg ist.lfun } in eval_tactic_ist ist (TacArg (CAst.make @@ TacCall (CAst.make (Locus.ArgVar f, [Reference (Locus.ArgVar x)])))) let register_tauto_tactic tac name0 args = let ids = List.map (fun id -> Id.of_string id) args in let ids = List.map (fun id -> Name id) ids in let name = { mltac_plugin = tauto_plugin; mltac_tactic = name0; } in let entry = { mltac_name = name; mltac_index = 0 } in let () = Tacenv.register_ml_tactic name [| tac |] in let tac = TacFun (ids, TacML (CAst.make (entry, []))) in let obj () = Tacenv.register_ltac true true (Id.of_string name0) tac in Mltop.declare_cache_obj obj tauto_plugin let () = register_tauto_tactic is_empty "is_empty" ["tauto_flags"; "X1"] let () = register_tauto_tactic is_unit_or_eq "is_unit_or_eq" ["tauto_flags"; "X1"] let () = register_tauto_tactic is_disj "is_disj" ["tauto_flags"; "X1"] let () = register_tauto_tactic is_conj "is_conj" ["tauto_flags"; "X1"] let () = register_tauto_tactic flatten_contravariant_disj "flatten_contravariant_disj" ["tauto_flags"; "X1"; "X2"; "id"] let () = register_tauto_tactic flatten_contravariant_conj "flatten_contravariant_conj" ["tauto_flags"; "X1"; "X2"; "id"] let () = register_tauto_tactic apply_nnpp "apply_nnpp" [] let () = register_tauto_tactic reduction_not_iff "reduction_not_iff" [] let () = register_tauto_tactic (with_flags tauto_uniform_unit_flags) "with_uniform_flags" ["f"] let () = register_tauto_tactic (with_flags tauto_power_flags) "with_power_flags" ["f"] coq-8.11.0/plugins/ltac/tacintern.mli0000644000175000017500000000417613612664533017374 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* glob_sign (** build an empty [glob_sign] using [Global.env()] as environment *) (** Main globalization functions *) val glob_tactic : raw_tactic_expr -> glob_tactic_expr val glob_tactic_env : Id.t list -> Environ.env -> raw_tactic_expr -> glob_tactic_expr (** Low-level variants *) val intern_pure_tactic : glob_sign -> raw_tactic_expr -> glob_tactic_expr val intern_tactic_or_tacarg : glob_sign -> raw_tactic_expr -> Tacexpr.glob_tactic_expr val intern_constr : glob_sign -> constr_expr -> glob_constr_and_expr val intern_constr_with_bindings : glob_sign -> constr_expr * constr_expr bindings -> glob_constr_and_expr * glob_constr_and_expr bindings val intern_hyp : glob_sign -> lident -> lident (** Adds a globalization function for extra generic arguments *) val intern_genarg : glob_sign -> raw_generic_argument -> glob_generic_argument (** printing *) val print_ltac : Libnames.qualid -> Pp.t (** Reduction expressions *) val intern_red_expr : glob_sign -> raw_red_expr -> glob_red_expr val dump_glob_red_expr : raw_red_expr -> unit (* Hooks *) val strict_check : bool ref coq-8.11.0/plugins/ltac/extraargs.mlg0000644000175000017500000002263313612664533017401 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Metasyntax.add_token_obj "<-"; Metasyntax.add_token_obj "->") "ltac_plugin" let pr_orient _prc _prlc _prt = function | true -> Pp.mt () | false -> Pp.str " <-" } ARGUMENT EXTEND orient TYPED AS bool PRINTED BY { pr_orient } | [ "->" ] -> { true } | [ "<-" ] -> { false } | [ ] -> { true } END { let pr_int _ _ _ i = Pp.int i let _natural = Pcoq.Prim.natural } ARGUMENT EXTEND natural TYPED AS int PRINTED BY { pr_int } | [ _natural(i) ] -> { i } END { let pr_orient = pr_orient () () () let pr_int_list = Pp.pr_sequence Pp.int let pr_int_list_full _prc _prlc _prt l = pr_int_list l let pr_occurrences _prc _prlc _prt l = match l with | ArgArg x -> pr_int_list x | ArgVar { CAst.loc = loc; v=id } -> Id.print id let occurrences_of = function | [] -> NoOccurrences | n::_ as nl when n < 0 -> AllOccurrencesBut (List.map abs nl) | nl -> if List.exists (fun n -> n < 0) nl then CErrors.user_err Pp.(str "Illegal negative occurrence number."); OnlyOccurrences nl let coerce_to_int v = match Value.to_int v with | None -> raise (CannotCoerceTo "an integer") | Some n -> n let int_list_of_VList v = match Value.to_list v with | Some l -> List.map (fun n -> coerce_to_int n) l | _ -> raise (CannotCoerceTo "an integer") let interp_occs ist gl l = match l with | ArgArg x -> x | ArgVar ({ CAst.v = id } as locid) -> (try int_list_of_VList (Id.Map.find id ist.lfun) with Not_found | CannotCoerceTo _ -> [interp_int ist locid]) let interp_occs ist gl l = Tacmach.project gl , interp_occs ist gl l let glob_occs ist l = l let subst_occs evm l = l } ARGUMENT EXTEND occurrences TYPED AS int list PRINTED BY { pr_int_list_full } INTERPRETED BY { interp_occs } GLOBALIZED BY { glob_occs } SUBSTITUTED BY { subst_occs } RAW_PRINTED BY { pr_occurrences } GLOB_PRINTED BY { pr_occurrences } | [ ne_integer_list(l) ] -> { ArgArg l } | [ var(id) ] -> { ArgVar id } END { let pr_occurrences = pr_occurrences () () () let pr_gen env sigma prc _prlc _prtac x = prc env sigma x let pr_globc env sigma _prc _prlc _prtac (_,glob) = Printer.pr_glob_constr_env env glob let interp_glob ist gl (t,_) = Tacmach.project gl , (ist,t) let glob_glob = Tacintern.intern_constr let pr_lconstr env sigma _ prc _ c = prc env sigma c let subst_glob = Tacsubst.subst_glob_constr_and_expr } ARGUMENT EXTEND glob PRINTED BY { pr_globc env sigma } INTERPRETED BY { interp_glob } GLOBALIZED BY { glob_glob } SUBSTITUTED BY { subst_glob } RAW_PRINTED BY { pr_gen env sigma } GLOB_PRINTED BY { pr_gen env sigma } | [ constr(c) ] -> { c } END { let l_constr = Pcoq.Constr.lconstr } ARGUMENT EXTEND lconstr TYPED AS constr PRINTED BY { pr_lconstr env sigma } | [ l_constr(c) ] -> { c } END ARGUMENT EXTEND lglob TYPED AS glob PRINTED BY { pr_globc env sigma } INTERPRETED BY { interp_glob } GLOBALIZED BY { glob_glob } SUBSTITUTED BY { subst_glob } RAW_PRINTED BY { pr_gen env sigma } GLOB_PRINTED BY { pr_gen env sigma } | [ lconstr(c) ] -> { c } END { let interp_casted_constr ist gl c = interp_constr_gen (Pretyping.OfType (pf_concl gl)) ist (pf_env gl) (project gl) c } ARGUMENT EXTEND casted_constr TYPED AS constr PRINTED BY { pr_gen env sigma } INTERPRETED BY { interp_casted_constr } | [ constr(c) ] -> { c } END { type 'id gen_place= ('id * hyp_location_flag,unit) location type loc_place = lident gen_place type place = Id.t gen_place let pr_gen_place pr_id = function ConclLocation () -> Pp.mt () | HypLocation (id,InHyp) -> str "in " ++ pr_id id | HypLocation (id,InHypTypeOnly) -> str "in (type of " ++ pr_id id ++ str ")" | HypLocation (id,InHypValueOnly) -> str "in (value of " ++ pr_id id ++ str ")" let pr_loc_place _ _ _ = pr_gen_place (fun { CAst.v = id } -> Id.print id) let pr_place _ _ _ = pr_gen_place Id.print let pr_hloc = pr_loc_place () () () let intern_place ist = function ConclLocation () -> ConclLocation () | HypLocation (id,hl) -> HypLocation (Tacintern.intern_hyp ist id,hl) let interp_place ist env sigma = function ConclLocation () -> ConclLocation () | HypLocation (id,hl) -> HypLocation (Tacinterp.interp_hyp ist env sigma id,hl) let interp_place ist gl p = Tacmach.project gl , interp_place ist (Tacmach.pf_env gl) (Tacmach.project gl) p let subst_place subst pl = pl let warn_deprecated_instantiate_syntax = CWarnings.create ~name:"deprecated-instantiate-syntax" ~category:"deprecated" (fun (v,v',id) -> let s = Id.to_string id in Pp.strbrk ("Syntax \"in (" ^ v ^ " of " ^ s ^ ")\" is deprecated; use \"in (" ^ v' ^ " of " ^ s ^ ")\".") ) } ARGUMENT EXTEND hloc PRINTED BY { pr_place } INTERPRETED BY { interp_place } GLOBALIZED BY { intern_place } SUBSTITUTED BY { subst_place } RAW_PRINTED BY { pr_loc_place } GLOB_PRINTED BY { pr_loc_place } | [ ] -> { ConclLocation () } | [ "in" "|-" "*" ] -> { ConclLocation () } | [ "in" ident(id) ] -> { HypLocation ((CAst.make id),InHyp) } | [ "in" "(" "Type" "of" ident(id) ")" ] -> { warn_deprecated_instantiate_syntax ("Type","type",id); HypLocation ((CAst.make id),InHypTypeOnly) } | [ "in" "(" "Value" "of" ident(id) ")" ] -> { warn_deprecated_instantiate_syntax ("Value","value",id); HypLocation ((CAst.make id),InHypValueOnly) } | [ "in" "(" "type" "of" ident(id) ")" ] -> { HypLocation ((CAst.make id),InHypTypeOnly) } | [ "in" "(" "value" "of" ident(id) ")" ] -> { HypLocation ((CAst.make id),InHypValueOnly) } END { let pr_rename _ _ _ (n, m) = Id.print n ++ str " into " ++ Id.print m } ARGUMENT EXTEND rename TYPED AS (ident * ident) PRINTED BY { pr_rename } | [ ident(n) "into" ident(m) ] -> { (n, m) } END (* Julien: Mise en commun des differentes version de replace with in by *) { let pr_by_arg_tac env sigma _prc _prlc prtac opt_c = match opt_c with | None -> mt () | Some t -> hov 2 (str "by" ++ spc () ++ prtac env sigma (3,Notation_gram.E) t) } ARGUMENT EXTEND by_arg_tac TYPED AS tactic option PRINTED BY { pr_by_arg_tac env sigma } | [ "by" tactic3(c) ] -> { Some c } | [ ] -> { None } END { let pr_by_arg_tac env sigma prtac opt_c = pr_by_arg_tac env sigma () () prtac opt_c let pr_in_clause _ _ _ cl = Pptactic.pr_in_clause Pputils.pr_lident cl let pr_in_top_clause _ _ _ cl = Pptactic.pr_in_clause Id.print cl let in_clause' = Pltac.in_clause } ARGUMENT EXTEND in_clause TYPED AS clause_dft_concl PRINTED BY { pr_in_top_clause } RAW_PRINTED BY { pr_in_clause } GLOB_PRINTED BY { pr_in_clause } | [ in_clause'(cl) ] -> { cl } END { let local_test_lpar_id_colon = let err () = raise Stream.Failure in Pcoq.Entry.of_parser "lpar_id_colon" (fun _ strm -> match Util.stream_nth 0 strm with | Tok.KEYWORD "(" -> (match Util.stream_nth 1 strm with | Tok.IDENT _ -> (match Util.stream_nth 2 strm with | Tok.KEYWORD ":" -> () | _ -> err ()) | _ -> err ()) | _ -> err ()) let pr_lpar_id_colon _ _ _ _ = mt () } ARGUMENT EXTEND test_lpar_id_colon TYPED AS unit PRINTED BY { pr_lpar_id_colon } | [ local_test_lpar_id_colon(x) ] -> { () } END coq-8.11.0/plugins/ltac/tacsubst.mli0000644000175000017500000000237013612664533017227 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* glob_tactic_expr -> glob_tactic_expr (** For generic arguments, we declare and store substitutions in a table *) val subst_genarg : substitution -> glob_generic_argument -> glob_generic_argument (** Misc *) val subst_glob_constr_and_expr : substitution -> glob_constr_and_expr -> glob_constr_and_expr val subst_glob_with_bindings : substitution -> glob_constr_and_expr with_bindings -> glob_constr_and_expr with_bindings coq-8.11.0/plugins/ltac/rewrite.ml0000644000175000017500000026140213612664533016712 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* >= fun ((polymorphic, program), locality) -> let global = not (Locality.make_section_locality locality) in Attributes.Notations.return { polymorphic; global } (** Constants used by the tactic. *) let classes_dirpath = Names.DirPath.make (List.map Id.of_string ["Classes";"Coq"]) let init_relation_classes () = if is_dirpath_prefix_of classes_dirpath (Lib.cwd ()) then () else Coqlib.check_required_library ["Coq";"Classes";"RelationClasses"] let init_setoid () = if is_dirpath_prefix_of classes_dirpath (Lib.cwd ()) then () else Coqlib.check_required_library ["Coq";"Setoids";"Setoid"] let find_reference dir s = Coqlib.find_reference "generalized rewriting" dir s [@@warning "-3"] let lazy_find_reference dir s = let gr = lazy (find_reference dir s) in fun () -> Lazy.force gr type evars = evar_map * Evar.Set.t (* goal evars, constraint evars *) let find_global dir s = let gr = lazy (find_reference dir s) in fun (evd,cstrs) -> let (evd, c) = Evarutil.new_global evd (Lazy.force gr) in (evd, cstrs), c (** Utility for dealing with polymorphic applications *) (** Global constants. *) let coq_eq_ref () = Coqlib.lib_ref "core.eq.type" let coq_eq = find_global ["Coq"; "Init"; "Logic"] "eq" let coq_f_equal = find_global ["Coq"; "Init"; "Logic"] "f_equal" let coq_all = find_global ["Coq"; "Init"; "Logic"] "all" let impl = find_global ["Coq"; "Program"; "Basics"] "impl" (** Bookkeeping which evars are constraints so that we can remove them at the end of the tactic. *) let goalevars evars = fst evars let cstrevars evars = snd evars let new_cstr_evar (evd,cstrs) env t = (* We handle the typeclass resolution of constraints ourselves *) let (evd', t) = Evarutil.new_evar env evd ~typeclass_candidate:false t in let ev, _ = destEvar evd' t in (evd', Evar.Set.add ev cstrs), t (** Building or looking up instances. *) let e_new_cstr_evar env evars t = let evd', t = new_cstr_evar !evars env t in evars := evd'; t (** Building or looking up instances. *) let extends_undefined evars evars' = let f ev evi found = found || not (Evd.mem evars ev) in fold_undefined f evars' false let app_poly_check env evars f args = let (evars, cstrs), fc = f evars in let evars, t = Typing.solve_evars env evars (mkApp (fc, args)) in (evars, cstrs), t let app_poly_nocheck env evars f args = let evars, fc = f evars in evars, mkApp (fc, args) let app_poly_sort b = if b then app_poly_nocheck else app_poly_check let find_class_proof proof_type proof_method env evars carrier relation = try let evars, goal = app_poly_check env evars proof_type [| carrier ; relation |] in let evars', c = Typeclasses.resolve_one_typeclass env (goalevars evars) goal in if extends_undefined (goalevars evars) evars' then raise Not_found else app_poly_check env (evars',cstrevars evars) proof_method [| carrier; relation; c |] with e when Logic.catchable_exception e -> raise Not_found (** Utility functions *) module GlobalBindings (M : sig val relation_classes : string list val morphisms : string list val relation : string list * string val app_poly : env -> evars -> (evars -> evars * constr) -> constr array -> evars * constr val arrow : evars -> evars * constr end) = struct open M open Context.Rel.Declaration let relation : evars -> evars * constr = find_global (fst relation) (snd relation) let reflexive_type = find_global relation_classes "Reflexive" let reflexive_proof = find_global relation_classes "reflexivity" let symmetric_type = find_global relation_classes "Symmetric" let symmetric_proof = find_global relation_classes "symmetry" let transitive_type = find_global relation_classes "Transitive" let transitive_proof = find_global relation_classes "transitivity" let forall_relation = find_global morphisms "forall_relation" let pointwise_relation = find_global morphisms "pointwise_relation" let forall_relation_ref = lazy_find_reference morphisms "forall_relation" let pointwise_relation_ref = lazy_find_reference morphisms "pointwise_relation" let respectful = find_global morphisms "respectful" let respectful_ref = lazy_find_reference morphisms "respectful" let default_relation = find_global ["Coq"; "Classes"; "SetoidTactics"] "DefaultRelation" let coq_forall = find_global morphisms "forall_def" let subrelation = find_global relation_classes "subrelation" let do_subrelation = find_global morphisms "do_subrelation" let apply_subrelation = find_global morphisms "apply_subrelation" let rewrite_relation_class = find_global relation_classes "RewriteRelation" let proper_class = let r = lazy (find_reference morphisms "Proper") in fun env sigma -> class_info env sigma (Lazy.force r) let proper_proxy_class = let r = lazy (find_reference morphisms "ProperProxy") in fun env sigma -> class_info env sigma (Lazy.force r) let proper_proj env sigma = mkConst (Option.get (pi3 (List.hd (proper_class env sigma).cl_projs))) let proper_type env (sigma,cstrs) = let l = (proper_class env sigma).cl_impl in let (sigma, c) = Evarutil.new_global sigma l in (sigma, cstrs), c let proper_proxy_type env (sigma,cstrs) = let l = (proper_proxy_class env sigma).cl_impl in let (sigma, c) = Evarutil.new_global sigma l in (sigma, cstrs), c let proper_proof env evars carrier relation x = let evars, goal = app_poly env evars (proper_proxy_type env) [| carrier ; relation; x |] in new_cstr_evar evars env goal let get_reflexive_proof env = find_class_proof reflexive_type reflexive_proof env let get_symmetric_proof env = find_class_proof symmetric_type symmetric_proof env let get_transitive_proof env = find_class_proof transitive_type transitive_proof env let mk_relation env evd a = app_poly env evd relation [| a |] (** Build an inferred signature from constraints on the arguments and expected output relation *) let build_signature evars env m (cstrs : (types * types option) option list) (finalcstr : (types * types option) option) = let mk_relty evars newenv ty obj = match obj with | None | Some (_, None) -> let evars, relty = mk_relation env evars ty in if closed0 (goalevars evars) ty then let env' = Environ.reset_with_named_context (Environ.named_context_val env) env in new_cstr_evar evars env' relty else new_cstr_evar evars newenv relty | Some (x, Some rel) -> evars, rel in let rec aux env evars ty l = let t = Reductionops.whd_all env (goalevars evars) ty in match EConstr.kind (goalevars evars) t, l with | Prod (na, ty, b), obj :: cstrs -> let b = Reductionops.nf_betaiota env (goalevars evars) b in if noccurn (goalevars evars) 1 b (* non-dependent product *) then let ty = Reductionops.nf_betaiota env (goalevars evars) ty in let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in let evars, relty = mk_relty evars env ty obj in let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs else let (evars, b, arg, cstrs) = aux (push_rel (LocalAssum (na, ty)) env) evars b cstrs in let ty = Reductionops.nf_betaiota env (goalevars evars) ty in let pred = mkLambda (na, ty, b) in let liftarg = mkLambda (na, ty, arg) in let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in if Option.is_empty obj then evars, mkProd(na, ty, b), arg', (ty, None) :: cstrs else user_err Pp.(str "build_signature: no constraint can apply on a dependent argument") | _, obj :: _ -> anomaly ~label:"build_signature" (Pp.str "not enough products.") | _, [] -> (match finalcstr with | None | Some (_, None) -> let t = Reductionops.nf_betaiota env (fst evars) ty in let evars, rel = mk_relty evars env t None in evars, t, rel, [t, Some rel] | Some (t, Some rel) -> evars, t, rel, [t, Some rel]) in aux env evars m cstrs (** Folding/unfolding of the tactic constants. *) let unfold_impl sigma t = match EConstr.kind sigma t with | App (arrow, [| a; b |])(* when eq_constr arrow (Lazy.force impl) *) -> mkProd (make_annot Anonymous Sorts.Relevant, a, lift 1 b) | _ -> assert false let unfold_all sigma t = match EConstr.kind sigma t with | App (id, [| a; b |]) (* when eq_constr id (Lazy.force coq_all) *) -> (match EConstr.kind sigma b with | Lambda (n, ty, b) -> mkProd (n, ty, b) | _ -> assert false) | _ -> assert false let unfold_forall sigma t = match EConstr.kind sigma t with | App (id, [| a; b |]) (* when eq_constr id (Lazy.force coq_all) *) -> (match EConstr.kind sigma b with | Lambda (n, ty, b) -> mkProd (n, ty, b) | _ -> assert false) | _ -> assert false let arrow_morphism env evd ta tb a b = let ap = is_Prop (goalevars evd) ta and bp = is_Prop (goalevars evd) tb in if ap && bp then app_poly env evd impl [| a; b |], unfold_impl else if ap then (* Domain in Prop, CoDomain in Type *) (app_poly env evd arrow [| a; b |]), unfold_impl (* (evd, mkProd (Anonymous, a, b)), (fun x -> x) *) else if bp then (* Dummy forall *) (app_poly env evd coq_all [| a; mkLambda (make_annot Anonymous Sorts.Relevant, a, lift 1 b) |]), unfold_forall else (* None in Prop, use arrow *) (app_poly env evd arrow [| a; b |]), unfold_impl let rec decomp_pointwise sigma n c = if Int.equal n 0 then c else match EConstr.kind sigma c with | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> decomp_pointwise sigma (pred n) relb | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> decomp_pointwise sigma (pred n) (Reductionops.beta_applist sigma (arelb, [mkRel 1])) | _ -> invalid_arg "decomp_pointwise" let rec apply_pointwise sigma rel = function | arg :: args -> (match EConstr.kind sigma rel with | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> apply_pointwise sigma relb args | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> apply_pointwise sigma (Reductionops.beta_applist sigma (arelb, [arg])) args | _ -> invalid_arg "apply_pointwise") | [] -> rel let pointwise_or_dep_relation env evd n t car rel = if noccurn (goalevars evd) 1 car && noccurn (goalevars evd) 1 rel then app_poly env evd pointwise_relation [| t; lift (-1) car; lift (-1) rel |] else app_poly env evd forall_relation [| t; mkLambda (make_annot n Sorts.Relevant, t, car); mkLambda (make_annot n Sorts.Relevant, t, rel) |] let lift_cstr env evars (args : constr list) c ty cstr = let start evars env car = match cstr with | None | Some (_, None) -> let evars, rel = mk_relation env evars car in new_cstr_evar evars env rel | Some (ty, Some rel) -> evars, rel in let rec aux evars env prod n = if Int.equal n 0 then start evars env prod else let sigma = goalevars evars in match EConstr.kind sigma (Reductionops.whd_all env sigma prod) with | Prod (na, ty, b) -> if noccurn sigma 1 b then let b' = lift (-1) b in let evars, rb = aux evars env b' (pred n) in app_poly env evars pointwise_relation [| ty; b'; rb |] else let evars, rb = aux evars (push_rel (LocalAssum (na, ty)) env) b (pred n) in app_poly env evars forall_relation [| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |] | _ -> raise Not_found in let rec find env c ty = function | [] -> None | arg :: args -> try let evars, found = aux evars env ty (succ (List.length args)) in Some (evars, found, c, ty, arg :: args) with Not_found -> let sigma = goalevars evars in let ty = Reductionops.whd_all env sigma ty in find env (mkApp (c, [| arg |])) (prod_applist sigma ty [arg]) args in find env c ty args let unlift_cstr env sigma = function | None -> None | Some codom -> Some (decomp_pointwise (goalevars sigma) 1 codom) (** Looking up declared rewrite relations (instances of [RewriteRelation]) *) let is_applied_rewrite_relation env sigma rels t = match EConstr.kind sigma t with | App (c, args) when Array.length args >= 2 -> let head = if isApp sigma c then fst (destApp sigma c) else c in if Termops.is_global sigma (coq_eq_ref ()) head then None else (try let params, args = Array.chop (Array.length args - 2) args in let env' = push_rel_context rels env in let (evars, (evar, _)) = Evarutil.new_type_evar env' sigma Evd.univ_flexible in let evars, inst = app_poly env (evars,Evar.Set.empty) rewrite_relation_class [| evar; mkApp (c, params) |] in let _ = Typeclasses.resolve_one_typeclass env' (goalevars evars) inst in Some (it_mkProd_or_LetIn t rels) with e when CErrors.noncritical e -> None) | _ -> None end (* let my_type_of env evars c = Typing.e_type_of env evars c *) (* let mytypeofkey = CProfile.declare_profile "my_type_of";; *) (* let my_type_of = CProfile.profile3 mytypeofkey my_type_of *) let type_app_poly env env evd f args = let evars, c = app_poly_nocheck env evd f args in let evd', t = Typing.type_of env (goalevars evars) c in (evd', cstrevars evars), c module PropGlobal = struct module Consts = struct let relation_classes = ["Coq"; "Classes"; "RelationClasses"] let morphisms = ["Coq"; "Classes"; "Morphisms"] let relation = ["Coq"; "Relations";"Relation_Definitions"], "relation" let app_poly = app_poly_nocheck let arrow = find_global ["Coq"; "Program"; "Basics"] "arrow" let coq_inverse = find_global ["Coq"; "Program"; "Basics"] "flip" end module G = GlobalBindings(Consts) include G include Consts let inverse env evd car rel = type_app_poly env env evd coq_inverse [| car ; car; mkProp; rel |] (* app_poly env evd coq_inverse [| car ; car; mkProp; rel |] *) end module TypeGlobal = struct module Consts = struct let relation_classes = ["Coq"; "Classes"; "CRelationClasses"] let morphisms = ["Coq"; "Classes"; "CMorphisms"] let relation = relation_classes, "crelation" let app_poly = app_poly_check let arrow = find_global ["Coq"; "Classes"; "CRelationClasses"] "arrow" let coq_inverse = find_global ["Coq"; "Classes"; "CRelationClasses"] "flip" end module G = GlobalBindings(Consts) include G include Consts let inverse env (evd,cstrs) car rel = let (evd, sort) = Evarutil.new_Type ~rigid:Evd.univ_flexible evd in app_poly_check env (evd,cstrs) coq_inverse [| car ; car; sort; rel |] end let sort_of_rel env evm rel = ESorts.kind evm (Reductionops.sort_of_arity env evm (Retyping.get_type_of env evm rel)) let is_applied_rewrite_relation = PropGlobal.is_applied_rewrite_relation (* let _ = *) (* Hook.set Equality.is_applied_rewrite_relation is_applied_rewrite_relation *) let split_head = function hd :: tl -> hd, tl | [] -> assert(false) let eq_pb (ty, env, x, y as pb) (ty', env', x', y' as pb') = let equal x y = Constr.equal (EConstr.Unsafe.to_constr x) (EConstr.Unsafe.to_constr y) in pb == pb' || (ty == ty' && equal x x' && equal y y') let problem_inclusion x y = List.for_all (fun pb -> List.exists (fun pb' -> eq_pb pb pb') y) x let evd_convertible env evd x y = try (* Unfortunately, the_conv_x might say they are unifiable even if some unsolvable constraints remain, so we check that this unification does not introduce any new problem. *) let _, pbs = Evd.extract_all_conv_pbs evd in let evd' = Evarconv.unify_delay env evd x y in let _, pbs' = Evd.extract_all_conv_pbs evd' in if evd' == evd || problem_inclusion pbs' pbs then Some evd' else None with e when CErrors.noncritical e -> None let convertible env evd x y = Reductionops.is_conv_leq env evd x y type hypinfo = { prf : constr; car : constr; rel : constr; sort : bool; (* true = Prop; false = Type *) c1 : constr; c2 : constr; holes : Clenv.hole list; } let get_symmetric_proof b = if b then PropGlobal.get_symmetric_proof else TypeGlobal.get_symmetric_proof let error_no_relation () = user_err Pp.(str "Cannot find a relation to rewrite.") let rec decompose_app_rel env evd t = (* Head normalize for compatibility with the old meta mechanism *) let t = Reductionops.whd_betaiota evd t in match EConstr.kind evd t with | App (f, [||]) -> assert false | App (f, [|arg|]) -> let (f', argl, argr) = decompose_app_rel env evd arg in let ty = Typing.unsafe_type_of env evd argl in let r = Retyping.relevance_of_type env evd ty in let f'' = mkLambda (make_annot (Name default_dependent_ident) r, ty, mkLambda (make_annot (Name (Id.of_string "y")) r, lift 1 ty, mkApp (lift 2 f, [| mkApp (lift 2 f', [| mkRel 2; mkRel 1 |]) |]))) in (f'', argl, argr) | App (f, args) -> let len = Array.length args in let fargs = Array.sub args 0 (Array.length args - 2) in let rel = mkApp (f, fargs) in rel, args.(len - 2), args.(len - 1) | _ -> error_no_relation () let decompose_app_rel env evd t = let (rel, t1, t2) = decompose_app_rel env evd t in let ty = Retyping.get_type_of env evd rel in let () = if not (Reductionops.is_arity env evd ty) then error_no_relation () in (rel, t1, t2) let decompose_applied_relation env sigma (c,l) = let open Context.Rel.Declaration in let ctype = Retyping.get_type_of env sigma c in let find_rel ty = let sigma, cl = Clenv.make_evar_clause env sigma ty in let sigma = Clenv.solve_evar_clause env sigma true cl l in let { Clenv.cl_holes = holes; Clenv.cl_concl = t } = cl in let (equiv, c1, c2) = decompose_app_rel env sigma t in let ty1 = Retyping.get_type_of env sigma c1 in let ty2 = Retyping.get_type_of env sigma c2 in match evd_convertible env sigma ty1 ty2 with | None -> None | Some sigma -> let sort = sort_of_rel env sigma equiv in let args = Array.map_of_list (fun h -> h.Clenv.hole_evar) holes in let value = mkApp (c, args) in Some (sigma, { prf=value; car=ty1; rel = equiv; sort = Sorts.is_prop sort; c1=c1; c2=c2; holes }) in match find_rel ctype with | Some c -> c | None -> let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *) match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with | Some c -> c | None -> user_err Pp.(str "Cannot find an homogeneous relation to rewrite.") let rewrite_db = "rewrite" let conv_transparent_state = TransparentState.cst_full let rewrite_transparent_state () = Hints.Hint_db.transparent_state (Hints.searchtable_map rewrite_db) let rewrite_core_unif_flags = { Unification.modulo_conv_on_closed_terms = None; Unification.use_metas_eagerly_in_conv_on_closed_terms = true; Unification.use_evars_eagerly_in_conv_on_closed_terms = true; Unification.modulo_delta = TransparentState.empty; Unification.modulo_delta_types = TransparentState.full; Unification.check_applied_meta_types = true; Unification.use_pattern_unification = true; Unification.use_meta_bound_pattern_unification = true; Unification.allowed_evars = Unification.AllowAll; Unification.restrict_conv_on_strict_subterms = false; Unification.modulo_betaiota = false; Unification.modulo_eta = true; } (* Flags used for the setoid variant of "rewrite" and for the strategies "hints"/"old_hints"/"terms" of "rewrite_strat", and for solving pre-existing evars in "rewrite" (see unify_abs) *) let rewrite_unif_flags = let flags = rewrite_core_unif_flags in { Unification.core_unify_flags = flags; Unification.merge_unify_flags = flags; Unification.subterm_unify_flags = flags; Unification.allow_K_in_toplevel_higher_order_unification = true; Unification.resolve_evars = true } let rewrite_core_conv_unif_flags = { rewrite_core_unif_flags with Unification.modulo_conv_on_closed_terms = Some conv_transparent_state; Unification.modulo_delta_types = conv_transparent_state; Unification.modulo_betaiota = true } (* Fallback flags for the setoid variant of "rewrite" *) let rewrite_conv_unif_flags = let flags = rewrite_core_conv_unif_flags in { Unification.core_unify_flags = flags; Unification.merge_unify_flags = flags; Unification.subterm_unify_flags = flags; Unification.allow_K_in_toplevel_higher_order_unification = true; Unification.resolve_evars = true } (* Flags for "setoid_rewrite c"/"rewrite_strat -> c" *) let general_rewrite_unif_flags () = let ts = rewrite_transparent_state () in let core_flags = { rewrite_core_unif_flags with Unification.modulo_conv_on_closed_terms = Some ts; Unification.use_evars_eagerly_in_conv_on_closed_terms = true; Unification.modulo_delta = ts; Unification.modulo_delta_types = TransparentState.full; Unification.modulo_betaiota = true } in { Unification.core_unify_flags = core_flags; Unification.merge_unify_flags = core_flags; Unification.subterm_unify_flags = { core_flags with Unification.modulo_delta = TransparentState.empty }; Unification.allow_K_in_toplevel_higher_order_unification = true; Unification.resolve_evars = true } let refresh_hypinfo env sigma (is, cb) = let sigma, cbl = Tacinterp.interp_open_constr_with_bindings is env sigma cb in let sigma, hypinfo = decompose_applied_relation env sigma cbl in let { c1; c2; car; rel; prf; sort; holes } = hypinfo in sigma, (car, rel, prf, c1, c2, holes, sort) (** FIXME: write this in the new monad interface *) let solve_remaining_by env sigma holes by = match by with | None -> sigma | Some tac -> let map h = if h.Clenv.hole_deps then None else match EConstr.kind sigma h.Clenv.hole_evar with | Evar (evk, _) -> Some evk | _ -> None in (* Only solve independent holes *) let indep = List.map_filter map holes in let ist = { Geninterp.lfun = Id.Map.empty ; poly = false ; extra = Geninterp.TacStore.empty } in let solve_tac = match tac with | Genarg.GenArg (Genarg.Glbwit tag, tac) -> Ftactic.run (Geninterp.interp tag ist tac) (fun _ -> Proofview.tclUNIT ()) in let solve_tac = tclCOMPLETE solve_tac in let solve sigma evk = let evi = try Some (Evd.find_undefined sigma evk) with Not_found -> None in match evi with | None -> sigma (* Evar should not be defined, but just in case *) | Some evi -> let env = Environ.reset_with_named_context evi.evar_hyps env in let ty = evi.evar_concl in let name, poly = Id.of_string "rewrite", false in let c, sigma = Pfedit.refine_by_tactic ~name ~poly env sigma ty solve_tac in Evd.define evk (EConstr.of_constr c) sigma in List.fold_left solve sigma indep let no_constraints cstrs = fun ev _ -> not (Evar.Set.mem ev cstrs) let poly_inverse sort = if sort then PropGlobal.inverse else TypeGlobal.inverse type rewrite_proof = | RewPrf of constr * constr (** A Relation (R : rew_car -> rew_car -> Prop) and a proof of R rew_from rew_to *) | RewCast of cast_kind (** A proof of convertibility (with casts) *) type rewrite_result_info = { rew_car : constr ; (** A type *) rew_from : constr ; (** A term of type rew_car *) rew_to : constr ; (** A term of type rew_car *) rew_prf : rewrite_proof ; (** A proof of rew_from == rew_to *) rew_evars : evars; } type rewrite_result = | Fail | Identity | Success of rewrite_result_info type 'a strategy_input = { state : 'a ; (* a parameter: for instance, a state *) env : Environ.env ; unfresh : Id.Set.t; (* Unfresh names *) term1 : constr ; ty1 : types ; (* first term and its type (convertible to rew_from) *) cstr : (bool (* prop *) * constr option) ; evars : evars } type 'a pure_strategy = { strategy : 'a strategy_input -> 'a * rewrite_result (* the updated state and the "result" *) } type strategy = unit pure_strategy let symmetry env sort rew = let { rew_evars = evars; rew_car = car; } = rew in let (rew_evars, rew_prf) = match rew.rew_prf with | RewCast _ -> (rew.rew_evars, rew.rew_prf) | RewPrf (rel, prf) -> try let evars, symprf = get_symmetric_proof sort env evars car rel in let prf = mkApp (symprf, [| rew.rew_from ; rew.rew_to ; prf |]) in (evars, RewPrf (rel, prf)) with Not_found -> let evars, rel = poly_inverse sort env evars car rel in (evars, RewPrf (rel, prf)) in { rew with rew_from = rew.rew_to; rew_to = rew.rew_from; rew_prf; rew_evars; } (* Matching/unifying the rewriting rule against [t] *) let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs) by t = try let left = if l2r then c1 else c2 in let sigma = Unification.w_unify ~flags env sigma CONV left t in let sigma = Typeclasses.resolve_typeclasses ~filter:(no_constraints cstrs) ~fail:true env sigma in let evd = solve_remaining_by env sigma holes by in let nf c = Reductionops.nf_evar evd (Reductionops.nf_meta evd c) in let c1 = nf c1 and c2 = nf c2 and rew_car = nf car and rel = nf rel and prf = nf prf in let ty1 = Retyping.get_type_of env evd c1 in let ty2 = Retyping.get_type_of env evd c2 in let () = if not (convertible env evd ty2 ty1) then raise Reduction.NotConvertible in let rew_evars = evd, cstrs in let rew_prf = RewPrf (rel, prf) in let rew = { rew_evars; rew_prf; rew_car; rew_from = c1; rew_to = c2; } in let rew = if l2r then rew else symmetry env sort rew in Some rew with | e when Class_tactics.catchable e -> None | Reduction.NotConvertible -> None let unify_abs (car, rel, prf, c1, c2) l2r sort env (sigma, cstrs) t = try let left = if l2r then c1 else c2 in (* The pattern is already instantiated, so the next w_unify is basically an eq_constr, except when preexisting evars occur in either the lemma or the goal, in which case the eq_constr also solved this evars *) let sigma = Unification.w_unify ~flags:rewrite_unif_flags env sigma CONV left t in let rew_evars = sigma, cstrs in let rew_prf = RewPrf (rel, prf) in let rew = { rew_car = car; rew_from = c1; rew_to = c2; rew_prf; rew_evars; } in let rew = if l2r then rew else symmetry env sort rew in Some rew with | e when Class_tactics.catchable e -> None | Reduction.NotConvertible -> None type rewrite_flags = { under_lambdas : bool; on_morphisms : bool } let default_flags = { under_lambdas = true; on_morphisms = true; } let get_opt_rew_rel = function RewPrf (rel, prf) -> Some rel | _ -> None let new_global (evars, cstrs) gr = let (sigma,c) = Evarutil.new_global evars gr in (sigma, cstrs), c let make_eq sigma = new_global sigma Coqlib.(lib_ref "core.eq.type") let make_eq_refl sigma = new_global sigma Coqlib.(lib_ref "core.eq.refl") let get_rew_prf evars r = match r.rew_prf with | RewPrf (rel, prf) -> evars, (rel, prf) | RewCast c -> let evars, eq = make_eq evars in let evars, eq_refl = make_eq_refl evars in let rel = mkApp (eq, [| r.rew_car |]) in evars, (rel, mkCast (mkApp (eq_refl, [| r.rew_car; r.rew_from |]), c, mkApp (rel, [| r.rew_from; r.rew_to |]))) let poly_subrelation sort = if sort then PropGlobal.subrelation else TypeGlobal.subrelation let resolve_subrelation env avoid car rel sort prf rel' res = if Termops.eq_constr (fst res.rew_evars) rel rel' then res else let evars, app = app_poly_check env res.rew_evars (poly_subrelation sort) [|car; rel; rel'|] in let evars, subrel = new_cstr_evar evars env app in let appsub = mkApp (subrel, [| res.rew_from ; res.rew_to ; prf |]) in { res with rew_prf = RewPrf (rel', appsub); rew_evars = evars } let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) evars = let evars, morph_instance, proj, sigargs, m', args, args' = let first = match (Array.findi (fun _ b -> not (Option.is_empty b)) args') with | Some i -> i | None -> invalid_arg "resolve_morphism" in let morphargs, morphobjs = Array.chop first args in let morphargs', morphobjs' = Array.chop first args' in let appm = mkApp(m, morphargs) in let appmtype = Typing.unsafe_type_of env (goalevars evars) appm in let cstrs = List.map (Option.map (fun r -> r.rew_car, get_opt_rew_rel r.rew_prf)) (Array.to_list morphobjs') in (* Desired signature *) let evars, appmtype', signature, sigargs = if b then PropGlobal.build_signature evars env appmtype cstrs cstr else TypeGlobal.build_signature evars env appmtype cstrs cstr in (* Actual signature found *) let cl_args = [| appmtype' ; signature ; appm |] in let evars, app = app_poly_sort b env evars (if b then PropGlobal.proper_type env else TypeGlobal.proper_type env) cl_args in let env' = let dosub, appsub = if b then PropGlobal.do_subrelation, PropGlobal.apply_subrelation else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation in EConstr.push_named (LocalDef (make_annot (Id.of_string "do_subrelation") Sorts.Relevant, snd (app_poly_sort b env evars dosub [||]), snd (app_poly_nocheck env evars appsub [||]))) env in let evars, morph = new_cstr_evar evars env' app in evars, morph, morph, sigargs, appm, morphobjs, morphobjs' in let projargs, subst, evars, respars, typeargs = Array.fold_left2 (fun (acc, subst, evars, sigargs, typeargs') x y -> let (carrier, relation), sigargs = split_head sigargs in match relation with | Some relation -> let carrier = substl subst carrier and relation = substl subst relation in (match y with | None -> let evars, proof = (if b then PropGlobal.proper_proof else TypeGlobal.proper_proof) env evars carrier relation x in [ proof ; x ; x ] @ acc, subst, evars, sigargs, x :: typeargs' | Some r -> let evars, proof = get_rew_prf evars r in [ snd proof; r.rew_to; x ] @ acc, subst, evars, sigargs, r.rew_to :: typeargs') | None -> if not (Option.is_empty y) then user_err Pp.(str "Cannot rewrite inside dependent arguments of a function"); x :: acc, x :: subst, evars, sigargs, x :: typeargs') ([], [], evars, sigargs, []) args args' in let proof = applist (proj, List.rev projargs) in let newt = applist (m', List.rev typeargs) in match respars with [ a, Some r ] -> evars, proof, substl subst a, substl subst r, oldt, fnewt newt | _ -> assert(false) let apply_constraint env avoid car rel prf cstr res = match snd cstr with | None -> res | Some r -> resolve_subrelation env avoid car rel (fst cstr) prf r res let coerce env avoid cstr res = let evars, (rel, prf) = get_rew_prf res.rew_evars res in let res = { res with rew_evars = evars } in apply_constraint env avoid res.rew_car rel prf cstr res let apply_rule unify loccs : int pure_strategy = let (nowhere_except_in,occs) = convert_occs loccs in let is_occ occ = if nowhere_except_in then List.mem occ occs else not (List.mem occ occs) in { strategy = fun { state = occ ; env ; unfresh ; term1 = t ; ty1 = ty ; cstr ; evars } -> let unif = if isEvar (goalevars evars) t then None else unify env evars t in match unif with | None -> (occ, Fail) | Some rew -> let occ = succ occ in if not (is_occ occ) then (occ, Fail) else if Termops.eq_constr (fst rew.rew_evars) t rew.rew_to then (occ, Identity) else let res = { rew with rew_car = ty } in let res = Success (coerce env unfresh cstr res) in (occ, res) } let apply_lemma l2r flags oc by loccs : strategy = { strategy = fun ({ state = () ; env ; term1 = t ; evars = (sigma, cstrs) } as input) -> let sigma, c = oc sigma in let sigma, hypinfo = decompose_applied_relation env sigma c in let { c1; c2; car; rel; prf; sort; holes } = hypinfo in let rew = (car, rel, prf, c1, c2, holes, sort) in let evars = (sigma, cstrs) in let unify env evars t = let rew = unify_eqn rew l2r flags env evars by t in match rew with | None -> None | Some rew -> Some rew in let _, res = (apply_rule unify loccs).strategy { input with state = 0 ; evars } in (), res } let e_app_poly env evars f args = let evars', c = app_poly_nocheck env !evars f args in evars := evars'; c let make_leibniz_proof env c ty r = let evars = ref r.rew_evars in let prf = match r.rew_prf with | RewPrf (rel, prf) -> let rel = e_app_poly env evars coq_eq [| ty |] in let prf = e_app_poly env evars coq_f_equal [| r.rew_car; ty; mkLambda (make_annot Anonymous Sorts.Relevant, r.rew_car, c); r.rew_from; r.rew_to; prf |] in RewPrf (rel, prf) | RewCast k -> r.rew_prf in { rew_car = ty; rew_evars = !evars; rew_from = subst1 r.rew_from c; rew_to = subst1 r.rew_to c; rew_prf = prf } let reset_env env = let env' = Global.env_of_context (Environ.named_context_val env) in Environ.push_rel_context (Environ.rel_context env) env' let fold_match ?(force=false) env sigma c = let (ci, p, c, brs) = destCase sigma c in let cty = Retyping.get_type_of env sigma c in let dep, pred, exists, (sk,eff) = let env', ctx, body = let ctx, pred = decompose_lam_assum sigma p in let env' = push_rel_context ctx env in env', ctx, pred in let sortp = Retyping.get_sort_family_of env' sigma body in let sortc = Retyping.get_sort_family_of env sigma cty in let dep = not (noccurn sigma 1 body) in let pred = if dep then p else it_mkProd_or_LetIn (subst1 mkProp body) (List.tl ctx) in let sk = if sortp == Sorts.InProp then if sortc == Sorts.InProp then if dep then case_dep_scheme_kind_from_prop else case_scheme_kind_from_prop else ( if dep then case_dep_scheme_kind_from_type_in_prop else case_scheme_kind_from_type) else ((* sortc <> InProp by typing *) if dep then case_dep_scheme_kind_from_type else case_scheme_kind_from_type) in let exists = Ind_tables.check_scheme sk ci.ci_ind in if exists || force then dep, pred, exists, Ind_tables.find_scheme sk ci.ci_ind else raise Not_found in let app = let ind, args = Inductiveops.find_mrectype env sigma cty in let pars, args = List.chop ci.ci_npar args in let meths = List.map (fun br -> br) (Array.to_list brs) in applist (mkConst sk, pars @ [pred] @ meths @ args @ [c]) in sk, (if exists then env else reset_env env), app, eff let unfold_match env sigma sk app = match EConstr.kind sigma app with | App (f', args) when Constant.equal (fst (destConst sigma f')) sk -> let v = Environ.constant_value_in (Global.env ()) (sk,Univ.Instance.empty)(*FIXME*) in let v = EConstr.of_constr v in Reductionops.whd_beta sigma (mkApp (v, args)) | _ -> app let is_rew_cast = function RewCast _ -> true | _ -> false let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let rec aux { state ; env ; unfresh ; term1 = t ; ty1 = ty ; cstr = (prop, cstr) ; evars } = let cstr' = Option.map (fun c -> (ty, Some c)) cstr in match EConstr.kind (goalevars evars) t with | App (m, args) -> let rewrite_args state success = let state, (args', evars', progress) = Array.fold_left (fun (state, (acc, evars, progress)) arg -> if not (Option.is_empty progress) && not all then state, (None :: acc, evars, progress) else let argty = Retyping.get_type_of env (goalevars evars) arg in let state, res = s.strategy { state ; env ; unfresh ; term1 = arg ; ty1 = argty ; cstr = (prop,None) ; evars } in let res' = match res with | Identity -> let progress = if Option.is_empty progress then Some false else progress in (None :: acc, evars, progress) | Success r -> (Some r :: acc, r.rew_evars, Some true) | Fail -> (None :: acc, evars, progress) in state, res') (state, ([], evars, success)) args in let res = match progress with | None -> Fail | Some false -> Identity | Some true -> let args' = Array.of_list (List.rev args') in if Array.exists (function | None -> false | Some r -> not (is_rew_cast r.rew_prf)) args' then let evars', prf, car, rel, c1, c2 = resolve_morphism env unfresh t m args args' (prop, cstr') evars' in let res = { rew_car = ty; rew_from = c1; rew_to = c2; rew_prf = RewPrf (rel, prf); rew_evars = evars' } in Success res else let args' = Array.map2 (fun aorig anew -> match anew with None -> aorig | Some r -> r.rew_to) args args' in let res = { rew_car = ty; rew_from = t; rew_to = mkApp (m, args'); rew_prf = RewCast DEFAULTcast; rew_evars = evars' } in Success res in state, res in if flags.on_morphisms then let mty = Retyping.get_type_of env (goalevars evars) m in let evars, cstr', m, mty, argsl, args = let argsl = Array.to_list args in let lift = if prop then PropGlobal.lift_cstr else TypeGlobal.lift_cstr in match lift env evars argsl m mty None with | Some (evars, cstr', m, mty, args) -> evars, Some cstr', m, mty, args, Array.of_list args | None -> evars, None, m, mty, argsl, args in let state, m' = s.strategy { state ; env ; unfresh ; term1 = m ; ty1 = mty ; cstr = (prop, cstr') ; evars } in match m' with | Fail -> rewrite_args state None (* Standard path, try rewrite on arguments *) | Identity -> rewrite_args state (Some false) | Success r -> (* We rewrote the function and get a proof of pointwise rel for the arguments. We just apply it. *) let prf = match r.rew_prf with | RewPrf (rel, prf) -> let app = if prop then PropGlobal.apply_pointwise else TypeGlobal.apply_pointwise in RewPrf (app (goalevars evars) rel argsl, mkApp (prf, args)) | x -> x in let res = { rew_car = Reductionops.hnf_prod_appvect env (goalevars evars) r.rew_car args; rew_from = mkApp(r.rew_from, args); rew_to = mkApp(r.rew_to, args); rew_prf = prf; rew_evars = r.rew_evars } in let res = match prf with | RewPrf (rel, prf) -> Success (apply_constraint env unfresh res.rew_car rel prf (prop,cstr) res) | _ -> Success res in state, res else rewrite_args state None | Prod (n, x, b) when noccurn (goalevars evars) 1 b -> let b = subst1 mkProp b in let tx = Retyping.get_type_of env (goalevars evars) x and tb = Retyping.get_type_of env (goalevars evars) b in let arr = if prop then PropGlobal.arrow_morphism else TypeGlobal.arrow_morphism in let (evars', mor), unfold = arr env evars tx tb x b in let state, res = aux { state ; env ; unfresh ; term1 = mor ; ty1 = ty ; cstr = (prop,cstr) ; evars = evars' } in let res = match res with | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } | Fail | Identity -> res in state, res (* if x' = None && flags.under_lambdas then *) (* let lam = mkLambda (n, x, b) in *) (* let lam', occ = aux env lam occ None in *) (* let res = *) (* match lam' with *) (* | None -> None *) (* | Some (prf, (car, rel, c1, c2)) -> *) (* Some (resolve_morphism env sigma t *) (* ~fnewt:unfold_all *) (* (Lazy.force coq_all) [| x ; lam |] [| None; lam' |] *) (* cstr evars) *) (* in res, occ *) (* else *) | Prod (n, dom, codom) -> let lam = mkLambda (n, dom, codom) in let (evars', app), unfold = if eq_constr (fst evars) ty mkProp then (app_poly_sort prop env evars coq_all [| dom; lam |]), TypeGlobal.unfold_all else let forall = if prop then PropGlobal.coq_forall else TypeGlobal.coq_forall in (app_poly_sort prop env evars forall [| dom; lam |]), TypeGlobal.unfold_forall in let state, res = aux { state ; env ; unfresh ; term1 = app ; ty1 = ty ; cstr = (prop,cstr) ; evars = evars' } in let res = match res with | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } | Fail | Identity -> res in state, res (* TODO: real rewriting under binders: introduce x x' (H : R x x') and rewrite with H at any occurrence of x. Ask for (R ==> R') for the lambda. Formalize this. B. Barras' idea is to have a context of relations, of length 1, with Σ for gluing dependent relations and using projections to get them out. *) (* | Lambda (n, t, b) when flags.under_lambdas -> *) (* let n' = name_app (fun id -> Tactics.fresh_id_in_env avoid id env) n in *) (* let n'' = name_app (fun id -> Tactics.fresh_id_in_env avoid id env) n' in *) (* let n''' = name_app (fun id -> Tactics.fresh_id_in_env avoid id env) n'' in *) (* let rel = new_cstr_evar cstr env (mkApp (Lazy.force coq_relation, [|t|])) in *) (* let env' = Environ.push_rel_context [(n'',None,lift 2 rel);(n'',None,lift 1 t);(n', None, t)] env in *) (* let b' = s env' avoid b (Typing.type_of env' (goalevars evars) (lift 2 b)) (unlift_cstr env (goalevars evars) cstr) evars in *) (* (match b' with *) (* | Some (Some r) -> *) (* let prf = match r.rew_prf with *) (* | RewPrf (rel, prf) -> *) (* let rel = pointwise_or_dep_relation n' t r.rew_car rel in *) (* let prf = mkLambda (n', t, prf) in *) (* RewPrf (rel, prf) *) (* | x -> x *) (* in *) (* Some (Some { r with *) (* rew_prf = prf; *) (* rew_car = mkProd (n, t, r.rew_car); *) (* rew_from = mkLambda(n, t, r.rew_from); *) (* rew_to = mkLambda (n, t, r.rew_to) }) *) (* | _ -> b') *) | Lambda (n, t, b) when flags.under_lambdas -> let n' = map_annot (Nameops.Name.map (fun id -> Tactics.fresh_id_in_env unfresh id env)) n in let open Context.Rel.Declaration in let env' = EConstr.push_rel (LocalAssum (n', t)) env in let bty = Retyping.get_type_of env' (goalevars evars) b in let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in let state, b' = s.strategy { state ; env = env' ; unfresh ; term1 = b ; ty1 = bty ; cstr = (prop, unlift env evars cstr) ; evars } in let res = match b' with | Success r -> let r = match r.rew_prf with | RewPrf (rel, prf) -> let point = if prop then PropGlobal.pointwise_or_dep_relation else TypeGlobal.pointwise_or_dep_relation in let evars, rel = point env r.rew_evars n'.binder_name t r.rew_car rel in let prf = mkLambda (n', t, prf) in { r with rew_prf = RewPrf (rel, prf); rew_evars = evars } | x -> r in Success { r with rew_car = mkProd (n, t, r.rew_car); rew_from = mkLambda(n, t, r.rew_from); rew_to = mkLambda (n, t, r.rew_to) } | Fail | Identity -> b' in state, res | Case (ci, p, c, brs) -> let cty = Retyping.get_type_of env (goalevars evars) c in let evars', eqty = app_poly_sort prop env evars coq_eq [| cty |] in let cstr' = Some eqty in let state, c' = s.strategy { state ; env ; unfresh ; term1 = c ; ty1 = cty ; cstr = (prop, cstr') ; evars = evars' } in let state, res = match c' with | Success r -> let case = mkCase (ci, lift 1 p, mkRel 1, Array.map (lift 1) brs) in let res = make_leibniz_proof env case ty r in state, Success (coerce env unfresh (prop,cstr) res) | Fail | Identity -> if Array.for_all (Int.equal 0) ci.ci_cstr_ndecls then let evars', eqty = app_poly_sort prop env evars coq_eq [| ty |] in let cstr = Some eqty in let state, found, brs' = Array.fold_left (fun (state, found, acc) br -> if not (Option.is_empty found) then (state, found, fun x -> lift 1 br :: acc x) else let state, res = s.strategy { state ; env ; unfresh ; term1 = br ; ty1 = ty ; cstr = (prop,cstr) ; evars } in match res with | Success r -> (state, Some r, fun x -> mkRel 1 :: acc x) | Fail | Identity -> (state, None, fun x -> lift 1 br :: acc x)) (state, None, fun x -> []) brs in match found with | Some r -> let ctxc = mkCase (ci, lift 1 p, lift 1 c, Array.of_list (List.rev (brs' c'))) in state, Success (make_leibniz_proof env ctxc ty r) | None -> state, c' else match try Some (fold_match env (goalevars evars) t) with Not_found -> None with | None -> state, c' | Some (cst, _, t', eff (*FIXME*)) -> let state, res = aux { state ; env ; unfresh ; term1 = t' ; ty1 = ty ; cstr = (prop,cstr) ; evars } in let res = match res with | Success prf -> Success { prf with rew_from = t; rew_to = unfold_match env (goalevars evars) cst prf.rew_to } | x' -> c' in state, res in let res = match res with | Success r -> Success (coerce env unfresh (prop,cstr) r) | Fail | Identity -> res in state, res | _ -> state, Fail in { strategy = aux } let all_subterms = subterm true default_flags let one_subterm = subterm false default_flags (** Requires transitivity of the rewrite step, if not a reduction. Not tail-recursive. *) let transitivity state env unfresh prop (res : rewrite_result_info) (next : 'a pure_strategy) : 'a * rewrite_result = let state, nextres = next.strategy { state ; env ; unfresh ; term1 = res.rew_to ; ty1 = res.rew_car ; cstr = (prop, get_opt_rew_rel res.rew_prf) ; evars = res.rew_evars } in let res = match nextres with | Fail -> Fail | Identity -> Success res | Success res' -> match res.rew_prf with | RewCast c -> Success { res' with rew_from = res.rew_from } | RewPrf (rew_rel, rew_prf) -> match res'.rew_prf with | RewCast _ -> Success { res with rew_to = res'.rew_to } | RewPrf (res'_rel, res'_prf) -> let trans = if prop then PropGlobal.transitive_type else TypeGlobal.transitive_type in let evars, prfty = app_poly_sort prop env res'.rew_evars trans [| res.rew_car; rew_rel |] in let evars, prf = new_cstr_evar evars env prfty in let prf = mkApp (prf, [|res.rew_from; res'.rew_from; res'.rew_to; rew_prf; res'_prf |]) in Success { res' with rew_from = res.rew_from; rew_evars = evars; rew_prf = RewPrf (res'_rel, prf) } in state, res (** Rewriting strategies. Inspired by ELAN's rewriting strategies: http://citeseerx.ist.psu.edu/viewdoc/summary?doi=10.1.1.21.4049 *) module Strategies = struct let fail : 'a pure_strategy = { strategy = fun { state } -> state, Fail } let id : 'a pure_strategy = { strategy = fun { state } -> state, Identity } let refl : 'a pure_strategy = { strategy = fun { state ; env ; term1 = t ; ty1 = ty ; cstr = (prop,cstr) ; evars } -> let evars, rel = match cstr with | None -> let mkr = if prop then PropGlobal.mk_relation else TypeGlobal.mk_relation in let evars, rty = mkr env evars ty in new_cstr_evar evars env rty | Some r -> evars, r in let evars, proof = let proxy = if prop then PropGlobal.proper_proxy_type env else TypeGlobal.proper_proxy_type env in let evars, mty = app_poly_sort prop env evars proxy [| ty ; rel; t |] in new_cstr_evar evars env mty in let res = Success { rew_car = ty; rew_from = t; rew_to = t; rew_prf = RewPrf (rel, proof); rew_evars = evars } in state, res } let progress (s : 'a pure_strategy) : 'a pure_strategy = { strategy = fun input -> let state, res = s.strategy input in match res with | Fail -> state, Fail | Identity -> state, Fail | Success r -> state, Success r } let seq first snd : 'a pure_strategy = { strategy = fun ({ env ; unfresh ; cstr } as input) -> let state, res = first.strategy input in match res with | Fail -> state, Fail | Identity -> snd.strategy { input with state } | Success res -> transitivity state env unfresh (fst cstr) res snd } let choice fst snd : 'a pure_strategy = { strategy = fun input -> let state, res = fst.strategy input in match res with | Fail -> snd.strategy { input with state } | Identity | Success _ -> state, res } let try_ str : 'a pure_strategy = choice str id let check_interrupt str input = Control.check_for_interrupt (); str input let fix (f : 'a pure_strategy -> 'a pure_strategy) : 'a pure_strategy = let rec aux input = (f { strategy = fun input -> check_interrupt aux input }).strategy input in { strategy = aux } let any (s : 'a pure_strategy) : 'a pure_strategy = fix (fun any -> try_ (seq s any)) let repeat (s : 'a pure_strategy) : 'a pure_strategy = seq s (any s) let bu (s : 'a pure_strategy) : 'a pure_strategy = fix (fun s' -> seq (choice (progress (all_subterms s')) s) (try_ s')) let td (s : 'a pure_strategy) : 'a pure_strategy = fix (fun s' -> seq (choice s (progress (all_subterms s'))) (try_ s')) let innermost (s : 'a pure_strategy) : 'a pure_strategy = fix (fun ins -> choice (one_subterm ins) s) let outermost (s : 'a pure_strategy) : 'a pure_strategy = fix (fun out -> choice s (one_subterm out)) let lemmas cs : 'a pure_strategy = List.fold_left (fun tac (l,l2r,by) -> choice tac (apply_lemma l2r rewrite_unif_flags l by AllOccurrences)) fail cs let inj_open hint = (); fun sigma -> let ctx = UState.of_context_set hint.Autorewrite.rew_ctx in let sigma = Evd.merge_universe_context sigma ctx in (sigma, (EConstr.of_constr hint.Autorewrite.rew_lemma, NoBindings)) let old_hints (db : string) : 'a pure_strategy = let rules = Autorewrite.find_rewrites db in lemmas (List.map (fun hint -> (inj_open hint, hint.Autorewrite.rew_l2r, hint.Autorewrite.rew_tac)) rules) let hints (db : string) : 'a pure_strategy = { strategy = fun ({ term1 = t } as input) -> let t = EConstr.Unsafe.to_constr t in let rules = Autorewrite.find_matches db t in let lemma hint = (inj_open hint, hint.Autorewrite.rew_l2r, hint.Autorewrite.rew_tac) in let lems = List.map lemma rules in (lemmas lems).strategy input } let reduce (r : Redexpr.red_expr) : 'a pure_strategy = { strategy = fun { state = state ; env = env ; term1 = t ; ty1 = ty ; cstr = cstr ; evars = evars } -> let rfn, ckind = Redexpr.reduction_of_red_expr env r in let sigma = goalevars evars in let (sigma, t') = rfn env sigma t in if Termops.eq_constr sigma t' t then state, Identity else state, Success { rew_car = ty; rew_from = t; rew_to = t'; rew_prf = RewCast ckind; rew_evars = sigma, cstrevars evars } } let fold_glob c : 'a pure_strategy = { strategy = fun { state ; env ; term1 = t ; ty1 = ty ; cstr ; evars } -> (* let sigma, (c,_) = Tacinterp.interp_open_constr_with_bindings is env (goalevars evars) c in *) let sigma, c = Pretyping.understand_tcc env (goalevars evars) c in let unfolded = try Tacred.try_red_product env sigma c with e when CErrors.noncritical e -> user_err Pp.(str "fold: the term is not unfoldable!") in try let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in let c' = Reductionops.nf_evar sigma c in state, Success { rew_car = ty; rew_from = t; rew_to = c'; rew_prf = RewCast DEFAULTcast; rew_evars = (sigma, snd evars) } with e when CErrors.noncritical e -> state, Fail } end (** The strategy for a single rewrite, dealing with occurrences. *) (** A dummy initial clauseenv to avoid generating initial evars before even finding a first application of the rewriting lemma, in setoid_rewrite mode *) let rewrite_with l2r flags c occs : strategy = { strategy = fun ({ state = () } as input) -> let unify env evars t = let (sigma, cstrs) = evars in let (sigma, rew) = refresh_hypinfo env sigma c in unify_eqn rew l2r flags env (sigma, cstrs) None t in let app = apply_rule unify occs in let strat = Strategies.fix (fun aux -> Strategies.choice app (subterm true default_flags aux)) in let _, res = strat.strategy { input with state = 0 } in ((), res) } let apply_strategy (s : strategy) env unfresh concl (prop, cstr) evars = let ty = Retyping.get_type_of env (goalevars evars) concl in let _, res = s.strategy { state = () ; env ; unfresh ; term1 = concl ; ty1 = ty ; cstr = (prop, Some cstr) ; evars } in res let solve_constraints env (evars,cstrs) = let oldtcs = Evd.get_typeclass_evars evars in let evars' = Evd.set_typeclass_evars evars cstrs in let evars' = Typeclasses.resolve_typeclasses env ~filter:all_evars ~split:false ~fail:true evars' in Evd.set_typeclass_evars evars' oldtcs let nf_zeta = Reductionops.clos_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) exception RewriteFailure of Pp.t type result = (evar_map * constr option * types) option option let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : result = let sigma, sort = Typing.sort_of env sigma concl in let evdref = ref sigma in let evars = (!evdref, Evar.Set.empty) in let evars, cstr = let prop, (evars, arrow) = if Sorts.is_prop sort then true, app_poly_sort true env evars impl [||] else false, app_poly_sort false env evars TypeGlobal.arrow [||] in match is_hyp with | None -> let evars, t = poly_inverse prop env evars (mkSort sort) arrow in evars, (prop, t) | Some _ -> evars, (prop, arrow) in let eq = apply_strategy strat env avoid concl cstr evars in match eq with | Fail -> None | Identity -> Some None | Success res -> let (_, cstrs) = res.rew_evars in let evars' = solve_constraints env res.rew_evars in let newt = Reductionops.nf_evar evars' res.rew_to in let evars = (* Keep only original evars (potentially instantiated) and goal evars, the rest has been defined and substituted already. *) Evar.Set.fold (fun ev acc -> if not (Evd.is_defined acc ev) then user_err ~hdr:"rewrite" (str "Unsolved constraint remaining: " ++ spc () ++ Termops.pr_evar_info env acc (Evd.find acc ev)) else Evd.remove acc ev) cstrs evars' in let res = match res.rew_prf with | RewCast c -> None | RewPrf (rel, p) -> let p = nf_zeta env evars' (Reductionops.nf_evar evars' p) in let term = match abs with | None -> p | Some (t, ty) -> let t = Reductionops.nf_evar evars' t in let ty = Reductionops.nf_evar evars' ty in mkApp (mkLambda (make_annot (Name (Id.of_string "lemma")) Sorts.Relevant, ty, p), [| t |]) in let proof = match is_hyp with | None -> term | Some id -> mkApp (term, [| mkVar id |]) in Some proof in Some (Some (evars, res, newt)) (** Insert a declaration after the last declaration it depends on *) let rec insert_dependent env sigma decl accu hyps = match hyps with | [] -> List.rev_append accu [decl] | ndecl :: rem -> if occur_var_in_decl env sigma (NamedDecl.get_id ndecl) decl then List.rev_append accu (decl :: hyps) else insert_dependent env sigma decl (ndecl :: accu) rem let assert_replacing id newt tac = let prf = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ctx = named_context env in let after, before = List.split_when (NamedDecl.get_id %> Id.equal id) ctx in let nc = match before with | [] -> assert false | d :: rem -> insert_dependent env sigma (LocalAssum (make_annot (NamedDecl.get_id d) Sorts.Relevant, newt)) [] after @ rem in let env' = Environ.reset_with_named_context (val_of_named_context nc) env in Refine.refine ~typecheck:true begin fun sigma -> let (sigma, ev) = Evarutil.new_evar env' sigma concl in let (sigma, ev') = Evarutil.new_evar env sigma newt in let map d = let n = NamedDecl.get_id d in if Id.equal n id then ev' else mkVar n in let (e, _) = destEvar sigma ev in (sigma, mkEvar (e, Array.map_of_list map nc)) end end in Proofview.tclTHEN prf (Proofview.tclFOCUS 2 2 tac) let newfail n s = Proofview.tclZERO (Refiner.FailError (n, lazy s)) let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = let open Proofview.Notations in (* For compatibility *) let beta = Tactics.reduct_in_concl ~check:false (Reductionops.nf_betaiota, DEFAULTcast) in let beta_hyp id = Tactics.reduct_in_hyp ~check:false ~reorder:false Reductionops.nf_betaiota (id, InHyp) in let treat sigma res = match res with | None -> newfail 0 (str "Nothing to rewrite") | Some None -> if progress then newfail 0 (str"Failed to progress") else Proofview.tclUNIT () | Some (Some res) -> let (undef, prf, newt) = res in let fold ev _ accu = if Evd.mem sigma ev then accu else ev :: accu in let gls = List.rev (Evd.fold_undefined fold undef []) in let gls = List.map Proofview.with_empty_state gls in match clause, prf with | Some id, Some p -> let tac = tclTHENLIST [ Refine.refine ~typecheck:true (fun h -> (h,p)); Proofview.Unsafe.tclNEWGOALS gls; ] in Proofview.Unsafe.tclEVARS undef <*> tclTHENFIRST (assert_replacing id newt tac) (beta_hyp id) | Some id, None -> Proofview.Unsafe.tclEVARS undef <*> convert_hyp ~check:false ~reorder:false (LocalAssum (make_annot id Sorts.Relevant, newt)) <*> beta_hyp id | None, Some p -> Proofview.Unsafe.tclEVARS undef <*> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let make = begin fun sigma -> let (sigma, ev) = Evarutil.new_evar env sigma newt in (sigma, mkApp (p, [| ev |])) end in Refine.refine ~typecheck:true make <*> Proofview.Unsafe.tclNEWGOALS gls end | None, None -> Proofview.Unsafe.tclEVARS undef <*> convert_concl ~check:false newt DEFAULTcast in Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ty = match clause with | None -> concl | Some id -> EConstr.of_constr (Environ.named_type id env) in let env = match clause with | None -> env | Some id -> (* Only consider variables not depending on [id] *) let ctx = named_context env in let filter decl = not (occur_var_in_decl env sigma id decl) in let nctx = List.filter filter ctx in Environ.reset_with_named_context (val_of_named_context nctx) env in try let res = cl_rewrite_clause_aux ?abs strat env Id.Set.empty sigma ty clause in let sigma = match origsigma with None -> sigma | Some sigma -> sigma in treat sigma res <*> (* For compatibility *) beta <*> Proofview.shelve_unifiable with | PretypeError (env, evd, (UnsatisfiableConstraints _ as e)) -> raise (RewriteFailure (Himsg.explain_pretype_error env evd e)) end let tactic_init_setoid () = try init_setoid (); Proofview.tclUNIT () with e when CErrors.noncritical e -> Tacticals.New.tclFAIL 0 (str"Setoid library not loaded") let cl_rewrite_clause_strat progress strat clause = tactic_init_setoid () <*> (if progress then Proofview.tclPROGRESS else fun x -> x) (Proofview.tclOR (cl_rewrite_clause_newtac ~progress strat clause) (fun (e, info) -> match e with | RewriteFailure e -> tclZEROMSG (str"setoid rewrite failed: " ++ e) | Refiner.FailError (n, pp) -> tclFAIL n (str"setoid rewrite failed: " ++ Lazy.force pp) | e -> Proofview.tclZERO ~info e)) (** Setoid rewriting when called with "setoid_rewrite" *) let cl_rewrite_clause l left2right occs clause = let strat = rewrite_with left2right (general_rewrite_unif_flags ()) l occs in cl_rewrite_clause_strat true strat clause (** Setoid rewriting when called with "rewrite_strat" *) let cl_rewrite_clause_strat strat clause = cl_rewrite_clause_strat false strat clause let apply_glob_constr c l2r occs = (); fun ({ state = () ; env = env } as input) -> let c sigma = let (sigma, c) = Pretyping.understand_tcc env sigma c in (sigma, (c, NoBindings)) in let flags = general_rewrite_unif_flags () in (apply_lemma l2r flags c None occs).strategy input let interp_glob_constr_list env = let make c = (); fun sigma -> let sigma, c = Pretyping.understand_tcc env sigma c in (sigma, (c, NoBindings)) in List.map (fun c -> make c, true, None) (* Syntax for rewriting with strategies *) type unary_strategy = Subterms | Subterm | Innermost | Outermost | Bottomup | Topdown | Progress | Try | Any | Repeat type binary_strategy = | Compose | Choice type ('constr,'redexpr) strategy_ast = | StratId | StratFail | StratRefl | StratUnary of unary_strategy * ('constr,'redexpr) strategy_ast | StratBinary of binary_strategy * ('constr,'redexpr) strategy_ast * ('constr,'redexpr) strategy_ast | StratConstr of 'constr * bool | StratTerms of 'constr list | StratHints of bool * string | StratEval of 'redexpr | StratFold of 'constr let rec map_strategy (f : 'a -> 'a2) (g : 'b -> 'b2) : ('a,'b) strategy_ast -> ('a2,'b2) strategy_ast = function | StratId | StratFail | StratRefl as s -> s | StratUnary (s, str) -> StratUnary (s, map_strategy f g str) | StratBinary (s, str, str') -> StratBinary (s, map_strategy f g str, map_strategy f g str') | StratConstr (c, b) -> StratConstr (f c, b) | StratTerms l -> StratTerms (List.map f l) | StratHints (b, id) -> StratHints (b, id) | StratEval r -> StratEval (g r) | StratFold c -> StratFold (f c) let pr_ustrategy = function | Subterms -> str "subterms" | Subterm -> str "subterm" | Innermost -> str "innermost" | Outermost -> str "outermost" | Bottomup -> str "bottomup" | Topdown -> str "topdown" | Progress -> str "progress" | Try -> str "try" | Any -> str "any" | Repeat -> str "repeat" let paren p = str "(" ++ p ++ str ")" let rec pr_strategy prc prr = function | StratId -> str "id" | StratFail -> str "fail" | StratRefl -> str "refl" | StratUnary (s, str) -> pr_ustrategy s ++ spc () ++ paren (pr_strategy prc prr str) | StratBinary (Choice, str1, str2) -> str "choice" ++ spc () ++ paren (pr_strategy prc prr str1) ++ spc () ++ paren (pr_strategy prc prr str2) | StratBinary (Compose, str1, str2) -> pr_strategy prc prr str1 ++ str ";" ++ spc () ++ pr_strategy prc prr str2 | StratConstr (c, true) -> prc c | StratConstr (c, false) -> str "<-" ++ spc () ++ prc c | StratTerms cl -> str "terms" ++ spc () ++ pr_sequence prc cl | StratHints (old, id) -> let cmd = if old then "old_hints" else "hints" in str cmd ++ spc () ++ str id | StratEval r -> str "eval" ++ spc () ++ prr r | StratFold c -> str "fold" ++ spc () ++ prc c let rec strategy_of_ast = function | StratId -> Strategies.id | StratFail -> Strategies.fail | StratRefl -> Strategies.refl | StratUnary (f, s) -> let s' = strategy_of_ast s in let f' = match f with | Subterms -> all_subterms | Subterm -> one_subterm | Innermost -> Strategies.innermost | Outermost -> Strategies.outermost | Bottomup -> Strategies.bu | Topdown -> Strategies.td | Progress -> Strategies.progress | Try -> Strategies.try_ | Any -> Strategies.any | Repeat -> Strategies.repeat in f' s' | StratBinary (f, s, t) -> let s' = strategy_of_ast s in let t' = strategy_of_ast t in let f' = match f with | Compose -> Strategies.seq | Choice -> Strategies.choice in f' s' t' | StratConstr (c, b) -> { strategy = apply_glob_constr (fst c) b AllOccurrences } | StratHints (old, id) -> if old then Strategies.old_hints id else Strategies.hints id | StratTerms l -> { strategy = (fun ({ state = () ; env } as input) -> let l' = interp_glob_constr_list env (List.map fst l) in (Strategies.lemmas l').strategy input) } | StratEval r -> { strategy = (fun ({ state = () ; env ; evars } as input) -> let (sigma,r_interp) = Tacinterp.interp_redexp env (goalevars evars) r in (Strategies.reduce r_interp).strategy { input with evars = (sigma,cstrevars evars) }) } | StratFold c -> Strategies.fold_glob (fst c) (* By default the strategy for "rewrite_db" is top-down *) let mkappc s l = CAst.make @@ CAppExpl ((None,qualid_of_ident (Id.of_string s),None),l) let declare_an_instance n s args = (((CAst.make @@ Name n),None), CAst.make @@ CAppExpl ((None, qualid_of_string s,None), args)) let declare_instance a aeq n s = declare_an_instance n s [a;aeq] let anew_instance atts binders (name,t) fields = let _id = Classes.new_instance ~poly:atts.polymorphic name binders t (true, CAst.make @@ CRecord (fields)) ~global:atts.global ~generalize:false Hints.empty_hint_info in () let declare_instance_refl atts binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Reflexive") "Coq.Classes.RelationClasses.Reflexive" in anew_instance atts binders instance [(qualid_of_ident (Id.of_string "reflexivity"),lemma)] let declare_instance_sym atts binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Symmetric") "Coq.Classes.RelationClasses.Symmetric" in anew_instance atts binders instance [(qualid_of_ident (Id.of_string "symmetry"),lemma)] let declare_instance_trans atts binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Transitive") "Coq.Classes.RelationClasses.Transitive" in anew_instance atts binders instance [(qualid_of_ident (Id.of_string "transitivity"),lemma)] let declare_relation atts ?(binders=[]) a aeq n refl symm trans = init_setoid (); let instance = declare_instance a aeq (add_suffix n "_relation") "Coq.Classes.RelationClasses.RewriteRelation" in let () = anew_instance atts binders instance [] in match (refl,symm,trans) with (None, None, None) -> () | (Some lemma1, None, None) -> declare_instance_refl atts binders a aeq n lemma1 | (None, Some lemma2, None) -> declare_instance_sym atts binders a aeq n lemma2 | (None, None, Some lemma3) -> declare_instance_trans atts binders a aeq n lemma3 | (Some lemma1, Some lemma2, None) -> let () = declare_instance_refl atts binders a aeq n lemma1 in declare_instance_sym atts binders a aeq n lemma2 | (Some lemma1, None, Some lemma3) -> let () = declare_instance_refl atts binders a aeq n lemma1 in let () = declare_instance_trans atts binders a aeq n lemma3 in let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PreOrder" in anew_instance atts binders instance [(qualid_of_ident (Id.of_string "PreOrder_Reflexive"), lemma1); (qualid_of_ident (Id.of_string "PreOrder_Transitive"),lemma3)] | (None, Some lemma2, Some lemma3) -> let () = declare_instance_sym atts binders a aeq n lemma2 in let () = declare_instance_trans atts binders a aeq n lemma3 in let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PER" in anew_instance atts binders instance [(qualid_of_ident (Id.of_string "PER_Symmetric"), lemma2); (qualid_of_ident (Id.of_string "PER_Transitive"),lemma3)] | (Some lemma1, Some lemma2, Some lemma3) -> let () = declare_instance_refl atts binders a aeq n lemma1 in let () = declare_instance_sym atts binders a aeq n lemma2 in let () = declare_instance_trans atts binders a aeq n lemma3 in let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in anew_instance atts binders instance [(qualid_of_ident (Id.of_string "Equivalence_Reflexive"), lemma1); (qualid_of_ident (Id.of_string "Equivalence_Symmetric"), lemma2); (qualid_of_ident (Id.of_string "Equivalence_Transitive"), lemma3)] let cHole = CAst.make @@ CHole (None, Namegen.IntroAnonymous, None) let proper_projection env sigma r ty = let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i)) in let ctx, inst = decompose_prod_assum sigma ty in let mor, args = destApp sigma inst in let instarg = mkApp (r, rel_vect 0 (List.length ctx)) in let app = mkApp (PropGlobal.proper_proj env sigma, Array.append args [| instarg |]) in it_mkLambda_or_LetIn app ctx let declare_projection n instance_id r = let poly = Global.is_polymorphic r in let env = Global.env () in let sigma = Evd.from_env env in let sigma,c = Evd.fresh_global env sigma r in let ty = Retyping.get_type_of env sigma c in let term = proper_projection env sigma c ty in let sigma, typ = Typing.type_of env sigma term in let ctx, typ = decompose_prod_assum sigma typ in let typ = let n = let rec aux t = match EConstr.kind sigma t with | App (f, [| a ; a' ; rel; rel' |]) when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> succ (aux rel') | _ -> 0 in let init = match EConstr.kind sigma typ with App (f, args) when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> mkApp (f, fst (Array.chop (Array.length args - 2) args)) | _ -> typ in aux init in let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) typ in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in let univs = Evd.univ_entry ~poly sigma in let typ = EConstr.to_constr sigma typ in let term = EConstr.to_constr sigma term in let cst = Declare.definition_entry ~types:typ ~univs term in let _ : Constant.t = Declare.declare_constant ~name:n ~kind:Decls.(IsDefinition Definition) (Declare.DefinitionEntry cst) in () let build_morphism_signature env sigma m = let m,ctx = Constrintern.interp_constr env sigma m in let sigma = Evd.from_ctx ctx in let t = Typing.unsafe_type_of env sigma m in let cstrs = let rec aux t = match EConstr.kind sigma t with | Prod (na, a, b) -> None :: aux b | _ -> [] in aux t in let evars, t', sig_, cstrs = PropGlobal.build_signature (sigma, Evar.Set.empty) env t cstrs None in let evd = ref evars in let _ = List.iter (fun (ty, rel) -> Option.iter (fun rel -> let default = e_app_poly env evd PropGlobal.default_relation [| ty; rel |] in ignore(e_new_cstr_evar env evd default)) rel) cstrs in let morph = e_app_poly env evd (PropGlobal.proper_type env) [| t; sig_; m |] in let evd = solve_constraints env !evd in let evd = Evd.minimize_universes evd in let m = Evarutil.nf_evars_universes evd (EConstr.Unsafe.to_constr morph) in Pretyping.check_evars env (Evd.from_env env) evd (EConstr.of_constr m); Evd.evar_universe_context evd, m let default_morphism sign m = let env = Global.env () in let sigma = Evd.from_env env in let t = Typing.unsafe_type_of env sigma m in let evars, _, sign, cstrs = PropGlobal.build_signature (sigma, Evar.Set.empty) env t (fst sign) (snd sign) in let evars, morph = app_poly_check env evars (PropGlobal.proper_type env) [| t; sign; m |] in let evars, mor = resolve_one_typeclass env (goalevars evars) morph in mor, proper_projection env sigma mor morph let add_setoid atts binders a aeq t n = init_setoid (); let () = declare_instance_refl atts binders a aeq n (mkappc "Seq_refl" [a;aeq;t]) in let () = declare_instance_sym atts binders a aeq n (mkappc "Seq_sym" [a;aeq;t]) in let () = declare_instance_trans atts binders a aeq n (mkappc "Seq_trans" [a;aeq;t]) in let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in anew_instance atts binders instance [(qualid_of_ident (Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]); (qualid_of_ident (Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]); (qualid_of_ident (Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])] let make_tactic name = let open Tacexpr in let tacqid = Libnames.qualid_of_string name in TacArg (CAst.make @@ (TacCall (CAst.make (tacqid, [])))) let add_morphism_as_parameter atts m n : unit = init_setoid (); let instance_id = add_suffix n "_Proper" in let env = Global.env () in let evd = Evd.from_env env in let uctx, instance = build_morphism_signature env evd m in let uctx = UState.univ_entry ~poly:atts.polymorphic uctx in let cst = Declare.declare_constant ~name:instance_id ~kind:Decls.(IsAssumption Logical) (Declare.ParameterEntry (None,(instance,uctx),None)) in Classes.add_instance (Classes.mk_instance (PropGlobal.proper_class env evd) Hints.empty_hint_info atts.global (GlobRef.ConstRef cst)); declare_projection n instance_id (GlobRef.ConstRef cst) let add_morphism_interactive atts m n : Lemmas.t = init_setoid (); let instance_id = add_suffix n "_Proper" in let env = Global.env () in let evd = Evd.from_env env in let uctx, instance = build_morphism_signature env evd m in let poly = atts.polymorphic in let kind = Decls.(IsDefinition Instance) in let tac = make_tactic "Coq.Classes.SetoidTactics.add_morphism_tactic" in let hook { DeclareDef.Hook.S.dref; _ } = dref |> function | GlobRef.ConstRef cst -> Classes.add_instance (Classes.mk_instance (PropGlobal.proper_class env evd) Hints.empty_hint_info atts.global (GlobRef.ConstRef cst)); declare_projection n instance_id (GlobRef.ConstRef cst) | _ -> assert false in let hook = DeclareDef.Hook.make hook in let info = Lemmas.Info.make ~hook ~kind () in Flags.silently (fun () -> let lemma = Lemmas.start_lemma ~name:instance_id ~poly ~info (Evd.from_ctx uctx) (EConstr.of_constr instance) in fst (Lemmas.by (Tacinterp.interp tac) lemma)) () let add_morphism atts binders m s n = init_setoid (); let instance_id = add_suffix n "_Proper" in let instance_name = (CAst.make @@ Name instance_id),None in let instance_t = CAst.make @@ CAppExpl ((None, Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper",None), [cHole; s; m]) in let tac = Tacinterp.interp (make_tactic "add_morphism_tactic") in let _id, lemma = Classes.new_instance_interactive ~global:atts.global ~poly:atts.polymorphic instance_name binders instance_t ~generalize:false ~tac ~hook:(declare_projection n instance_id) Hints.empty_hint_info None in lemma (* no instance body -> always open proof *) (** Bind to "rewrite" too *) (** Taken from original setoid_replace, to emulate the old rewrite semantics where lemmas are first instantiated and then rewrite proceeds. *) let check_evar_map_of_evars_defs env evd = let metas = Evd.meta_list evd in let check_freemetas_is_empty rebus = Evd.Metaset.iter (fun m -> if Evd.meta_defined evd m then () else begin raise (Logic.RefinerError (env, evd, Logic.UnresolvedBindings [Evd.meta_name evd m])) end) in List.iter (fun (_,binding) -> match binding with Evd.Cltyp (_,{Evd.rebus=rebus; Evd.freemetas=freemetas}) -> check_freemetas_is_empty rebus freemetas | Evd.Clval (_,({Evd.rebus=rebus1; Evd.freemetas=freemetas1},_), {Evd.rebus=rebus2; Evd.freemetas=freemetas2}) -> check_freemetas_is_empty rebus1 freemetas1 ; check_freemetas_is_empty rebus2 freemetas2 ) metas (* Find a subterm which matches the pattern to rewrite for "rewrite" *) let unification_rewrite l2r c1 c2 sigma prf car rel but env = let (sigma,c') = try (* ~flags:(false,true) to allow to mark occurrences that must not be rewritten simply by replacing them with let-defined definitions in the context *) Unification.w_unify_to_subterm ~flags:rewrite_unif_flags env sigma ((if l2r then c1 else c2),but) with | ex when Pretype_errors.precatchable_exception ex -> (* ~flags:(true,true) to make Ring work (since it really exploits conversion) *) Unification.w_unify_to_subterm ~flags:rewrite_conv_unif_flags env sigma ((if l2r then c1 else c2),but) in let nf c = Reductionops.nf_evar sigma c in let c1 = if l2r then nf c' else nf c1 and c2 = if l2r then nf c2 else nf c' and car = nf car and rel = nf rel in check_evar_map_of_evars_defs env sigma; let prf = nf prf in let prfty = nf (Retyping.get_type_of env sigma prf) in let sort = sort_of_rel env sigma but in let abs = prf, prfty in let prf = mkRel 1 in let res = (car, rel, prf, c1, c2) in abs, sigma, res, Sorts.is_prop sort let get_hyp gl (c,l) clause l2r = let evars = Tacmach.New.project gl in let env = Tacmach.New.pf_env gl in let sigma, hi = decompose_applied_relation env evars (c,l) in let but = match clause with | Some id -> Tacmach.New.pf_get_hyp_typ id gl | None -> Reductionops.nf_evar evars (Tacmach.New.pf_concl gl) in unification_rewrite l2r hi.c1 hi.c2 sigma hi.prf hi.car hi.rel but env let general_rewrite_flags = { under_lambdas = false; on_morphisms = true } (* let rewriteclaustac_key = CProfile.declare_profile "cl_rewrite_clause_tac";; *) (* let cl_rewrite_clause_tac = CProfile.profile5 rewriteclaustac_key cl_rewrite_clause_tac *) (** Setoid rewriting when called with "rewrite" *) let general_s_rewrite cl l2r occs (c,l) ~new_goals = Proofview.Goal.enter begin fun gl -> let abs, evd, res, sort = get_hyp gl (c,l) cl l2r in let unify env evars t = unify_abs res l2r sort env evars t in let app = apply_rule unify occs in let recstrat aux = Strategies.choice app (subterm true general_rewrite_flags aux) in let substrat = Strategies.fix recstrat in let strat = { strategy = fun ({ state = () } as input) -> let _, res = substrat.strategy { input with state = 0 } in (), res } in let origsigma = Tacmach.New.project gl in tactic_init_setoid () <*> Proofview.tclOR (tclPROGRESS (tclTHEN (Proofview.Unsafe.tclEVARS evd) (cl_rewrite_clause_newtac ~progress:true ~abs:(Some abs) ~origsigma strat cl))) (fun (e, info) -> match e with | RewriteFailure e -> tclFAIL 0 (str"setoid rewrite failed: " ++ e) | e -> Proofview.tclZERO ~info e) end let _ = Hook.set Equality.general_setoid_rewrite_clause general_s_rewrite (** [setoid_]{reflexivity,symmetry,transitivity} tactics *) let not_declared env sigma ty rel = tclFAIL 0 (str" The relation " ++ Printer.pr_econstr_env env sigma rel ++ str" is not a declared " ++ str ty ++ str" relation. Maybe you need to require the Coq.Classes.RelationClasses library") let setoid_proof ty fn fallback = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in Proofview.tclORELSE begin try let rel, _, _ = decompose_app_rel env sigma concl in let (sigma, t) = Typing.type_of env sigma rel in let car = snd (List.hd (fst (Reductionops.splay_prod env sigma t))) in (try init_relation_classes () with _ -> raise Not_found); fn env sigma car rel with e -> Proofview.tclZERO e end begin function | e -> Proofview.tclORELSE fallback begin function (e', info) -> match e' with | Hipattern.NoEquationFound -> begin match e with | (Not_found, _) -> let rel, _, _ = decompose_app_rel env sigma concl in not_declared env sigma ty rel | (e, info) -> Proofview.tclZERO ~info e end | e' -> Proofview.tclZERO ~info e' end end end let tac_open ((evm,_), c) tac = (tclTHEN (Proofview.Unsafe.tclEVARS evm) (tac c)) let poly_proof getp gett env evm car rel = if Sorts.is_prop (sort_of_rel env evm rel) then getp env (evm,Evar.Set.empty) car rel else gett env (evm,Evar.Set.empty) car rel let setoid_reflexivity = setoid_proof "reflexive" (fun env evm car rel -> tac_open (poly_proof PropGlobal.get_reflexive_proof TypeGlobal.get_reflexive_proof env evm car rel) (fun c -> tclCOMPLETE (apply c))) (reflexivity_red true) let setoid_symmetry = setoid_proof "symmetric" (fun env evm car rel -> tac_open (poly_proof PropGlobal.get_symmetric_proof TypeGlobal.get_symmetric_proof env evm car rel) (fun c -> apply c)) (symmetry_red true) let setoid_transitivity c = setoid_proof "transitive" (fun env evm car rel -> tac_open (poly_proof PropGlobal.get_transitive_proof TypeGlobal.get_transitive_proof env evm car rel) (fun proof -> match c with | None -> eapply proof | Some c -> apply_with_bindings (proof,ImplicitBindings [ c ]))) (transitivity_red true c) let setoid_symmetry_in id = let open Tacmach.New in Proofview.Goal.enter begin fun gl -> let sigma = project gl in let ctype = pf_unsafe_type_of gl (mkVar id) in let binders,concl = decompose_prod_assum sigma ctype in let (equiv, args) = decompose_app sigma concl in let rec split_last_two = function | [c1;c2] -> [],(c1, c2) | x::y::z -> let l,res = split_last_two (y::z) in x::l, res | _ -> user_err Pp.(str "Cannot find an equivalence relation to rewrite.") in let others,(c1,c2) = split_last_two args in let he,c1,c2 = mkApp (equiv, Array.of_list others),c1,c2 in let new_hyp' = mkApp (he, [| c2 ; c1 |]) in let new_hyp = it_mkProd_or_LetIn new_hyp' binders in (tclTHENLAST (Tactics.assert_after_replacing id new_hyp) (tclTHENLIST [ intros; setoid_symmetry; apply (mkVar id); Tactics.assumption ])) end let _ = Hook.set Tactics.setoid_reflexivity setoid_reflexivity let _ = Hook.set Tactics.setoid_symmetry setoid_symmetry let _ = Hook.set Tactics.setoid_symmetry_in setoid_symmetry_in let _ = Hook.set Tactics.setoid_transitivity setoid_transitivity let get_lemma_proof f env evm x y = let (evm, _), c = f env (evm,Evar.Set.empty) x y in evm, c let get_reflexive_proof = get_lemma_proof PropGlobal.get_reflexive_proof let get_symmetric_proof = get_lemma_proof PropGlobal.get_symmetric_proof let get_transitive_proof = get_lemma_proof PropGlobal.get_transitive_proof coq-8.11.0/plugins/ltac/extratactics.mli0000644000175000017500000000170213612664533020073 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit Proofview.tactic val injHyp : Names.Id.t -> unit Proofview.tactic (* val refine_tac : Evd.open_constr -> unit Proofview.tactic *) val onSomeWithHoles : ('a option -> unit Proofview.tactic) -> 'a Tactypes.delayed_open option -> unit Proofview.tactic coq-8.11.0/plugins/ltac/Ltac.v0000644000175000017500000000000013612664533015732 0ustar treinentreinencoq-8.11.0/plugins/ltac/taccoerce.mli0000644000175000017500000000770513612664533017336 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t val to_constr : t -> constr option val of_uconstr : Ltac_pretype.closed_glob_constr -> t val to_uconstr : t -> Ltac_pretype.closed_glob_constr option val of_int : int -> t val to_int : t -> int option val to_list : t -> t list option val to_option : t -> t option option val to_pair : t -> (t * t) option val cast : 'a typed_abstract_argument_type -> Geninterp.Val.t -> 'a end (** {5 Coercion functions} *) val coerce_to_constr_context : Value.t -> constr val coerce_var_to_ident : bool -> Environ.env -> Evd.evar_map -> Value.t -> Id.t val coerce_to_ident_not_fresh : Evd.evar_map -> Value.t -> Id.t val coerce_to_intro_pattern : Evd.evar_map -> Value.t -> delayed_open_constr intro_pattern_expr val coerce_to_intro_pattern_naming : Evd.evar_map -> Value.t -> Namegen.intro_pattern_naming_expr val coerce_to_hint_base : Value.t -> string val coerce_to_int : Value.t -> int val coerce_to_constr : Environ.env -> Value.t -> Ltac_pretype.constr_under_binders val coerce_to_uconstr : Value.t -> Ltac_pretype.closed_glob_constr val coerce_to_closed_constr : Environ.env -> Value.t -> constr val coerce_to_evaluable_ref : Environ.env -> Evd.evar_map -> Value.t -> evaluable_global_reference val coerce_to_constr_list : Environ.env -> Value.t -> constr list val coerce_to_intro_pattern_list : ?loc:Loc.t -> Evd.evar_map -> Value.t -> Tacexpr.intro_patterns val coerce_to_hyp : Environ.env -> Evd.evar_map -> Value.t -> Id.t val coerce_to_hyp_list : Environ.env -> Evd.evar_map -> Value.t -> Id.t list val coerce_to_reference : Evd.evar_map -> Value.t -> GlobRef.t val coerce_to_quantified_hypothesis : Evd.evar_map -> Value.t -> quantified_hypothesis val coerce_to_decl_or_quant_hyp : Evd.evar_map -> Value.t -> quantified_hypothesis val coerce_to_int_or_var_list : Value.t -> int Locus.or_var list (** {5 Missing generic arguments} *) val wit_constr_context : (Empty.t, Empty.t, EConstr.constr) genarg_type val wit_constr_under_binders : (Empty.t, Empty.t, Ltac_pretype.constr_under_binders) genarg_type val error_ltac_variable : ?loc:Loc.t -> Id.t -> (Environ.env * Evd.evar_map) option -> Value.t -> string -> 'a (** Abstract application, to print ltac functions *) type appl = | UnnamedAppl (** For generic applications: nothing is printed *) | GlbAppl of (Names.KerName.t * Val.t list) list (** For calls to global constants, some may alias other. *) type tacvalue = | VFun of appl*Tacexpr.ltac_trace * Val.t Id.Map.t * Name.t list * Tacexpr.glob_tactic_expr | VRec of Val.t Id.Map.t ref * Tacexpr.glob_tactic_expr val wit_tacvalue : (Empty.t, tacvalue, tacvalue) Genarg.genarg_type val pr_value : (Environ.env * Evd.evar_map) option -> Geninterp.Val.t -> Pp.t coq-8.11.0/plugins/ltac/plugin_base.dune0000644000175000017500000000045213612664533020040 0ustar treinentreinen(library (name ltac_plugin) (public_name coq.plugins.ltac) (synopsis "Coq's LTAC tactic language") (modules :standard \ tauto) (libraries coq.stm)) (library (name tauto_plugin) (public_name coq.plugins.tauto) (synopsis "Coq's tauto tactic") (modules tauto) (libraries coq.plugins.ltac)) coq-8.11.0/plugins/ltac/coretactics.mlg0000644000175000017500000002447613612664533017713 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Tactics.intros_reflexivity } END TACTIC EXTEND exact | [ "exact" casted_constr(c) ] -> { Tactics.exact_no_check c } END TACTIC EXTEND assumption | [ "assumption" ] -> { Tactics.assumption } END TACTIC EXTEND etransitivity | [ "etransitivity" ] -> { Tactics.intros_transitivity None } END TACTIC EXTEND cut | [ "cut" constr(c) ] -> { Tactics.cut c } END TACTIC EXTEND exact_no_check | [ "exact_no_check" constr(c) ] -> { Tactics.exact_no_check c } END TACTIC EXTEND vm_cast_no_check | [ "vm_cast_no_check" constr(c) ] -> { Tactics.vm_cast_no_check c } END TACTIC EXTEND native_cast_no_check | [ "native_cast_no_check" constr(c) ] -> { Tactics.native_cast_no_check c } END TACTIC EXTEND casetype | [ "casetype" constr(c) ] -> { Tactics.case_type c } END TACTIC EXTEND elimtype | [ "elimtype" constr(c) ] -> { Tactics.elim_type c } END TACTIC EXTEND lapply | [ "lapply" constr(c) ] -> { Tactics.cut_and_apply c } END TACTIC EXTEND transitivity | [ "transitivity" constr(c) ] -> { Tactics.intros_transitivity (Some c) } END (** Left *) TACTIC EXTEND left | [ "left" ] -> { Tactics.left_with_bindings false NoBindings } END TACTIC EXTEND eleft | [ "eleft" ] -> { Tactics.left_with_bindings true NoBindings } END TACTIC EXTEND left_with | [ "left" "with" bindings(bl) ] -> { Tacticals.New.tclDELAYEDWITHHOLES false bl (fun bl -> Tactics.left_with_bindings false bl) } END TACTIC EXTEND eleft_with | [ "eleft" "with" bindings(bl) ] -> { Tacticals.New.tclDELAYEDWITHHOLES true bl (fun bl -> Tactics.left_with_bindings true bl) } END (** Right *) TACTIC EXTEND right | [ "right" ] -> { Tactics.right_with_bindings false NoBindings } END TACTIC EXTEND eright | [ "eright" ] -> { Tactics.right_with_bindings true NoBindings } END TACTIC EXTEND right_with | [ "right" "with" bindings(bl) ] -> { Tacticals.New.tclDELAYEDWITHHOLES false bl (fun bl -> Tactics.right_with_bindings false bl) } END TACTIC EXTEND eright_with | [ "eright" "with" bindings(bl) ] -> { Tacticals.New.tclDELAYEDWITHHOLES true bl (fun bl -> Tactics.right_with_bindings true bl) } END (** Constructor *) TACTIC EXTEND constructor | [ "constructor" ] -> { Tactics.any_constructor false None } | [ "constructor" int_or_var(i) ] -> { Tactics.constructor_tac false None i NoBindings } | [ "constructor" int_or_var(i) "with" bindings(bl) ] -> { let tac bl = Tactics.constructor_tac false None i bl in Tacticals.New.tclDELAYEDWITHHOLES false bl tac } END TACTIC EXTEND econstructor | [ "econstructor" ] -> { Tactics.any_constructor true None } | [ "econstructor" int_or_var(i) ] -> { Tactics.constructor_tac true None i NoBindings } | [ "econstructor" int_or_var(i) "with" bindings(bl) ] -> { let tac bl = Tactics.constructor_tac true None i bl in Tacticals.New.tclDELAYEDWITHHOLES true bl tac } END (** Specialize *) TACTIC EXTEND specialize | [ "specialize" constr_with_bindings(c) ] -> { Tacticals.New.tclDELAYEDWITHHOLES false c (fun c -> Tactics.specialize c None) } | [ "specialize" constr_with_bindings(c) "as" simple_intropattern(ipat) ] -> { Tacticals.New.tclDELAYEDWITHHOLES false c (fun c -> Tactics.specialize c (Some ipat)) } END TACTIC EXTEND symmetry | [ "symmetry" ] -> { Tactics.intros_symmetry {onhyps=Some[];concl_occs=AllOccurrences} } END TACTIC EXTEND symmetry_in | [ "symmetry" "in" in_clause(cl) ] -> { Tactics.intros_symmetry cl } END (** Split *) { let rec delayed_list = function | [] -> fun _ sigma -> (sigma, []) | x :: l -> fun env sigma -> let (sigma, x) = x env sigma in let (sigma, l) = delayed_list l env sigma in (sigma, x :: l) } TACTIC EXTEND split | [ "split" ] -> { Tactics.split_with_bindings false [NoBindings] } END TACTIC EXTEND esplit | [ "esplit" ] -> { Tactics.split_with_bindings true [NoBindings] } END TACTIC EXTEND split_with | [ "split" "with" bindings(bl) ] -> { Tacticals.New.tclDELAYEDWITHHOLES false bl (fun bl -> Tactics.split_with_bindings false [bl]) } END TACTIC EXTEND esplit_with | [ "esplit" "with" bindings(bl) ] -> { Tacticals.New.tclDELAYEDWITHHOLES true bl (fun bl -> Tactics.split_with_bindings true [bl]) } END TACTIC EXTEND exists | [ "exists" ] -> { Tactics.split_with_bindings false [NoBindings] } | [ "exists" ne_bindings_list_sep(bll, ",") ] -> { Tacticals.New.tclDELAYEDWITHHOLES false (delayed_list bll) (fun bll -> Tactics.split_with_bindings false bll) } END TACTIC EXTEND eexists | [ "eexists" ] -> { Tactics.split_with_bindings true [NoBindings] } | [ "eexists" ne_bindings_list_sep(bll, ",") ] -> { Tacticals.New.tclDELAYEDWITHHOLES true (delayed_list bll) (fun bll -> Tactics.split_with_bindings true bll) } END (** Intro *) TACTIC EXTEND intros_until | [ "intros" "until" quantified_hypothesis(h) ] -> { Tactics.intros_until h } END TACTIC EXTEND intro | [ "intro" ] -> { Tactics.intro_move None MoveLast } | [ "intro" ident(id) ] -> { Tactics.intro_move (Some id) MoveLast } | [ "intro" ident(id) "at" "top" ] -> { Tactics.intro_move (Some id) MoveFirst } | [ "intro" ident(id) "at" "bottom" ] -> { Tactics.intro_move (Some id) MoveLast } | [ "intro" ident(id) "after" hyp(h) ] -> { Tactics.intro_move (Some id) (MoveAfter h) } | [ "intro" ident(id) "before" hyp(h) ] -> { Tactics.intro_move (Some id) (MoveBefore h) } | [ "intro" "at" "top" ] -> { Tactics.intro_move None MoveFirst } | [ "intro" "at" "bottom" ] -> { Tactics.intro_move None MoveLast } | [ "intro" "after" hyp(h) ] -> { Tactics.intro_move None (MoveAfter h) } | [ "intro" "before" hyp(h) ] -> { Tactics.intro_move None (MoveBefore h) } END (** Move *) TACTIC EXTEND move | [ "move" hyp(id) "at" "top" ] -> { Tactics.move_hyp id MoveFirst } | [ "move" hyp(id) "at" "bottom" ] -> { Tactics.move_hyp id MoveLast } | [ "move" hyp(id) "after" hyp(h) ] -> { Tactics.move_hyp id (MoveAfter h) } | [ "move" hyp(id) "before" hyp(h) ] -> { Tactics.move_hyp id (MoveBefore h) } END (** Rename *) TACTIC EXTEND rename | [ "rename" ne_rename_list_sep(ids, ",") ] -> { Tactics.rename_hyp ids } END (** Revert *) TACTIC EXTEND revert | [ "revert" ne_hyp_list(hl) ] -> { Tactics.revert hl } END (** Simple induction / destruct *) { let simple_induct h = Tacticals.New.tclTHEN (Tactics.intros_until h) (Tacticals.New.onLastHyp Tactics.simplest_elim) } TACTIC EXTEND simple_induction | [ "simple" "induction" quantified_hypothesis(h) ] -> { simple_induct h } END { let simple_destruct h = Tacticals.New.tclTHEN (Tactics.intros_until h) (Tacticals.New.onLastHyp Tactics.simplest_case) } TACTIC EXTEND simple_destruct | [ "simple" "destruct" quantified_hypothesis(h) ] -> { simple_destruct h } END (** Double induction *) TACTIC EXTEND double_induction | [ "double" "induction" quantified_hypothesis(h1) quantified_hypothesis(h2) ] -> { Elim.h_double_induction h1 h2 } END (* Admit *) TACTIC EXTEND admit |[ "admit" ] -> { Proofview.give_up } END (* Fix *) TACTIC EXTEND fix | [ "fix" ident(id) natural(n) ] -> { Tactics.fix id n } END (* Cofix *) TACTIC EXTEND cofix | [ "cofix" ident(id) ] -> { Tactics.cofix id } END (* Clear *) TACTIC EXTEND clear | [ "clear" hyp_list(ids) ] -> { if List.is_empty ids then Tactics.keep [] else Tactics.clear ids } | [ "clear" "-" ne_hyp_list(ids) ] -> { Tactics.keep ids } END (* Clearbody *) TACTIC EXTEND clearbody | [ "clearbody" ne_hyp_list(ids) ] -> { Tactics.clear_body ids } END (* Generalize dependent *) TACTIC EXTEND generalize_dependent | [ "generalize" "dependent" constr(c) ] -> { Tactics.generalize_dep c } END (* Table of "pervasives" macros tactics (e.g. auto, simpl, etc.) *) { open Tacexpr let initial_atomic () = let nocl = {onhyps=Some[];concl_occs=AllOccurrences} in let iter (s, t) = let body = TacAtom (CAst.make t) in Tacenv.register_ltac false false (Names.Id.of_string s) body in let () = List.iter iter [ "red", TacReduce(Red false,nocl); "hnf", TacReduce(Hnf,nocl); "simpl", TacReduce(Simpl (Redops.all_flags,None),nocl); "compute", TacReduce(Cbv Redops.all_flags,nocl); "intros", TacIntroPattern (false,[]); ] in let iter (s, t) = Tacenv.register_ltac false false (Names.Id.of_string s) t in List.iter iter [ "idtac",TacId []; "fail", TacFail(TacLocal,ArgArg 0,[]); "fresh", TacArg(CAst.make @@ TacFreshId []) ] let () = Mltop.declare_cache_obj initial_atomic "ltac_plugin" (* First-class Ltac access to primitive blocks *) let initial_name s = { mltac_plugin = "ltac_plugin"; mltac_tactic = s; } let initial_entry s = { mltac_name = initial_name s; mltac_index = 0; } let register_list_tactical name f = let tac args ist = match args with | [v] -> begin match Tacinterp.Value.to_list v with | None -> Tacticals.New.tclZEROMSG (Pp.str "Expected a list") | Some tacs -> let tacs = List.map (fun tac -> Tacinterp.tactic_of_value ist tac) tacs in f tacs end | _ -> assert false in Tacenv.register_ml_tactic (initial_name name) [|tac|] let () = register_list_tactical "first" Tacticals.New.tclFIRST let () = register_list_tactical "solve" Tacticals.New.tclSOLVE let initial_tacticals () = let idn n = Id.of_string (Printf.sprintf "_%i" n) in let varn n = Reference (ArgVar (CAst.make (idn n))) in let iter (s, t) = Tacenv.register_ltac false false (Id.of_string s) t in List.iter iter [ "first", TacFun ([Name (idn 0)], TacML (CAst.make (initial_entry "first", [varn 0]))); "solve", TacFun ([Name (idn 0)], TacML (CAst.make (initial_entry "solve", [varn 0]))); ] let () = Mltop.declare_cache_obj initial_tacticals "ltac_plugin" } coq-8.11.0/plugins/extraction/0000755000175000017500000000000013612664533016127 5ustar treinentreinencoq-8.11.0/plugins/extraction/ExtrHaskellString.v0000644000175000017500000001071313612664533021735 0ustar treinentreinen(** * Special handling of ascii and strings for extraction to Haskell. *) Require Coq.extraction.Extraction. Require Import Ascii. Require Import String. Require Import Coq.Strings.Byte. (** * At the moment, Coq's extraction has no way to add extra import * statements to the extracted Haskell code. You will have to * manually add: * * import qualified Data.Bits * import qualified Data.Char *) Extract Inductive ascii => "Prelude.Char" [ "(\b0 b1 b2 b3 b4 b5 b6 b7 -> Data.Char.chr ( (if b0 then Data.Bits.shiftL 1 0 else 0) Prelude.+ (if b1 then Data.Bits.shiftL 1 1 else 0) Prelude.+ (if b2 then Data.Bits.shiftL 1 2 else 0) Prelude.+ (if b3 then Data.Bits.shiftL 1 3 else 0) Prelude.+ (if b4 then Data.Bits.shiftL 1 4 else 0) Prelude.+ (if b5 then Data.Bits.shiftL 1 5 else 0) Prelude.+ (if b6 then Data.Bits.shiftL 1 6 else 0) Prelude.+ (if b7 then Data.Bits.shiftL 1 7 else 0)))" ] "(\f a -> f (Data.Bits.testBit (Data.Char.ord a) 0) (Data.Bits.testBit (Data.Char.ord a) 1) (Data.Bits.testBit (Data.Char.ord a) 2) (Data.Bits.testBit (Data.Char.ord a) 3) (Data.Bits.testBit (Data.Char.ord a) 4) (Data.Bits.testBit (Data.Char.ord a) 5) (Data.Bits.testBit (Data.Char.ord a) 6) (Data.Bits.testBit (Data.Char.ord a) 7))". Extract Inlined Constant Ascii.ascii_dec => "(Prelude.==)". Extract Inlined Constant Ascii.eqb => "(Prelude.==)". Extract Inductive string => "Prelude.String" [ "([])" "(:)" ]. Extract Inlined Constant String.string_dec => "(Prelude.==)". Extract Inlined Constant String.eqb => "(Prelude.==)". (* python -c 'print(" ".join(r""" "%s" """.strip() % (r"'"'\''"'" if chr(i) == "'"'"'" else repr(""" "" """.strip()) if chr(i) == """ " """.strip() else repr(chr(i))) for i in range(256)))' # " to satisfy Coq's comment parser *) Extract Inductive byte => "Prelude.Char" ["'\x00'" "'\x01'" "'\x02'" "'\x03'" "'\x04'" "'\x05'" "'\x06'" "'\x07'" "'\x08'" "'\t'" "'\n'" "'\x0b'" "'\x0c'" "'\r'" "'\x0e'" "'\x0f'" "'\x10'" "'\x11'" "'\x12'" "'\x13'" "'\x14'" "'\x15'" "'\x16'" "'\x17'" "'\x18'" "'\x19'" "'\x1a'" "'\x1b'" "'\x1c'" "'\x1d'" "'\x1e'" "'\x1f'" "' '" "'!'" "'""'" "'#'" "'$'" "'%'" "'&'" "'\''" "'('" "')'" "'*'" "'+'" "','" "'-'" "'.'" "'/'" "'0'" "'1'" "'2'" "'3'" "'4'" "'5'" "'6'" "'7'" "'8'" "'9'" "':'" "';'" "'<'" "'='" "'>'" "'?'" "'@'" "'A'" "'B'" "'C'" "'D'" "'E'" "'F'" "'G'" "'H'" "'I'" "'J'" "'K'" "'L'" "'M'" "'N'" "'O'" "'P'" "'Q'" "'R'" "'S'" "'T'" "'U'" "'V'" "'W'" "'X'" "'Y'" "'Z'" "'['" "'\\'" "']'" "'^'" "'_'" "'`'" "'a'" "'b'" "'c'" "'d'" "'e'" "'f'" "'g'" "'h'" "'i'" "'j'" "'k'" "'l'" "'m'" "'n'" "'o'" "'p'" "'q'" "'r'" "'s'" "'t'" "'u'" "'v'" "'w'" "'x'" "'y'" "'z'" "'{'" "'|'" "'}'" "'~'" "'\x7f'" "'\x80'" "'\x81'" "'\x82'" "'\x83'" "'\x84'" "'\x85'" "'\x86'" "'\x87'" "'\x88'" "'\x89'" "'\x8a'" "'\x8b'" "'\x8c'" "'\x8d'" "'\x8e'" "'\x8f'" "'\x90'" "'\x91'" "'\x92'" "'\x93'" "'\x94'" "'\x95'" "'\x96'" "'\x97'" "'\x98'" "'\x99'" "'\x9a'" "'\x9b'" "'\x9c'" "'\x9d'" "'\x9e'" "'\x9f'" "'\xa0'" "'\xa1'" "'\xa2'" "'\xa3'" "'\xa4'" "'\xa5'" "'\xa6'" "'\xa7'" "'\xa8'" "'\xa9'" "'\xaa'" "'\xab'" "'\xac'" "'\xad'" "'\xae'" "'\xaf'" "'\xb0'" "'\xb1'" "'\xb2'" "'\xb3'" "'\xb4'" "'\xb5'" "'\xb6'" "'\xb7'" "'\xb8'" "'\xb9'" "'\xba'" "'\xbb'" "'\xbc'" "'\xbd'" "'\xbe'" "'\xbf'" "'\xc0'" "'\xc1'" "'\xc2'" "'\xc3'" "'\xc4'" "'\xc5'" "'\xc6'" "'\xc7'" "'\xc8'" "'\xc9'" "'\xca'" "'\xcb'" "'\xcc'" "'\xcd'" "'\xce'" "'\xcf'" "'\xd0'" "'\xd1'" "'\xd2'" "'\xd3'" "'\xd4'" "'\xd5'" "'\xd6'" "'\xd7'" "'\xd8'" "'\xd9'" "'\xda'" "'\xdb'" "'\xdc'" "'\xdd'" "'\xde'" "'\xdf'" "'\xe0'" "'\xe1'" "'\xe2'" "'\xe3'" "'\xe4'" "'\xe5'" "'\xe6'" "'\xe7'" "'\xe8'" "'\xe9'" "'\xea'" "'\xeb'" "'\xec'" "'\xed'" "'\xee'" "'\xef'" "'\xf0'" "'\xf1'" "'\xf2'" "'\xf3'" "'\xf4'" "'\xf5'" "'\xf6'" "'\xf7'" "'\xf8'" "'\xf9'" "'\xfa'" "'\xfb'" "'\xfc'" "'\xfd'" "'\xfe'" "'\xff'"]. Extract Inlined Constant Byte.eqb => "(Prelude.==)". Extract Inlined Constant Byte.byte_eq_dec => "(Prelude.==)". Extract Inlined Constant Ascii.ascii_of_byte => "(\x -> x)". Extract Inlined Constant Ascii.byte_of_ascii => "(\x -> x)". (* Require Import ExtrHaskellBasic. Definition test := "ceci est un test"%string. Definition test2 := List.map (option_map Byte.to_nat) (List.map Byte.of_nat (List.seq 0 256)). Definition test3 := List.map ascii_of_nat (List.seq 0 256). Extraction Language Haskell. Recursive Extraction test Ascii.zero Ascii.one test2 test3 byte_rect. *) coq-8.11.0/plugins/extraction/haskell.ml0000644000175000017500000003437313612664533020116 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Id.Set.add (Id.of_string s)) [ "Any"; "case"; "class"; "data"; "default"; "deriving"; "do"; "else"; "if"; "import"; "in"; "infix"; "infixl"; "infixr"; "instance"; "let"; "module"; "newtype"; "of"; "then"; "type"; "where"; "_"; "__"; "as"; "qualified"; "hiding" ; "unit" ; "unsafeCoerce" ] Id.Set.empty let pp_comment s = str "-- " ++ s ++ fnl () let pp_bracket_comment s = str"{- " ++ hov 0 s ++ str" -}" (* Note: do not shorten [str "foo" ++ fnl ()] into [str "foo\n"], the '\n' character interacts badly with the Format boxing mechanism *) let preamble mod_name comment used_modules usf = let pp_import mp = str ("import qualified "^ string_of_modfile mp) ++ fnl () in (if not (usf.magic || usf.tunknown) then mt () else str "{-# OPTIONS_GHC -cpp -XMagicHash #-}" ++ fnl () ++ str "{- For Hugs, use the option -F\"cpp -P -traditional\" -}" ++ fnl2 ()) ++ (match comment with | None -> mt () | Some com -> pp_bracket_comment com ++ fnl2 ()) ++ str "module " ++ pr_upper_id mod_name ++ str " where" ++ fnl2 () ++ str "import qualified Prelude" ++ fnl () ++ prlist pp_import used_modules ++ fnl () ++ (if not (usf.magic || usf.tunknown) then mt () else str "#ifdef __GLASGOW_HASKELL__" ++ fnl () ++ str "import qualified GHC.Base" ++ fnl () ++ str "#else" ++ fnl () ++ str "-- HUGS" ++ fnl () ++ str "import qualified IOExts" ++ fnl () ++ str "#endif" ++ fnl2 ()) ++ (if not usf.magic then mt () else str "#ifdef __GLASGOW_HASKELL__" ++ fnl () ++ str "unsafeCoerce :: a -> b" ++ fnl () ++ str "unsafeCoerce = GHC.Base.unsafeCoerce#" ++ fnl () ++ str "#else" ++ fnl () ++ str "-- HUGS" ++ fnl () ++ str "unsafeCoerce :: a -> b" ++ fnl () ++ str "unsafeCoerce = IOExts.unsafeCoerce" ++ fnl () ++ str "#endif" ++ fnl2 ()) ++ (if not usf.tunknown then mt () else str "#ifdef __GLASGOW_HASKELL__" ++ fnl () ++ str "type Any = GHC.Base.Any" ++ fnl () ++ str "#else" ++ fnl () ++ str "-- HUGS" ++ fnl () ++ str "type Any = ()" ++ fnl () ++ str "#endif" ++ fnl2 ()) ++ (if not usf.mldummy then mt () else str "__ :: any" ++ fnl () ++ str "__ = Prelude.error \"Logical or arity value used\"" ++ fnl2 ()) let pp_abst = function | [] -> (mt ()) | l -> (str "\\" ++ prlist_with_sep (fun () -> (str " ")) Id.print l ++ str " ->" ++ spc ()) (*s The pretty-printer for haskell syntax *) let pp_global k r = if is_inline_custom r then str (find_custom r) else str (Common.pp_global k r) (*s Pretty-printing of types. [par] is a boolean indicating whether parentheses are needed or not. *) let rec pp_type par vl t = let rec pp_rec par = function | Tmeta _ | Tvar' _ -> assert false | Tvar i -> (try Id.print (List.nth vl (pred i)) with Failure _ -> (str "a" ++ int i)) | Tglob (r,[]) -> pp_global Type r | Tglob (GlobRef.IndRef(kn,0),l) when not (keep_singleton ()) && MutInd.equal kn (mk_ind "Coq.Init.Specif" "sig") -> pp_type true vl (List.hd l) | Tglob (r,l) -> pp_par par (pp_global Type r ++ spc () ++ prlist_with_sep spc (pp_type true vl) l) | Tarr (t1,t2) -> pp_par par (pp_rec true t1 ++ spc () ++ str "->" ++ spc () ++ pp_rec false t2) | Tdummy _ -> str "()" | Tunknown -> str "Any" | Taxiom -> str "() -- AXIOM TO BE REALIZED" ++ fnl () in hov 0 (pp_rec par t) (*s Pretty-printing of expressions. [par] indicates whether parentheses are needed or not. [env] is the list of names for the de Bruijn variables. [args] is the list of collected arguments (already pretty-printed). *) let expr_needs_par = function | MLlam _ -> true | MLcase _ -> false (* now that we use the case ... of { ... } syntax *) | _ -> false let rec pp_expr par env args = let apply st = pp_apply st par args and apply2 st = pp_apply2 st par args in function | MLrel n -> let id = get_db_name n env in (* Try to survive to the occurrence of a Dummy rel. TODO: we should get rid of this hack (cf. BZ#592) *) let id = if Id.equal id dummy_name then Id.of_string "__" else id in apply (Id.print id) | MLapp (f,args') -> let stl = List.map (pp_expr true env []) args' in pp_expr par env (stl @ args) f | MLlam _ as a -> let fl,a' = collect_lams a in let fl,env' = push_vars (List.map id_of_mlid fl) env in let st = (pp_abst (List.rev fl) ++ pp_expr false env' [] a') in apply2 st | MLletin (id,a1,a2) -> let i,env' = push_vars [id_of_mlid id] env in let pp_id = Id.print (List.hd i) and pp_a1 = pp_expr false env [] a1 and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in let pp_def = str "let {" ++ cut () ++ hov 1 (pp_id ++ str " = " ++ pp_a1 ++ str "}") in apply2 (hv 0 (hv 0 (hv 1 pp_def ++ spc () ++ str "in") ++ spc () ++ hov 0 pp_a2)) | MLglob r -> apply (pp_global Term r) | MLcons (_,r,a) as c -> assert (List.is_empty args); begin match a with | _ when is_native_char c -> pp_native_char c | [] -> pp_global Cons r | [a] -> pp_par par (pp_global Cons r ++ spc () ++ pp_expr true env [] a) | _ -> pp_par par (pp_global Cons r ++ spc () ++ prlist_with_sep spc (pp_expr true env []) a) end | MLtuple l -> assert (List.is_empty args); pp_boxed_tuple (pp_expr true env []) l | MLcase (_,t, pv) when is_custom_match pv -> if not (is_regular_match pv) then user_err Pp.(str "Cannot mix yet user-given match and general patterns."); let mkfun (ids,_,e) = if not (List.is_empty ids) then named_lams (List.rev ids) e else dummy_lams (ast_lift 1 e) 1 in let pp_branch tr = pp_expr true env [] (mkfun tr) ++ fnl () in let inner = str (find_custom_match pv) ++ fnl () ++ prvect pp_branch pv ++ pp_expr true env [] t in apply2 (hov 2 inner) | MLcase (typ,t,pv) -> apply2 (v 0 (str "case " ++ pp_expr false env [] t ++ str " of {" ++ fnl () ++ pp_pat env pv)) | MLfix (i,ids,defs) -> let ids',env' = push_vars (List.rev (Array.to_list ids)) env in pp_fix par env' i (Array.of_list (List.rev ids'),defs) args | MLexn s -> (* An [MLexn] may be applied, but I don't really care. *) pp_par par (str "Prelude.error" ++ spc () ++ qs s) | MLdummy k -> (* An [MLdummy] may be applied, but I don't really care. *) (match msg_of_implicit k with | "" -> str "__" | s -> str "__" ++ spc () ++ pp_bracket_comment (str s)) | MLmagic a -> pp_apply (str "unsafeCoerce") par (pp_expr true env [] a :: args) | MLaxiom -> pp_par par (str "Prelude.error \"AXIOM TO BE REALIZED\"") | MLuint _ -> pp_par par (str "Prelude.error \"EXTRACTION OF UINT NOT IMPLEMENTED\"") | MLfloat _ -> pp_par par (str "Prelude.error \"EXTRACTION OF FLOAT NOT IMPLEMENTED\"") and pp_cons_pat par r ppl = pp_par par (pp_global Cons r ++ space_if (not (List.is_empty ppl)) ++ prlist_with_sep spc identity ppl) and pp_gen_pat par ids env = function | Pcons (r,l) -> pp_cons_pat par r (List.map (pp_gen_pat true ids env) l) | Pusual r -> pp_cons_pat par r (List.map Id.print ids) | Ptuple l -> pp_boxed_tuple (pp_gen_pat false ids env) l | Pwild -> str "_" | Prel n -> Id.print (get_db_name n env) and pp_one_pat env (ids,p,t) = let ids',env' = push_vars (List.rev_map id_of_mlid ids) env in hov 2 (str " " ++ pp_gen_pat false (List.rev ids') env' p ++ str " ->" ++ spc () ++ pp_expr (expr_needs_par t) env' [] t) and pp_pat env pv = prvecti (fun i x -> pp_one_pat env pv.(i) ++ if Int.equal i (Array.length pv - 1) then str "}" else (str ";" ++ fnl ())) pv (*s names of the functions ([ids]) are already pushed in [env], and passed here just for convenience. *) and pp_fix par env i (ids,bl) args = pp_par par (v 0 (v 1 (str "let {" ++ fnl () ++ prvect_with_sep (fun () -> str ";" ++ fnl ()) (fun (fi,ti) -> pp_function env (Id.print fi) ti) (Array.map2 (fun a b -> a,b) ids bl) ++ str "}") ++ fnl () ++ str "in " ++ pp_apply (Id.print ids.(i)) false args)) and pp_function env f t = let bl,t' = collect_lams t in let bl,env' = push_vars (List.map id_of_mlid bl) env in (f ++ pr_binding (List.rev bl) ++ str " =" ++ fnl () ++ str " " ++ hov 2 (pp_expr false env' [] t')) (*s Pretty-printing of inductive types declaration. *) let pp_logical_ind packet = pp_comment (Id.print packet.ip_typename ++ str " : logical inductive") ++ pp_comment (str "with constructors : " ++ prvect_with_sep spc Id.print packet.ip_consnames) let pp_singleton kn packet = let name = pp_global Type (GlobRef.IndRef (kn,0)) in let l = rename_tvars keywords packet.ip_vars in hov 2 (str "type " ++ name ++ spc () ++ prlist_with_sep spc Id.print l ++ (if not (List.is_empty l) then str " " else mt ()) ++ str "=" ++ spc () ++ pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ pp_comment (str "singleton inductive, whose constructor was " ++ Id.print packet.ip_consnames.(0))) let pp_one_ind ip pl cv = let pl = rename_tvars keywords pl in let pp_constructor (r,l) = (pp_global Cons r ++ match l with | [] -> (mt ()) | _ -> (str " " ++ prlist_with_sep (fun () -> (str " ")) (pp_type true pl) l)) in str (if Array.is_empty cv then "type " else "data ") ++ pp_global Type (GlobRef.IndRef ip) ++ prlist_strict (fun id -> str " " ++ pr_lower_id id) pl ++ str " =" ++ if Array.is_empty cv then str " () -- empty inductive" else (fnl () ++ str " " ++ v 0 (str " " ++ prvect_with_sep (fun () -> fnl () ++ str "| ") pp_constructor (Array.mapi (fun i c -> GlobRef.ConstructRef (ip,i+1),c) cv))) let rec pp_ind first kn i ind = if i >= Array.length ind.ind_packets then if first then mt () else fnl () else let ip = (kn,i) in let p = ind.ind_packets.(i) in if is_custom (GlobRef.IndRef (kn,i)) then pp_ind first kn (i+1) ind else if p.ip_logical then pp_logical_ind p ++ pp_ind first kn (i+1) ind else pp_one_ind ip p.ip_vars p.ip_types ++ fnl () ++ pp_ind false kn (i+1) ind (*s Pretty-printing of a declaration. *) let pp_decl = function | Dind (kn,i) when i.ind_kind == Singleton -> pp_singleton kn i.ind_packets.(0) ++ fnl () | Dind (kn,i) -> hov 0 (pp_ind true kn 0 i) | Dtype (r, l, t) -> if is_inline_custom r then mt () else let l = rename_tvars keywords l in let st = try let ids,s = find_type_custom r in prlist (fun id -> str (id^" ")) ids ++ str "=" ++ spc () ++ str s with Not_found -> prlist (fun id -> Id.print id ++ str " ") l ++ if t == Taxiom then str "= () -- AXIOM TO BE REALIZED" ++ fnl () else str "=" ++ spc () ++ pp_type false l t in hov 2 (str "type " ++ pp_global Type r ++ spc () ++ st) ++ fnl2 () | Dfix (rv, defs, typs) -> let names = Array.map (fun r -> if is_inline_custom r then mt () else pp_global Term r) rv in prvecti (fun i r -> let void = is_inline_custom r || (not (is_custom r) && match defs.(i) with MLexn "UNUSED" -> true | _ -> false) in if void then mt () else hov 2 (names.(i) ++ str " :: " ++ pp_type false [] typs.(i)) ++ fnl () ++ (if is_custom r then (names.(i) ++ str " = " ++ str (find_custom r)) else (pp_function (empty_env ()) names.(i) defs.(i))) ++ fnl2 ()) rv | Dterm (r, a, t) -> if is_inline_custom r then mt () else let e = pp_global Term r in hov 2 (e ++ str " :: " ++ pp_type false [] t) ++ fnl () ++ if is_custom r then hov 0 (e ++ str " = " ++ str (find_custom r) ++ fnl2 ()) else hov 0 (pp_function (empty_env ()) e a ++ fnl2 ()) let rec pp_structure_elem = function | (l,SEdecl d) -> pp_decl d | (l,SEmodule m) -> pp_module_expr m.ml_mod_expr | (l,SEmodtype m) -> mt () (* for the moment we simply discard module type *) and pp_module_expr = function | MEstruct (mp,sel) -> prlist_strict pp_structure_elem sel | MEfunctor _ -> mt () (* for the moment we simply discard unapplied functors *) | MEident _ | MEapply _ -> assert false (* should be expanded in extract_env *) let pp_struct = let pp_sel (mp,sel) = push_visible mp []; let p = prlist_strict pp_structure_elem sel in pop_visible (); p in prlist_strict pp_sel let haskell_descr = { keywords = keywords; file_suffix = ".hs"; file_naming = string_of_modfile; preamble = preamble; pp_struct = pp_struct; sig_suffix = None; sig_preamble = (fun _ _ _ _ -> mt ()); pp_sig = (fun _ -> mt ()); pp_decl = pp_decl; } coq-8.11.0/plugins/extraction/modutil.ml0000644000175000017500000003505613612664533020147 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* mp | MTsig(mp,_) -> mp | MTwith(mt,_)-> msid_of_mt mt | MTfunsig _ -> assert false (* A functor cannot be inside a MTwith *) (*s Apply some functions upon all [ml_decl] and [ml_spec] found in a [ml_structure]. *) let se_iter do_decl do_spec do_mp = let rec mt_iter = function | MTident mp -> do_mp mp | MTfunsig (_,mt,mt') -> mt_iter mt; mt_iter mt' | MTwith (mt,ML_With_type(idl,l,t))-> let mp_mt = msid_of_mt mt in let l',idl' = List.sep_last idl in let mp_w = List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl' in let r = GlobRef.ConstRef (Constant.make2 mp_w (Label.of_id l')) in mt_iter mt; do_spec (Stype(r,l,Some t)) | MTwith (mt,ML_With_module(idl,mp))-> let mp_mt = msid_of_mt mt in let mp_w = List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl in mt_iter mt; do_mp mp_w; do_mp mp | MTsig (_, sign) -> List.iter spec_iter sign and spec_iter = function | (_,Spec s) -> do_spec s | (_,Smodule mt) -> mt_iter mt | (_,Smodtype mt) -> mt_iter mt in let rec se_iter = function | (_,SEdecl d) -> do_decl d | (_,SEmodule m) -> me_iter m.ml_mod_expr; mt_iter m.ml_mod_type | (_,SEmodtype m) -> mt_iter m and me_iter = function | MEident mp -> do_mp mp | MEfunctor (_,mt,me) -> me_iter me; mt_iter mt | MEapply (me,me') -> me_iter me; me_iter me' | MEstruct (msid, sel) -> List.iter se_iter sel in se_iter let struct_iter do_decl do_spec do_mp s = List.iter (function (_,sel) -> List.iter (se_iter do_decl do_spec do_mp) sel) s (*s Apply some fonctions upon all references in [ml_type], [ml_ast], [ml_decl], [ml_spec] and [ml_structure]. *) type do_ref = GlobRef.t -> unit let record_iter_references do_term = function | Record l -> List.iter (Option.iter do_term) l | _ -> () let type_iter_references do_type t = let rec iter = function | Tglob (r,l) -> do_type r; List.iter iter l | Tarr (a,b) -> iter a; iter b | _ -> () in iter t let patt_iter_references do_cons p = let rec iter = function | Pcons (r,l) -> do_cons r; List.iter iter l | Pusual r -> do_cons r | Ptuple l -> List.iter iter l | Prel _ | Pwild -> () in iter p let ast_iter_references do_term do_cons do_type a = let rec iter a = ast_iter iter a; match a with | MLglob r -> do_term r | MLcons (_,r,_) -> do_cons r | MLcase (ty,_,v) -> type_iter_references do_type ty; Array.iter (fun (_,p,_) -> patt_iter_references do_cons p) v | MLrel _ | MLlam _ | MLapp _ | MLletin _ | MLtuple _ | MLfix _ | MLexn _ | MLdummy _ | MLaxiom | MLmagic _ | MLuint _ | MLfloat _ -> () in iter a let ind_iter_references do_term do_cons do_type kn ind = let type_iter = type_iter_references do_type in let cons_iter cp l = do_cons (GlobRef.ConstructRef cp); List.iter type_iter l in let packet_iter ip p = do_type (GlobRef.IndRef ip); if lang () == Ocaml then (match ind.ind_equiv with | Miniml.Equiv kne -> do_type (GlobRef.IndRef (MutInd.make1 kne, snd ip)); | _ -> ()); Array.iteri (fun j -> cons_iter (ip,j+1)) p.ip_types in if lang () == Ocaml then record_iter_references do_term ind.ind_kind; Array.iteri (fun i -> packet_iter (kn,i)) ind.ind_packets let decl_iter_references do_term do_cons do_type = let type_iter = type_iter_references do_type and ast_iter = ast_iter_references do_term do_cons do_type in function | Dind (kn,ind) -> ind_iter_references do_term do_cons do_type kn ind | Dtype (r,_,t) -> do_type r; type_iter t | Dterm (r,a,t) -> do_term r; ast_iter a; type_iter t | Dfix(rv,c,t) -> Array.iter do_term rv; Array.iter ast_iter c; Array.iter type_iter t let spec_iter_references do_term do_cons do_type = function | Sind (kn,ind) -> ind_iter_references do_term do_cons do_type kn ind | Stype (r,_,ot) -> do_type r; Option.iter (type_iter_references do_type) ot | Sval (r,t) -> do_term r; type_iter_references do_type t (*s Searching occurrences of a particular term (no lifting done). *) exception Found let rec ast_search f a = if f a then raise Found else ast_iter (ast_search f) a let decl_ast_search f = function | Dterm (_,a,_) -> ast_search f a | Dfix (_,c,_) -> Array.iter (ast_search f) c | _ -> () let struct_ast_search f s = try struct_iter (decl_ast_search f) (fun _ -> ()) (fun _ -> ()) s; false with Found -> true let rec type_search f = function | Tarr (a,b) -> type_search f a; type_search f b | Tglob (r,l) -> List.iter (type_search f) l | u -> if f u then raise Found let decl_type_search f = function | Dind (_,{ind_packets=p}) -> Array.iter (fun {ip_types=v} -> Array.iter (List.iter (type_search f)) v) p | Dterm (_,_,u) -> type_search f u | Dfix (_,_,v) -> Array.iter (type_search f) v | Dtype (_,_,u) -> type_search f u let spec_type_search f = function | Sind (_,{ind_packets=p}) -> Array.iter (fun {ip_types=v} -> Array.iter (List.iter (type_search f)) v) p | Stype (_,_,ot) -> Option.iter (type_search f) ot | Sval (_,u) -> type_search f u let struct_type_search f s = try struct_iter (decl_type_search f) (spec_type_search f) (fun _ -> ()) s; false with Found -> true (*s Generating the signature. *) let rec msig_of_ms = function | [] -> [] | (l,SEdecl (Dind (kn,i))) :: ms -> (l,Spec (Sind (kn,i))) :: (msig_of_ms ms) | (l,SEdecl (Dterm (r,_,t))) :: ms -> (l,Spec (Sval (r,t))) :: (msig_of_ms ms) | (l,SEdecl (Dtype (r,v,t))) :: ms -> (l,Spec (Stype (r,v,Some t))) :: (msig_of_ms ms) | (l,SEdecl (Dfix (rv,_,tv))) :: ms -> let msig = ref (msig_of_ms ms) in for i = Array.length rv - 1 downto 0 do msig := (l,Spec (Sval (rv.(i),tv.(i))))::!msig done; !msig | (l,SEmodule m) :: ms -> (l,Smodule m.ml_mod_type) :: (msig_of_ms ms) | (l,SEmodtype m) :: ms -> (l,Smodtype m) :: (msig_of_ms ms) let signature_of_structure s = List.map (fun (mp,ms) -> mp,msig_of_ms ms) s let rec mtyp_of_mexpr = function | MEfunctor (id,ty,e) -> MTfunsig (id,ty, mtyp_of_mexpr e) | MEstruct (mp,str) -> MTsig (mp, msig_of_ms str) | _ -> assert false (*s Searching one [ml_decl] in a [ml_structure] by its [global_reference] *) let is_modular = function | SEdecl _ -> false | SEmodule _ | SEmodtype _ -> true let rec search_structure l m = function | [] -> raise Not_found | (lab,d)::_ when Label.equal lab l && (is_modular d : bool) == m -> d | _::fields -> search_structure l m fields let get_decl_in_structure r struc = try let base_mp,ll = labels_of_ref r in if not (at_toplevel base_mp) then error_not_visible r; let sel = List.assoc_f ModPath.equal base_mp struc in let rec go ll sel = match ll with | [] -> assert false | l :: ll -> match search_structure l (not (List.is_empty ll)) sel with | SEdecl d -> d | SEmodtype m -> assert false | SEmodule m -> match m.ml_mod_expr with | MEstruct (_,sel) -> go ll sel | _ -> error_not_visible r in go ll sel with Not_found -> anomaly (Pp.str "reference not found in extracted structure.") (*s Optimization of a [ml_structure]. *) (* Some transformations of ML terms. [optimize_struct] simplify all beta redexes (when the argument does not occur, it is just thrown away; when it occurs exactly once it is substituted; otherwise a let-in redex is created for clarity) and iota redexes, plus some other optimizations. *) let dfix_to_mlfix rv av i = let rec make_subst n s = if n < 0 then s else make_subst (n-1) (Refmap'.add rv.(n) (n+1) s) in let s = make_subst (Array.length rv - 1) Refmap'.empty in let rec subst n t = match t with | MLglob ((GlobRef.ConstRef kn) as refe) -> (try MLrel (n + (Refmap'.find refe s)) with Not_found -> t) | _ -> ast_map_lift subst n t in let ids = Array.map (fun r -> Label.to_id (label_of_r r)) rv in let c = Array.map (subst 0) av in MLfix(i, ids, c) (* [optim_se] applies the [normalize] function everywhere and does the inlining of code. The inlined functions are kept for the moment in order to preserve the global interface, later [depcheck_se] will get rid of them if possible *) let rec optim_se top to_appear s = function | [] -> [] | (l,SEdecl (Dterm (r,a,t))) :: lse -> let a = normalize (ast_glob_subst !s a) in let i = inline r a in if i then s := Refmap'.add r a !s; let d = match dump_unused_vars (optimize_fix a) with | MLfix (0, _, [|c|]) -> Dfix ([|r|], [|ast_subst (MLglob r) c|], [|t|]) | a -> Dterm (r, a, t) in (l,SEdecl d) :: (optim_se top to_appear s lse) | (l,SEdecl (Dfix (rv,av,tv))) :: lse -> let av = Array.map (fun a -> normalize (ast_glob_subst !s a)) av in (* This fake body ensures that no fixpoint will be auto-inlined. *) let fake_body = MLfix (0,[||],[||]) in for i = 0 to Array.length rv - 1 do if inline rv.(i) fake_body then s := Refmap'.add rv.(i) (dfix_to_mlfix rv av i) !s done; let av' = Array.map dump_unused_vars av in (l,SEdecl (Dfix (rv, av', tv))) :: (optim_se top to_appear s lse) | (l,SEmodule m) :: lse -> let m = { m with ml_mod_expr = optim_me to_appear s m.ml_mod_expr} in (l,SEmodule m) :: (optim_se top to_appear s lse) | se :: lse -> se :: (optim_se top to_appear s lse) and optim_me to_appear s = function | MEstruct (msid, lse) -> MEstruct (msid, optim_se false to_appear s lse) | MEident mp as me -> me | MEapply (me, me') -> MEapply (optim_me to_appear s me, optim_me to_appear s me') | MEfunctor (mbid,mt,me) -> MEfunctor (mbid,mt, optim_me to_appear s me) (* After these optimisations, some dependencies may not be needed anymore. For non-library extraction, we recompute a minimal set of dependencies for first-level definitions (no module pruning yet). *) let base_r = let open GlobRef in function | ConstRef c as r -> r | IndRef (kn,_) -> IndRef (kn,0) | ConstructRef ((kn,_),_) -> IndRef (kn,0) | _ -> assert false let reset_needed, add_needed, add_needed_mp, found_needed, is_needed = let needed = ref Refset'.empty and needed_mps = ref MPset.empty in ((fun () -> needed := Refset'.empty; needed_mps := MPset.empty), (fun r -> needed := Refset'.add (base_r r) !needed), (fun mp -> needed_mps := MPset.add mp !needed_mps), (fun r -> needed := Refset'.remove (base_r r) !needed), (fun r -> let r = base_r r in Refset'.mem r !needed || MPset.mem (modpath_of_r r) !needed_mps)) let declared_refs = function | Dind (kn,_) -> [GlobRef.IndRef (kn,0)] | Dtype (r,_,_) -> [r] | Dterm (r,_,_) -> [r] | Dfix (rv,_,_) -> Array.to_list rv (* Computes the dependencies of a declaration, except in case of custom extraction. *) let compute_deps_decl = function | Dind (kn,ind) -> (* Todo Later : avoid dependencies when Extract Inductive *) ind_iter_references add_needed add_needed add_needed kn ind | Dtype (r,ids,t) -> if not (is_custom r) then type_iter_references add_needed t | Dterm (r,u,t) -> type_iter_references add_needed t; if not (is_custom r) then ast_iter_references add_needed add_needed add_needed u | Dfix _ as d -> decl_iter_references add_needed add_needed add_needed d let compute_deps_spec = function | Sind (kn,ind) -> (* Todo Later : avoid dependencies when Extract Inductive *) ind_iter_references add_needed add_needed add_needed kn ind | Stype (r,ids,t) -> if not (is_custom r) then Option.iter (type_iter_references add_needed) t | Sval (r,t) -> type_iter_references add_needed t let rec depcheck_se = function | [] -> [] | ((l,SEdecl d) as t) :: se -> let se' = depcheck_se se in let refs = declared_refs d in let refs' = List.filter is_needed refs in if List.is_empty refs' then (List.iter remove_info_axiom refs; List.iter remove_opaque refs; se') else begin List.iter found_needed refs'; (* Hack to avoid extracting unused part of a Dfix *) match d with | Dfix (rv,trms,tys) when (List.for_all is_custom refs') -> let trms' = Array.make (Array.length rv) (MLexn "UNUSED") in ((l,SEdecl (Dfix (rv,trms',tys))) :: se') | _ -> (compute_deps_decl d; t::se') end | t :: se -> let se' = depcheck_se se in se_iter compute_deps_decl compute_deps_spec add_needed_mp t; t :: se' let rec depcheck_struct = function | [] -> [] | (mp,lse)::struc -> let struc' = depcheck_struct struc in let lse' = depcheck_se lse in if List.is_empty lse' then struc' else (mp,lse')::struc' exception RemainingImplicit of kill_reason let check_for_remaining_implicits struc = let check = function | MLdummy (Kimplicit _ as k) -> raise (RemainingImplicit k) | _ -> false in try ignore (struct_ast_search check struc) with RemainingImplicit k -> err_or_warn_remaining_implicit k let optimize_struct to_appear struc = let subst = ref (Refmap'.empty : ml_ast Refmap'.t) in let opt_struc = List.map (fun (mp,lse) -> (mp, optim_se true (fst to_appear) subst lse)) struc in let mini_struc = if library () then List.filter (fun (_,lse) -> not (List.is_empty lse)) opt_struc else begin reset_needed (); List.iter add_needed (fst to_appear); List.iter add_needed_mp (snd to_appear); depcheck_struct opt_struc end in let () = check_for_remaining_implicits mini_struc in mini_struc coq-8.11.0/plugins/extraction/extract_env.ml0000644000175000017500000006624213612664533021015 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let mp,l = KerName.repr kn in begin match Libobject.object_tag o with | "CONSTANT" -> let constant = Global.lookup_constant (Constant.make1 kn) in Some (l, SFBconst constant) | "INDUCTIVE" -> let inductive = Global.lookup_mind (MutInd.make1 kn) in Some (l, SFBmind inductive) | _ -> None end | (_,kn), Lib.Leaf Libobject.ModuleObject _ -> let mp,l = KerName.repr kn in let modl = Global.lookup_module (MPdot (mp, l)) in Some (l, SFBmodule modl) | (_,kn), Lib.Leaf Libobject.ModuleTypeObject _ -> let mp,l = KerName.repr kn in let modtype = Global.lookup_modtype (MPdot (mp, l)) in Some (l, SFBmodtype modtype) | (_,kn), Lib.Leaf Libobject.IncludeObject _ -> user_err Pp.(str "No extraction of toplevel Include yet.") | _ -> None in List.rev (List.map_filter get_reference (Lib.contents ())) let environment_until dir_opt = let rec parse = function | [] when Option.is_empty dir_opt -> [Lib.current_mp (), toplevel_env ()] | [] -> [] | d :: l -> let meb = Modops.destr_nofunctor (Global.lookup_module (MPfile d)).mod_type in match dir_opt with | Some d' when DirPath.equal d d' -> [MPfile d, meb] | _ -> (MPfile d, meb) :: (parse l) in parse (Library.loaded_libraries ()) (*s Visit: a structure recording the needed dependencies for the current extraction *) module type VISIT = sig (* Reset the dependencies by emptying the visit lists *) val reset : unit -> unit (* Add the module_path and all its prefixes to the mp visit list. We'll keep all fields of these modules. *) val add_mp_all : ModPath.t -> unit (* Add reference / ... in the visit lists. These functions silently add the mp of their arg in the mp list *) val add_ref : GlobRef.t -> unit val add_kn : KerName.t -> unit val add_decl_deps : ml_decl -> unit val add_spec_deps : ml_spec -> unit (* Test functions: is a particular object a needed dependency for the current extraction ? *) val needed_ind : MutInd.t -> bool val needed_cst : Constant.t -> bool val needed_mp : ModPath.t -> bool val needed_mp_all : ModPath.t -> bool end module Visit : VISIT = struct type must_visit = { mutable kn : KNset.t; mutable mp : MPset.t; mutable mp_all : MPset.t } (* the imperative internal visit lists *) let v = { kn = KNset.empty; mp = MPset.empty; mp_all = MPset.empty } (* the accessor functions *) let reset () = v.kn <- KNset.empty; v.mp <- MPset.empty; v.mp_all <- MPset.empty let needed_ind i = KNset.mem (MutInd.user i) v.kn let needed_cst c = KNset.mem (Constant.user c) v.kn let needed_mp mp = MPset.mem mp v.mp || MPset.mem mp v.mp_all let needed_mp_all mp = MPset.mem mp v.mp_all let add_mp mp = check_loaded_modfile mp; v.mp <- MPset.union (prefixes_mp mp) v.mp let add_mp_all mp = check_loaded_modfile mp; v.mp <- MPset.union (prefixes_mp mp) v.mp; v.mp_all <- MPset.add mp v.mp_all let add_kn kn = v.kn <- KNset.add kn v.kn; add_mp (KerName.modpath kn) let add_ref = let open GlobRef in function | ConstRef c -> add_kn (Constant.user c) | IndRef (ind,_) | ConstructRef ((ind,_),_) -> add_kn (MutInd.user ind) | VarRef _ -> assert false let add_decl_deps = decl_iter_references add_ref add_ref add_ref let add_spec_deps = spec_iter_references add_ref add_ref add_ref end let add_field_label mp = function | (lab, (SFBconst _|SFBmind _)) -> Visit.add_kn (KerName.make mp lab) | (lab, (SFBmodule _|SFBmodtype _)) -> Visit.add_mp_all (MPdot (mp,lab)) let rec add_labels mp = function | MoreFunctor (_,_,m) -> add_labels mp m | NoFunctor sign -> List.iter (add_field_label mp) sign exception Impossible let check_arity env cb = let t = cb.const_type in if Reduction.is_arity env t then raise Impossible let get_body lbody = EConstr.of_constr (Mod_subst.force_constr lbody) let check_fix env sg cb i = match cb.const_body with | Def lbody -> (match EConstr.kind sg (get_body lbody) with | Fix ((_,j),recd) when Int.equal i j -> check_arity env cb; (true,recd) | CoFix (j,recd) when Int.equal i j -> check_arity env cb; (false,recd) | _ -> raise Impossible) | Undef _ | OpaqueDef _ | Primitive _ -> raise Impossible let prec_declaration_equal sg (na1, ca1, ta1) (na2, ca2, ta2) = Array.equal (Context.eq_annot Name.equal) na1 na2 && Array.equal (EConstr.eq_constr sg) ca1 ca2 && Array.equal (EConstr.eq_constr sg) ta1 ta2 let factor_fix env sg l cb msb = let _,recd as check = check_fix env sg cb 0 in let n = Array.length (let fi,_,_ = recd in fi) in if Int.equal n 1 then [|l|], recd, msb else begin if List.length msb < n-1 then raise Impossible; let msb', msb'' = List.chop (n-1) msb in let labels = Array.make n l in List.iteri (fun j -> function | (l,SFBconst cb') -> let check' = check_fix env sg cb' (j+1) in if not ((fst check : bool) == (fst check') && prec_declaration_equal sg (snd check) (snd check')) then raise Impossible; labels.(j+1) <- l; | _ -> raise Impossible) msb'; labels, recd, msb'' end (** Expanding a [module_alg_expr] into a version without abbreviations or functor applications. This is done via a detour to entries (hack proposed by Elie) *) let expand_mexpr env mpo me = let inl = Some (Flags.get_inline_level()) in Mod_typing.translate_mse env mpo inl me let expand_modtype env mp me = let inl = Some (Flags.get_inline_level()) in Mod_typing.translate_modtype env mp inl ([],me) let no_delta = Mod_subst.empty_delta_resolver let flatten_modtype env mp me_alg struc_opt = match struc_opt with | Some me -> me, no_delta | None -> let mtb = expand_modtype env mp me_alg in mtb.mod_type, mtb.mod_delta (** Ad-hoc update of environment, inspired by [Mod_typing.check_with_aux_def]. *) let env_for_mtb_with_def env mp me reso idl = let struc = Modops.destr_nofunctor me in let l = Label.of_id (List.hd idl) in let spot = function (l',SFBconst _) -> Label.equal l l' | _ -> false in let before = fst (List.split_when spot struc) in Modops.add_structure mp before reso env let make_cst resolver mp l = Mod_subst.constant_of_delta_kn resolver (KerName.make mp l) let make_mind resolver mp l = Mod_subst.mind_of_delta_kn resolver (KerName.make mp l) (* From a [structure_body] (i.e. a list of [structure_field_body]) to specifications. *) let rec extract_structure_spec env mp reso = function | [] -> [] | (l,SFBconst cb) :: msig -> let c = make_cst reso mp l in let s = extract_constant_spec env c cb in let specs = extract_structure_spec env mp reso msig in if logical_spec s then specs else begin Visit.add_spec_deps s; (l,Spec s) :: specs end | (l,SFBmind _) :: msig -> let mind = make_mind reso mp l in let s = Sind (mind, extract_inductive env mind) in let specs = extract_structure_spec env mp reso msig in if logical_spec s then specs else begin Visit.add_spec_deps s; (l,Spec s) :: specs end | (l,SFBmodule mb) :: msig -> let specs = extract_structure_spec env mp reso msig in let spec = extract_mbody_spec env mb.mod_mp mb in (l,Smodule spec) :: specs | (l,SFBmodtype mtb) :: msig -> let specs = extract_structure_spec env mp reso msig in let spec = extract_mbody_spec env mtb.mod_mp mtb in (l,Smodtype spec) :: specs (* From [module_expression] to specifications *) (* Invariant: the [me_alg] given to [extract_mexpr_spec] and [extract_mexpression_spec] should come from a [mod_type_alg] field. This way, any encountered [MEident] should be a true module type. *) and extract_mexpr_spec env mp1 (me_struct_o,me_alg) = match me_alg with | MEident mp -> Visit.add_mp_all mp; MTident mp | MEwith(me',WithDef(idl,(c,ctx)))-> let me_struct,delta = flatten_modtype env mp1 me' me_struct_o in let env' = env_for_mtb_with_def env mp1 me_struct delta idl in let mt = extract_mexpr_spec env mp1 (None,me') in let sg = Evd.from_env env in (match extract_with_type env' sg (EConstr.of_constr c) with (* cb may contain some kn *) | None -> mt | Some (vl,typ) -> type_iter_references Visit.add_ref typ; MTwith(mt,ML_With_type(idl,vl,typ))) | MEwith(me',WithMod(idl,mp))-> Visit.add_mp_all mp; MTwith(extract_mexpr_spec env mp1 (None,me'), ML_With_module(idl,mp)) | MEapply _ -> (* No higher-order module type in OCaml : we use the expanded version *) let me_struct,delta = flatten_modtype env mp1 me_alg me_struct_o in extract_msignature_spec env mp1 delta me_struct and extract_mexpression_spec env mp1 (me_struct,me_alg) = match me_alg with | MoreFunctor (mbid, mtb, me_alg') -> let me_struct' = match me_struct with | MoreFunctor (mbid',_,me') when MBId.equal mbid' mbid -> me' | _ -> assert false in let mp = MPbound mbid in let env' = Modops.add_module_type mp mtb env in MTfunsig (mbid, extract_mbody_spec env mp mtb, extract_mexpression_spec env' mp1 (me_struct',me_alg')) | NoFunctor m -> extract_mexpr_spec env mp1 (Some me_struct,m) and extract_msignature_spec env mp1 reso = function | NoFunctor struc -> let env' = Modops.add_structure mp1 struc reso env in MTsig (mp1, extract_structure_spec env' mp1 reso struc) | MoreFunctor (mbid, mtb, me) -> let mp = MPbound mbid in let env' = Modops.add_module_type mp mtb env in MTfunsig (mbid, extract_mbody_spec env mp mtb, extract_msignature_spec env' mp1 reso me) and extract_mbody_spec : 'a. _ -> _ -> 'a generic_module_body -> _ = fun env mp mb -> match mb.mod_type_alg with | Some ty -> extract_mexpression_spec env mp (mb.mod_type,ty) | None -> extract_msignature_spec env mp mb.mod_delta mb.mod_type (* From a [structure_body] (i.e. a list of [structure_field_body]) to implementations. NB: when [all=false], the evaluation order of the list is important: last to first ensures correct dependencies. *) let rec extract_structure env mp reso ~all = function | [] -> [] | (l,SFBconst cb) :: struc -> (try let sg = Evd.from_env env in let vl,recd,struc = factor_fix env sg l cb struc in let vc = Array.map (make_cst reso mp) vl in let ms = extract_structure env mp reso ~all struc in let b = Array.exists Visit.needed_cst vc in if all || b then let d = extract_fixpoint env sg vc recd in if (not b) && (logical_decl d) then ms else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end else ms with Impossible -> let ms = extract_structure env mp reso ~all struc in let c = make_cst reso mp l in let b = Visit.needed_cst c in if all || b then let d = extract_constant env c cb in if (not b) && (logical_decl d) then ms else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end else ms) | (l,SFBmind mib) :: struc -> let ms = extract_structure env mp reso ~all struc in let mind = make_mind reso mp l in let b = Visit.needed_ind mind in if all || b then let d = Dind (mind, extract_inductive env mind) in if (not b) && (logical_decl d) then ms else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end else ms | (l,SFBmodule mb) :: struc -> let ms = extract_structure env mp reso ~all struc in let mp = MPdot (mp,l) in let all' = all || Visit.needed_mp_all mp in if all' || Visit.needed_mp mp then (l,SEmodule (extract_module env mp ~all:all' mb)) :: ms else ms | (l,SFBmodtype mtb) :: struc -> let ms = extract_structure env mp reso ~all struc in let mp = MPdot (mp,l) in if all || Visit.needed_mp mp then (l,SEmodtype (extract_mbody_spec env mp mtb)) :: ms else ms (* From [module_expr] and [module_expression] to implementations *) and extract_mexpr env mp = function | MEwith _ -> assert false (* no 'with' syntax for modules *) | me when lang () != Ocaml || Table.is_extrcompute () -> (* In Haskell/Scheme, we expand everything. For now, we also extract everything, dead code will be removed later (see [Modutil.optimize_struct]. *) let sign,_,delta,_ = expand_mexpr env (Some mp) me in extract_msignature env mp delta ~all:true sign | MEident mp -> if is_modfile mp && not (modular ()) then error_MPfile_as_mod mp false; Visit.add_mp_all mp; Miniml.MEident mp | MEapply (me, arg) -> Miniml.MEapply (extract_mexpr env mp me, extract_mexpr env mp (MEident arg)) and extract_mexpression env mp = function | NoFunctor me -> extract_mexpr env mp me | MoreFunctor (mbid, mtb, me) -> let mp1 = MPbound mbid in let env' = Modops.add_module_type mp1 mtb env in Miniml.MEfunctor (mbid, extract_mbody_spec env mp1 mtb, extract_mexpression env' mp me) and extract_msignature env mp reso ~all = function | NoFunctor struc -> let env' = Modops.add_structure mp struc reso env in Miniml.MEstruct (mp,extract_structure env' mp reso ~all struc) | MoreFunctor (mbid, mtb, me) -> let mp1 = MPbound mbid in let env' = Modops.add_module_type mp1 mtb env in Miniml.MEfunctor (mbid, extract_mbody_spec env mp1 mtb, extract_msignature env' mp reso ~all me) and extract_module env mp ~all mb = (* A module has an empty [mod_expr] when : - it is a module variable (for instance X inside a Module F [X:SIG]) - it is a module assumption (Declare Module). Since we look at modules from outside, we shouldn't have variables. But a Declare Module at toplevel seems legal (cf #2525). For the moment we don't support this situation. *) let impl = match mb.mod_expr with | Abstract -> error_no_module_expr mp | Algebraic me -> extract_mexpression env mp me | Struct sign -> (* This module has a signature, otherwise it would be FullStruct. We extract just the elements required by this signature. *) let () = add_labels mp mb.mod_type in extract_msignature env mp mb.mod_delta ~all:false sign | FullStruct -> extract_msignature env mp mb.mod_delta ~all mb.mod_type in (* Slight optimization: for modules without explicit signatures ([FullStruct] case), we build the type out of the extracted implementation *) let typ = match mb.mod_expr with | FullStruct -> assert (Option.is_empty mb.mod_type_alg); mtyp_of_mexpr impl | _ -> extract_mbody_spec env mp mb in { ml_mod_expr = impl; ml_mod_type = typ } let mono_environment refs mpl = Visit.reset (); List.iter Visit.add_ref refs; List.iter Visit.add_mp_all mpl; let env = Global.env () in let l = List.rev (environment_until None) in List.rev_map (fun (mp,struc) -> mp, extract_structure env mp no_delta ~all:(Visit.needed_mp_all mp) struc) l (**************************************) (*S Part II : Input/Output primitives *) (**************************************) let descr () = match lang () with | Ocaml -> Ocaml.ocaml_descr | Haskell -> Haskell.haskell_descr | Scheme -> Scheme.scheme_descr | JSON -> Json.json_descr (* From a filename string "foo.ml" or "foo", builds "foo.ml" and "foo.mli" Works similarly for the other languages. *) let default_id = Id.of_string "Main" let mono_filename f = let d = descr () in match f with | None -> None, None, default_id | Some f -> let f = if Filename.check_suffix f d.file_suffix then Filename.chop_suffix f d.file_suffix else f in let id = if lang () != Haskell then default_id else try Id.of_string (Filename.basename f) with UserError _ -> user_err Pp.(str "Extraction: provided filename is not a valid identifier") in Some (f^d.file_suffix), Option.map ((^) f) d.sig_suffix, id (* Builds a suitable filename from a module id *) let module_filename mp = let f = file_of_modfile mp in let d = descr () in let p = d.file_naming mp ^ d.file_suffix in Some p, Option.map ((^) f) d.sig_suffix, Id.of_string f (*s Extraction of one decl to stdout. *) let print_one_decl struc mp decl = let d = descr () in reset_renaming_tables AllButExternal; set_phase Pre; ignore (d.pp_struct struc); set_phase Impl; push_visible mp []; let ans = d.pp_decl decl in pop_visible (); v 0 ans (*s Extraction of a ml struct to a file. *) (** For Recursive Extraction, writing directly on stdout won't work with coqide, we use a buffer instead *) let buf = Buffer.create 1000 let formatter dry file = let ft = if dry then Format.make_formatter (fun _ _ _ -> ()) (fun _ -> ()) else match file with | Some f -> Topfmt.with_output_to f | None -> Format.formatter_of_buffer buf in (* XXX: Fixme, this shouldn't depend on Topfmt *) (* We never want to see ellipsis ... in extracted code *) Format.pp_set_max_boxes ft max_int; (* We reuse the width information given via "Set Printing Width" *) (match Topfmt.get_margin () with | None -> () | Some i -> Format.pp_set_margin ft i; Format.pp_set_max_indent ft (i-10)); (* note: max_indent should be < margin above, otherwise it's ignored *) ft let get_comment () = let s = file_comment () in if String.is_empty s then None else let split_comment = Str.split (Str.regexp "[ \t\n]+") s in Some (prlist_with_sep spc str split_comment) let print_structure_to_file (fn,si,mo) dry struc = Buffer.clear buf; let d = descr () in reset_renaming_tables AllButExternal; let unsafe_needs = { mldummy = struct_ast_search Mlutil.isMLdummy struc; tdummy = struct_type_search Mlutil.isTdummy struc; tunknown = struct_type_search ((==) Tunknown) struc; magic = if lang () != Haskell then false else struct_ast_search (function MLmagic _ -> true | _ -> false) struc } in (* First, a dry run, for computing objects to rename or duplicate *) set_phase Pre; ignore (d.pp_struct struc); let opened = opened_libraries () in (* Print the implementation *) let cout = if dry then None else Option.map open_out fn in let ft = formatter dry cout in let comment = get_comment () in begin try (* The real printing of the implementation *) set_phase Impl; pp_with ft (d.preamble mo comment opened unsafe_needs); pp_with ft (d.pp_struct struc); Format.pp_print_flush ft (); Option.iter close_out cout; with reraise -> Format.pp_print_flush ft (); Option.iter close_out cout; raise reraise end; if not dry then Option.iter info_file fn; (* Now, let's print the signature *) Option.iter (fun si -> let cout = open_out si in let ft = formatter false (Some cout) in begin try set_phase Intf; pp_with ft (d.sig_preamble mo comment opened unsafe_needs); pp_with ft (d.pp_sig (signature_of_structure struc)); Format.pp_print_flush ft (); close_out cout; with reraise -> Format.pp_print_flush ft (); close_out cout; raise reraise end; info_file si) (if dry then None else si); (* Print the buffer content via Coq standard formatter (ok with coqide). *) if not (Int.equal (Buffer.length buf) 0) then begin Feedback.msg_notice (str (Buffer.contents buf)); Buffer.reset buf end (*********************************************) (*s Part III: the actual extraction commands *) (*********************************************) let reset () = Visit.reset (); reset_tables (); reset_renaming_tables Everything let init ?(compute=false) ?(inner=false) modular library = if not inner then (check_inside_section (); check_inside_module ()); set_keywords (descr ()).keywords; set_modular modular; set_library library; set_extrcompute compute; reset (); if modular && lang () == Scheme then error_scheme () let warns () = warning_opaques (access_opaque ()); warning_axioms () (* From a list of [reference], let's retrieve whether they correspond to modules or [global_reference]. Warn the user if both is possible. *) let rec locate_ref = function | [] -> [],[] | qid::l -> let mpo = try Some (Nametab.locate_module qid) with Not_found -> None and ro = try Some (Smartlocate.global_with_alias qid) with Nametab.GlobalizationError _ | UserError _ -> None in match mpo, ro with | None, None -> Nametab.error_global_not_found qid | None, Some r -> let refs,mps = locate_ref l in r::refs,mps | Some mp, None -> let refs,mps = locate_ref l in refs,mp::mps | Some mp, Some r -> warning_ambiguous_name (qid,mp,r); let refs,mps = locate_ref l in refs,mp::mps (*s Recursive extraction in the Coq toplevel. The vernacular command is \verb!Recursive Extraction! [qualid1] ... [qualidn]. Also used when extracting to a file with the command: \verb!Extraction "file"! [qualid1] ... [qualidn]. *) let full_extr f (refs,mps) = init false false; List.iter (fun mp -> if is_modfile mp then error_MPfile_as_mod mp true) mps; let struc = optimize_struct (refs,mps) (mono_environment refs mps) in warns (); print_structure_to_file (mono_filename f) false struc; reset () let full_extraction f lr = full_extr f (locate_ref lr) (*s Separate extraction is similar to recursive extraction, with the output decomposed in many files, one per Coq .v file *) let separate_extraction lr = init true false; let refs,mps = locate_ref lr in let struc = optimize_struct (refs,mps) (mono_environment refs mps) in warns (); let print = function | (MPfile dir as mp, sel) as e -> print_structure_to_file (module_filename mp) false [e] | _ -> assert false in List.iter print struc; reset () (*s Simple extraction in the Coq toplevel. The vernacular command is \verb!Extraction! [qualid]. *) let simple_extraction r = match locate_ref [r] with | ([], [mp]) as p -> full_extr None p | [r],[] -> init false false; let struc = optimize_struct ([r],[]) (mono_environment [r] []) in let d = get_decl_in_structure r struc in warns (); let flag = if is_custom r then str "(** User defined extraction *)" ++ fnl() else mt () in let ans = flag ++ print_one_decl struc (modpath_of_r r) d in reset (); Feedback.msg_notice ans | _ -> assert false (*s (Recursive) Extraction of a library. The vernacular command is \verb!(Recursive) Extraction Library! [M]. *) let extraction_library is_rec m = init true true; let dir_m = let q = qualid_of_ident m in try Nametab.full_name_module q with Not_found -> error_unknown_module q in Visit.add_mp_all (MPfile dir_m); let env = Global.env () in let l = List.rev (environment_until (Some dir_m)) in let select l (mp,struc) = if Visit.needed_mp mp then (mp, extract_structure env mp no_delta ~all:true struc) :: l else l in let struc = List.fold_left select [] l in let struc = optimize_struct ([],[]) struc in warns (); let print = function | (MPfile dir as mp, sel) as e -> let dry = not is_rec && not (DirPath.equal dir dir_m) in print_structure_to_file (module_filename mp) dry [e] | _ -> assert false in List.iter print struc; reset () (** For extraction compute, we flatten all the module structure, getting rid of module types or unapplied functors *) let flatten_structure struc = let rec flatten_elem (lab,elem) = match elem with |SEdecl d -> [d] |SEmodtype _ -> [] |SEmodule m -> match m.ml_mod_expr with |MEfunctor _ -> [] |MEident _ | MEapply _ -> assert false (* should be expanded *) |MEstruct (_,elems) -> flatten_elems elems and flatten_elems l = List.flatten (List.map flatten_elem l) in flatten_elems (List.flatten (List.map snd struc)) let structure_for_compute env sg c = init false false ~compute:true; let ast, mlt = Extraction.extract_constr env sg c in let ast = Mlutil.normalize ast in let refs = ref GlobRef.Set.empty in let add_ref r = refs := GlobRef.Set.add r !refs in let () = ast_iter_references add_ref add_ref add_ref ast in let refs = GlobRef.Set.elements !refs in let struc = optimize_struct (refs,[]) (mono_environment refs []) in (flatten_structure struc), ast, mlt (* For the test-suite : extraction to a temporary file + run ocamlc on it *) let compile f = try let args = ["ocamlc";"-I";Filename.dirname f;"-c";f^"i";f] in let res = CUnix.sys_command (Envars.ocamlfind ()) args in match res with | Unix.WEXITED 0 -> () | Unix.WEXITED n | Unix.WSIGNALED n | Unix.WSTOPPED n -> CErrors.user_err Pp.(str "Compilation of file " ++ str f ++ str " failed with exit code " ++ int n) with Unix.Unix_error (e,_,_) -> CErrors.user_err Pp.(str "Compilation of file " ++ str f ++ str " failed with error " ++ str (Unix.error_message e)) let remove f = if Sys.file_exists f then Sys.remove f let extract_and_compile l = if lang () != Ocaml then CErrors.user_err (Pp.str "This command only works with OCaml extraction"); let f = Filename.temp_file "testextraction" ".ml" in let () = full_extraction (Some f) l in let () = compile f in let () = remove f; remove (f^"i") in let base = Filename.chop_suffix f ".ml" in let () = remove (base^".cmo"); remove (base^".cmi") in Feedback.msg_notice (str "Extracted code successfully compiled") (* Show the extraction of the current ongoing proof *) let show_extraction ~pstate = init ~inner:true false false; let prf = Proof_global.get_proof pstate in let sigma, env = Pfedit.get_current_context pstate in let trms = Proof.partial_proof prf in let extr_term t = let ast, ty = extract_constr env sigma t in let mp = Lib.current_mp () in let l = Label.of_id (Proof_global.get_proof_name pstate) in let fake_ref = GlobRef.ConstRef (Constant.make2 mp l) in let decl = Dterm (fake_ref, ast, ty) in print_one_decl [] mp decl in Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl extr_term trms) coq-8.11.0/plugins/extraction/ExtrOcamlBasic.v0000644000175000017500000000312113612664533021153 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool [ true false ]. Extract Inductive option => option [ Some None ]. Extract Inductive unit => unit [ "()" ]. Extract Inductive list => list [ "[]" "( :: )" ]. Extract Inductive prod => "( * )" [ "" ]. (** NB: The "" above is a hack, but produce nicer code than "(,)" *) (** Mapping sumbool to bool and sumor to option is not always nicer, but it helps when realizing stuff like [lt_eq_lt_dec] *) Extract Inductive sumbool => bool [ true false ]. Extract Inductive sumor => option [ Some None ]. (** Restore laziness of andb, orb. NB: without these Extract Constant, andb/orb would be inlined by extraction in order to have laziness, producing inelegant (if ... then ... else false) and (if ... then true else ...). *) Extract Inlined Constant andb => "(&&)". Extract Inlined Constant orb => "(||)". coq-8.11.0/plugins/extraction/haskell.mli0000644000175000017500000000131713612664533020257 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bigint. Parameter bigint_opp : bigint -> bigint. Parameter bigint_twice : bigint -> bigint. Extract Inlined Constant bigint => "Big.big_int". Extract Inlined Constant bigint_zero => "Big.zero". Extract Inlined Constant bigint_succ => "Big.succ". Extract Inlined Constant bigint_opp => "Big.opp". Extract Inlined Constant bigint_twice => "Big.double". Definition bigint_of_nat : nat -> bigint := (fix loop acc n := match n with | O => acc | S n => loop (bigint_succ acc) n end) bigint_zero. Fixpoint bigint_of_pos p := match p with | xH => bigint_succ bigint_zero | xO p => bigint_twice (bigint_of_pos p) | xI p => bigint_succ (bigint_twice (bigint_of_pos p)) end. Fixpoint bigint_of_z z := match z with | Z0 => bigint_zero | Zpos p => bigint_of_pos p | Zneg p => bigint_opp (bigint_of_pos p) end. Fixpoint bigint_of_n n := match n with | N0 => bigint_zero | Npos p => bigint_of_pos p end. (** NB: as for [pred] or [minus], [nat_of_bigint], [n_of_bigint] and [pos_of_bigint] are total and return zero (resp. one) for non-positive inputs. *) Parameter bigint_natlike_rec : forall A, A -> (A->A) -> bigint -> A. Extract Constant bigint_natlike_rec => "Big.nat_rec". Definition nat_of_bigint : bigint -> nat := bigint_natlike_rec _ O S. Parameter bigint_poslike_rec : forall A, (A->A) -> (A->A) -> A -> bigint -> A. Extract Constant bigint_poslike_rec => "Big.positive_rec". Definition pos_of_bigint : bigint -> positive := bigint_poslike_rec _ xI xO xH. Parameter bigint_zlike_case : forall A, A -> (bigint->A) -> (bigint->A) -> bigint -> A. Extract Constant bigint_zlike_case => "Big.z_rec". Definition z_of_bigint : bigint -> Z := bigint_zlike_case _ Z0 (fun i => Zpos (pos_of_bigint i)) (fun i => Zneg (pos_of_bigint i)). Definition n_of_bigint : bigint -> N := bigint_zlike_case _ N0 (fun i => Npos (pos_of_bigint i)) (fun _ => N0). (* Tests: Definition small := 1234%nat. Definition big := 12345678901234567890%positive. Definition nat_0 := nat_of_bigint (bigint_of_nat 0). Definition nat_1 := nat_of_bigint (bigint_of_nat small). Definition pos_1 := pos_of_bigint (bigint_of_pos 1). Definition pos_2 := pos_of_bigint (bigint_of_pos big). Definition n_0 := n_of_bigint (bigint_of_n 0). Definition n_1 := n_of_bigint (bigint_of_n 1). Definition n_2 := n_of_bigint (bigint_of_n (Npos big)). Definition z_0 := z_of_bigint (bigint_of_z 0). Definition z_1 := z_of_bigint (bigint_of_z 1). Definition z_2 := z_of_bigint (bigint_of_z (Zpos big)). Definition z_m1 := z_of_bigint (bigint_of_z (-1)). Definition z_m2 := z_of_bigint (bigint_of_z (Zneg big)). Definition test := (nat_0, nat_1, pos_1, pos_2, n_0, n_1, n_2, z_0, z_1, z_2, z_m1, z_m2). Definition check := (O, small, xH, big, 0%N, 1%N, Npos big, 0%Z, 1%Z, Zpos big, (-1)%Z, Zneg big). Extraction "/tmp/test.ml" check test. ... and we check that test=check *) coq-8.11.0/plugins/extraction/big.ml0000644000175000017500000001354113612664533017226 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 0 then fp z else fn (opp z) let compare_case e l g x y = let s = compare x y in if s = 0 then e else if s<0 then l else g let nat_rec fO fS = let rec loop acc n = if sign n <= 0 then acc else loop (fS acc) (pred n) in loop fO let positive_rec f2p1 f2p f1 = let rec loop n = if le n one then f1 else let (q,r) = quomod n two in if eq r zero then f2p (loop q) else f2p1 (loop q) in loop let z_rec fO fp fn = z_case (fun _ -> fO) fp fn coq-8.11.0/plugins/extraction/ocaml.ml0000644000175000017500000007002313612664533017556 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* mt () | l -> str "fun " ++ prlist_with_sep (fun () -> str " ") Id.print l ++ str " ->" ++ spc () let pp_parameters l = (pp_boxed_tuple pp_tvar l ++ space_if (not (List.is_empty l))) let pp_string_parameters l = (pp_boxed_tuple str l ++ space_if (not (List.is_empty l))) let pp_letin pat def body = let fstline = str "let " ++ pat ++ str " =" ++ spc () ++ def in hv 0 (hv 0 (hov 2 fstline ++ spc () ++ str "in") ++ spc () ++ hov 0 body) (*s Ocaml renaming issues. *) let keywords = List.fold_right (fun s -> Id.Set.add (Id.of_string s)) [ "and"; "as"; "assert"; "begin"; "class"; "constraint"; "do"; "done"; "downto"; "else"; "end"; "exception"; "external"; "false"; "for"; "fun"; "function"; "functor"; "if"; "in"; "include"; "inherit"; "initializer"; "lazy"; "let"; "match"; "method"; "module"; "mutable"; "new"; "object"; "of"; "open"; "or"; "parser"; "private"; "rec"; "sig"; "struct"; "then"; "to"; "true"; "try"; "type"; "val"; "virtual"; "when"; "while"; "with"; "mod"; "land"; "lor"; "lxor"; "lsl"; "lsr"; "asr" ; "unit" ; "_" ; "__" ] Id.Set.empty (* Note: do not shorten [str "foo" ++ fnl ()] into [str "foo\n"], the '\n' character interacts badly with the Format boxing mechanism *) let pp_open mp = str ("open "^ string_of_modfile mp) ++ fnl () let pp_comment s = str "(* " ++ hov 0 s ++ str " *)" let pp_header_comment = function | None -> mt () | Some com -> pp_comment com ++ fnl2 () let then_nl pp = if Pp.ismt pp then mt () else pp ++ fnl () let pp_tdummy usf = if usf.tdummy || usf.tunknown then str "type __ = Obj.t" ++ fnl () else mt () let pp_mldummy usf = if usf.mldummy then str "let __ = let rec f _ = Obj.repr f in Obj.repr f" ++ fnl () else mt () let preamble _ comment used_modules usf = pp_header_comment comment ++ then_nl (prlist pp_open used_modules) ++ then_nl (pp_tdummy usf ++ pp_mldummy usf) let sig_preamble _ comment used_modules usf = pp_header_comment comment ++ then_nl (prlist pp_open used_modules) ++ then_nl (pp_tdummy usf) (*s The pretty-printer for Ocaml syntax*) (* Beware of the side-effects of [pp_global] and [pp_modname]. They are used to update table of content for modules. Many [let] below should not be altered since they force evaluation order. *) let str_global k r = if is_inline_custom r then find_custom r else Common.pp_global k r let pp_global k r = str (str_global k r) let pp_modname mp = str (Common.pp_module mp) (* grammar from OCaml 4.06 manual, "Prefix and infix symbols" *) let infix_symbols = ['=' ; '<' ; '>' ; '@' ; '^' ; ';' ; '&' ; '+' ; '-' ; '*' ; '/' ; '$' ; '%' ] let operator_chars = [ '!' ; '$' ; '%' ; '&' ; '*' ; '+' ; '-' ; '.' ; '/' ; ':' ; '<' ; '=' ; '>' ; '?' ; '@' ; '^' ; '|' ; '~' ] (* infix ops in OCaml, but disallowed by preceding grammar *) let builtin_infixes = [ "::" ; "," ] let substring_all_opchars s start stop = let rec check_char i = if i >= stop then true else List.mem s.[i] operator_chars && check_char (i+1) in check_char start let is_infix r = is_inline_custom r && (let s = find_custom r in let len = String.length s in len >= 3 && (* parenthesized *) (s.[0] == '(' && s.[len-1] == ')' && let inparens = String.trim (String.sub s 1 (len - 2)) in let inparens_len = String.length inparens in (* either, begins with infix symbol, any remainder is all operator chars *) (List.mem inparens.[0] infix_symbols && substring_all_opchars inparens 1 inparens_len) || (* or, starts with #, at least one more char, all are operator chars *) (inparens.[0] == '#' && inparens_len >= 2 && substring_all_opchars inparens 1 inparens_len) || (* or, is an OCaml built-in infix *) (List.mem inparens builtin_infixes))) let get_infix r = let s = find_custom r in String.sub s 1 (String.length s - 2) let get_ind = let open GlobRef in function | IndRef _ as r -> r | ConstructRef (ind,_) -> IndRef ind | _ -> assert false let pp_one_field r i = function | Some r -> pp_global Term r | None -> pp_global Type (get_ind r) ++ str "__" ++ int i let pp_field r fields i = pp_one_field r i (List.nth fields i) let pp_fields r fields = List.map_i (pp_one_field r) 0 fields (*s Pretty-printing of types. [par] is a boolean indicating whether parentheses are needed or not. *) let pp_type par vl t = let rec pp_rec par = function | Tmeta _ | Tvar' _ | Taxiom -> assert false | Tvar i -> (try pp_tvar (List.nth vl (pred i)) with Failure _ -> (str "'a" ++ int i)) | Tglob (r,[a1;a2]) when is_infix r -> pp_par par (pp_rec true a1 ++ str (get_infix r) ++ pp_rec true a2) | Tglob (r,[]) -> pp_global Type r | Tglob (GlobRef.IndRef(kn,0),l) when not (keep_singleton ()) && MutInd.equal kn (mk_ind "Coq.Init.Specif" "sig") -> pp_tuple_light pp_rec l | Tglob (r,l) -> pp_tuple_light pp_rec l ++ spc () ++ pp_global Type r | Tarr (t1,t2) -> pp_par par (pp_rec true t1 ++ spc () ++ str "->" ++ spc () ++ pp_rec false t2) | Tdummy _ -> str "__" | Tunknown -> str "__" in hov 0 (pp_rec par t) (*s Pretty-printing of expressions. [par] indicates whether parentheses are needed or not. [env] is the list of names for the de Bruijn variables. [args] is the list of collected arguments (already pretty-printed). *) let is_bool_patt p s = try let r = match p with | Pusual r -> r | Pcons (r,[]) -> r | _ -> raise Not_found in String.equal (find_custom r) s with Not_found -> false let is_ifthenelse = function | [|([],p1,_);([],p2,_)|] -> is_bool_patt p1 "true" && is_bool_patt p2 "false" | _ -> false let expr_needs_par = function | MLlam _ -> true | MLcase (_,_,[|_|]) -> false | MLcase (_,_,pv) -> not (is_ifthenelse pv) | _ -> false let rec pp_expr par env args = let apply st = pp_apply st par args and apply2 st = pp_apply2 st par args in function | MLrel n -> let id = get_db_name n env in (* Try to survive to the occurrence of a Dummy rel. TODO: we should get rid of this hack (cf. #592) *) let id = if Id.equal id dummy_name then Id.of_string "__" else id in apply (Id.print id) | MLapp (f,args') -> let stl = List.map (pp_expr true env []) args' in pp_expr par env (stl @ args) f | MLlam _ as a -> let fl,a' = collect_lams a in let fl = List.map id_of_mlid fl in let fl,env' = push_vars fl env in let st = pp_abst (List.rev fl) ++ pp_expr false env' [] a' in apply2 st | MLletin (id,a1,a2) -> let i,env' = push_vars [id_of_mlid id] env in let pp_id = Id.print (List.hd i) and pp_a1 = pp_expr false env [] a1 and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in hv 0 (apply2 (pp_letin pp_id pp_a1 pp_a2)) | MLglob r -> apply (pp_global Term r) | MLfix (i,ids,defs) -> let ids',env' = push_vars (List.rev (Array.to_list ids)) env in pp_fix par env' i (Array.of_list (List.rev ids'),defs) args | MLexn s -> (* An [MLexn] may be applied, but I don't really care. *) pp_par par (str "assert false" ++ spc () ++ str ("(* "^s^" *)")) | MLdummy k -> (* An [MLdummy] may be applied, but I don't really care. *) (match msg_of_implicit k with | "" -> str "__" | s -> str "__" ++ spc () ++ str ("(* "^s^" *)")) | MLmagic a -> pp_apply (str "Obj.magic") par (pp_expr true env [] a :: args) | MLaxiom -> pp_par par (str "failwith \"AXIOM TO BE REALIZED\"") | MLcons (_,r,a) as c -> assert (List.is_empty args); begin match a with | _ when is_native_char c -> pp_native_char c | [a1;a2] when is_infix r -> let pp = pp_expr true env [] in pp_par par (pp a1 ++ str (get_infix r) ++ pp a2) | _ when is_coinductive r -> let ne = not (List.is_empty a) in let tuple = space_if ne ++ pp_tuple (pp_expr true env []) a in pp_par par (str "lazy " ++ pp_par ne (pp_global Cons r ++ tuple)) | [] -> pp_global Cons r | _ -> let fds = get_record_fields r in if not (List.is_empty fds) then pp_record_pat (pp_fields r fds, List.map (pp_expr true env []) a) else let tuple = pp_tuple (pp_expr true env []) a in if String.is_empty (str_global Cons r) (* hack Extract Inductive prod *) then tuple else pp_par par (pp_global Cons r ++ spc () ++ tuple) end | MLtuple l -> assert (List.is_empty args); pp_boxed_tuple (pp_expr true env []) l | MLcase (_, t, pv) when is_custom_match pv -> if not (is_regular_match pv) then user_err Pp.(str "Cannot mix yet user-given match and general patterns."); let mkfun (ids,_,e) = if not (List.is_empty ids) then named_lams (List.rev ids) e else dummy_lams (ast_lift 1 e) 1 in let pp_branch tr = pp_expr true env [] (mkfun tr) ++ fnl () in let inner = str (find_custom_match pv) ++ fnl () ++ prvect pp_branch pv ++ pp_expr true env [] t in apply2 (hov 2 inner) | MLcase (typ, t, pv) -> let head = if not (is_coinductive_type typ) then pp_expr false env [] t else (str "Lazy.force" ++ spc () ++ pp_expr true env [] t) in (* First, can this match be printed as a mere record projection ? *) (try pp_record_proj par env typ t pv args with Impossible -> (* Second, can this match be printed as a let-in ? *) if Int.equal (Array.length pv) 1 then let s1,s2 = pp_one_pat env pv.(0) in hv 0 (apply2 (pp_letin s1 head s2)) else (* Third, can this match be printed as [if ... then ... else] ? *) (try apply2 (pp_ifthenelse env head pv) with Not_found -> (* Otherwise, standard match *) apply2 (v 0 (str "match " ++ head ++ str " with" ++ fnl () ++ pp_pat env pv)))) | MLuint i -> assert (args=[]); str "(" ++ str (Uint63.compile i) ++ str ")" | MLfloat f -> assert (args=[]); str "(" ++ str (Float64.compile f) ++ str ")" and pp_record_proj par env typ t pv args = (* Can a match be printed as a mere record projection ? *) let fields = record_fields_of_type typ in if List.is_empty fields then raise Impossible; if not (Int.equal (Array.length pv) 1) then raise Impossible; if has_deep_pattern pv then raise Impossible; let (ids,pat,body) = pv.(0) in let n = List.length ids in let no_patvar a = not (List.exists (ast_occurs_itvl 1 n) a) in let rel_i,a = match body with | MLrel i | MLmagic(MLrel i) when i <= n -> i,[] | MLapp(MLrel i, a) | MLmagic(MLapp(MLrel i, a)) | MLapp(MLmagic(MLrel i), a) when i<=n && no_patvar a -> i,a | _ -> raise Impossible in let magic = match body with MLmagic _ | MLapp(MLmagic _, _) -> true | _ -> false in let rec lookup_rel i idx = function | Prel j :: l -> if Int.equal i j then idx else lookup_rel i (idx+1) l | Pwild :: l -> lookup_rel i (idx+1) l | _ -> raise Impossible in let r,idx = match pat with | Pusual r -> r, n-rel_i | Pcons (r,l) -> r, lookup_rel rel_i 0 l | _ -> raise Impossible in if is_infix r then raise Impossible; let env' = snd (push_vars (List.rev_map id_of_mlid ids) env) in let pp_args = (List.map (pp_expr true env' []) a) @ args in let pp_head = pp_expr true env [] t ++ str "." ++ pp_field r fields idx in if magic then pp_apply (str "Obj.magic") par (pp_head :: pp_args) else pp_apply pp_head par pp_args and pp_record_pat (fields, args) = str "{ " ++ prlist_with_sep (fun () -> str ";" ++ spc ()) (fun (f,a) -> f ++ str " =" ++ spc () ++ a) (List.combine fields args) ++ str " }" and pp_cons_pat r ppl = if is_infix r && Int.equal (List.length ppl) 2 then List.hd ppl ++ str (get_infix r) ++ List.hd (List.tl ppl) else let fields = get_record_fields r in if not (List.is_empty fields) then pp_record_pat (pp_fields r fields, ppl) else if String.is_empty (str_global Cons r) then pp_boxed_tuple identity ppl (* Hack Extract Inductive prod *) else pp_global Cons r ++ space_if (not (List.is_empty ppl)) ++ pp_boxed_tuple identity ppl and pp_gen_pat ids env = function | Pcons (r, l) -> pp_cons_pat r (List.map (pp_gen_pat ids env) l) | Pusual r -> pp_cons_pat r (List.map Id.print ids) | Ptuple l -> pp_boxed_tuple (pp_gen_pat ids env) l | Pwild -> str "_" | Prel n -> Id.print (get_db_name n env) and pp_ifthenelse env expr pv = match pv with | [|([],tru,the);([],fal,els)|] when (is_bool_patt tru "true") && (is_bool_patt fal "false") -> hv 0 (hov 2 (str "if " ++ expr) ++ spc () ++ hov 2 (str "then " ++ hov 2 (pp_expr (expr_needs_par the) env [] the)) ++ spc () ++ hov 2 (str "else " ++ hov 2 (pp_expr (expr_needs_par els) env [] els))) | _ -> raise Not_found and pp_one_pat env (ids,p,t) = let ids',env' = push_vars (List.rev_map id_of_mlid ids) env in pp_gen_pat (List.rev ids') env' p, pp_expr (expr_needs_par t) env' [] t and pp_pat env pv = prvecti (fun i x -> let s1,s2 = pp_one_pat env x in hv 2 (hov 4 (str "| " ++ s1 ++ str " ->") ++ spc () ++ hov 2 s2) ++ if Int.equal i (Array.length pv - 1) then mt () else fnl ()) pv and pp_function env t = let bl,t' = collect_lams t in let bl,env' = push_vars (List.map id_of_mlid bl) env in match t' with | MLcase(Tglob(r,_),MLrel 1,pv) when not (is_coinductive r) && List.is_empty (get_record_fields r) && not (is_custom_match pv) -> if not (ast_occurs 1 (MLcase(Tunknown,MLaxiom,pv))) then pr_binding (List.rev (List.tl bl)) ++ str " = function" ++ fnl () ++ v 0 (pp_pat env' pv) else pr_binding (List.rev bl) ++ str " = match " ++ Id.print (List.hd bl) ++ str " with" ++ fnl () ++ v 0 (pp_pat env' pv) | _ -> pr_binding (List.rev bl) ++ str " =" ++ fnl () ++ str " " ++ hov 2 (pp_expr false env' [] t') (*s names of the functions ([ids]) are already pushed in [env], and passed here just for convenience. *) and pp_fix par env i (ids,bl) args = pp_par par (v 0 (str "let rec " ++ prvect_with_sep (fun () -> fnl () ++ str "and ") (fun (fi,ti) -> Id.print fi ++ pp_function env ti) (Array.map2 (fun id b -> (id,b)) ids bl) ++ fnl () ++ hov 2 (str "in " ++ pp_apply (Id.print ids.(i)) false args))) (* Ad-hoc double-newline in v boxes, with enough negative whitespace to avoid indenting the intermediate blank line *) let cut2 () = brk (0,-100000) ++ brk (0,0) let pp_val e typ = hov 4 (str "(** val " ++ e ++ str " :" ++ spc () ++ pp_type false [] typ ++ str " **)") ++ cut2 () (*s Pretty-printing of [Dfix] *) let pp_Dfix (rv,c,t) = let names = Array.map (fun r -> if is_inline_custom r then mt () else pp_global Term r) rv in let rec pp init i = if i >= Array.length rv then mt () else let void = is_inline_custom rv.(i) || (not (is_custom rv.(i)) && match c.(i) with MLexn "UNUSED" -> true | _ -> false) in if void then pp init (i+1) else let def = if is_custom rv.(i) then str " = " ++ str (find_custom rv.(i)) else pp_function (empty_env ()) c.(i) in (if init then mt () else cut2 ()) ++ pp_val names.(i) t.(i) ++ str (if init then "let rec " else "and ") ++ names.(i) ++ def ++ pp false (i+1) in pp true 0 (*s Pretty-printing of inductive types declaration. *) let pp_equiv param_list name = function | NoEquiv, _ -> mt () | Equiv kn, i -> str " = " ++ pp_parameters param_list ++ pp_global Type (GlobRef.IndRef (MutInd.make1 kn,i)) | RenEquiv ren, _ -> str " = " ++ pp_parameters param_list ++ str (ren^".") ++ name let pp_one_ind prefix ip_equiv pl name cnames ctyps = let pl = rename_tvars keywords pl in let pp_constructor i typs = (if Int.equal i 0 then mt () else fnl ()) ++ hov 3 (str "| " ++ cnames.(i) ++ (if List.is_empty typs then mt () else str " of ") ++ prlist_with_sep (fun () -> spc () ++ str "* ") (pp_type true pl) typs) in pp_parameters pl ++ str prefix ++ name ++ pp_equiv pl name ip_equiv ++ str " =" ++ if Int.equal (Array.length ctyps) 0 then str " unit (* empty inductive *)" else fnl () ++ v 0 (prvecti pp_constructor ctyps) let pp_logical_ind packet = pp_comment (Id.print packet.ip_typename ++ str " : logical inductive") ++ fnl () ++ pp_comment (str "with constructors : " ++ prvect_with_sep spc Id.print packet.ip_consnames) ++ fnl () let pp_singleton kn packet = let name = pp_global Type (GlobRef.IndRef (kn,0)) in let l = rename_tvars keywords packet.ip_vars in hov 2 (str "type " ++ pp_parameters l ++ name ++ str " =" ++ spc () ++ pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ pp_comment (str "singleton inductive, whose constructor was " ++ Id.print packet.ip_consnames.(0))) let pp_record kn fields ip_equiv packet = let ind = GlobRef.IndRef (kn,0) in let name = pp_global Type ind in let fieldnames = pp_fields ind fields in let l = List.combine fieldnames packet.ip_types.(0) in let pl = rename_tvars keywords packet.ip_vars in str "type " ++ pp_parameters pl ++ name ++ pp_equiv pl name ip_equiv ++ str " = { "++ hov 0 (prlist_with_sep (fun () -> str ";" ++ spc ()) (fun (p,t) -> p ++ str " : " ++ pp_type true pl t) l) ++ str " }" let pp_coind pl name = let pl = rename_tvars keywords pl in pp_parameters pl ++ name ++ str " = " ++ pp_parameters pl ++ str "__" ++ name ++ str " Lazy.t" ++ fnl() ++ str "and " let pp_ind co kn ind = let prefix = if co then "__" else "" in let initkwd = str "type " in let nextkwd = fnl () ++ str "and " in let names = Array.mapi (fun i p -> if p.ip_logical then mt () else pp_global Type (GlobRef.IndRef (kn,i))) ind.ind_packets in let cnames = Array.mapi (fun i p -> if p.ip_logical then [||] else Array.mapi (fun j _ -> pp_global Cons (GlobRef.ConstructRef ((kn,i),j+1))) p.ip_types) ind.ind_packets in let rec pp i kwd = if i >= Array.length ind.ind_packets then mt () else let ip = (kn,i) in let ip_equiv = ind.ind_equiv, i in let p = ind.ind_packets.(i) in if is_custom (GlobRef.IndRef ip) then pp (i+1) kwd else if p.ip_logical then pp_logical_ind p ++ pp (i+1) kwd else kwd ++ (if co then pp_coind p.ip_vars names.(i) else mt ()) ++ pp_one_ind prefix ip_equiv p.ip_vars names.(i) cnames.(i) p.ip_types ++ pp (i+1) nextkwd in pp 0 initkwd (*s Pretty-printing of a declaration. *) let pp_mind kn i = match i.ind_kind with | Singleton -> pp_singleton kn i.ind_packets.(0) | Coinductive -> pp_ind true kn i | Record fields -> pp_record kn fields (i.ind_equiv,0) i.ind_packets.(0) | Standard -> pp_ind false kn i let pp_decl = function | Dtype (r,_,_) when is_inline_custom r -> mt () | Dterm (r,_,_) when is_inline_custom r -> mt () | Dind (kn,i) -> pp_mind kn i | Dtype (r, l, t) -> let name = pp_global Type r in let l = rename_tvars keywords l in let ids, def = try let ids,s = find_type_custom r in pp_string_parameters ids, str " =" ++ spc () ++ str s with Not_found -> pp_parameters l, if t == Taxiom then str " (* AXIOM TO BE REALIZED *)" else str " =" ++ spc () ++ pp_type false l t in hov 2 (str "type " ++ ids ++ name ++ def) | Dterm (r, a, t) -> let def = if is_custom r then str (" = " ^ find_custom r) else pp_function (empty_env ()) a in let name = pp_global Term r in pp_val name t ++ hov 0 (str "let " ++ name ++ def ++ mt ()) | Dfix (rv,defs,typs) -> pp_Dfix (rv,defs,typs) let pp_spec = function | Sval (r,_) when is_inline_custom r -> mt () | Stype (r,_,_) when is_inline_custom r -> mt () | Sind (kn,i) -> pp_mind kn i | Sval (r,t) -> let def = pp_type false [] t in let name = pp_global Term r in hov 2 (str "val " ++ name ++ str " :" ++ spc () ++ def) | Stype (r,vl,ot) -> let name = pp_global Type r in let l = rename_tvars keywords vl in let ids, def = try let ids, s = find_type_custom r in pp_string_parameters ids, str " =" ++ spc () ++ str s with Not_found -> let ids = pp_parameters l in match ot with | None -> ids, mt () | Some Taxiom -> ids, str " (* AXIOM TO BE REALIZED *)" | Some t -> ids, str " =" ++ spc () ++ pp_type false l t in hov 2 (str "type " ++ ids ++ name ++ def) let rec pp_specif = function | (_,Spec (Sval _ as s)) -> pp_spec s | (l,Spec s) -> (match Common.get_duplicate (top_visible_mp ()) l with | None -> pp_spec s | Some ren -> hov 1 (str ("module "^ren^" : sig") ++ fnl () ++ pp_spec s) ++ fnl () ++ str "end" ++ fnl () ++ str ("include module type of struct include "^ren^" end")) | (l,Smodule mt) -> let def = pp_module_type [] mt in let name = pp_modname (MPdot (top_visible_mp (), l)) in hov 1 (str "module " ++ name ++ str " :" ++ fnl () ++ def) ++ (match Common.get_duplicate (top_visible_mp ()) l with | None -> Pp.mt () | Some ren -> fnl () ++ hov 1 (str ("module "^ren^" :") ++ spc () ++ str "module type of struct include " ++ name ++ str " end")) | (l,Smodtype mt) -> let def = pp_module_type [] mt in let name = pp_modname (MPdot (top_visible_mp (), l)) in hov 1 (str "module type " ++ name ++ str " =" ++ fnl () ++ def) ++ (match Common.get_duplicate (top_visible_mp ()) l with | None -> Pp.mt () | Some ren -> fnl () ++ str ("module type "^ren^" = ") ++ name) and pp_module_type params = function | MTident kn -> pp_modname kn | MTfunsig (mbid, mt, mt') -> let typ = pp_module_type [] mt in let name = pp_modname (MPbound mbid) in let def = pp_module_type (MPbound mbid :: params) mt' in str "functor (" ++ name ++ str ":" ++ typ ++ str ") ->" ++ fnl () ++ def | MTsig (mp, sign) -> push_visible mp params; let try_pp_specif l x = let px = pp_specif x in if Pp.ismt px then l else px::l in (* We cannot use fold_right here due to side effects in pp_specif *) let l = List.fold_left try_pp_specif [] sign in let l = List.rev l in pop_visible (); str "sig" ++ fnl () ++ (if List.is_empty l then mt () else v 1 (str " " ++ prlist_with_sep cut2 identity l) ++ fnl ()) ++ str "end" | MTwith(mt,ML_With_type(idl,vl,typ)) -> let ids = pp_parameters (rename_tvars keywords vl) in let mp_mt = msid_of_mt mt in let l,idl' = List.sep_last idl in let mp_w = List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl' in let r = GlobRef.ConstRef (Constant.make2 mp_w (Label.of_id l)) in push_visible mp_mt []; let pp_w = str " with type " ++ ids ++ pp_global Type r in pop_visible(); pp_module_type [] mt ++ pp_w ++ str " = " ++ pp_type false vl typ | MTwith(mt,ML_With_module(idl,mp)) -> let mp_mt = msid_of_mt mt in let mp_w = List.fold_left (fun mp id -> MPdot(mp,Label.of_id id)) mp_mt idl in push_visible mp_mt []; let pp_w = str " with module " ++ pp_modname mp_w in pop_visible (); pp_module_type [] mt ++ pp_w ++ str " = " ++ pp_modname mp let is_short = function MEident _ | MEapply _ -> true | _ -> false let rec pp_structure_elem = function | (l,SEdecl d) -> (match Common.get_duplicate (top_visible_mp ()) l with | None -> pp_decl d | Some ren -> hov 1 (str ("module "^ren^" = struct") ++ fnl () ++ pp_decl d) ++ fnl () ++ str "end" ++ fnl () ++ str ("include "^ren)) | (l,SEmodule m) -> let typ = (* virtual printing of the type, in order to have a correct mli later*) if Common.get_phase () == Pre then str ": " ++ pp_module_type [] m.ml_mod_type else mt () in let def = pp_module_expr [] m.ml_mod_expr in let name = pp_modname (MPdot (top_visible_mp (), l)) in hov 1 (str "module " ++ name ++ typ ++ str " =" ++ (if is_short m.ml_mod_expr then spc () else fnl ()) ++ def) ++ (match Common.get_duplicate (top_visible_mp ()) l with | Some ren -> fnl () ++ str ("module "^ren^" = ") ++ name | None -> mt ()) | (l,SEmodtype m) -> let def = pp_module_type [] m in let name = pp_modname (MPdot (top_visible_mp (), l)) in hov 1 (str "module type " ++ name ++ str " =" ++ fnl () ++ def) ++ (match Common.get_duplicate (top_visible_mp ()) l with | None -> mt () | Some ren -> fnl () ++ str ("module type "^ren^" = ") ++ name) and pp_module_expr params = function | MEident mp -> pp_modname mp | MEapply (me, me') -> pp_module_expr [] me ++ str "(" ++ pp_module_expr [] me' ++ str ")" | MEfunctor (mbid, mt, me) -> let name = pp_modname (MPbound mbid) in let typ = pp_module_type [] mt in let def = pp_module_expr (MPbound mbid :: params) me in str "functor (" ++ name ++ str ":" ++ typ ++ str ") ->" ++ fnl () ++ def | MEstruct (mp, sel) -> push_visible mp params; let try_pp_structure_elem l x = let px = pp_structure_elem x in if Pp.ismt px then l else px::l in (* We cannot use fold_right here due to side effects in pp_structure_elem *) let l = List.fold_left try_pp_structure_elem [] sel in let l = List.rev l in pop_visible (); str "struct" ++ fnl () ++ (if List.is_empty l then mt () else v 1 (str " " ++ prlist_with_sep cut2 identity l) ++ fnl ()) ++ str "end" let rec prlist_sep_nonempty sep f = function | [] -> mt () | [h] -> f h | h::t -> let e = f h in let r = prlist_sep_nonempty sep f t in if Pp.ismt e then r else e ++ sep () ++ r let do_struct f s = let ppl (mp,sel) = push_visible mp []; let p = prlist_sep_nonempty cut2 f sel in (* for monolithic extraction, we try to simulate the unavailability of [MPfile] in names by artificially nesting these [MPfile] *) (if modular () then pop_visible ()); p in let p = prlist_sep_nonempty cut2 ppl s in (if not (modular ()) then repeat (List.length s) pop_visible ()); v 0 p ++ fnl () let pp_struct s = do_struct pp_structure_elem s let pp_signature s = do_struct pp_specif s let ocaml_descr = { keywords = keywords; file_suffix = ".ml"; file_naming = file_of_modfile; preamble = preamble; pp_struct = pp_struct; sig_suffix = Some ".mli"; sig_preamble = sig_preamble; pp_sig = pp_signature; pp_decl = pp_decl; } coq-8.11.0/plugins/extraction/common.mli0000644000175000017500000000555513612664533020134 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Pp.t val fnl2 : unit -> Pp.t val space_if : bool -> Pp.t val pp_par : bool -> Pp.t -> Pp.t (** [pp_apply] : a head part applied to arguments, possibly with parenthesis *) val pp_apply : Pp.t -> bool -> Pp.t list -> Pp.t (** Same as [pp_apply], but with also protection of the head by parenthesis *) val pp_apply2 : Pp.t -> bool -> Pp.t list -> Pp.t val pp_tuple_light : (bool -> 'a -> Pp.t) -> 'a list -> Pp.t val pp_tuple : ('a -> Pp.t) -> 'a list -> Pp.t val pp_boxed_tuple : ('a -> Pp.t) -> 'a list -> Pp.t val pr_binding : Id.t list -> Pp.t val rename_id : Id.t -> Id.Set.t -> Id.t type env = Id.t list * Id.Set.t val empty_env : unit -> env val rename_vars: Id.Set.t -> Id.t list -> env val rename_tvars: Id.Set.t -> Id.t list -> Id.t list val push_vars : Id.t list -> env -> Id.t list * env val get_db_name : int -> env -> Id.t type phase = Pre | Impl | Intf val set_phase : phase -> unit val get_phase : unit -> phase val opened_libraries : unit -> ModPath.t list type kind = Term | Type | Cons | Mod val pp_global : kind -> GlobRef.t -> string val pp_module : ModPath.t -> string val top_visible_mp : unit -> ModPath.t (* In [push_visible], the [module_path list] corresponds to module parameters, the innermost one coming first in the list *) val push_visible : ModPath.t -> ModPath.t list -> unit val pop_visible : unit -> unit val get_duplicate : ModPath.t -> Label.t -> string option type reset_kind = AllButExternal | Everything val reset_renaming_tables : reset_kind -> unit val set_keywords : Id.Set.t -> unit (** For instance: [mk_ind "Coq.Init.Datatypes" "nat"] *) val mk_ind : string -> string -> MutInd.t (** Special hack for constants of type Ascii.ascii : if an [Extract Inductive ascii => char] has been declared, then the constants are directly turned into chars *) val is_native_char : ml_ast -> bool val get_native_char : ml_ast -> char val pp_native_char : ml_ast -> Pp.t coq-8.11.0/plugins/extraction/extraction.mli0000644000175000017500000000276113612664533021020 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constant.t -> Opaqueproof.opaque constant_body -> ml_decl val extract_constant_spec : env -> Constant.t -> 'a constant_body -> ml_spec (** For extracting "module ... with ..." declaration *) val extract_with_type : env -> evar_map -> EConstr.t -> ( Id.t list * ml_type ) option val extract_fixpoint : env -> evar_map -> Constant.t array -> (EConstr.t, EConstr.types) Constr.prec_declaration -> ml_decl val extract_inductive : env -> MutInd.t -> ml_ind (** For Extraction Compute and Show Extraction *) val extract_constr : env -> evar_map -> EConstr.t -> ml_ast * ml_type (*s Is a [ml_decl] or a [ml_spec] logical ? *) val logical_decl : ml_decl -> bool val logical_spec : ml_spec -> bool coq-8.11.0/plugins/extraction/ExtrOCamlInt63.v0000644000175000017500000000452613612664533021007 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool [ true false ]. Extract Inductive prod => "( * )" [ "" ]. Extract Inductive comparison => int [ "0" "(-1)" "1" ]. Extract Inductive DoubleType.carry => "Uint63.carry" [ "Uint63.C0" "Uint63.C1" ]. (** Primitive types and operators. *) Extract Constant Int63.int => "Uint63.t". Extraction Inline Int63.int. (* Otherwise, the name conflicts with the primitive OCaml type [int] *) Extract Constant Int63.lsl => "Uint63.l_sl". Extract Constant Int63.lsr => "Uint63.l_sr". Extract Constant Int63.land => "Uint63.l_and". Extract Constant Int63.lor => "Uint63.l_or". Extract Constant Int63.lxor => "Uint63.l_xor". Extract Constant Int63.add => "Uint63.add". Extract Constant Int63.sub => "Uint63.sub". Extract Constant Int63.mul => "Uint63.mul". Extract Constant Int63.mulc => "Uint63.mulc". Extract Constant Int63.div => "Uint63.div". Extract Constant Int63.mod => "Uint63.rem". Extract Constant Int63.eqb => "Uint63.equal". Extract Constant Int63.ltb => "Uint63.lt". Extract Constant Int63.leb => "Uint63.le". Extract Constant Int63.addc => "Uint63.addc". Extract Constant Int63.addcarryc => "Uint63.addcarryc". Extract Constant Int63.subc => "Uint63.subc". Extract Constant Int63.subcarryc => "Uint63.subcarryc". Extract Constant Int63.diveucl => "Uint63.diveucl". Extract Constant Int63.diveucl_21 => "Uint63.div21". Extract Constant Int63.addmuldiv => "Uint63.addmuldiv". Extract Constant Int63.compare => "Uint63.compare". Extract Constant Int63.head0 => "Uint63.head0". Extract Constant Int63.tail0 => "Uint63.tail0". coq-8.11.0/plugins/extraction/ExtrOCamlFloats.v0000644000175000017500000000562613612664533021336 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool [ true false ]. Extract Inductive prod => "( * )" [ "" ]. Extract Inductive FloatClass.float_class => "Float64.float_class" [ "PNormal" "NNormal" "PSubn" "NSubn" "PZero" "NZero" "PInf" "NInf" "NaN" ]. Extract Inductive PrimFloat.float_comparison => "Float64.float_comparison" [ "FEq" "FLt" "FGt" "FNotComparable" ]. (** Primitive types and operators. *) Extract Constant PrimFloat.float => "Float64.t". Extraction Inline PrimFloat.float. (* Otherwise, the name conflicts with the primitive OCaml type [float] *) Extract Constant PrimFloat.classify => "Float64.classify". Extract Constant PrimFloat.abs => "Float64.abs". Extract Constant PrimFloat.sqrt => "Float64.sqrt". Extract Constant PrimFloat.opp => "Float64.opp". Extract Constant PrimFloat.eqb => "Float64.eq". Extract Constant PrimFloat.ltb => "Float64.lt". Extract Constant PrimFloat.leb => "Float64.le". Extract Constant PrimFloat.compare => "Float64.compare". Extract Constant PrimFloat.mul => "Float64.mul". Extract Constant PrimFloat.add => "Float64.add". Extract Constant PrimFloat.sub => "Float64.sub". Extract Constant PrimFloat.div => "Float64.div". Extract Constant PrimFloat.of_int63 => "Float64.of_int63". Extract Constant PrimFloat.normfr_mantissa => "Float64.normfr_mantissa". Extract Constant PrimFloat.frshiftexp => "Float64.frshiftexp". Extract Constant PrimFloat.ldshiftexp => "Float64.ldshiftexp". Extract Constant PrimFloat.next_up => "Float64.next_up". Extract Constant PrimFloat.next_down => "Float64.next_down". coq-8.11.0/plugins/extraction/ExtrHaskellBasic.v0000644000175000017500000000140713612664533021510 0ustar treinentreinen(** Extraction to Haskell : use of basic Haskell types *) Require Coq.extraction.Extraction. Extract Inductive bool => "Prelude.Bool" [ "Prelude.True" "Prelude.False" ]. Extract Inductive option => "Prelude.Maybe" [ "Prelude.Just" "Prelude.Nothing" ]. Extract Inductive unit => "()" [ "()" ]. Extract Inductive list => "([])" [ "([])" "(:)" ]. Extract Inductive prod => "(,)" [ "(,)" ]. Extract Inductive sumbool => "Prelude.Bool" [ "Prelude.True" "Prelude.False" ]. Extract Inductive sumor => "Prelude.Maybe" [ "Prelude.Just" "Prelude.Nothing" ]. Extract Inductive sum => "Prelude.Either" [ "Prelude.Left" "Prelude.Right" ]. Extract Inlined Constant andb => "(Prelude.&&)". Extract Inlined Constant orb => "(Prelude.||)". Extract Inlined Constant negb => "Prelude.not". coq-8.11.0/plugins/extraction/ExtrHaskellZInteger.v0000644000175000017500000000152513612664533022217 0ustar treinentreinen(** Extraction of [Z] into Haskell's [Integer] *) Require Coq.extraction.Extraction. Require Import ZArith. Require Import ExtrHaskellZNum. (** Disclaimer: trying to obtain efficient certified programs by extracting [Z] into [Integer] isn't necessarily a good idea. See comments in [ExtrOcamlNatInt.v]. *) Extract Inductive positive => "Prelude.Integer" [ "(\x -> 2 Prelude.* x Prelude.+ 1)" "(\x -> 2 Prelude.* x)" "1" ] "(\fI fO fH n -> if n Prelude.== 1 then fH () else if Prelude.odd n then fI (n `Prelude.div` 2) else fO (n `Prelude.div` 2))". Extract Inductive Z => "Prelude.Integer" [ "0" "(\x -> x)" "Prelude.negate" ] "(\fO fP fN n -> if n Prelude.== 0 then fO () else if n Prelude.> 0 then fP n else fN (Prelude.negate n))". coq-8.11.0/plugins/extraction/table.ml0000644000175000017500000007337313612664533017565 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* MutInd.equal kn kn' | ConstRef _ | VarRef _ -> false let repr_of_r = let open GlobRef in function | ConstRef kn -> Constant.repr2 kn | IndRef (kn,_) | ConstructRef ((kn,_),_) -> MutInd.repr2 kn | VarRef v -> KerName.repr (Lib.make_kn v) let modpath_of_r r = let mp,_ = repr_of_r r in mp let label_of_r r = let _,l = repr_of_r r in l let rec base_mp = function | MPdot (mp,l) -> base_mp mp | mp -> mp let is_modfile = function | MPfile _ -> true | _ -> false let raw_string_of_modfile = function | MPfile f -> String.capitalize_ascii (Id.to_string (List.hd (DirPath.repr f))) | _ -> assert false let is_toplevel mp = ModPath.equal mp ModPath.initial || ModPath.equal mp (Lib.current_mp ()) let at_toplevel mp = is_modfile mp || is_toplevel mp let mp_length mp = let mp0 = Lib.current_mp () in let rec len = function | mp when ModPath.equal mp mp0 -> 1 | MPdot (mp,_) -> 1 + len mp | _ -> 1 in len mp let rec prefixes_mp mp = match mp with | MPdot (mp',_) -> MPset.add mp (prefixes_mp mp') | _ -> MPset.singleton mp let rec get_nth_label_mp n = function | MPdot (mp,l) -> if Int.equal n 1 then l else get_nth_label_mp (n-1) mp | _ -> failwith "get_nth_label: not enough MPdot" let common_prefix_from_list mp0 mpl = let prefixes = prefixes_mp mp0 in let rec f = function | [] -> None | mp :: l -> if MPset.mem mp prefixes then Some mp else f l in f mpl let rec parse_labels2 ll mp1 = function | mp when ModPath.equal mp1 mp -> mp,ll | MPdot (mp,l) -> parse_labels2 (l::ll) mp1 mp | mp -> mp,ll let labels_of_ref r = let mp_top = Lib.current_mp () in let mp,l = repr_of_r r in parse_labels2 [l] mp_top mp (*S The main tables: constants, inductives, records, ... *) (* These tables are not registered within coq save/undo mechanism since we reset their contents at each run of Extraction *) (* We use [constant_body] (resp. [mutual_inductive_body]) as checksum to ensure that the table contents aren't outdated. *) (*s Constants tables. *) let typedefs = ref (Cmap_env.empty : (Opaqueproof.opaque constant_body * ml_type) Cmap_env.t) let init_typedefs () = typedefs := Cmap_env.empty let add_typedef kn cb t = typedefs := Cmap_env.add kn (cb,t) !typedefs let lookup_typedef kn cb = try let (cb0,t) = Cmap_env.find kn !typedefs in if cb0 == cb then Some t else None with Not_found -> None let cst_types = ref (Cmap_env.empty : (Opaqueproof.opaque constant_body * ml_schema) Cmap_env.t) let init_cst_types () = cst_types := Cmap_env.empty let add_cst_type kn cb s = cst_types := Cmap_env.add kn (cb,s) !cst_types let lookup_cst_type kn cb = try let (cb0,s) = Cmap_env.find kn !cst_types in if cb0 == cb then Some s else None with Not_found -> None (*s Inductives table. *) let inductives = ref (Mindmap_env.empty : (mutual_inductive_body * ml_ind) Mindmap_env.t) let init_inductives () = inductives := Mindmap_env.empty let add_ind kn mib ml_ind = inductives := Mindmap_env.add kn (mib,ml_ind) !inductives let lookup_ind kn mib = try let (mib0,ml_ind) = Mindmap_env.find kn !inductives in if mib == mib0 then Some ml_ind else None with Not_found -> None let unsafe_lookup_ind kn = snd (Mindmap_env.find kn !inductives) let inductive_kinds = ref (Mindmap_env.empty : inductive_kind Mindmap_env.t) let init_inductive_kinds () = inductive_kinds := Mindmap_env.empty let add_inductive_kind kn k = inductive_kinds := Mindmap_env.add kn k !inductive_kinds let is_coinductive r = let kn = let open GlobRef in match r with | ConstructRef ((kn,_),_) -> kn | IndRef (kn,_) -> kn | _ -> assert false in try Mindmap_env.find kn !inductive_kinds == Coinductive with Not_found -> false let is_coinductive_type = function | Tglob (r,_) -> is_coinductive r | _ -> false let get_record_fields r = let kn = let open GlobRef in match r with | ConstructRef ((kn,_),_) -> kn | IndRef (kn,_) -> kn | _ -> assert false in try match Mindmap_env.find kn !inductive_kinds with | Record f -> f | _ -> [] with Not_found -> [] let record_fields_of_type = function | Tglob (r,_) -> get_record_fields r | _ -> [] (*s Recursors table. *) (* NB: here we can use the equivalence between canonical and user constant names. *) let recursors = ref KNset.empty let init_recursors () = recursors := KNset.empty let add_recursors env ind = let kn = MutInd.canonical ind in let mk_kn id = KerName.make (KerName.modpath kn) (Label.of_id id) in let mib = Environ.lookup_mind ind env in Array.iter (fun mip -> let id = mip.mind_typename in let kn_rec = mk_kn (Nameops.add_suffix id "_rec") and kn_rect = mk_kn (Nameops.add_suffix id "_rect") in recursors := KNset.add kn_rec (KNset.add kn_rect !recursors)) mib.mind_packets let is_recursor = function | GlobRef.ConstRef c -> KNset.mem (Constant.canonical c) !recursors | _ -> false (*s Record tables. *) (* NB: here, working modulo name equivalence is ok *) let projs = ref (GlobRef.Map.empty : (inductive*int) GlobRef.Map.t) let init_projs () = projs := GlobRef.Map.empty let add_projection n kn ip = projs := GlobRef.Map.add (GlobRef.ConstRef kn) (ip,n) !projs let is_projection r = GlobRef.Map.mem r !projs let projection_arity r = snd (GlobRef.Map.find r !projs) let projection_info r = GlobRef.Map.find r !projs (*s Table of used axioms *) let info_axioms = ref Refset'.empty let log_axioms = ref Refset'.empty let init_axioms () = info_axioms := Refset'.empty; log_axioms := Refset'.empty let add_info_axiom r = info_axioms := Refset'.add r !info_axioms let remove_info_axiom r = info_axioms := Refset'.remove r !info_axioms let add_log_axiom r = log_axioms := Refset'.add r !log_axioms let opaques = ref Refset'.empty let init_opaques () = opaques := Refset'.empty let add_opaque r = opaques := Refset'.add r !opaques let remove_opaque r = opaques := Refset'.remove r !opaques (*s Extraction modes: modular or monolithic, library or minimal ? Nota: - Recursive Extraction : monolithic, minimal - Separate Extraction : modular, minimal - Extraction Library : modular, library *) let modular_ref = ref false let library_ref = ref false let set_modular b = modular_ref := b let modular () = !modular_ref let set_library b = library_ref := b let library () = !library_ref let extrcompute = ref false let set_extrcompute b = extrcompute := b let is_extrcompute () = !extrcompute (*s Printing. *) (* The following functions work even on objects not in [Global.env ()]. Warning: for inductive objects, this only works if an [extract_inductive] have been done earlier, otherwise we can only ask the Nametab about currently visible objects. *) let safe_basename_of_global r = let last_chance r = try Nametab.basename_of_global r with Not_found -> anomaly (Pp.str "Inductive object unknown to extraction and not globally visible.") in let open GlobRef in match r with | ConstRef kn -> Label.to_id (Constant.label kn) | IndRef (kn,0) -> Label.to_id (MutInd.label kn) | IndRef (kn,i) -> (try (unsafe_lookup_ind kn).ind_packets.(i).ip_typename with Not_found -> last_chance r) | ConstructRef ((kn,i),j) -> (try (unsafe_lookup_ind kn).ind_packets.(i).ip_consnames.(j-1) with Not_found -> last_chance r) | VarRef v -> v let string_of_global r = try string_of_qualid (Nametab.shortest_qualid_of_global Id.Set.empty r) with Not_found -> Id.to_string (safe_basename_of_global r) let safe_pr_global r = str (string_of_global r) (* idem, but with qualification, and only for constants. *) let safe_pr_long_global r = try Printer.pr_global r with Not_found -> match r with | GlobRef.ConstRef kn -> let mp,l = Constant.repr2 kn in str ((ModPath.to_string mp)^"."^(Label.to_string l)) | _ -> assert false let pr_long_mp mp = let lid = DirPath.repr (Nametab.dirpath_of_module mp) in str (String.concat "." (List.rev_map Id.to_string lid)) let pr_long_global ref = pr_path (Nametab.path_of_global ref) (*S Warning and Error messages. *) let err s = user_err ~hdr:"Extraction" s let warn_extraction_axiom_to_realize = CWarnings.create ~name:"extraction-axiom-to-realize" ~category:"extraction" (fun axioms -> let s = if Int.equal (List.length axioms) 1 then "axiom" else "axioms" in strbrk ("The following "^s^" must be realized in the extracted code:") ++ hov 1 (spc () ++ prlist_with_sep spc safe_pr_global axioms) ++ str "." ++ fnl ()) let warn_extraction_logical_axiom = CWarnings.create ~name:"extraction-logical-axiom" ~category:"extraction" (fun axioms -> let s = if Int.equal (List.length axioms) 1 then "axiom was" else "axioms were" in (strbrk ("The following logical "^s^" encountered:") ++ hov 1 (spc () ++ prlist_with_sep spc safe_pr_global axioms ++ str ".\n") ++ strbrk "Having invalid logical axiom in the environment when extracting" ++ spc () ++ strbrk "may lead to incorrect or non-terminating ML terms." ++ fnl ())) let warning_axioms () = let info_axioms = Refset'.elements !info_axioms in if not (List.is_empty info_axioms) then warn_extraction_axiom_to_realize info_axioms; let log_axioms = Refset'.elements !log_axioms in if not (List.is_empty log_axioms) then warn_extraction_logical_axiom log_axioms let warn_extraction_opaque_accessed = CWarnings.create ~name:"extraction-opaque-accessed" ~category:"extraction" (fun lst -> strbrk "The extraction is currently set to bypass opacity, " ++ strbrk "the following opaque constant bodies have been accessed :" ++ lst ++ str "." ++ fnl ()) let warn_extraction_opaque_as_axiom = CWarnings.create ~name:"extraction-opaque-as-axiom" ~category:"extraction" (fun lst -> strbrk "The extraction now honors the opacity constraints by default, " ++ strbrk "the following opaque constants have been extracted as axioms :" ++ lst ++ str "." ++ fnl () ++ strbrk "If necessary, use \"Set Extraction AccessOpaque\" to change this." ++ fnl ()) let warning_opaques accessed = let opaques = Refset'.elements !opaques in if not (List.is_empty opaques) then let lst = hov 1 (spc () ++ prlist_with_sep spc safe_pr_global opaques) in if accessed then warn_extraction_opaque_accessed lst else warn_extraction_opaque_as_axiom lst let warning_ambiguous_name = CWarnings.create ~name:"extraction-ambiguous-name" ~category:"extraction" (fun (q,mp,r) -> strbrk "The name " ++ pr_qualid q ++ strbrk " is ambiguous, " ++ strbrk "do you mean module " ++ pr_long_mp mp ++ strbrk " or object " ++ pr_long_global r ++ str " ?" ++ fnl () ++ strbrk "First choice is assumed, for the second one please use " ++ strbrk "fully qualified name." ++ fnl ()) let error_axiom_scheme r i = err (str "The type scheme axiom " ++ spc () ++ safe_pr_global r ++ spc () ++ str "needs " ++ int i ++ str " type variable(s).") let warn_extraction_inside_module = CWarnings.create ~name:"extraction-inside-module" ~category:"extraction" (fun () -> strbrk "Extraction inside an opened module is experimental." ++ strbrk "In case of problem, close it first.") let check_inside_module () = if Lib.is_modtype () then err (str "You can't do that within a Module Type." ++ fnl () ++ str "Close it and try again.") else if Lib.is_module () then warn_extraction_inside_module () let check_inside_section () = if Global.sections_are_opened () then err (str "You can't do that within a section." ++ fnl () ++ str "Close it and try again.") let warn_extraction_reserved_identifier = CWarnings.create ~name:"extraction-reserved-identifier" ~category:"extraction" (fun s -> strbrk ("The identifier "^s^ " contains __ which is reserved for the extraction")) let warning_id s = warn_extraction_reserved_identifier s let error_constant r = err (safe_pr_global r ++ str " is not a constant.") let error_inductive r = err (safe_pr_global r ++ spc () ++ str "is not an inductive type.") let error_nb_cons () = err (str "Not the right number of constructors.") let error_module_clash mp1 mp2 = err (str "The Coq modules " ++ pr_long_mp mp1 ++ str " and " ++ pr_long_mp mp2 ++ str " have the same ML name.\n" ++ str "This is not supported yet. Please do some renaming first.") let error_no_module_expr mp = err (str "The module " ++ pr_long_mp mp ++ str " has no body, it probably comes from\n" ++ str "some Declare Module outside any Module Type.\n" ++ str "This situation is currently unsupported by the extraction.") let error_singleton_become_prop id og = let loc = match og with | Some g -> fnl () ++ str "in " ++ safe_pr_global g ++ str " (or in its mutual block)" | None -> mt () in err (str "The informative inductive type " ++ Id.print id ++ str " has a Prop instance" ++ loc ++ str "." ++ fnl () ++ str "This happens when a sort-polymorphic singleton inductive type\n" ++ str "has logical parameters, such as (I,I) : (True * True) : Prop.\n" ++ str "The Ocaml extraction cannot handle this situation yet.\n" ++ str "Instead, use a sort-monomorphic type such as (True /\\ True)\n" ++ str "or extract to Haskell.") let error_unknown_module m = err (str "Module" ++ spc () ++ pr_qualid m ++ spc () ++ str "not found.") let error_scheme () = err (str "No Scheme modular extraction available yet.") let error_not_visible r = err (safe_pr_global r ++ str " is not directly visible.\n" ++ str "For example, it may be inside an applied functor.\n" ++ str "Use Recursive Extraction to get the whole environment.") let error_MPfile_as_mod mp b = let s1 = if b then "asked" else "required" in let s2 = if b then "extract some objects of this module or\n" else "" in err (str ("Extraction of file "^(raw_string_of_modfile mp)^ ".v as a module is "^s1^".\n"^ "Monolithic Extraction cannot deal with this situation.\n"^ "Please "^s2^"use (Recursive) Extraction Library instead.\n")) let argnames_of_global r = let env = Global.env () in let typ, _ = Typeops.type_of_global_in_context env r in let rels,_ = decompose_prod (Reduction.whd_all env typ) in List.rev_map (fun x -> Context.binder_name (fst x)) rels let msg_of_implicit = function | Kimplicit (r,i) -> let name = match (List.nth (argnames_of_global r) (i-1)) with | Anonymous -> "" | Name id -> "(" ^ Id.to_string id ^ ") " in (String.ordinal i)^" argument "^name^"of "^(string_of_global r) | Ktype | Kprop -> "" let error_remaining_implicit k = let s = msg_of_implicit k in err (str ("An implicit occurs after extraction : "^s^".") ++ fnl () ++ str "Please check your Extraction Implicit declarations." ++ fnl() ++ str "You might also try Unset Extraction SafeImplicits to force" ++ fnl() ++ str "the extraction of unsafe code and review it manually.") let warn_extraction_remaining_implicit = CWarnings.create ~name:"extraction-remaining-implicit" ~category:"extraction" (fun s -> strbrk ("At least an implicit occurs after extraction : "^s^".") ++ fnl () ++ strbrk "Extraction SafeImplicits is unset, extracting nonetheless," ++ strbrk "but this code is potentially unsafe, please review it manually.") let warning_remaining_implicit k = let s = msg_of_implicit k in warn_extraction_remaining_implicit s let check_loaded_modfile mp = match base_mp mp with | MPfile dp -> if not (Library.library_is_loaded dp) then begin match base_mp (Lib.current_mp ()) with | MPfile dp' when not (DirPath.equal dp dp') -> err (str "Please load library " ++ DirPath.print dp ++ str " first.") | _ -> () end | _ -> () let info_file f = Flags.if_verbose Feedback.msg_info (str ("The file "^f^" has been created by extraction.")) (*S The Extraction auxiliary commands *) (* The objects defined below should survive an arbitrary time, so we register them to coq save/undo mechanism. *) let my_bool_option name initval = let flag = ref initval in let access = fun () -> !flag in let () = declare_bool_option {optdepr = false; optname = "Extraction "^name; optkey = ["Extraction"; name]; optread = access; optwrite = (:=) flag } in access (*s Extraction AccessOpaque *) let access_opaque = my_bool_option "AccessOpaque" true (*s Extraction AutoInline *) let auto_inline = my_bool_option "AutoInline" false (*s Extraction TypeExpand *) let type_expand = my_bool_option "TypeExpand" true (*s Extraction KeepSingleton *) let keep_singleton = my_bool_option "KeepSingleton" false (*s Extraction Optimize *) type opt_flag = { opt_kill_dum : bool; (* 1 *) opt_fix_fun : bool; (* 2 *) opt_case_iot : bool; (* 4 *) opt_case_idr : bool; (* 8 *) opt_case_idg : bool; (* 16 *) opt_case_cst : bool; (* 32 *) opt_case_fun : bool; (* 64 *) opt_case_app : bool; (* 128 *) opt_let_app : bool; (* 256 *) opt_lin_let : bool; (* 512 *) opt_lin_beta : bool } (* 1024 *) let kth_digit n k = not (Int.equal (n land (1 lsl k)) 0) let flag_of_int n = { opt_kill_dum = kth_digit n 0; opt_fix_fun = kth_digit n 1; opt_case_iot = kth_digit n 2; opt_case_idr = kth_digit n 3; opt_case_idg = kth_digit n 4; opt_case_cst = kth_digit n 5; opt_case_fun = kth_digit n 6; opt_case_app = kth_digit n 7; opt_let_app = kth_digit n 8; opt_lin_let = kth_digit n 9; opt_lin_beta = kth_digit n 10 } (* For the moment, we allow by default everything except : - the type-unsafe optimization [opt_case_idg], which anyway cannot be activated currently (cf [Mlutil.branch_as_fun]) - the linear let and beta reduction [opt_lin_let] and [opt_lin_beta] (may lead to complexity blow-up, subsumed by finer reductions when inlining recursors). *) let int_flag_init = 1 + 2 + 4 + 8 (*+ 16*) + 32 + 64 + 128 + 256 (*+ 512 + 1024*) let int_flag_ref = ref int_flag_init let opt_flag_ref = ref (flag_of_int int_flag_init) let chg_flag n = int_flag_ref := n; opt_flag_ref := flag_of_int n let optims () = !opt_flag_ref let () = declare_bool_option {optdepr = false; optname = "Extraction Optimize"; optkey = ["Extraction"; "Optimize"]; optread = (fun () -> not (Int.equal !int_flag_ref 0)); optwrite = (fun b -> chg_flag (if b then int_flag_init else 0))} let () = declare_int_option { optdepr = false; optname = "Extraction Flag"; optkey = ["Extraction";"Flag"]; optread = (fun _ -> Some !int_flag_ref); optwrite = (function | None -> chg_flag 0 | Some i -> chg_flag (max i 0))} (* This option controls whether "dummy lambda" are removed when a toplevel constant is defined. *) let conservative_types_ref = ref false let conservative_types () = !conservative_types_ref let () = declare_bool_option {optdepr = false; optname = "Extraction Conservative Types"; optkey = ["Extraction"; "Conservative"; "Types"]; optread = (fun () -> !conservative_types_ref); optwrite = (fun b -> conservative_types_ref := b) } (* Allows to print a comment at the beginning of the output files *) let file_comment_ref = ref "" let file_comment () = !file_comment_ref let () = declare_string_option {optdepr = false; optname = "Extraction File Comment"; optkey = ["Extraction"; "File"; "Comment"]; optread = (fun () -> !file_comment_ref); optwrite = (fun s -> file_comment_ref := s) } (*s Extraction Lang *) type lang = Ocaml | Haskell | Scheme | JSON let lang_ref = Summary.ref Ocaml ~name:"ExtrLang" let lang () = !lang_ref let extr_lang : lang -> obj = declare_object @@ superglobal_object_nodischarge "Extraction Lang" ~cache:(fun (_,l) -> lang_ref := l) ~subst:None let extraction_language x = Lib.add_anonymous_leaf (extr_lang x) (*s Extraction Inline/NoInline *) let empty_inline_table = (Refset'.empty,Refset'.empty) let inline_table = Summary.ref empty_inline_table ~name:"ExtrInline" let to_inline r = Refset'.mem r (fst !inline_table) let to_keep r = Refset'.mem r (snd !inline_table) let add_inline_entries b l = let f b = if b then Refset'.add else Refset'.remove in let i,k = !inline_table in inline_table := (List.fold_right (f b) l i), (List.fold_right (f (not b)) l k) (* Registration of operations for rollback. *) let inline_extraction : bool * GlobRef.t list -> obj = declare_object @@ superglobal_object "Extraction Inline" ~cache:(fun (_,(b,l)) -> add_inline_entries b l) ~subst:(Some (fun (s,(b,l)) -> (b,(List.map (fun x -> fst (subst_global s x)) l)))) ~discharge:(fun (_,x) -> Some x) (* Grammar entries. *) let extraction_inline b l = let refs = List.map Smartlocate.global_with_alias l in List.iter (fun r -> match r with | GlobRef.ConstRef _ -> () | _ -> error_constant r) refs; Lib.add_anonymous_leaf (inline_extraction (b,refs)) (* Printing part *) let print_extraction_inline () = let (i,n)= !inline_table in let i'= Refset'.filter (function GlobRef.ConstRef _ -> true | _ -> false) i in (str "Extraction Inline:" ++ fnl () ++ Refset'.fold (fun r p -> (p ++ str " " ++ safe_pr_long_global r ++ fnl ())) i' (mt ()) ++ str "Extraction NoInline:" ++ fnl () ++ Refset'.fold (fun r p -> (p ++ str " " ++ safe_pr_long_global r ++ fnl ())) n (mt ())) (* Reset part *) let reset_inline : unit -> obj = declare_object @@ superglobal_object_nodischarge "Reset Extraction Inline" ~cache:(fun (_,_)-> inline_table := empty_inline_table) ~subst:None let reset_extraction_inline () = Lib.add_anonymous_leaf (reset_inline ()) (*s Extraction Implicit *) let safe_implicit = my_bool_option "SafeImplicits" true let err_or_warn_remaining_implicit k = if safe_implicit () then error_remaining_implicit k else warning_remaining_implicit k type int_or_id = ArgInt of int | ArgId of Id.t let implicits_table = Summary.ref Refmap'.empty ~name:"ExtrImplicit" let implicits_of_global r = try Refmap'.find r !implicits_table with Not_found -> Int.Set.empty let add_implicits r l = let names = argnames_of_global r in let n = List.length names in let add_arg s = function | ArgInt i -> if 1 <= i && i <= n then Int.Set.add i s else err (int i ++ str " is not a valid argument number for " ++ safe_pr_global r) | ArgId id -> try let i = List.index Name.equal (Name id) names in Int.Set.add i s with Not_found -> err (str "No argument " ++ Id.print id ++ str " for " ++ safe_pr_global r) in let ints = List.fold_left add_arg Int.Set.empty l in implicits_table := Refmap'.add r ints !implicits_table (* Registration of operations for rollback. *) let implicit_extraction : GlobRef.t * int_or_id list -> obj = declare_object @@ superglobal_object_nodischarge "Extraction Implicit" ~cache:(fun (_,(r,l)) -> add_implicits r l) ~subst:(Some (fun (s,(r,l)) -> (fst (subst_global s r), l))) (* Grammar entries. *) let extraction_implicit r l = check_inside_section (); Lib.add_anonymous_leaf (implicit_extraction (Smartlocate.global_with_alias r,l)) (*s Extraction Blacklist of filenames not to use while extracting *) let blacklist_table = Summary.ref Id.Set.empty ~name:"ExtrBlacklist" let modfile_ids = ref Id.Set.empty let modfile_mps = ref MPmap.empty let reset_modfile () = modfile_ids := !blacklist_table; modfile_mps := MPmap.empty let string_of_modfile mp = try MPmap.find mp !modfile_mps with Not_found -> let id = Id.of_string (raw_string_of_modfile mp) in let id' = next_ident_away id !modfile_ids in let s' = Id.to_string id' in modfile_ids := Id.Set.add id' !modfile_ids; modfile_mps := MPmap.add mp s' !modfile_mps; s' (* same as [string_of_modfile], but preserves the capital/uncapital 1st char *) let file_of_modfile mp = let s0 = match mp with | MPfile f -> Id.to_string (List.hd (DirPath.repr f)) | _ -> assert false in String.mapi (fun i c -> if i = 0 then s0.[0] else c) (string_of_modfile mp) let add_blacklist_entries l = blacklist_table := List.fold_right (fun s -> Id.Set.add (Id.of_string (String.capitalize_ascii s))) l !blacklist_table (* Registration of operations for rollback. *) let blacklist_extraction : string list -> obj = declare_object @@ superglobal_object_nodischarge "Extraction Blacklist" ~cache:(fun (_,l) -> add_blacklist_entries l) ~subst:None (* Grammar entries. *) let extraction_blacklist l = let l = List.rev_map Id.to_string l in Lib.add_anonymous_leaf (blacklist_extraction l) (* Printing part *) let print_extraction_blacklist () = prlist_with_sep fnl Id.print (Id.Set.elements !blacklist_table) (* Reset part *) let reset_blacklist : unit -> obj = declare_object @@ superglobal_object_nodischarge "Reset Extraction Blacklist" ~cache:(fun (_,_)-> blacklist_table := Id.Set.empty) ~subst:None let reset_extraction_blacklist () = Lib.add_anonymous_leaf (reset_blacklist ()) (*s Extract Constant/Inductive. *) (* UGLY HACK: to be defined in [extraction.ml] *) let (use_type_scheme_nb_args, type_scheme_nb_args_hook) = Hook.make () let customs = Summary.ref Refmap'.empty ~name:"ExtrCustom" let add_custom r ids s = customs := Refmap'.add r (ids,s) !customs let is_custom r = Refmap'.mem r !customs let is_inline_custom r = (is_custom r) && (to_inline r) let find_custom r = snd (Refmap'.find r !customs) let find_type_custom r = Refmap'.find r !customs let custom_matchs = Summary.ref Refmap'.empty ~name:"ExtrCustomMatchs" let add_custom_match r s = custom_matchs := Refmap'.add r s !custom_matchs let indref_of_match pv = if Array.is_empty pv then raise Not_found; let (_,pat,_) = pv.(0) in match pat with | Pusual (GlobRef.ConstructRef (ip,_)) -> GlobRef.IndRef ip | Pcons (GlobRef.ConstructRef (ip,_),_) -> GlobRef.IndRef ip | _ -> raise Not_found let is_custom_match pv = try Refmap'.mem (indref_of_match pv) !custom_matchs with Not_found -> false let find_custom_match pv = Refmap'.find (indref_of_match pv) !custom_matchs (* Registration of operations for rollback. *) let in_customs : GlobRef.t * string list * string -> obj = declare_object @@ superglobal_object_nodischarge "ML extractions" ~cache:(fun (_,(r,ids,s)) -> add_custom r ids s) ~subst:(Some (fun (s,(r,ids,str)) -> (fst (subst_global s r), ids, str))) let in_custom_matchs : GlobRef.t * string -> obj = declare_object @@ superglobal_object_nodischarge "ML extractions custom matches" ~cache:(fun (_,(r,s)) -> add_custom_match r s) ~subst:(Some (fun (subs,(r,s)) -> (fst (subst_global subs r), s))) (* Grammar entries. *) let extract_constant_inline inline r ids s = check_inside_section (); let g = Smartlocate.global_with_alias r in match g with | GlobRef.ConstRef kn -> let env = Global.env () in let typ, _ = Typeops.type_of_global_in_context env (GlobRef.ConstRef kn) in let typ = Reduction.whd_all env typ in if Reduction.is_arity env typ then begin let nargs = Hook.get use_type_scheme_nb_args env typ in if not (Int.equal (List.length ids) nargs) then error_axiom_scheme g nargs end; Lib.add_anonymous_leaf (inline_extraction (inline,[g])); Lib.add_anonymous_leaf (in_customs (g,ids,s)) | _ -> error_constant g let extract_inductive r s l optstr = check_inside_section (); let g = Smartlocate.global_with_alias r in Dumpglob.add_glob ?loc:r.CAst.loc g; match g with | GlobRef.IndRef ((kn,i) as ip) -> let mib = Global.lookup_mind kn in let n = Array.length mib.mind_packets.(i).mind_consnames in if not (Int.equal n (List.length l)) then error_nb_cons (); Lib.add_anonymous_leaf (inline_extraction (true,[g])); Lib.add_anonymous_leaf (in_customs (g,[],s)); Option.iter (fun s -> Lib.add_anonymous_leaf (in_custom_matchs (g,s))) optstr; List.iteri (fun j s -> let g = GlobRef.ConstructRef (ip,succ j) in Lib.add_anonymous_leaf (inline_extraction (true,[g])); Lib.add_anonymous_leaf (in_customs (g,[],s))) l | _ -> error_inductive g (*s Tables synchronization. *) let reset_tables () = init_typedefs (); init_cst_types (); init_inductives (); init_inductive_kinds (); init_recursors (); init_projs (); init_axioms (); init_opaques (); reset_modfile () coq-8.11.0/plugins/extraction/ExtrHaskellZInt.v0000644000175000017500000000151313612664533021351 0ustar treinentreinen(** Extraction of [Z] into Haskell's [Int] *) Require Coq.extraction.Extraction. Require Import ZArith. Require Import ExtrHaskellZNum. (** * Disclaimer: trying to obtain efficient certified programs * by extracting [Z] into [Int] is definitively *not* a good idea. * See comments in [ExtrOcamlNatInt.v]. *) Extract Inductive positive => "Prelude.Int" [ "(\x -> 2 Prelude.* x Prelude.+ 1)" "(\x -> 2 Prelude.* x)" "1" ] "(\fI fO fH n -> if n Prelude.== 1 then fH () else if Prelude.odd n then fI (n `Prelude.div` 2) else fO (n `Prelude.div` 2))". Extract Inductive Z => "Prelude.Int" [ "0" "(\x -> x)" "Prelude.negate" ] "(\fO fP fN n -> if n Prelude.== 0 then fO () else if n Prelude.> 0 then fP n else fN (Prelude.negate n))". coq-8.11.0/plugins/extraction/ExtrOcamlZBigInt.v0000644000175000017500000000711413612664533021446 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* "Big.big_int" [ "Big.doubleplusone" "Big.double" "Big.one" ] "Big.positive_case". Extract Inductive Z => "Big.big_int" [ "Big.zero" "" "Big.opp" ] "Big.z_case". Extract Inductive N => "Big.big_int" [ "Big.zero" "" ] "Big.n_case". (** Nota: the "" above is used as an identity function "(fun p->p)" *) (** Efficient (but uncertified) versions for usual functions *) Extract Constant Pos.add => "Big.add". Extract Constant Pos.succ => "Big.succ". Extract Constant Pos.pred => "fun n -> Big.max Big.one (Big.pred n)". Extract Constant Pos.sub => "fun n m -> Big.max Big.one (Big.sub n m)". Extract Constant Pos.mul => "Big.mult". Extract Constant Pos.min => "Big.min". Extract Constant Pos.max => "Big.max". Extract Constant Pos.compare => "fun x y -> Big.compare_case Eq Lt Gt x y". Extract Constant Pos.compare_cont => "fun c x y -> Big.compare_case c Lt Gt x y". Extract Constant N.add => "Big.add". Extract Constant N.succ => "Big.succ". Extract Constant N.pred => "fun n -> Big.max Big.zero (Big.pred n)". Extract Constant N.sub => "fun n m -> Big.max Big.zero (Big.sub n m)". Extract Constant N.mul => "Big.mult". Extract Constant N.min => "Big.min". Extract Constant N.max => "Big.max". Extract Constant N.div => "fun a b -> if Big.eq b Big.zero then Big.zero else Big.div a b". Extract Constant N.modulo => "fun a b -> if Big.eq b Big.zero then Big.zero else Big.modulo a b". Extract Constant N.compare => "Big.compare_case Eq Lt Gt". Extract Constant Z.add => "Big.add". Extract Constant Z.succ => "Big.succ". Extract Constant Z.pred => "Big.pred". Extract Constant Z.sub => "Big.sub". Extract Constant Z.mul => "Big.mult". Extract Constant Z.opp => "Big.opp". Extract Constant Z.abs => "Big.abs". Extract Constant Z.min => "Big.min". Extract Constant Z.max => "Big.max". Extract Constant Z.compare => "Big.compare_case Eq Lt Gt". Extract Constant Z.of_N => "fun p -> p". Extract Constant Z.abs_N => "Big.abs". (** Z.div and Z.modulo are quite complex to define in terms of (/) and (mod). For the moment we don't even try *) (** Test: Require Import ZArith NArith. Extraction "/tmp/test.ml" Pos.add Pos.pred Pos.sub Pos.mul Pos.compare N.pred N.sub N.div N.modulo N.compare Z.add Z.mul Z.compare Z.of_N Z.abs_N Z.div Z.modulo. *) coq-8.11.0/plugins/extraction/extract_env.mli0000644000175000017500000000311413612664533021153 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit val full_extraction : string option -> qualid list -> unit val separate_extraction : qualid list -> unit val extraction_library : bool -> Id.t -> unit (* For the test-suite : extraction to a temporary file + ocamlc on it *) val extract_and_compile : qualid list -> unit (* For debug / external output via coqtop.byte + Drop : *) val mono_environment : GlobRef.t list -> ModPath.t list -> Miniml.ml_structure (* Used by the Relation Extraction plugin *) val print_one_decl : Miniml.ml_structure -> ModPath.t -> Miniml.ml_decl -> Pp.t (* Used by Extraction Compute *) val structure_for_compute : Environ.env -> Evd.evar_map -> EConstr.t -> Miniml.ml_decl list * Miniml.ml_ast * Miniml.ml_type (* Show the extraction of the current ongoing proof *) val show_extraction : pstate:Proof_global.t -> unit coq-8.11.0/plugins/extraction/mlutil.mli0000644000175000017500000001051613612664533020143 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit val new_meta : 'a -> ml_type val type_subst_list : ml_type list -> ml_type -> ml_type val type_subst_vect : ml_type array -> ml_type -> ml_type val instantiation : ml_schema -> ml_type val needs_magic : ml_type * ml_type -> bool val put_magic_if : bool -> ml_ast -> ml_ast val put_magic : ml_type * ml_type -> ml_ast -> ml_ast val generalizable : ml_ast -> bool (*s ML type environment. *) module Mlenv : sig type t val empty : t (* get the n-th more recently entered schema and instantiate it. *) val get : t -> int -> ml_type (* Adding a type in an environment, after generalizing free meta *) val push_gen : t -> ml_type -> t (* Adding a type with no [Tvar] *) val push_type : t -> ml_type -> t (* Adding a type with no [Tvar] nor [Tmeta] *) val push_std_type : t -> ml_type -> t end (*s Utility functions over ML types without meta *) val type_mem_kn : MutInd.t -> ml_type -> bool val type_maxvar : ml_type -> int val type_decomp : ml_type -> ml_type list * ml_type val type_recomp : ml_type list * ml_type -> ml_type val var2var' : ml_type -> ml_type type abbrev_map = GlobRef.t -> ml_type option val type_expand : abbrev_map -> ml_type -> ml_type val type_simpl : ml_type -> ml_type val type_to_sign : abbrev_map -> ml_type -> sign val type_to_signature : abbrev_map -> ml_type -> signature val type_expunge : abbrev_map -> ml_type -> ml_type val type_expunge_from_sign : abbrev_map -> signature -> ml_type -> ml_type val eq_ml_type : ml_type -> ml_type -> bool val isTdummy : ml_type -> bool val isMLdummy : ml_ast -> bool val isKill : sign -> bool val case_expunge : signature -> ml_ast -> ml_ident list * ml_ast val term_expunge : signature -> ml_ident list * ml_ast -> ml_ast (*s Special identifiers. [dummy_name] is to be used for dead code and will be printed as [_] in concrete (Caml) code. *) val anonymous_name : Id.t val dummy_name : Id.t val id_of_name : Name.t -> Id.t val id_of_mlid : ml_ident -> Id.t val tmp_id : ml_ident -> ml_ident (*s [collect_lambda MLlam(id1,...MLlam(idn,t)...)] returns the list [idn;...;id1] and the term [t]. *) val collect_lams : ml_ast -> ml_ident list * ml_ast val collect_n_lams : int -> ml_ast -> ml_ident list * ml_ast val remove_n_lams : int -> ml_ast -> ml_ast val nb_lams : ml_ast -> int val named_lams : ml_ident list -> ml_ast -> ml_ast val dummy_lams : ml_ast -> int -> ml_ast val anonym_or_dummy_lams : ml_ast -> signature -> ml_ast val eta_args_sign : int -> signature -> ml_ast list (*s Utility functions over ML terms. *) val mlapp : ml_ast -> ml_ast list -> ml_ast val ast_map : (ml_ast -> ml_ast) -> ml_ast -> ml_ast val ast_map_lift : (int -> ml_ast -> ml_ast) -> int -> ml_ast -> ml_ast val ast_iter : (ml_ast -> unit) -> ml_ast -> unit val ast_occurs : int -> ml_ast -> bool val ast_occurs_itvl : int -> int -> ml_ast -> bool val ast_lift : int -> ml_ast -> ml_ast val ast_pop : ml_ast -> ml_ast val ast_subst : ml_ast -> ml_ast -> ml_ast val ast_glob_subst : ml_ast Refmap'.t -> ml_ast -> ml_ast val dump_unused_vars : ml_ast -> ml_ast val normalize : ml_ast -> ml_ast val optimize_fix : ml_ast -> ml_ast val inline : GlobRef.t -> ml_ast -> bool val is_basic_pattern : ml_pattern -> bool val has_deep_pattern : ml_branch array -> bool val is_regular_match : ml_branch array -> bool exception Impossible (* Classification of signatures *) type sign_kind = | EmptySig | NonLogicalSig (* at least a [Keep] *) | SafeLogicalSig (* only [Kill Ktype] *) | UnsafeLogicalSig (* No [Keep], not all [Kill Ktype] *) val sign_kind : signature -> sign_kind val sign_no_final_keeps : signature -> signature coq-8.11.0/plugins/extraction/ExtrOcamlIntConv.v0000644000175000017500000000613413612664533021521 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int. Parameter int_opp : int -> int. Parameter int_twice : int -> int. Extract Inlined Constant int => int. Extract Inlined Constant int_zero => "0". Extract Inlined Constant int_succ => "succ". Extract Inlined Constant int_opp => "-". Extract Inlined Constant int_twice => "2 *". Definition int_of_nat : nat -> int := (fix loop acc n := match n with | O => acc | S n => loop (int_succ acc) n end) int_zero. Fixpoint int_of_pos p := match p with | xH => int_succ int_zero | xO p => int_twice (int_of_pos p) | xI p => int_succ (int_twice (int_of_pos p)) end. Fixpoint int_of_z z := match z with | Z0 => int_zero | Zpos p => int_of_pos p | Zneg p => int_opp (int_of_pos p) end. Fixpoint int_of_n n := match n with | N0 => int_zero | Npos p => int_of_pos p end. (** NB: as for [pred] or [minus], [nat_of_int], [n_of_int] and [pos_of_int] are total and return zero (resp. one) for non-positive inputs. *) Parameter int_natlike_rec : forall A, A -> (A->A) -> int -> A. Extract Constant int_natlike_rec => "fun fO fS -> let rec loop acc i = if i <= 0 then acc else loop (fS acc) (i-1) in loop fO". Definition nat_of_int : int -> nat := int_natlike_rec _ O S. Parameter int_poslike_rec : forall A, A -> (A->A) -> (A->A) -> int -> A. Extract Constant int_poslike_rec => "fun f1 f2x f2x1 -> let rec loop i = if i <= 1 then f1 else if i land 1 = 0 then f2x (loop (i lsr 1)) else f2x1 (loop (i lsr 1)) in loop". Definition pos_of_int : int -> positive := int_poslike_rec _ xH xO xI. Parameter int_zlike_case : forall A, A -> (int->A) -> (int->A) -> int -> A. Extract Constant int_zlike_case => "fun f0 fpos fneg i -> if i = 0 then f0 else if i>0 then fpos i else fneg (-i)". Definition z_of_int : int -> Z := int_zlike_case _ Z0 (fun i => Zpos (pos_of_int i)) (fun i => Zneg (pos_of_int i)). Definition n_of_int : int -> N := int_zlike_case _ N0 (fun i => Npos (pos_of_int i)) (fun _ => N0). (** Warning: [z_of_int] is currently wrong for Ocaml's [min_int], since [min_int] has no positive opposite ([-min_int = min_int]). *) (* Extraction "/tmp/test.ml" nat_of_int int_of_nat pos_of_int int_of_pos z_of_int int_of_z n_of_int int_of_n. *) coq-8.11.0/plugins/extraction/extraction.ml0000644000175000017500000013747713612664533020664 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Logic | InSet | InType -> Info let info_of_sort s = info_of_family (Sorts.family s) let rec flag_of_type env sg t : flag = let t = whd_all env sg t in match EConstr.kind sg t with | Prod (x,t,c) -> flag_of_type (EConstr.push_rel (LocalAssum (x,t)) env) sg c | Sort s -> (info_of_sort (EConstr.ESorts.kind sg s),TypeScheme) | _ -> (info_of_family (sort_of env sg t),Default) (*s Two particular cases of [flag_of_type]. *) let is_default env sg t = match flag_of_type env sg t with | (Info, Default) -> true | _ -> false exception NotDefault of kill_reason let check_default env sg t = match flag_of_type env sg t with | _,TypeScheme -> raise (NotDefault Ktype) | Logic,_ -> raise (NotDefault Kprop) | _ -> () let is_info_scheme env sg t = match flag_of_type env sg t with | (Info, TypeScheme) -> true | _ -> false let push_rel_assum (n, t) env = EConstr.push_rel (LocalAssum (n, t)) env let push_rels_assum assums = EConstr.push_rel_context (List.map (fun (x,t) -> LocalAssum (x,t)) assums) let get_body lconstr = EConstr.of_constr (Mod_subst.force_constr lconstr) let get_opaque env c = EConstr.of_constr (fst (Opaqueproof.force_proof Library.indirect_accessor (Environ.opaque_tables env) c)) let applistc c args = EConstr.mkApp (c, Array.of_list args) (* Same as [Environ.push_rec_types], but for [EConstr.t] *) let push_rec_types (lna,typarray,_) env = let ctxt = Array.map2_i (fun i na t -> LocalAssum (na, EConstr.Vars.lift i t)) lna typarray in Array.fold_left (fun e assum -> EConstr.push_rel assum e) env ctxt (* Same as [Termops.nb_lam], but for [EConstr.t] *) let nb_lam sg c = List.length (fst (EConstr.decompose_lam sg c)) (* Same as [Term.decompose_lam_n] but for [EConstr.t] *) let decompose_lam_n sg n = let rec lamdec_rec l n c = if n <= 0 then l,c else match EConstr.kind sg c with | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c | Cast (c,_,_) -> lamdec_rec l n c | _ -> raise Not_found in lamdec_rec [] n (*s [type_sign] gernerates a signature aimed at treating a type application. *) let rec type_sign env sg c = match EConstr.kind sg (whd_all env sg c) with | Prod (n,t,d) -> (if is_info_scheme env sg t then Keep else Kill Kprop) :: (type_sign (push_rel_assum (n,t) env) sg d) | _ -> [] let rec type_scheme_nb_args env sg c = match EConstr.kind sg (whd_all env sg c) with | Prod (n,t,d) -> let n = type_scheme_nb_args (push_rel_assum (n,t) env) sg d in if is_info_scheme env sg t then n+1 else n | _ -> 0 let type_scheme_nb_args' env c = type_scheme_nb_args env (Evd.from_env env) (EConstr.of_constr c) let _ = Hook.set type_scheme_nb_args_hook type_scheme_nb_args' (*s [type_sign_vl] does the same, plus a type var list. *) (* When generating type variables, we avoid any ' in their names (otherwise this may cause a lexer conflict in ocaml with 'a'). We also get rid of unicode characters. Anyway, since type variables are local, the created name is just a matter of taste... See also Bug #3227 *) let make_typvar n vl = let id = id_of_name n in let id' = let s = Id.to_string id in if not (String.contains s '\'') && Unicode.is_basic_ascii s then id else id_of_name Anonymous in let vl = Id.Set.of_list vl in next_ident_away id' vl let rec type_sign_vl env sg c = match EConstr.kind sg (whd_all env sg c) with | Prod (n,t,d) -> let s,vl = type_sign_vl (push_rel_assum (n,t) env) sg d in if not (is_info_scheme env sg t) then Kill Kprop::s, vl else Keep::s, (make_typvar n.binder_name vl) :: vl | _ -> [],[] let rec nb_default_params env sg c = match EConstr.kind sg (whd_all env sg c) with | Prod (n,t,d) -> let n = nb_default_params (push_rel_assum (n,t) env) sg d in if is_default env sg t then n+1 else n | _ -> 0 (* Enriching a signature with implicit information *) let sign_with_implicits r s nb_params = let implicits = implicits_of_global r in let rec add_impl i = function | [] -> [] | Keep::s when Int.Set.mem i implicits -> Kill (Kimplicit (r,i)) :: add_impl (i+1) s | sign::s -> sign :: add_impl (i+1) s in add_impl (1+nb_params) s (*S Management of type variable contexts. *) (* A De Bruijn variable context (db) is a context for translating Coq [Rel] into ML type [Tvar]. *) (*s From a type signature toward a type variable context (db). *) let db_from_sign s = let rec make i acc = function | [] -> acc | Keep :: l -> make (i+1) (i::acc) l | Kill _ :: l -> make i (0::acc) l in make 1 [] s (*s Create a type variable context from indications taken from an inductive type (see just below). *) let rec db_from_ind dbmap i = if Int.equal i 0 then [] else (try Int.Map.find i dbmap with Not_found -> 0)::(db_from_ind dbmap (i-1)) (*s [parse_ind_args] builds a map: [i->j] iff the i-th Coq argument of a constructor corresponds to the j-th type var of the ML inductive. *) (* \begin{itemize} \item [si] : signature of the inductive \item [i] : counter of Coq args for [(I args)] \item [j] : counter of ML type vars \item [relmax] : total args number of the constructor \end{itemize} *) let parse_ind_args si args relmax = let rec parse i j = function | [] -> Int.Map.empty | Kill _ :: s -> parse (i+1) j s | Keep :: s -> (match Constr.kind args.(i-1) with | Rel k -> Int.Map.add (relmax+1-k) j (parse (i+1) (j+1) s) | _ -> parse (i+1) (j+1) s) in parse 1 1 si (*S Extraction of a type. *) (* [extract_type env db c args] is used to produce an ML type from the coq term [(c args)], which is supposed to be a Coq type. *) (* [db] is a context for translating Coq [Rel] into ML type [Tvar]. *) (* [j] stands for the next ML type var. [j=0] means we do not generate ML type var anymore (in subterms for example). *) let rec extract_type env sg db j c args = match EConstr.kind sg (whd_betaiotazeta sg c) with | App (d, args') -> (* We just accumulate the arguments. *) extract_type env sg db j d (Array.to_list args' @ args) | Lambda (_,_,d) -> (match args with | [] -> assert false (* A lambda cannot be a type. *) | a :: args -> extract_type env sg db j (EConstr.Vars.subst1 a d) args) | Prod (n,t,d) -> assert (List.is_empty args); let env' = push_rel_assum (n,t) env in (match flag_of_type env sg t with | (Info, Default) -> (* Standard case: two [extract_type] ... *) let mld = extract_type env' sg (0::db) j d [] in (match expand env mld with | Tdummy d -> Tdummy d | _ -> Tarr (extract_type env sg db 0 t [], mld)) | (Info, TypeScheme) when j > 0 -> (* A new type var. *) let mld = extract_type env' sg (j::db) (j+1) d [] in (match expand env mld with | Tdummy d -> Tdummy d | _ -> Tarr (Tdummy Ktype, mld)) | _,lvl -> let mld = extract_type env' sg (0::db) j d [] in (match expand env mld with | Tdummy d -> Tdummy d | _ -> let reason = if lvl == TypeScheme then Ktype else Kprop in Tarr (Tdummy reason, mld))) | Sort _ -> Tdummy Ktype (* The two logical cases. *) | _ when sort_of env sg (applistc c args) == InProp -> Tdummy Kprop | Rel n -> (match EConstr.lookup_rel n env with | LocalDef (_,t,_) -> extract_type env sg db j (EConstr.Vars.lift n t) args | _ -> (* Asks [db] a translation for [n]. *) if n > List.length db then Tunknown else let n' = List.nth db (n-1) in if Int.equal n' 0 then Tunknown else Tvar n') | Const (kn,u) -> let r = GlobRef.ConstRef kn in let typ = type_of env sg (EConstr.mkConstU (kn,u)) in (match flag_of_type env sg typ with | (Logic,_) -> assert false (* Cf. logical cases above *) | (Info, TypeScheme) -> let mlt = extract_type_app env sg db (r, type_sign env sg typ) args in (match (lookup_constant kn env).const_body with | Undef _ | OpaqueDef _ | Primitive _ -> mlt | Def _ when is_custom (GlobRef.ConstRef kn) -> mlt | Def lbody -> let newc = applistc (get_body lbody) args in let mlt' = extract_type env sg db j newc [] in (* ML type abbreviations interact badly with Coq *) (* reduction, so [mlt] and [mlt'] might be different: *) (* The more precise is [mlt'], extracted after reduction *) (* The shortest is [mlt], which use abbreviations *) (* If possible, we take [mlt], otherwise [mlt']. *) if eq_ml_type (expand env mlt) (expand env mlt') then mlt else mlt') | (Info, Default) -> (* Not an ML type, for example [(c:forall X, X->X) Type nat] *) (match (lookup_constant kn env).const_body with | Undef _ | OpaqueDef _ | Primitive _ -> Tunknown (* Brutal approx ... *) | Def lbody -> (* We try to reduce. *) let newc = applistc (get_body lbody) args in extract_type env sg db j newc [])) | Ind ((kn,i),u) -> let s = (extract_ind env kn).ind_packets.(i).ip_sign in extract_type_app env sg db (GlobRef.IndRef (kn,i),s) args | Proj (p,t) -> (* Let's try to reduce, if it hasn't already been done. *) if Projection.unfolded p then Tunknown else extract_type env sg db j (EConstr.mkProj (Projection.unfold p, t)) args | Case _ | Fix _ | CoFix _ -> Tunknown | Evar _ | Meta _ -> Taxiom (* only possible during Show Extraction *) | Var v -> (* For Show Extraction *) let open Context.Named.Declaration in (match EConstr.lookup_named v env with | LocalDef (_,body,_) -> extract_type env sg db j (EConstr.applist (body,args)) [] | LocalAssum (_,ty) -> let r = GlobRef.VarRef v in (match flag_of_type env sg ty with | (Logic,_) -> assert false (* Cf. logical cases above *) | (Info, TypeScheme) -> extract_type_app env sg db (r, type_sign env sg ty) args | (Info, Default) -> Tunknown)) | Cast _ | LetIn _ | Construct _ | Int _ | Float _ -> assert false (*s Auxiliary function dealing with type application. Precondition: [r] is a type scheme represented by the signature [s], and is completely applied: [List.length args = List.length s]. *) and extract_type_app env sg db (r,s) args = let ml_args = List.fold_right (fun (b,c) a -> if b == Keep then let p = List.length (fst (splay_prod env sg (type_of env sg c))) in let db = iterate (fun l -> 0 :: l) p db in (extract_type_scheme env sg db c p) :: a else a) (List.combine s args) [] in Tglob (r, ml_args) (*S Extraction of a type scheme. *) (* [extract_type_scheme env db c p] works on a Coq term [c] which is an informative type scheme. It means that [c] is not a Coq type, but will be when applied to sufficiently many arguments ([p] in fact). This function decomposes p lambdas, with eta-expansion if needed. *) (* [db] is a context for translating Coq [Rel] into ML type [Tvar]. *) and extract_type_scheme env sg db c p = if Int.equal p 0 then extract_type env sg db 0 c [] else let c = whd_betaiotazeta sg c in match EConstr.kind sg c with | Lambda (n,t,d) -> extract_type_scheme (push_rel_assum (n,t) env) sg db d (p-1) | _ -> let rels = fst (splay_prod env sg (type_of env sg c)) in let env = push_rels_assum rels env in let eta_args = List.rev_map EConstr.mkRel (List.interval 1 p) in extract_type env sg db 0 (EConstr.Vars.lift p c) eta_args (*S Extraction of an inductive type. *) (* First, a version with cache *) and extract_ind env kn = (* kn is supposed to be in long form *) let mib = Environ.lookup_mind kn env in match lookup_ind kn mib with | Some ml_ind -> ml_ind | None -> try extract_really_ind env kn mib with SingletonInductiveBecomesProp id -> (* TODO : which inductive is concerned in the block ? *) error_singleton_become_prop id (Some (GlobRef.IndRef (kn,0))) (* Then the real function *) and extract_really_ind env kn mib = (* First, if this inductive is aliased via a Module, we process the original inductive if possible. When at toplevel of the monolithic case, we cannot do much (cf Vector and bug #2570) *) let equiv = if lang () != Ocaml || (not (modular ()) && at_toplevel (MutInd.modpath kn)) || KerName.equal (MutInd.canonical kn) (MutInd.user kn) then NoEquiv else begin ignore (extract_ind env (MutInd.make1 (MutInd.canonical kn))); Equiv (MutInd.canonical kn) end in (* Everything concerning parameters. *) (* We do that first, since they are common to all the [mib]. *) let mip0 = mib.mind_packets.(0) in let npar = mib.mind_nparams in let epar = push_rel_context mib.mind_params_ctxt env in let sg = Evd.from_env env in (* First pass: we store inductive signatures together with *) (* their type var list. *) let packets = Array.mapi (fun i mip -> let (_,u),_ = UnivGen.fresh_inductive_instance env (kn,i) in let ar = Inductive.type_of_inductive env ((mib,mip),u) in let ar = EConstr.of_constr ar in let info = (fst (flag_of_type env sg ar) = Info) in let s,v = if info then type_sign_vl env sg ar else [],[] in let t = Array.make (Array.length mip.mind_nf_lc) [] in { ip_typename = mip.mind_typename; ip_consnames = mip.mind_consnames; ip_logical = not info; ip_sign = s; ip_vars = v; ip_types = t }, u) mib.mind_packets in add_ind kn mib {ind_kind = Standard; ind_nparams = npar; ind_packets = Array.map fst packets; ind_equiv = equiv }; (* Second pass: we extract constructors *) for i = 0 to mib.mind_ntypes - 1 do let p,u = packets.(i) in if not p.ip_logical then let types = arities_of_constructors env ((kn,i),u) in for j = 0 to Array.length types - 1 do let t = snd (decompose_prod_n npar types.(j)) in let prods,head = dest_prod epar t in let nprods = List.length prods in let args = match Constr.kind head with | App (f,args) -> args (* [Constr.kind f = Ind ip] *) | _ -> [||] in let dbmap = parse_ind_args p.ip_sign args (nprods + npar) in let db = db_from_ind dbmap npar in p.ip_types.(j) <- extract_type_cons epar sg db dbmap (EConstr.of_constr t) (npar+1) done done; (* Third pass: we determine special cases. *) let ind_info = try let ip = (kn, 0) in let r = GlobRef.IndRef ip in if is_custom r then raise (I Standard); if mib.mind_finite == CoFinite then raise (I Coinductive); if not (Int.equal mib.mind_ntypes 1) then raise (I Standard); let p,u = packets.(0) in if p.ip_logical then raise (I Standard); if not (Int.equal (Array.length p.ip_types) 1) then raise (I Standard); let typ = p.ip_types.(0) in let l = List.filter (fun t -> not (isTdummy (expand env t))) typ in if not (keep_singleton ()) && Int.equal (List.length l) 1 && not (type_mem_kn kn (List.hd l)) then raise (I Singleton); if List.is_empty l then raise (I Standard); if mib.mind_record == Declarations.NotRecord then raise (I Standard); (* Now we're sure it's a record. *) (* First, we find its field names. *) let rec names_prod t = match Constr.kind t with | Prod(n,_,t) -> n::(names_prod t) | LetIn(_,_,_,t) -> names_prod t | Cast(t,_,_) -> names_prod t | _ -> [] in let field_names = List.skipn mib.mind_nparams (names_prod mip0.mind_user_lc.(0)) in assert (Int.equal (List.length field_names) (List.length typ)); let projs = ref Cset.empty in let mp = MutInd.modpath kn in let rec select_fields l typs = match l,typs with | [],[] -> [] | _::l, typ::typs when isTdummy (expand env typ) -> select_fields l typs | {binder_name=Anonymous}::l, typ::typs -> None :: (select_fields l typs) | {binder_name=Name id}::l, typ::typs -> let knp = Constant.make2 mp (Label.of_id id) in (* Is it safe to use [id] for projections [foo.id] ? *) if List.for_all ((==) Keep) (type2signature env typ) then projs := Cset.add knp !projs; Some (GlobRef.ConstRef knp) :: (select_fields l typs) | _ -> assert false in let field_glob = select_fields field_names typ in (* Is this record officially declared with its projections ? *) (* If so, we use this information. *) begin try let ty = Inductive.type_of_inductive env ((mib,mip0),u) in let n = nb_default_params env sg (EConstr.of_constr ty) in let check_proj kn = if Cset.mem kn !projs then add_projection n kn ip in List.iter (Option.iter check_proj) (lookup_projections ip) with Not_found -> () end; Record field_glob with (I info) -> info in let i = {ind_kind = ind_info; ind_nparams = npar; ind_packets = Array.map fst packets; ind_equiv = equiv } in add_ind kn mib i; add_inductive_kind kn i.ind_kind; i (*s [extract_type_cons] extracts the type of an inductive constructor toward the corresponding list of ML types. - [db] is a context for translating Coq [Rel] into ML type [Tvar] - [dbmap] is a translation map (produced by a call to [parse_in_args]) - [i] is the rank of the current product (initially [params_nb+1]) *) and extract_type_cons env sg db dbmap c i = match EConstr.kind sg (whd_all env sg c) with | Prod (n,t,d) -> let env' = push_rel_assum (n,t) env in let db' = (try Int.Map.find i dbmap with Not_found -> 0) :: db in let l = extract_type_cons env' sg db' dbmap d (i+1) in (extract_type env sg db 0 t []) :: l | _ -> [] (*s Recording the ML type abbreviation of a Coq type scheme constant. *) and mlt_env env r = let open GlobRef in match r with | IndRef _ | ConstructRef _ | VarRef _ -> None | ConstRef kn -> let cb = Environ.lookup_constant kn env in match cb.const_body with | Undef _ | OpaqueDef _ | Primitive _ -> None | Def l_body -> match lookup_typedef kn cb with | Some _ as o -> o | None -> let sg = Evd.from_env env in let typ = EConstr.of_constr cb.const_type (* FIXME not sure if we should instantiate univs here *) in match flag_of_type env sg typ with | Info,TypeScheme -> let body = get_body l_body in let s = type_sign env sg typ in let db = db_from_sign s in let t = extract_type_scheme env sg db body (List.length s) in add_typedef kn cb t; Some t | _ -> None and expand env = type_expand (mlt_env env) and type2signature env = type_to_signature (mlt_env env) let type2sign env = type_to_sign (mlt_env env) let type_expunge env = type_expunge (mlt_env env) let type_expunge_from_sign env = type_expunge_from_sign (mlt_env env) (*s Extraction of the type of a constant. *) let record_constant_type env sg kn opt_typ = let cb = lookup_constant kn env in match lookup_cst_type kn cb with | Some schema -> schema | None -> let typ = match opt_typ with | None -> EConstr.of_constr cb.const_type | Some typ -> typ in let mlt = extract_type env sg [] 1 typ [] in let schema = (type_maxvar mlt, mlt) in let () = add_cst_type kn cb schema in schema (*S Extraction of a term. *) (* Precondition: [(c args)] is not a type scheme, and is informative. *) (* [mle] is a ML environment [Mlenv.t]. *) (* [mlt] is the ML type we want our extraction of [(c args)] to have. *) let rec extract_term env sg mle mlt c args = match EConstr.kind sg c with | App (f,a) -> extract_term env sg mle mlt f (Array.to_list a @ args) | Lambda (n, t, d) -> let id = map_annot id_of_name n in let idna = map_annot Name.mk_name id in (match args with | a :: l -> (* We make as many [LetIn] as possible. *) let l' = List.map (EConstr.Vars.lift 1) l in let d' = EConstr.mkLetIn (idna,a,t,applistc d l') in extract_term env sg mle mlt d' [] | [] -> let env' = push_rel_assum (idna, t) env in let id, a = try check_default env sg t; Id id.binder_name, new_meta() with NotDefault d -> Dummy, Tdummy d in let b = new_meta () in (* If [mlt] cannot be unified with an arrow type, then magic! *) let magic = needs_magic (mlt, Tarr (a, b)) in let d' = extract_term env' sg (Mlenv.push_type mle a) b d [] in put_magic_if magic (MLlam (id, d'))) | LetIn (n, c1, t1, c2) -> let id = map_annot id_of_name n in let env' = EConstr.push_rel (LocalDef (map_annot Name.mk_name id, c1, t1)) env in (* We directly push the args inside the [LetIn]. TODO: the opt_let_app flag is supposed to prevent that *) let args' = List.map (EConstr.Vars.lift 1) args in (try check_default env sg t1; let a = new_meta () in let c1' = extract_term env sg mle a c1 [] in (* The type of [c1'] is generalized and stored in [mle]. *) let mle' = if generalizable c1' then Mlenv.push_gen mle a else Mlenv.push_type mle a in MLletin (Id id.binder_name, c1', extract_term env' sg mle' mlt c2 args') with NotDefault d -> let mle' = Mlenv.push_std_type mle (Tdummy d) in ast_pop (extract_term env' sg mle' mlt c2 args')) | Const (kn,_) -> extract_cst_app env sg mle mlt kn args | Construct (cp,_) -> extract_cons_app env sg mle mlt cp args | Proj (p, c) -> let term = Retyping.expand_projection env (Evd.from_env env) p c [] in extract_term env sg mle mlt term args | Rel n -> (* As soon as the expected [mlt] for the head is known, *) (* we unify it with an fresh copy of the stored type of [Rel n]. *) let extract_rel mlt = put_magic (mlt, Mlenv.get mle n) (MLrel n) in extract_app env sg mle mlt extract_rel args | Case ({ci_ind=ip},_,c0,br) -> extract_app env sg mle mlt (extract_case env sg mle (ip,c0,br)) args | Fix ((_,i),recd) -> extract_app env sg mle mlt (extract_fix env sg mle i recd) args | CoFix (i,recd) -> extract_app env sg mle mlt (extract_fix env sg mle i recd) args | Cast (c,_,_) -> extract_term env sg mle mlt c args | Evar _ | Meta _ -> MLaxiom | Var v -> (* Only during Show Extraction *) let open Context.Named.Declaration in let ty = match EConstr.lookup_named v env with | LocalAssum (_,ty) -> ty | LocalDef (_,_,ty) -> ty in let vty = extract_type env sg [] 0 ty [] in let extract_var mlt = put_magic (mlt,vty) (MLglob (GlobRef.VarRef v)) in extract_app env sg mle mlt extract_var args | Int i -> assert (args = []); MLuint i | Float f -> assert (args = []); MLfloat f | Ind _ | Prod _ | Sort _ -> assert false (*s [extract_maybe_term] is [extract_term] for usual terms, else [MLdummy] *) and extract_maybe_term env sg mle mlt c = try check_default env sg (type_of env sg c); extract_term env sg mle mlt c [] with NotDefault d -> put_magic (mlt, Tdummy d) (MLdummy d) (*s Generic way to deal with an application. *) (* We first type all arguments starting with unknown meta types. This gives us the expected type of the head. Then we use the [mk_head] to produce the ML head from this type. *) and extract_app env sg mle mlt mk_head args = let metas = List.map new_meta args in let type_head = type_recomp (metas, mlt) in let mlargs = List.map2 (extract_maybe_term env sg mle) metas args in mlapp (mk_head type_head) mlargs (*s Auxiliary function used to extract arguments of constant or constructor. *) and make_mlargs env sg e s args typs = let rec f = function | [], [], _ -> [] | a::la, t::lt, [] -> extract_maybe_term env sg e t a :: (f (la,lt,[])) | a::la, t::lt, Keep::s -> extract_maybe_term env sg e t a :: (f (la,lt,s)) | _::la, _::lt, _::s -> f (la,lt,s) | _ -> assert false in f (args,typs,s) (*s Extraction of a constant applied to arguments. *) and extract_cst_app env sg mle mlt kn args = (* First, the [ml_schema] of the constant, in expanded version. *) let nb,t = record_constant_type env sg kn None in let schema = nb, expand env t in (* Can we instantiate types variables for this constant ? *) (* In Ocaml, inside the definition of this constant, the answer is no. *) let instantiated = if lang () == Ocaml && List.mem_f Constant.equal kn !current_fixpoints then var2var' (snd schema) else instantiation schema in (* Then the expected type of this constant. *) let a = new_meta () in (* We compare stored and expected types in two steps. *) (* First, can [kn] be applied to all args ? *) let metas = List.map new_meta args in let magic1 = needs_magic (type_recomp (metas, a), instantiated) in (* Second, is the resulting type compatible with the expected type [mlt] ? *) let magic2 = needs_magic (a, mlt) in (* The internal head receives a magic if [magic1] *) let head = put_magic_if magic1 (MLglob (GlobRef.ConstRef kn)) in (* Now, the extraction of the arguments. *) let s_full = type2signature env (snd schema) in let s_full = sign_with_implicits (GlobRef.ConstRef kn) s_full 0 in let s = sign_no_final_keeps s_full in let ls = List.length s in let la = List.length args in (* The ml arguments, already expunged from known logical ones *) let mla = make_mlargs env sg mle s args metas in (* For strict languages, purely logical signatures lead to a dummy lam (except when [Kill Ktype] everywhere). So a [MLdummy] is left accordingly. *) let optdummy = match sign_kind s_full with | UnsafeLogicalSig when lang () != Haskell -> [MLdummy Kprop] | _ -> [] in (* Different situations depending of the number of arguments: *) if la >= ls then (* Enough args, cleanup already done in [mla], we only add the additional dummy if needed. *) put_magic_if (magic2 && not magic1) (mlapp head (optdummy @ mla)) else (* Partially applied function with some logical arg missing. We complete via eta and expunge logical args. *) let ls' = ls-la in let s' = List.skipn la s in let mla = (List.map (ast_lift ls') mla) @ (eta_args_sign ls' s') in let e = anonym_or_dummy_lams (mlapp head mla) s' in put_magic_if magic2 (remove_n_lams (List.length optdummy) e) (*s Extraction of an inductive constructor applied to arguments. *) (* \begin{itemize} \item In ML, constructor arguments are uncurryfied. \item We managed to suppress logical parts inside inductive definitions, but they must appears outside (for partial applications for instance) \item We also suppressed all Coq parameters to the inductives, since they are fixed, and thus are not used for the computation. \end{itemize} *) and extract_cons_app env sg mle mlt (((kn,i) as ip,j) as cp) args = (* First, we build the type of the constructor, stored in small pieces. *) let mi = extract_ind env kn in let params_nb = mi.ind_nparams in let oi = mi.ind_packets.(i) in let nb_tvars = List.length oi.ip_vars and types = List.map (expand env) oi.ip_types.(j-1) in let list_tvar = List.map (fun i -> Tvar i) (List.interval 1 nb_tvars) in let type_cons = type_recomp (types, Tglob (GlobRef.IndRef ip, list_tvar)) in let type_cons = instantiation (nb_tvars, type_cons) in (* Then, the usual variables [s], [ls], [la], ... *) let s = List.map (type2sign env) types in let s = sign_with_implicits (GlobRef.ConstructRef cp) s params_nb in let ls = List.length s in let la = List.length args in assert (la <= ls + params_nb); let la' = max 0 (la - params_nb) in let args' = List.lastn la' args in (* Now, we build the expected type of the constructor *) let metas = List.map new_meta args' in (* If stored and expected types differ, then magic! *) let a = new_meta () in let magic1 = needs_magic (type_cons, type_recomp (metas, a)) in let magic2 = needs_magic (a, mlt) in let head mla = if mi.ind_kind == Singleton then put_magic_if magic1 (List.hd mla) (* assert (List.length mla = 1) *) else let typeargs = match snd (type_decomp type_cons) with | Tglob (_,l) -> List.map type_simpl l | _ -> assert false in let typ = Tglob(GlobRef.IndRef ip, typeargs) in put_magic_if magic1 (MLcons (typ, GlobRef.ConstructRef cp, mla)) in (* Different situations depending of the number of arguments: *) if la < params_nb then let head' = head (eta_args_sign ls s) in put_magic_if magic2 (dummy_lams (anonym_or_dummy_lams head' s) (params_nb - la)) else let mla = make_mlargs env sg mle s args' metas in if Int.equal la (ls + params_nb) then put_magic_if (magic2 && not magic1) (head mla) else (* [ params_nb <= la <= ls + params_nb ] *) let ls' = params_nb + ls - la in let s' = List.lastn ls' s in let mla = (List.map (ast_lift ls') mla) @ (eta_args_sign ls' s') in put_magic_if magic2 (anonym_or_dummy_lams (head mla) s') (*S Extraction of a case. *) and extract_case env sg mle ((kn,i) as ip,c,br) mlt = (* [br]: bodies of each branch (in functional form) *) (* [ni]: number of arguments without parameters in each branch *) let ni = constructors_nrealargs env ip in let br_size = Array.length br in assert (Int.equal (Array.length ni) br_size); if Int.equal br_size 0 then begin add_recursors env kn; (* May have passed unseen if logical ... *) MLexn "absurd case" end else (* [c] has an inductive type, and is not a type scheme type. *) let t = type_of env sg c in (* The only non-informative case: [c] is of sort [Prop] *) if (sort_of env sg t) == InProp then begin add_recursors env kn; (* May have passed unseen if logical ... *) (* Logical singleton case: *) (* [match c with C i j k -> t] becomes [t'] *) assert (Int.equal br_size 1); let s = iterate (fun l -> Kill Kprop :: l) ni.(0) [] in let mlt = iterate (fun t -> Tarr (Tdummy Kprop, t)) ni.(0) mlt in let e = extract_maybe_term env sg mle mlt br.(0) in snd (case_expunge s e) end else let mi = extract_ind env kn in let oi = mi.ind_packets.(i) in let metas = Array.init (List.length oi.ip_vars) new_meta in (* The extraction of the head. *) let type_head = Tglob (GlobRef.IndRef ip, Array.to_list metas) in let a = extract_term env sg mle type_head c [] in (* The extraction of each branch. *) let extract_branch i = let r = GlobRef.ConstructRef (ip,i+1) in (* The types of the arguments of the corresponding constructor. *) let f t = type_subst_vect metas (expand env t) in let l = List.map f oi.ip_types.(i) in (* the corresponding signature *) let s = List.map (type2sign env) oi.ip_types.(i) in let s = sign_with_implicits r s mi.ind_nparams in (* Extraction of the branch (in functional form). *) let e = extract_maybe_term env sg mle (type_recomp (l,mlt)) br.(i) in (* We suppress dummy arguments according to signature. *) let ids,e = case_expunge s e in (List.rev ids, Pusual r, e) in if mi.ind_kind == Singleton then begin (* Informative singleton case: *) (* [match c with C i -> t] becomes [let i = c' in t'] *) assert (Int.equal br_size 1); let (ids,_,e') = extract_branch 0 in assert (Int.equal (List.length ids) 1); MLletin (tmp_id (List.hd ids),a,e') end else (* Standard case: we apply [extract_branch]. *) let typs = List.map type_simpl (Array.to_list metas) in let typ = Tglob (GlobRef.IndRef ip,typs) in MLcase (typ, a, Array.init br_size extract_branch) (*s Extraction of a (co)-fixpoint. *) and extract_fix env sg mle i (fi,ti,ci as recd) mlt = let env = push_rec_types recd env in let metas = Array.map new_meta fi in metas.(i) <- mlt; let mle = Array.fold_left Mlenv.push_type mle metas in let ei = Array.map2 (extract_maybe_term env sg mle) metas ci in MLfix (i, Array.map (fun x -> id_of_name x.binder_name) fi, ei) (*S ML declarations. *) (* [decomp_lams_eta env c t] finds the number [n] of products in the type [t], and decompose the term [c] in [n] lambdas, with eta-expansion if needed. *) let decomp_lams_eta_n n m env sg c t = let rels = fst (splay_prod_n env sg n t) in let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,c)) rels in let rels',c = EConstr.decompose_lam sg c in let d = n - m in (* we'd better keep rels' as long as possible. *) let rels = (List.firstn d rels) @ rels' in let eta_args = List.rev_map EConstr.mkRel (List.interval 1 d) in rels, applistc (EConstr.Vars.lift d c) eta_args (* Let's try to identify some situation where extracted code will allow generalisation of type variables *) let rec gentypvar_ok sg c = match EConstr.kind sg c with | Lambda _ | Const _ -> true | App (c,v) -> (* if all arguments are variables, these variables will disappear after extraction (see [empty_s] below) *) Array.for_all (EConstr.isRel sg) v && gentypvar_ok sg c | Cast (c,_,_) -> gentypvar_ok sg c | _ -> false (*s From a constant to a ML declaration. *) let extract_std_constant env sg kn body typ = reset_meta_count (); (* The short type [t] (i.e. possibly with abbreviations). *) let t = snd (record_constant_type env sg kn (Some typ)) in (* The real type [t']: without head products, expanded, *) (* and with [Tvar] translated to [Tvar'] (not instantiable). *) let l,t' = type_decomp (expand env (var2var' t)) in let s = List.map (type2sign env) l in (* Check for user-declared implicit information *) let s = sign_with_implicits (GlobRef.ConstRef kn) s 0 in (* Decomposing the top level lambdas of [body]. If there isn't enough, it's ok, as long as remaining args aren't to be pruned (and initial lambdas aren't to be all removed if the target language is strict). In other situations, eta-expansions create artificially enough lams (but that may break user's clever let-ins and partial applications). *) let rels, c = let n = List.length s and m = nb_lam sg body in if n <= m then decompose_lam_n sg n body else let s,s' = List.chop m s in if List.for_all ((==) Keep) s' && (lang () == Haskell || sign_kind s != UnsafeLogicalSig) then decompose_lam_n sg m body else decomp_lams_eta_n n m env sg body typ in (* Should we do one eta-expansion to avoid non-generalizable '_a ? *) let rels, c = let n = List.length rels in let s,s' = List.chop n s in let k = sign_kind s in let empty_s = (k == EmptySig || k == SafeLogicalSig) in if lang () == Ocaml && empty_s && not (gentypvar_ok sg c) && not (List.is_empty s') && not (Int.equal (type_maxvar t) 0) then decomp_lams_eta_n (n+1) n env sg body typ else rels,c in let n = List.length rels in let s = List.firstn n s in let l,l' = List.chop n l in let t' = type_recomp (l',t') in (* The initial ML environment. *) let mle = List.fold_left Mlenv.push_std_type Mlenv.empty l in (* The lambdas names. *) let ids = List.map (fun (n,_) -> Id (id_of_name n.binder_name)) rels in (* The according Coq environment. *) let env = push_rels_assum rels env in (* The real extraction: *) let e = extract_term env sg mle t' c [] in (* Expunging term and type from dummy lambdas. *) let trm = term_expunge s (ids,e) in trm, type_expunge_from_sign env s t (* Extracts the type of an axiom, honors the Extraction Implicit declaration. *) let extract_axiom env sg kn typ = reset_meta_count (); (* The short type [t] (i.e. possibly with abbreviations). *) let t = snd (record_constant_type env sg kn (Some typ)) in (* The real type [t']: without head products, expanded, *) (* and with [Tvar] translated to [Tvar'] (not instantiable). *) let l,_ = type_decomp (expand env (var2var' t)) in let s = List.map (type2sign env) l in (* Check for user-declared implicit information *) let s = sign_with_implicits (GlobRef.ConstRef kn) s 0 in type_expunge_from_sign env s t let extract_fixpoint env sg vkn (fi,ti,ci) = let n = Array.length vkn in let types = Array.make n (Tdummy Kprop) and terms = Array.make n (MLdummy Kprop) in let kns = Array.to_list vkn in current_fixpoints := kns; (* for replacing recursive calls [Rel ..] by the corresponding [Const]: *) let sub = List.rev_map EConstr.mkConst kns in for i = 0 to n-1 do if sort_of env sg ti.(i) != InProp then try let e,t = extract_std_constant env sg vkn.(i) (EConstr.Vars.substl sub ci.(i)) ti.(i) in terms.(i) <- e; types.(i) <- t; with SingletonInductiveBecomesProp id -> error_singleton_become_prop id (Some (GlobRef.ConstRef vkn.(i))) done; current_fixpoints := []; Dfix (Array.map (fun kn -> GlobRef.ConstRef kn) vkn, terms, types) (** Because of automatic unboxing the easy way [mk_def c] on the constant body of primitive projections doesn't work. We pretend that they are implemented by matches until someone figures out how to clean it up (test with #4710 when working on this). *) let fake_match_projection env p = let ind = Projection.Repr.inductive p in let proj_arg = Projection.Repr.arg p in let mib, mip = Inductive.lookup_mind_specif env ind in let u = Univ.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in let indu = mkIndU (ind,u) in let ctx, paramslet = let subst = List.init mib.mind_ntypes (fun i -> mkIndU ((fst ind, mib.mind_ntypes - i - 1), u)) in let (ctx, cty) = mip.mind_nf_lc.(0) in let cty = Term.it_mkProd_or_LetIn cty ctx in let rctx, _ = decompose_prod_assum (Vars.substl subst cty) in List.chop mip.mind_consnrealdecls.(0) rctx in let ci_pp_info = { ind_tags = []; cstr_tags = [|Context.Rel.to_tags ctx|]; style = LetStyle } in let ci = { ci_ind = ind; ci_npar = mib.mind_nparams; ci_cstr_ndecls = mip.mind_consnrealdecls; ci_cstr_nargs = mip.mind_consnrealargs; ci_relevance = Declareops.relevance_of_projection_repr mib p; ci_pp_info; } in let x = match mib.mind_record with | NotRecord | FakeRecord -> assert false | PrimRecord info -> let x, _, _, _ = info.(snd ind) in make_annot (Name x) mip.mind_relevance in let indty = mkApp (indu, Context.Rel.to_extended_vect mkRel 0 paramslet) in let rec fold arg j subst = function | [] -> assert false | LocalAssum (na,ty) :: rem -> let ty = Vars.substl subst (liftn 1 j ty) in if arg != proj_arg then let lab = match na.binder_name with Name id -> Label.of_id id | Anonymous -> assert false in let kn = Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg:arg lab in fold (arg+1) (j+1) (mkProj (Projection.make kn false, mkRel 1)::subst) rem else let p = mkLambda (x, lift 1 indty, liftn 1 2 ty) in let branch = lift 1 (it_mkLambda_or_LetIn (mkRel (List.length ctx - (j-1))) ctx) in let body = mkCase (ci, p, mkRel 1, [|branch|]) in it_mkLambda_or_LetIn (mkLambda (x,indty,body)) mib.mind_params_ctxt | LocalDef (_,c,t) :: rem -> let c = liftn 1 j c in let c1 = Vars.substl subst c in fold arg (j+1) (c1::subst) rem in fold 0 1 [] (List.rev ctx) let extract_constant env kn cb = let sg = Evd.from_env env in let r = GlobRef.ConstRef kn in let typ = EConstr.of_constr cb.const_type in let warn_info () = if not (is_custom r) then add_info_axiom r in let warn_log () = if not (constant_has_body cb) then add_log_axiom r in let mk_typ_ax () = let n = type_scheme_nb_args env sg typ in let ids = iterate (fun l -> anonymous_name::l) n [] in Dtype (r, ids, Taxiom) in let mk_typ c = let s,vl = type_sign_vl env sg typ in let db = db_from_sign s in let t = extract_type_scheme env sg db c (List.length s) in Dtype (r, vl, t) in let mk_ax () = let t = extract_axiom env sg kn typ in Dterm (r, MLaxiom, t) in let mk_def c = let e,t = extract_std_constant env sg kn c typ in Dterm (r,e,t) in try match flag_of_type env sg typ with | (Logic,TypeScheme) -> warn_log (); Dtype (r, [], Tdummy Ktype) | (Logic,Default) -> warn_log (); Dterm (r, MLdummy Kprop, Tdummy Kprop) | (Info,TypeScheme) -> (match cb.const_body with | Primitive _ | Undef _ -> warn_info (); mk_typ_ax () | Def c -> (match Recordops.find_primitive_projection kn with | None -> mk_typ (get_body c) | Some p -> let body = fake_match_projection env p in mk_typ (EConstr.of_constr body)) | OpaqueDef c -> add_opaque r; if access_opaque () then mk_typ (get_opaque env c) else mk_typ_ax ()) | (Info,Default) -> (match cb.const_body with | Primitive _ | Undef _ -> warn_info (); mk_ax () | Def c -> (match Recordops.find_primitive_projection kn with | None -> mk_def (get_body c) | Some p -> let body = fake_match_projection env p in mk_def (EConstr.of_constr body)) | OpaqueDef c -> add_opaque r; if access_opaque () then mk_def (get_opaque env c) else mk_ax ()) with SingletonInductiveBecomesProp id -> error_singleton_become_prop id (Some (GlobRef.ConstRef kn)) let extract_constant_spec env kn cb = let sg = Evd.from_env env in let r = GlobRef.ConstRef kn in let typ = EConstr.of_constr cb.const_type in try match flag_of_type env sg typ with | (Logic, TypeScheme) -> Stype (r, [], Some (Tdummy Ktype)) | (Logic, Default) -> Sval (r, Tdummy Kprop) | (Info, TypeScheme) -> let s,vl = type_sign_vl env sg typ in (match cb.const_body with | Undef _ | OpaqueDef _ | Primitive _ -> Stype (r, vl, None) | Def body -> let db = db_from_sign s in let body = get_body body in let t = extract_type_scheme env sg db body (List.length s) in Stype (r, vl, Some t)) | (Info, Default) -> let t = snd (record_constant_type env sg kn (Some typ)) in Sval (r, type_expunge env t) with SingletonInductiveBecomesProp id -> error_singleton_become_prop id (Some (GlobRef.ConstRef kn)) let extract_with_type env sg c = try let typ = type_of env sg c in match flag_of_type env sg typ with | (Info, TypeScheme) -> let s,vl = type_sign_vl env sg typ in let db = db_from_sign s in let t = extract_type_scheme env sg db c (List.length s) in Some (vl, t) | _ -> None with SingletonInductiveBecomesProp id -> error_singleton_become_prop id None let extract_constr env sg c = reset_meta_count (); try let typ = type_of env sg c in match flag_of_type env sg typ with | (_,TypeScheme) -> MLdummy Ktype, Tdummy Ktype | (Logic,_) -> MLdummy Kprop, Tdummy Kprop | (Info,Default) -> let mlt = extract_type env sg [] 1 typ [] in extract_term env sg Mlenv.empty mlt c [], mlt with SingletonInductiveBecomesProp id -> error_singleton_become_prop id None let extract_inductive env kn = let ind = extract_ind env kn in add_recursors env kn; let f i j l = let implicits = implicits_of_global (GlobRef.ConstructRef ((kn,i),j+1)) in let rec filter i = function | [] -> [] | t::l -> let l' = filter (succ i) l in if isTdummy (expand env t) || Int.Set.mem i implicits then l' else t::l' in filter (1+ind.ind_nparams) l in let packets = Array.mapi (fun i p -> { p with ip_types = Array.mapi (f i) p.ip_types }) ind.ind_packets in { ind with ind_packets = packets } (*s Is a [ml_decl] logical ? *) let logical_decl = function | Dterm (_,MLdummy _,Tdummy _) -> true | Dtype (_,[],Tdummy _) -> true | Dfix (_,av,tv) -> (Array.for_all isMLdummy av) && (Array.for_all isTdummy tv) | Dind (_,i) -> Array.for_all (fun ip -> ip.ip_logical) i.ind_packets | _ -> false (*s Is a [ml_spec] logical ? *) let logical_spec = function | Stype (_, [], Some (Tdummy _)) -> true | Sval (_,Tdummy _) -> true | Sind (_,i) -> Array.for_all (fun ip -> ip.ip_logical) i.ind_packets | _ -> false coq-8.11.0/plugins/extraction/miniml.mli0000644000175000017500000001434513612664533020126 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* string; (* the second argument is a comment to add to the preamble *) preamble : Id.t -> Pp.t option -> ModPath.t list -> unsafe_needs -> Pp.t; pp_struct : ml_structure -> Pp.t; (* Concerning a possible interface file *) sig_suffix : string option; (* the second argument is a comment to add to the preamble *) sig_preamble : Id.t -> Pp.t option -> ModPath.t list -> unsafe_needs -> Pp.t; pp_sig : ml_signature -> Pp.t; (* for an isolated declaration print *) pp_decl : ml_decl -> Pp.t; } coq-8.11.0/plugins/extraction/common.ml0000644000175000017500000005205113612664533017754 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* true | _ -> false (*s Some pretty-print utility functions. *) let pp_par par st = if par then str "(" ++ st ++ str ")" else st (** [pp_apply] : a head part applied to arguments, possibly with parenthesis *) let pp_apply st par args = match args with | [] -> st | _ -> hov 2 (pp_par par (st ++ spc () ++ prlist_with_sep spc identity args)) (** Same as [pp_apply], but with also protection of the head by parenthesis *) let pp_apply2 st par args = let par' = not (List.is_empty args) || par in pp_apply (pp_par par' st) par args let pr_binding = function | [] -> mt () | l -> str " " ++ prlist_with_sep (fun () -> str " ") Id.print l let pp_tuple_light f = function | [] -> mt () | [x] -> f true x | l -> pp_par true (prlist_with_sep (fun () -> str "," ++ spc ()) (f false) l) let pp_tuple f = function | [] -> mt () | [x] -> f x | l -> pp_par true (prlist_with_sep (fun () -> str "," ++ spc ()) f l) let pp_boxed_tuple f = function | [] -> mt () | [x] -> f x | l -> pp_par true (hov 0 (prlist_with_sep (fun () -> str "," ++ spc ()) f l)) (** By default, in module Format, you can do horizontal placing of blocks even if they include newlines, as long as the number of chars in the blocks is less that a line length. To avoid this awkward situation, we attach a big virtual size to [fnl] newlines. *) (* EG: This looks quite suspicious... but beware of bugs *) (* let fnl () = stras (1000000,"") ++ fnl () *) let fnl () = fnl () let fnl2 () = fnl () ++ fnl () let space_if = function true -> str " " | false -> mt () let begins_with s prefix = let len = String.length prefix in String.length s >= len && String.equal (String.sub s 0 len) prefix let begins_with_CoqXX s = let n = String.length s in n >= 4 && s.[0] == 'C' && s.[1] == 'o' && s.[2] == 'q' && let i = ref 3 in try while !i < n do match s.[!i] with | '_' -> i:=n (*Stop*) | '0'..'9' -> incr i | _ -> raise Not_found done; true with Not_found -> false let unquote s = if lang () != Scheme then s else String.map (fun c -> if c == '\'' then '~' else c) s let rec qualify delim = function | [] -> assert false | [s] -> s | ""::l -> qualify delim l | s::l -> s^delim^(qualify delim l) let dottify = qualify "." let pseudo_qualify = qualify "__" (*s Uppercase/lowercase renamings. *) let is_upper s = match s.[0] with 'A' .. 'Z' -> true | _ -> false let is_lower s = match s.[0] with 'a' .. 'z' | '_' -> true | _ -> false let lowercase_id id = Id.of_string (String.uncapitalize_ascii (ascii_of_id id)) let uppercase_id id = let s = ascii_of_id id in assert (not (String.is_empty s)); if s.[0] == '_' then Id.of_string ("Coq_"^s) else Id.of_string (String.capitalize_ascii s) type kind = Term | Type | Cons | Mod module KOrd = struct type t = kind * string let compare (k1, s1) (k2, s2) = let c = pervasives_compare k1 k2 (* OK *) in if c = 0 then String.compare s1 s2 else c end module KMap = Map.Make(KOrd) let upperkind = function | Type -> lang () == Haskell | Term -> false | Cons | Mod -> true let kindcase_id k id = if upperkind k then uppercase_id id else lowercase_id id (*s de Bruijn environments for programs *) type env = Id.t list * Id.Set.t (*s Generic renaming issues for local variable names. *) let rec rename_id id avoid = if Id.Set.mem id avoid then rename_id (increment_subscript id) avoid else id let rec rename_vars avoid = function | [] -> [], avoid | id :: idl when id == dummy_name -> (* we don't rename dummy binders *) let (idl', avoid') = rename_vars avoid idl in (id :: idl', avoid') | id :: idl -> let (idl, avoid) = rename_vars avoid idl in let id = rename_id (lowercase_id id) avoid in (id :: idl, Id.Set.add id avoid) let rename_tvars avoid l = let rec rename avoid = function | [] -> [],avoid | id :: idl -> let id = rename_id (lowercase_id id) avoid in let idl, avoid = rename (Id.Set.add id avoid) idl in (id :: idl, avoid) in fst (rename avoid l) let push_vars ids (db,avoid) = let ids',avoid' = rename_vars avoid ids in ids', (ids' @ db, avoid') let get_db_name n (db,_) = List.nth db (pred n) (*S Renamings of global objects. *) (*s Tables of global renamings *) let register_cleanup, do_cleanup = let funs = ref [] in (fun f -> funs:=f::!funs), (fun () -> List.iter (fun f -> f ()) !funs) type phase = Pre | Impl | Intf let set_phase, get_phase = let ph = ref Impl in ((:=) ph), (fun () -> !ph) let set_keywords, get_keywords = let k = ref Id.Set.empty in ((:=) k), (fun () -> !k) let add_global_ids, get_global_ids = let ids = ref Id.Set.empty in register_cleanup (fun () -> ids := get_keywords ()); let add s = ids := Id.Set.add s !ids and get () = !ids in (add,get) let empty_env () = [], get_global_ids () (* We might have built [global_reference] whose canonical part is inaccurate. We must hence compare only the user part, hence using a Hashtbl might be incorrect *) let mktable_id autoclean = let m = ref Id.Map.empty in let clear () = m := Id.Map.empty in if autoclean then register_cleanup clear; (fun r v -> m := Id.Map.add r v !m), (fun r -> Id.Map.find r !m), clear let mktable_ref autoclean = let m = ref Refmap'.empty in let clear () = m := Refmap'.empty in if autoclean then register_cleanup clear; (fun r v -> m := Refmap'.add r v !m), (fun r -> Refmap'.find r !m), clear let mktable_modpath autoclean = let m = ref MPmap.empty in let clear () = m := MPmap.empty in if autoclean then register_cleanup clear; (fun r v -> m := MPmap.add r v !m), (fun r -> MPmap.find r !m), clear (* A table recording objects in the first level of all MPfile *) let add_mpfiles_content,get_mpfiles_content,clear_mpfiles_content = mktable_modpath false let get_mpfiles_content mp = try get_mpfiles_content mp with Not_found -> failwith "get_mpfiles_content" (*s The list of external modules that will be opened initially *) let mpfiles_add, mpfiles_mem, mpfiles_list, mpfiles_clear = let m = ref MPset.empty in let add mp = m:=MPset.add mp !m and mem mp = MPset.mem mp !m and list () = MPset.elements !m and clear () = m:=MPset.empty in register_cleanup clear; (add,mem,list,clear) (*s List of module parameters that we should alpha-rename *) let params_ren_add, params_ren_mem = let m = ref MPset.empty in let add mp = m:=MPset.add mp !m and mem mp = MPset.mem mp !m and clear () = m:=MPset.empty in register_cleanup clear; (add,mem) (*s table indicating the visible horizon at a precise moment, i.e. the stack of structures we are inside. - The sequence of [mp] parts should have the following form: a [MPfile] at the beginning, and then more and more [MPdot] over this [MPfile], or [MPbound] when inside the type of a module parameter. - the [params] are the [MPbound] when [mp] is a functor, the innermost [MPbound] coming first in the list. - The [content] part is used to record all the names already seen at this level. *) type visible_layer = { mp : ModPath.t; params : ModPath.t list; mutable content : Label.t KMap.t; } let pop_visible, push_visible, get_visible = let vis = ref [] in register_cleanup (fun () -> vis := []); let pop () = match !vis with | [] -> assert false | v :: vl -> vis := vl; (* we save the 1st-level-content of MPfile for later use *) if get_phase () == Impl && modular () && is_modfile v.mp then add_mpfiles_content v.mp v.content and push mp mps = vis := { mp = mp; params = mps; content = KMap.empty } :: !vis and get () = !vis in (pop,push,get) let get_visible_mps () = List.map (function v -> v.mp) (get_visible ()) let top_visible () = match get_visible () with [] -> assert false | v::_ -> v let top_visible_mp () = (top_visible ()).mp let add_visible ks l = let visible = top_visible () in visible.content <- KMap.add ks l visible.content (* table of local module wrappers used to provide non-ambiguous names *) module DupOrd = struct type t = ModPath.t * Label.t let compare (mp1, l1) (mp2, l2) = let c = Label.compare l1 l2 in if Int.equal c 0 then ModPath.compare mp1 mp2 else c end module DupMap = Map.Make(DupOrd) let add_duplicate, get_duplicate = let index = ref 0 and dups = ref DupMap.empty in register_cleanup (fun () -> index := 0; dups := DupMap.empty); let add mp l = incr index; let ren = "Coq__" ^ string_of_int !index in dups := DupMap.add (mp,l) ren !dups and get mp l = try Some (DupMap.find (mp, l) !dups) with Not_found -> None in (add,get) type reset_kind = AllButExternal | Everything let reset_renaming_tables flag = do_cleanup (); if flag == Everything then clear_mpfiles_content () (*S Renaming functions *) (* This function creates from [id] a correct uppercase/lowercase identifier. This is done by adding a [Coq_] or [coq_] prefix. To avoid potential clashes with previous [Coq_id] variable, these prefixes are duplicated if already existing. *) let modular_rename k id = let s = ascii_of_id id in let prefix,is_ok = if upperkind k then "Coq_",is_upper else "coq_",is_lower in if not (is_ok s) || Id.Set.mem id (get_keywords ()) || begins_with s prefix then prefix ^ s else s (*s For monolithic extraction, first-level modules might have to be renamed with unique numbers *) let modfstlev_rename = let add_index,get_index,_ = mktable_id true in fun l -> let id = Label.to_id l in try let n = get_index id in add_index id (n+1); let s = if n == 0 then "" else string_of_int (n-1) in "Coq"^s^"_"^(ascii_of_id id) with Not_found -> let s = ascii_of_id id in if is_lower s || begins_with_CoqXX s then (add_index id 1; "Coq_"^s) else (add_index id 0; s) (*s Creating renaming for a [module_path] : first, the real function ... *) let rec mp_renaming_fun mp = match mp with | _ when not (modular ()) && at_toplevel mp -> [""] | MPdot (mp,l) -> let lmp = mp_renaming mp in let mp = match lmp with | [""] -> modfstlev_rename l | _ -> modular_rename Mod (Label.to_id l) in mp ::lmp | MPbound mbid -> let s = modular_rename Mod (MBId.to_id mbid) in if not (params_ren_mem mp) then [s] else let i,_,_ = MBId.repr mbid in [s^"__"^string_of_int i] | MPfile _ -> assert (modular ()); (* see [at_toplevel] above *) assert (get_phase () == Pre); let current_mpfile = (List.last (get_visible ())).mp in if not (ModPath.equal mp current_mpfile) then mpfiles_add mp; [string_of_modfile mp] (* ... and its version using a cache *) and mp_renaming = let add,get,_ = mktable_modpath true in fun x -> try if is_mp_bound (base_mp x) then raise Not_found; get x with Not_found -> let y = mp_renaming_fun x in add x y; y (*s Renamings creation for a [global_reference]: we build its fully-qualified name in a [string list] form (head is the short name). *) let ref_renaming_fun (k,r) = let mp = modpath_of_r r in let l = mp_renaming mp in let l = if lang () != Ocaml && not (modular ()) then [""] else l in let s = let idg = safe_basename_of_global r in match l with | [""] -> (* this happens only at toplevel of the monolithic case *) let globs = get_global_ids () in let id = next_ident_away (kindcase_id k idg) globs in Id.to_string id | _ -> modular_rename k idg in add_global_ids (Id.of_string s); s::l (* Cached version of the last function *) let ref_renaming = let add,get,_ = mktable_ref true in fun ((k,r) as x) -> try if is_mp_bound (base_mp (modpath_of_r r)) then raise Not_found; get r with Not_found -> let y = ref_renaming_fun x in add r y; y (* [visible_clash mp0 (k,s)] checks if [mp0-s] of kind [k] can be printed as [s] in the current context of visible modules. More precisely, we check if there exists a visible [mp] that contains [s]. The verification stops if we encounter [mp=mp0]. *) let rec clash mem mp0 ks = function | [] -> false | mp :: _ when ModPath.equal mp mp0 -> false | mp :: _ when mem mp ks -> true | _ :: mpl -> clash mem mp0 ks mpl let mpfiles_clash mp0 ks = clash (fun mp k -> KMap.mem k (get_mpfiles_content mp)) mp0 ks (List.rev (mpfiles_list ())) let rec params_lookup mp0 ks = function | [] -> false | param :: _ when ModPath.equal mp0 param -> true | param :: params -> let () = match ks with | (Mod, mp) when String.equal (List.hd (mp_renaming param)) mp -> params_ren_add param | _ -> () in params_lookup mp0 ks params let visible_clash mp0 ks = let rec clash = function | [] -> false | v :: _ when ModPath.equal v.mp mp0 -> false | v :: vis -> let b = KMap.mem ks v.content in if b && not (is_mp_bound mp0) then true else begin if b then params_ren_add mp0; if params_lookup mp0 ks v.params then false else clash vis end in clash (get_visible ()) (* Same, but with verbose output (and mp0 shouldn't be a MPbound) *) let visible_clash_dbg mp0 ks = let rec clash = function | [] -> None | v :: _ when ModPath.equal v.mp mp0 -> None | v :: vis -> try Some (v.mp,KMap.find ks v.content) with Not_found -> if params_lookup mp0 ks v.params then None else clash vis in clash (get_visible ()) (* After the 1st pass, we can decide which modules will be opened initially *) let opened_libraries () = if not (modular ()) then [] else let used_files = mpfiles_list () in let used_ks = List.map (fun mp -> Mod,string_of_modfile mp) used_files in (* By default, we open all used files. Ambiguities will be resolved later by using qualified names. Nonetheless, we don't open any file A that contains an immediate submodule A.B hiding another file B : otherwise, after such an open, there's no unambiguous way to refer to objects of B. *) let to_open = List.filter (fun mp -> not (List.exists (fun k -> KMap.mem k (get_mpfiles_content mp)) used_ks)) used_files in mpfiles_clear (); List.iter mpfiles_add to_open; mpfiles_list () (*s On-the-fly qualification issues for both monolithic or modular extraction. *) (* [pp_ocaml_gen] below is a function that factorize the printing of both [global_reference] and module names for ocaml. When [k=Mod] then [olab=None], otherwise it contains the label of the reference to print. [rls] is the string list giving the qualified name, short name at the end. *) (* In Coq, we can qualify [M.t] even if we are inside [M], but in Ocaml we cannot do that. So, if [t] gets hidden and we need a long name for it, we duplicate the _definition_ of t in a Coq__XXX module, and similarly for a sub-module [M.N] *) let pp_duplicate k' prefix mp rls olab = let rls', lbl = if k' != Mod then (* Here rls=[s], the ref to print is ., and olab<>None *) rls, Option.get olab else (* Here rls=s::rls', we search the label for s inside mp *) List.tl rls, get_nth_label_mp (mp_length mp - mp_length prefix) mp in match get_duplicate prefix lbl with | Some ren -> dottify (ren :: rls') | None -> assert (get_phase () == Pre); (* otherwise it's too late *) add_duplicate prefix lbl; dottify rls let fstlev_ks k = function | [] -> assert false | [s] -> k,s | s::_ -> Mod,s (* [pp_ocaml_local] : [mp] has something in common with [top_visible ()] but isn't equal to it *) let pp_ocaml_local k prefix mp rls olab = (* what is the largest prefix of [mp] that belongs to [visible]? *) assert (k != Mod || not (ModPath.equal mp prefix)); (* mp as whole module isn't in itself *) let rls' = List.skipn (mp_length prefix) rls in let k's = fstlev_ks k rls' in (* Reference r / module path mp is of the form [.s.<...>]. *) if not (visible_clash prefix k's) then dottify rls' else pp_duplicate (fst k's) prefix mp rls' olab (* [pp_ocaml_bound] : [mp] starts with a [MPbound], and we are not inside (i.e. we are not printing the type of the module parameter) *) let pp_ocaml_bound base rls = (* clash with a MPbound will be detected and fixed by renaming this MPbound *) if get_phase () == Pre then ignore (visible_clash base (Mod,List.hd rls)); dottify rls (* [pp_ocaml_extern] : [mp] isn't local, it is defined in another [MPfile]. *) let pp_ocaml_extern k base rls = match rls with | [] -> assert false | base_s :: rls' -> if (not (modular ())) (* Pseudo qualification with "" *) || (List.is_empty rls') (* Case of a file A.v used as a module later *) || (not (mpfiles_mem base)) (* Module not opened *) || (mpfiles_clash base (fstlev_ks k rls')) (* Conflict in opened files *) || (visible_clash base (fstlev_ks k rls')) (* Local conflict *) then (* We need to fully qualify. Last clash situation is unsupported *) match visible_clash_dbg base (Mod,base_s) with | None -> dottify rls | Some (mp,l) -> error_module_clash base (MPdot (mp,l)) else (* Standard situation : object in an opened file *) dottify rls' (* [pp_ocaml_gen] : choosing between [pp_ocaml_local] or [pp_ocaml_extern] *) let pp_ocaml_gen k mp rls olab = match common_prefix_from_list mp (get_visible_mps ()) with | Some prefix -> pp_ocaml_local k prefix mp rls olab | None -> let base = base_mp mp in if is_mp_bound base then pp_ocaml_bound base rls else pp_ocaml_extern k base rls (* For Haskell, things are simpler: we have removed (almost) all structures *) let pp_haskell_gen k mp rls = match rls with | [] -> assert false | s::rls' -> let str = pseudo_qualify rls' in let str = if is_upper str && not (upperkind k) then ("_"^str) else str in if ModPath.equal (base_mp mp) (top_visible_mp ()) then str else s^"."^str (* Main name printing function for a reference *) let pp_global k r = let ls = ref_renaming (k,r) in assert (List.length ls > 1); let s = List.hd ls in let mp,l = repr_of_r r in if ModPath.equal mp (top_visible_mp ()) then (* simplest situation: definition of r (or use in the same context) *) (* we update the visible environment *) (add_visible (k,s) l; unquote s) else let rls = List.rev ls in (* for what come next it's easier this way *) match lang () with | Scheme -> unquote s (* no modular Scheme extraction... *) | JSON -> dottify (List.map unquote rls) | Haskell -> if modular () then pp_haskell_gen k mp rls else s | Ocaml -> pp_ocaml_gen k mp rls (Some l) (* The next function is used only in Ocaml extraction...*) let pp_module mp = let ls = mp_renaming mp in match mp with | MPdot (mp0,l) when ModPath.equal mp0 (top_visible_mp ()) -> (* simplest situation: definition of mp (or use in the same context) *) (* we update the visible environment *) let s = List.hd ls in add_visible (Mod,s) l; s | _ -> pp_ocaml_gen Mod mp (List.rev ls) None (** Special hack for constants of type Ascii.ascii : if an [Extract Inductive ascii => char] has been declared, then the constants are directly turned into chars *) let mk_ind path s = MutInd.make2 (MPfile (dirpath_of_string path)) (Label.make s) let ind_ascii = mk_ind "Coq.Strings.Ascii" "ascii" let check_extract_ascii () = try let char_type = match lang () with | Ocaml -> "char" | Haskell -> "Prelude.Char" | _ -> raise Not_found in String.equal (find_custom (GlobRef.IndRef (ind_ascii, 0))) (char_type) with Not_found -> false let is_list_cons l = List.for_all (function MLcons (_,GlobRef.ConstructRef(_,_),[]) -> true | _ -> false) l let is_native_char = function | MLcons(_,GlobRef.ConstructRef ((kn,0),1),l) -> MutInd.equal kn ind_ascii && check_extract_ascii () && is_list_cons l | _ -> false let get_native_char c = let rec cumul = function | [] -> 0 | MLcons(_,GlobRef.ConstructRef(_,j),[])::l -> (2-j) + 2 * (cumul l) | _ -> assert false in let l = match c with MLcons(_,_,l) -> l | _ -> assert false in Char.chr (cumul l) let pp_native_char c = str ("'"^Char.escaped (get_native_char c)^"'") coq-8.11.0/plugins/extraction/CHANGES0000644000175000017500000003342213612664533017126 0ustar treinentreinen8.0 -> today See the main CHANGES file in the archive 7.4 -> 8.0 No revolution this time. Mostly "behind-the-scene" clean-up and bug-fixes, but also a few steps toward a more user-friendly extraction: * syntax of extraction: - The old (Recursive) Extraction Module M. is now (Recursive) Extraction Library M. The old name was misleading since this command only works with M being a library M.v, and not a module produced by interactive command Module M. - The other commands Extraction foo. Recursive Extraction foo bar. Extraction "myfile.ml" foo bar. now accept that foo can be a module name instead of just a constant name. * Support of type scheme axioms (i.e. axiom whose type is an arity (x1:X1)...(xn:Xn)s with s a sort). For example: Axiom myprod : Set -> Set -> Set. Extract Constant myprod "'a" "'b" => "'a * 'b". Recursive Extraction myprod. -------> type ('a,'b) myprod = 'a * 'b * More flexible support of axioms. When an axiom isn't realized via Extract Constant before extraction, a warning is produced (instead of an error), and the extracted code must be completed later by hand. To find what needs to be completed, search for the following string: AXIOM TO BE REALIZED * Cosmetics: When extraction produces a file, it tells it. * (Experimental) It is allowed to extract under a opened interactive module (but still outside sections). Feature to be used with caution. * A problem has been identified concerning .v files used as normal interactive modules, like in Definition foo :=O. Require A. Module M:=A Extraction M. I might try to support that in the future. In the meanwhile, the current behaviour of extraction is to forbid this. * bug fixes: - many concerning Records. - a Stack Overflow with mutual inductive (BZ#320) - some optimizations have been removed since they were not type-safe: For example if e has type: type 'x a = A Then: match e with A -> A -----X----> e To be investigated further. 7.3 -> 7.4 * The two main new features: - Automatic generation of Obj.magic when the extracted code in Ocaml is not directly typable. - An experimental extraction of Coq's new modules to Ocaml modules. * Concerning those Obj.magic: - The extraction now computes the expected type of any terms. Then it compares it with the actual type of the produced code. And when a mismatch is found, a Obj.magic is inserted. - As a rule, any extracted development that was compiling out of the box should not contain any Obj.magic. At the other hand, generation of Obj.magic is not optimized yet: there might be several of them at a place were one would have been enough. - Examples of code needing those Obj.magic: * plugins/extraction/test_extraction.v in the Coq source * in the users' contributions: Lannion Lyon/CIRCUITS Rocq/HIGMAN - As a side-effect of this Obj.magic feature, we now print the types of the extracted terms, both in .ml files as commented documentation and in interfaces .mli files - This feature hasn't been ported yet to Haskell. We are aware of some unsafe casting functions like "unsafeCoerce" on some Haskell implems. So it will eventually be done. * Concerning the extraction of Coq's new modules: - Taking in account the new Coq's modules system has implied a *huge* rewrite of most of the extraction code. - The extraction core (translation from Coq to an abstract mini-ML) is now complete and fairly stable, and supports modules, modules type and functors and all that stuff. - The ocaml pretty-print part, especially the renaming issue, is clearly weaker, and certainly still contains bugs. - Nothing done for translating these Coq Modules to Haskell. - A temporary drawback of this module extraction implementation is that efficiency (especially extraction speed) has been somehow neglected. To improve ... - As an interesting side-effect, definitions are now printed according to the user's original order. No more of this "dependency-correct but weird" order. In particular realized axioms via Extract Constant are now at their right place, and not at the beginning. * Other news: - Records are now printed using the Ocaml record syntax - Syntax output toward Scheme. Quite funny, but quite experimental and not documented. I recommend using the bigloo compiler since it contains natively some pattern matching. - the dummy constant "__" have changed. see README - a few bug-fixes (BZ#191 and others) 7.2 -> 7.3 * Improved documentation in the Reference Manual. * Theoretical bad news: - a naughty example (see the end of test_extraction.v) forced me to stop eliminating lambdas and arguments corresponding to so-called "arity" in the general case. - The dummy constant used in extraction ( let prop = () in ocaml ) may in some cases be applied to arguments. This problem is dealt by generating sufficient abstraction before the (). * Theoretical good news: - there is now a mechanism that remove useless prop/arity lambdas at the top of function declarations. If your function had signature nat -> prop -> nat in the previous extraction, it will now be nat -> nat. So the extractions of common terms should look very much like the old V6.2 one, except in some particular cases (functions as parameters, partial applications, etc). In particular the bad news above have nearly no impact... * By the way there is no more "let prop = ()" in ocaml. Those () are directly inlined. And in Haskell the dummy constant is now __ (two underscore) and is defined by __ = Prelude.error "Logical or arity value used" This dummy constant should never be evaluated when computing an informative value, thanks to the lazy strategy. Hence the error message. * Syntax changes, see Documentation for details: Extraction Language Ocaml. Extraction Language Haskell. Extraction Language Toplevel. That fixes the target language of extraction. Default is Ocaml, even in the coq toplevel: you can now do copy-paste from the coq toplevel without renaming problems. Toplevel language is the ocaml pseudo-language used previously used inside the coq toplevel: coq names are printed with the coq way, i.e. with no renaming. So there is no more particular commands for Haskell, like Haskell Extraction "file" id. Just set your favourite language and go... * Haskell extraction has been tested at last (and corrected...). See specificities in Documentation. * Extraction of CoInductive in Ocaml language is now correct: it uses the Lazy.force and lazy features of Ocaml. * Modular extraction in Ocaml is now far more readable: instead of qualifying everywhere (A.foo), there are now some "open" at the beginning of files. Possible clashes are dealt with. * By default, any recursive function associated with an inductive type (foo_rec and foo_rect when foo is inductive type) will now be inlined in extracted code. * A few constants are explicitly declared to be inlined in extracted code. For the moment there are: Wf.Acc_rec Wf.Acc_rect Wf.well_founded_induction Wf.well_founded_induction_type Those constants does not match the auto-inlining criterion based on strictness. Of course, you can still override this behaviour via some Extraction NoInline. * There is now a web page showing the extraction of all standard theories: http://www.lri.fr/~letouzey/extraction 7.1 -> 7.2 : * Syntax changes, see Documentation for more details: Set/Unset Extraction Optimize. Default is Set. This control all optimizations made on the ML terms (mostly reduction of dummy beta/iota redexes, but also simplications on Cases, etc). Put this option to Unset if you what a ML term as close as possible to the Coq term. Set/Unset Extraction AutoInline. Default in Set, so by default, the extraction mechanism feels free to inline the bodies of some defined constants, according to some heuristics like size of bodies, useness of some arguments, etc. Those heuristics are not always perfect, you may want to disable this feature, do it by Unset. Extraction Inline toto foo. Extraction NoInline titi faa bor. In addition to the automatic inline feature, you can now tell precisely to inline some more constants by the Extraction Inline command. Conversely, you can forbid the inlining of some specific constants by automatic inlining. Those two commands enable a precise control of what is inlined and what is not. Print Extraction Inline. Sum up the current state of the table recording the custom inlinings (Extraction (No)Inline). Reset Extraction Inline. Put the table recording the custom inlinings back to empty. As a consequence, there is no more need for options inside the commands of extraction: Extraction foo. Recursive Extraction foo bar. Extraction "file" foo bar. Extraction Module Mymodule. Recursive Extraction Module Mymodule. New: The last syntax extracts the module Mymodule and all the modules it depends on. You can also try the Haskell versions (not tested yet): Haskell Extraction foo. Haskell Recursive Extraction foo bar. Haskell Extraction "file" foo bar. Haskell Extraction Module Mymodule. Haskell Recursive Extraction Module Mymodule. And there's still the realization syntax: Extract Constant coq_bla => "caml_bla". Extract Inlined Constant coq_bla => "caml_bla". Extract Inductive myinductive => mycamlind [my_caml_constr1 ... ]. Note that now, the Extract Inlined Constant command is sugar for an Extract Constant followed by a Extraction Inline. So be careful with Reset Extraction Inline. * Lot of works around optimization of produced code. Should make code more readable. - fixpoint definitions : there should be no more stupid printings like let foo x = let rec f x = .... (f y) .... in f x but rather let rec foo x = .... (foo y) .... - generalized iota (in particular iota and permutation cases/cases): A generalized iota redex is a "Cases e of ...." where e is ok. And the recursive predicate "ok" is given by: e is ok if e is a Constructor or a Cases where all branches are ok. In the case of generalized iota redex, it might be good idea to reduce it, so we do it. Example: match (match t with O -> Left | S n -> match n with O -> Right | S m -> Left) with Left -> blabla | Right -> bloblo After simplification, that gives: match t with O -> blabla | S n -> match n with O -> bloblo | S n -> blabla As shown on the example, code duplication can occur. In practice it seems not to happen frequently. - "constant" case: In V7.1 we used to simplify cases where all branches are the same. In V7.2 we can simplify in addition terms like cases e of C1 x y -> f (C x y) | C2 z -> f (C2 z) If x y z don't occur in f, we can produce (f e). - permutation cases/fun: extracted code has frequenty functions in branches of cases: let foo x = match x with O -> fun _ -> .... | S y -> fun _ -> .... the optimization consist in lifting the common "fun _ ->", and that gives let foo x _ = match x with O -> ..... | S y -> .... * Some bug corrections (many thanks in particular to Michel Levy). * Testing in coq contributions: If you are interested in extraction, you can look at the extraction tests I'have put in the following coq contributions Bordeaux/Additions computation of fibonacci(2000) Bordeaux/EXCEPTIONS multiplication using exception. Bordeaux/SearchTrees list -> binary tree. maximum. Dyade/BDDS boolean tautology checker. Lyon/CIRCUITS multiplication via a modelization of a circuit. Lyon/FIRING-SQUAD print the states of the firing squad. Marseille/CIRCUITS compares integers via a modelization of a circuit. Nancy/FOUnify unification of two first-order terms. Rocq/ARITH/Chinese computation of the chinese remainder. Rocq/COC small coc typechecker. (test by B. Barras, not by me) Rocq/HIGMAN run the proof on one example. Rocq/GRAPHS linear constraints checker in Z. Sophia-Antipolis/Stalmarck boolean tautology checker. Suresnes/BDD boolean tautology checker. Just do "make" in those contributions, the extraction test is integrated. More tests will follow on more contributions. 7.0 -> 7.1 : mostly bug corrections. No theoretical problems dealed with. * The semantics of Extract Constant changed: If you provide a extraction for p by Extract Constant p => "0", your generated ML file will begin by a let p = 0. The old semantics, which was to replace p everywhere by the provided terms, is still available via the Extract Inlined Constant p => "0" syntax. * There are more optimizations applied to the generated code: - identity cases: match e with P x y -> P x y | Q z -> Q z | ... is simplified into e. Especially interesting with the sumbool terms: there will be no more match ... with Left -> Left | Right -> Right - constant cases: match e with P x y -> c | Q z -> c | ... is simplified into c as soon as x, y, z do not occur in c. So no more match ... with Left -> Left | Right -> Left. * the extraction at Toplevel (Extraction foo and Recursive Extraction foo), which was only a development tool at the beginning, is now closer to the real extraction to a file. In particular optimizations are done, and constants like recursors ( ..._rec ) are expanded. * the singleton optimization is now protected against circular type. ( Remind : this optimization is the one that simplify type 'a sig = Exists of 'a into type 'a sig = 'a and match e with (Exists c) -> d into let c = e in d ) * Fixed one bug concerning casted code * The inductives generated should now have always correct type-var list ('a,'b,'c...) * Code cleanup until three days before release. Messing-up code in the last three days before release. 6.x -> 7.0 : Everything changed. See README coq-8.11.0/plugins/extraction/table.mli0000644000175000017500000001571213612664533017727 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Id.t (*s Warning and Error messages. *) val warning_axioms : unit -> unit val warning_opaques : bool -> unit val warning_ambiguous_name : ?loc:Loc.t -> qualid * ModPath.t * GlobRef.t -> unit val warning_id : string -> unit val error_axiom_scheme : GlobRef.t -> int -> 'a val error_constant : GlobRef.t -> 'a val error_inductive : GlobRef.t -> 'a val error_nb_cons : unit -> 'a val error_module_clash : ModPath.t -> ModPath.t -> 'a val error_no_module_expr : ModPath.t -> 'a val error_singleton_become_prop : Id.t -> GlobRef.t option -> 'a val error_unknown_module : qualid -> 'a val error_scheme : unit -> 'a val error_not_visible : GlobRef.t -> 'a val error_MPfile_as_mod : ModPath.t -> bool -> 'a val check_inside_module : unit -> unit val check_inside_section : unit -> unit val check_loaded_modfile : ModPath.t -> unit val msg_of_implicit : kill_reason -> string val err_or_warn_remaining_implicit : kill_reason -> unit val info_file : string -> unit (*s utilities about [module_path] and [kernel_names] and [GlobRef.t] *) val occur_kn_in_ref : MutInd.t -> GlobRef.t -> bool val repr_of_r : GlobRef.t -> ModPath.t * Label.t val modpath_of_r : GlobRef.t -> ModPath.t val label_of_r : GlobRef.t -> Label.t val base_mp : ModPath.t -> ModPath.t val is_modfile : ModPath.t -> bool val string_of_modfile : ModPath.t -> string val file_of_modfile : ModPath.t -> string val is_toplevel : ModPath.t -> bool val at_toplevel : ModPath.t -> bool val mp_length : ModPath.t -> int val prefixes_mp : ModPath.t -> MPset.t val common_prefix_from_list : ModPath.t -> ModPath.t list -> ModPath.t option val get_nth_label_mp : int -> ModPath.t -> Label.t val labels_of_ref : GlobRef.t -> ModPath.t * Label.t list (*s Some table-related operations *) (* For avoiding repeated extraction of the same constant or inductive, we use cache functions below. Indexing by constant name isn't enough, due to modules we could have a same constant name but different content. So we check that the [constant_body] hasn't changed from recording time to retrieving time. Same for inductive : we store [mutual_inductive_body] as checksum. In both case, we should ideally also check the env *) val add_typedef : Constant.t -> Opaqueproof.opaque constant_body -> ml_type -> unit val lookup_typedef : Constant.t -> Opaqueproof.opaque constant_body -> ml_type option val add_cst_type : Constant.t -> Opaqueproof.opaque constant_body -> ml_schema -> unit val lookup_cst_type : Constant.t -> Opaqueproof.opaque constant_body -> ml_schema option val add_ind : MutInd.t -> mutual_inductive_body -> ml_ind -> unit val lookup_ind : MutInd.t -> mutual_inductive_body -> ml_ind option val add_inductive_kind : MutInd.t -> inductive_kind -> unit val is_coinductive : GlobRef.t -> bool val is_coinductive_type : ml_type -> bool (* What are the fields of a record (empty for a non-record) *) val get_record_fields : GlobRef.t -> GlobRef.t option list val record_fields_of_type : ml_type -> GlobRef.t option list val add_recursors : Environ.env -> MutInd.t -> unit val is_recursor : GlobRef.t -> bool val add_projection : int -> Constant.t -> inductive -> unit val is_projection : GlobRef.t -> bool val projection_arity : GlobRef.t -> int val projection_info : GlobRef.t -> inductive * int (* arity *) val add_info_axiom : GlobRef.t -> unit val remove_info_axiom : GlobRef.t -> unit val add_log_axiom : GlobRef.t -> unit val add_opaque : GlobRef.t -> unit val remove_opaque : GlobRef.t -> unit val reset_tables : unit -> unit (*s AccessOpaque parameter *) val access_opaque : unit -> bool (*s AutoInline parameter *) val auto_inline : unit -> bool (*s TypeExpand parameter *) val type_expand : unit -> bool (*s KeepSingleton parameter *) val keep_singleton : unit -> bool (*s Optimize parameter *) type opt_flag = { opt_kill_dum : bool; (* 1 *) opt_fix_fun : bool; (* 2 *) opt_case_iot : bool; (* 4 *) opt_case_idr : bool; (* 8 *) opt_case_idg : bool; (* 16 *) opt_case_cst : bool; (* 32 *) opt_case_fun : bool; (* 64 *) opt_case_app : bool; (* 128 *) opt_let_app : bool; (* 256 *) opt_lin_let : bool; (* 512 *) opt_lin_beta : bool } (* 1024 *) val optims : unit -> opt_flag (*s Controls whether dummy lambda are removed *) val conservative_types : unit -> bool (*s A comment to print at the beginning of the files *) val file_comment : unit -> string (*s Target language. *) type lang = Ocaml | Haskell | Scheme | JSON val lang : unit -> lang (*s Extraction modes: modular or monolithic, library or minimal ? Nota: - Recursive Extraction : monolithic, minimal - Separate Extraction : modular, minimal - Extraction Library : modular, library *) val set_modular : bool -> unit val modular : unit -> bool val set_library : bool -> unit val library : unit -> bool val set_extrcompute : bool -> unit val is_extrcompute : unit -> bool (*s Table for custom inlining *) val to_inline : GlobRef.t -> bool val to_keep : GlobRef.t -> bool (*s Table for implicits arguments *) val implicits_of_global : GlobRef.t -> Int.Set.t (*s Table for user-given custom ML extractions. *) (* UGLY HACK: registration of a function defined in [extraction.ml] *) val type_scheme_nb_args_hook : (Environ.env -> Constr.t -> int) Hook.t val is_custom : GlobRef.t -> bool val is_inline_custom : GlobRef.t -> bool val find_custom : GlobRef.t -> string val find_type_custom : GlobRef.t -> string list * string val is_custom_match : ml_branch array -> bool val find_custom_match : ml_branch array -> string (*s Extraction commands. *) val extraction_language : lang -> unit val extraction_inline : bool -> qualid list -> unit val print_extraction_inline : unit -> Pp.t val reset_extraction_inline : unit -> unit val extract_constant_inline : bool -> qualid -> string list -> string -> unit val extract_inductive : qualid -> string -> string list -> string option -> unit type int_or_id = ArgInt of int | ArgId of Id.t val extraction_implicit : qualid -> int_or_id list -> unit (*s Table of blacklisted filenames *) val extraction_blacklist : Id.t list -> unit val reset_extraction_blacklist : unit -> unit val print_extraction_blacklist : unit -> Pp.t coq-8.11.0/plugins/extraction/json.ml0000644000175000017500000002032613612664533017435 0ustar treinentreinenopen Pp open Util open Names open Table open Miniml open Mlutil open Common let json_str s = qs s let json_int i = int i let json_bool b = if b then str "true" else str "false" let json_global typ ref = if is_custom ref then json_str (find_custom ref) else json_str (Common.pp_global typ ref) let json_id id = json_str (Id.to_string id) let json_dict_one (k, v) = json_str k ++ str (": ") ++ v let json_dict_open l = str ("{") ++ fnl () ++ str (" ") ++ hov 0 (prlist_with_sep pr_comma json_dict_one l) let json_dict l = json_dict_open l ++ fnl () ++ str ("}") let json_list l = str ("[") ++ fnl () ++ str (" ") ++ hov 0 (prlist_with_sep pr_comma (fun x -> x) l) ++ fnl () ++ str ("]") let json_listarr a = str ("[") ++ fnl () ++ str (" ") ++ hov 0 (prvect_with_sep pr_comma (fun x -> x) a) ++ fnl () ++ str ("]") let preamble mod_name comment used_modules usf = (match comment with | None -> mt () | Some s -> str "/* " ++ hov 0 s ++ str " */" ++ fnl ()) ++ json_dict_open [ ("what", json_str "module"); ("name", json_id mod_name); ("need_magic", json_bool (usf.magic)); ("need_dummy", json_bool (usf.mldummy)); ("used_modules", json_list (List.map (fun mf -> json_str (file_of_modfile mf)) used_modules)) ] (*s Pretty-printing of types. *) let rec json_type vl = function | Tmeta _ | Tvar' _ -> assert false | Tvar i -> (try let varid = List.nth vl (pred i) in json_dict [ ("what", json_str "type:var"); ("name", json_id varid) ] with Failure _ -> json_dict [ ("what", json_str "type:varidx"); ("name", json_int i) ]) | Tglob (r, l) -> json_dict [ ("what", json_str "type:glob"); ("name", json_global Type r); ("args", json_list (List.map (json_type vl) l)) ] | Tarr (t1,t2) -> json_dict [ ("what", json_str "type:arrow"); ("left", json_type vl t1); ("right", json_type vl t2) ] | Tdummy _ -> json_dict [("what", json_str "type:dummy")] | Tunknown -> json_dict [("what", json_str "type:unknown")] | Taxiom -> json_dict [("what", json_str "type:axiom")] (*s Pretty-printing of expressions. *) let rec json_expr env = function | MLrel n -> json_dict [ ("what", json_str "expr:rel"); ("name", json_id (get_db_name n env)) ] | MLapp (f, args) -> json_dict [ ("what", json_str "expr:apply"); ("func", json_expr env f); ("args", json_list (List.map (json_expr env) args)) ] | MLlam _ as a -> let fl, a' = collect_lams a in let fl, env' = push_vars (List.map id_of_mlid fl) env in json_dict [ ("what", json_str "expr:lambda"); ("argnames", json_list (List.map json_id (List.rev fl))); ("body", json_expr env' a') ] | MLletin (id, a1, a2) -> let i, env' = push_vars [id_of_mlid id] env in json_dict [ ("what", json_str "expr:let"); ("name", json_id (List.hd i)); ("nameval", json_expr env a1); ("body", json_expr env' a2) ] | MLglob r -> json_dict [ ("what", json_str "expr:global"); ("name", json_global Term r) ] | MLcons (_, r, a) -> json_dict [ ("what", json_str "expr:constructor"); ("name", json_global Cons r); ("args", json_list (List.map (json_expr env) a)) ] | MLtuple l -> json_dict [ ("what", json_str "expr:tuple"); ("items", json_list (List.map (json_expr env) l)) ] | MLcase (typ, t, pv) -> json_dict [ ("what", json_str "expr:case"); ("expr", json_expr env t); ("cases", json_listarr (Array.map (fun x -> json_one_pat env x) pv)) ] | MLfix (i, ids, defs) -> let ids', env' = push_vars (List.rev (Array.to_list ids)) env in let ids' = Array.of_list (List.rev ids') in json_dict [ ("what", json_str "expr:fix"); ("funcs", json_listarr (Array.map (fun (fi, ti) -> json_dict [ ("what", json_str "fix:item"); ("name", json_id fi); ("body", json_function env' ti) ]) (Array.map2 (fun a b -> a,b) ids' defs))); ("for", json_int i); ] | MLexn s -> json_dict [ ("what", json_str "expr:exception"); ("msg", json_str s) ] | MLdummy _ -> json_dict [("what", json_str "expr:dummy")] | MLmagic a -> json_dict [ ("what", json_str "expr:coerce"); ("value", json_expr env a) ] | MLaxiom -> json_dict [("what", json_str "expr:axiom")] | MLuint i -> json_dict [ ("what", json_str "expr:int"); ("int", json_str (Uint63.to_string i)) ] | MLfloat f -> json_dict [ ("what", json_str "expr:float"); ("float", json_str (Float64.to_string f)) ] and json_one_pat env (ids,p,t) = let ids', env' = push_vars (List.rev_map id_of_mlid ids) env in json_dict [ ("what", json_str "case"); ("pat", json_gen_pat (List.rev ids') env' p); ("body", json_expr env' t) ] and json_gen_pat ids env = function | Pcons (r, l) -> json_cons_pat r (List.map (json_gen_pat ids env) l) | Pusual r -> json_cons_pat r (List.map json_id ids) | Ptuple l -> json_dict [ ("what", json_str "pat:tuple"); ("items", json_list (List.map (json_gen_pat ids env) l)) ] | Pwild -> json_dict [("what", json_str "pat:wild")] | Prel n -> json_dict [ ("what", json_str "pat:rel"); ("name", json_id (get_db_name n env)) ] and json_cons_pat r ppl = json_dict [ ("what", json_str "pat:constructor"); ("name", json_global Cons r); ("argnames", json_list ppl) ] and json_function env t = let bl, t' = collect_lams t in let bl, env' = push_vars (List.map id_of_mlid bl) env in json_dict [ ("what", json_str "expr:lambda"); ("argnames", json_list (List.map json_id (List.rev bl))); ("body", json_expr env' t') ] (*s Pretty-printing of inductive types declaration. *) let json_ind ip pl cv = json_dict [ ("what", json_str "decl:ind"); ("name", json_global Type (GlobRef.IndRef ip)); ("argnames", json_list (List.map json_id pl)); ("constructors", json_listarr (Array.mapi (fun idx c -> json_dict [ ("name", json_global Cons (GlobRef.ConstructRef (ip, idx+1))); ("argtypes", json_list (List.map (json_type pl) c)) ]) cv)) ] (*s Pretty-printing of a declaration. *) let pp_decl = function | Dind (kn, defs) -> prvecti_with_sep pr_comma (fun i p -> if p.ip_logical then str "" else json_ind (kn, i) p.ip_vars p.ip_types) defs.ind_packets | Dtype (r, l, t) -> json_dict [ ("what", json_str "decl:type"); ("name", json_global Type r); ("argnames", json_list (List.map json_id l)); ("value", json_type l t) ] | Dfix (rv, defs, typs) -> json_dict [ ("what", json_str "decl:fixgroup"); ("fixlist", json_listarr (Array.mapi (fun i r -> json_dict [ ("what", json_str "fixgroup:item"); ("name", json_global Term rv.(i)); ("type", json_type [] typs.(i)); ("value", json_function (empty_env ()) defs.(i)) ]) rv)) ] | Dterm (r, a, t) -> json_dict [ ("what", json_str "decl:term"); ("name", json_global Term r); ("type", json_type [] t); ("value", json_function (empty_env ()) a) ] let rec pp_structure_elem = function | (l,SEdecl d) -> [ pp_decl d ] | (l,SEmodule m) -> pp_module_expr m.ml_mod_expr | (l,SEmodtype m) -> [] (* for the moment we simply discard module type *) and pp_module_expr = function | MEstruct (mp,sel) -> List.concat (List.map pp_structure_elem sel) | MEfunctor _ -> [] (* for the moment we simply discard unapplied functors *) | MEident _ | MEapply _ -> assert false (* should be expansed in extract_env *) let pp_struct mls = let pp_sel (mp,sel) = push_visible mp []; let p = prlist_with_sep pr_comma identity (List.concat (List.map pp_structure_elem sel)) in pop_visible (); p in str "," ++ fnl () ++ str " " ++ qs "declarations" ++ str ": [" ++ fnl () ++ str " " ++ hov 0 (prlist_with_sep pr_comma pp_sel mls) ++ fnl () ++ str " ]" ++ fnl () ++ str "}" ++ fnl () let json_descr = { keywords = Id.Set.empty; file_suffix = ".json"; file_naming = file_of_modfile; preamble = preamble; pp_struct = pp_struct; sig_suffix = None; sig_preamble = (fun _ _ _ _ -> mt ()); pp_sig = (fun _ -> mt ()); pp_decl = pp_decl; } coq-8.11.0/plugins/extraction/ExtrOcamlZInt.v0000644000175000017500000000652113612664533021025 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int [ "(fun p->1+2*p)" "(fun p->2*p)" "1" ] "(fun f2p1 f2p f1 p -> if p<=1 then f1 () else if p mod 2 = 0 then f2p (p/2) else f2p1 (p/2))". Extract Inductive Z => int [ "0" "" "(~-)" ] "(fun f0 fp fn z -> if z=0 then f0 () else if z>0 then fp z else fn (-z))". Extract Inductive N => int [ "0" "" ] "(fun f0 fp n -> if n=0 then f0 () else fp n)". (** Nota: the "" above is used as an identity function "(fun p->p)" *) (** Efficient (but uncertified) versions for usual functions *) Extract Constant Pos.add => "(+)". Extract Constant Pos.succ => "Pervasives.succ". Extract Constant Pos.pred => "fun n -> Pervasives.max 1 (n-1)". Extract Constant Pos.sub => "fun n m -> Pervasives.max 1 (n-m)". Extract Constant Pos.mul => "( * )". Extract Constant Pos.min => "Pervasives.min". Extract Constant Pos.max => "Pervasives.max". Extract Constant Pos.compare => "fun x y -> if x=y then Eq else if x "fun c x y -> if x=y then c else if x "(+)". Extract Constant N.succ => "Pervasives.succ". Extract Constant N.pred => "fun n -> Pervasives.max 0 (n-1)". Extract Constant N.sub => "fun n m -> Pervasives.max 0 (n-m)". Extract Constant N.mul => "( * )". Extract Constant N.min => "Pervasives.min". Extract Constant N.max => "Pervasives.max". Extract Constant N.div => "fun a b -> if b=0 then 0 else a/b". Extract Constant N.modulo => "fun a b -> if b=0 then a else a mod b". Extract Constant N.compare => "fun x y -> if x=y then Eq else if x "(+)". Extract Constant Z.succ => "Pervasives.succ". Extract Constant Z.pred => "Pervasives.pred". Extract Constant Z.sub => "(-)". Extract Constant Z.mul => "( * )". Extract Constant Z.opp => "(~-)". Extract Constant Z.abs => "Pervasives.abs". Extract Constant Z.min => "Pervasives.min". Extract Constant Z.max => "Pervasives.max". Extract Constant Z.compare => "fun x y -> if x=y then Eq else if x "fun p -> p". Extract Constant Z.abs_N => "Pervasives.abs". (** Z.div and Z.modulo are quite complex to define in terms of (/) and (mod). For the moment we don't even try *) coq-8.11.0/plugins/extraction/ExtrHaskellNatNum.v0000644000175000017500000000346113612664533021673 0ustar treinentreinen(** * Efficient (but uncertified) extraction of usual [nat] functions * into equivalent versions in Haskell's Prelude that are defined * for any [Num] typeclass instances. Useful in combination with * [Extract Inductive nat] that maps [nat] onto a Haskell type that * implements [Num]. *) Require Coq.extraction.Extraction. Require Import Arith. Require Import EqNat. Extract Inlined Constant Nat.add => "(Prelude.+)". Extract Inlined Constant Nat.mul => "(Prelude.*)". Extract Inlined Constant Nat.max => "Prelude.max". Extract Inlined Constant Nat.min => "Prelude.min". Extract Inlined Constant Init.Nat.add => "(Prelude.+)". Extract Inlined Constant Init.Nat.mul => "(Prelude.*)". Extract Inlined Constant Init.Nat.max => "Prelude.max". Extract Inlined Constant Init.Nat.min => "Prelude.min". Extract Inlined Constant Compare_dec.lt_dec => "(Prelude.<)". Extract Inlined Constant Compare_dec.leb => "(Prelude.<=)". Extract Inlined Constant Compare_dec.le_lt_dec => "(Prelude.<=)". Extract Inlined Constant EqNat.beq_nat => "(Prelude.==)". Extract Inlined Constant EqNat.eq_nat_decide => "(Prelude.==)". Extract Inlined Constant Peano_dec.eq_nat_dec => "(Prelude.==)". Extract Constant Nat.pred => "(\n -> Prelude.max 0 (Prelude.pred n))". Extract Constant Nat.sub => "(\n m -> Prelude.max 0 (n Prelude.- m))". Extract Constant Init.Nat.pred => "(\n -> Prelude.max 0 (Prelude.pred n))". Extract Constant Init.Nat.sub => "(\n m -> Prelude.max 0 (n Prelude.- m))". Extract Constant Nat.div => "(\n m -> if m Prelude.== 0 then 0 else Prelude.div n m)". Extract Constant Nat.modulo => "(\n m -> if m Prelude.== 0 then 0 else Prelude.mod n m)". Extract Constant Init.Nat.div => "(\n m -> if m Prelude.== 0 then 0 else Prelude.div n m)". Extract Constant Init.Nat.modulo => "(\n m -> if m Prelude.== 0 then 0 else Prelude.mod n m)". coq-8.11.0/plugins/extraction/ocaml.mli0000644000175000017500000000131513612664533017725 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { id } | [ string(s) ] -> { s } END { let pr_int_or_id _ _ _ = function | ArgInt i -> int i | ArgId id -> Id.print id } ARGUMENT EXTEND int_or_id PRINTED BY { pr_int_or_id } | [ preident(id) ] -> { ArgId (Id.of_string id) } | [ integer(i) ] -> { ArgInt i } END { let pr_language = function | Ocaml -> str "OCaml" | Haskell -> str "Haskell" | Scheme -> str "Scheme" | JSON -> str "JSON" let warn_deprecated_ocaml_spelling = CWarnings.create ~name:"deprecated-ocaml-spelling" ~category:"deprecated" (fun () -> strbrk ("The spelling \"OCaml\" should be used instead of \"Ocaml\".")) } VERNAC ARGUMENT EXTEND language PRINTED BY { pr_language } | [ "Ocaml" ] -> { let _ = warn_deprecated_ocaml_spelling () in Ocaml } | [ "OCaml" ] -> { Ocaml } | [ "Haskell" ] -> { Haskell } | [ "Scheme" ] -> { Scheme } | [ "JSON" ] -> { JSON } END (* Extraction commands *) VERNAC COMMAND EXTEND Extraction CLASSIFIED AS QUERY (* Extraction in the Coq toplevel *) | [ "Extraction" global(x) ] -> { simple_extraction x } | [ "Recursive" "Extraction" ne_global_list(l) ] -> { full_extraction None l } (* Monolithic extraction to a file *) | [ "Extraction" string(f) ne_global_list(l) ] -> { full_extraction (Some f) l } (* Extraction to a temporary file and OCaml compilation *) | [ "Extraction" "TestCompile" ne_global_list(l) ] -> { extract_and_compile l } END VERNAC COMMAND EXTEND SeparateExtraction CLASSIFIED AS QUERY (* Same, with content split in several files *) | [ "Separate" "Extraction" ne_global_list(l) ] -> { separate_extraction l } END (* Modular extraction (one Coq library = one ML module) *) VERNAC COMMAND EXTEND ExtractionLibrary CLASSIFIED AS QUERY | [ "Extraction" "Library" ident(m) ] -> { extraction_library false m } END VERNAC COMMAND EXTEND RecursiveExtractionLibrary CLASSIFIED AS QUERY | [ "Recursive" "Extraction" "Library" ident(m) ] -> { extraction_library true m } END (* Target Language *) VERNAC COMMAND EXTEND ExtractionLanguage CLASSIFIED AS SIDEFF | [ "Extraction" "Language" language(l) ] -> { extraction_language l } END VERNAC COMMAND EXTEND ExtractionInline CLASSIFIED AS SIDEFF (* Custom inlining directives *) | [ "Extraction" "Inline" ne_global_list(l) ] -> { extraction_inline true l } END VERNAC COMMAND EXTEND ExtractionNoInline CLASSIFIED AS SIDEFF | [ "Extraction" "NoInline" ne_global_list(l) ] -> { extraction_inline false l } END VERNAC COMMAND EXTEND PrintExtractionInline CLASSIFIED AS QUERY | [ "Print" "Extraction" "Inline" ] -> {Feedback.msg_notice (print_extraction_inline ()) } END VERNAC COMMAND EXTEND ResetExtractionInline CLASSIFIED AS SIDEFF | [ "Reset" "Extraction" "Inline" ] -> { reset_extraction_inline () } END VERNAC COMMAND EXTEND ExtractionImplicit CLASSIFIED AS SIDEFF (* Custom implicit arguments of some csts/inds/constructors *) | [ "Extraction" "Implicit" global(r) "[" int_or_id_list(l) "]" ] -> { extraction_implicit r l } END VERNAC COMMAND EXTEND ExtractionBlacklist CLASSIFIED AS SIDEFF (* Force Extraction to not use some filenames *) | [ "Extraction" "Blacklist" ne_ident_list(l) ] -> { extraction_blacklist l } END VERNAC COMMAND EXTEND PrintExtractionBlacklist CLASSIFIED AS QUERY | [ "Print" "Extraction" "Blacklist" ] -> { Feedback.msg_notice (print_extraction_blacklist ()) } END VERNAC COMMAND EXTEND ResetExtractionBlacklist CLASSIFIED AS SIDEFF | [ "Reset" "Extraction" "Blacklist" ] -> { reset_extraction_blacklist () } END (* Overriding of a Coq object by an ML one *) VERNAC COMMAND EXTEND ExtractionConstant CLASSIFIED AS SIDEFF | [ "Extract" "Constant" global(x) string_list(idl) "=>" mlname(y) ] -> { extract_constant_inline false x idl y } END VERNAC COMMAND EXTEND ExtractionInlinedConstant CLASSIFIED AS SIDEFF | [ "Extract" "Inlined" "Constant" global(x) "=>" mlname(y) ] -> { extract_constant_inline true x [] y } END VERNAC COMMAND EXTEND ExtractionInductive CLASSIFIED AS SIDEFF | [ "Extract" "Inductive" global(x) "=>" mlname(id) "[" mlname_list(idl) "]" string_opt(o) ] -> { extract_inductive x id idl o } END (* Show the extraction of the current proof *) VERNAC COMMAND EXTEND ShowExtraction CLASSIFIED AS QUERY STATE proof_query | [ "Show" "Extraction" ] -> { show_extraction } END coq-8.11.0/plugins/extraction/ExtrHaskellZNum.v0000644000175000017500000000160713612664533021362 0ustar treinentreinen(** * Efficient (but uncertified) extraction of usual [Z] functions * into equivalent versions in Haskell's Prelude that are defined * for any [Num] typeclass instances. Useful in combination with * [Extract Inductive Z] that maps [Z] onto a Haskell type that * implements [Num]. *) Require Coq.extraction.Extraction. Require Import ZArith. Require Import EqNat. Extract Inlined Constant Z.add => "(Prelude.+)". Extract Inlined Constant Z.sub => "(Prelude.-)". Extract Inlined Constant Z.mul => "(Prelude.*)". Extract Inlined Constant Z.max => "Prelude.max". Extract Inlined Constant Z.min => "Prelude.min". Extract Inlined Constant Z_ge_lt_dec => "(Prelude.>=)". Extract Inlined Constant Z_gt_le_dec => "(Prelude.>)". Extract Constant Z.div => "(\n m -> if m Prelude.== 0 then 0 else Prelude.div n m)". Extract Constant Z.modulo => "(\n m -> if m Prelude.== 0 then 0 else Prelude.mod n m)". coq-8.11.0/plugins/extraction/mlutil.ml0000644000175000017500000014543613612664533020004 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* anonymous_name | Name.Name id when Id.equal id dummy_name -> anonymous_name | Name.Name id -> id let id_of_mlid = function | Dummy -> dummy_name | Id id -> id | Tmp id -> id let tmp_id = function | Id id -> Tmp id | a -> a let is_tmp = function Tmp _ -> true | _ -> false (*S Operations upon ML types (with meta). *) let meta_count = ref 0 let reset_meta_count () = meta_count := 0 let new_meta _ = incr meta_count; Tmeta {id = !meta_count; contents = None} let rec eq_ml_type t1 t2 = match t1, t2 with | Tarr (tl1, tr1), Tarr (tl2, tr2) -> eq_ml_type tl1 tl2 && eq_ml_type tr1 tr2 | Tglob (gr1, t1), Tglob (gr2, t2) -> GlobRef.equal gr1 gr2 && List.equal eq_ml_type t1 t2 | Tvar i1, Tvar i2 -> Int.equal i1 i2 | Tvar' i1, Tvar' i2 -> Int.equal i1 i2 | Tmeta m1, Tmeta m2 -> eq_ml_meta m1 m2 | Tdummy k1, Tdummy k2 -> k1 == k2 | Tunknown, Tunknown -> true | Taxiom, Taxiom -> true | (Tarr _ | Tglob _ | Tvar _ | Tvar' _ | Tmeta _ | Tdummy _ | Tunknown | Taxiom), _ -> false and eq_ml_meta m1 m2 = Int.equal m1.id m2.id && Option.equal eq_ml_type m1.contents m2.contents (* Simultaneous substitution of [[Tvar 1; ... ; Tvar n]] by [l] in a ML type. *) let type_subst_list l t = let rec subst t = match t with | Tvar j -> List.nth l (j-1) | Tmeta {contents=None} -> t | Tmeta {contents=Some u} -> subst u | Tarr (a,b) -> Tarr (subst a, subst b) | Tglob (r, l) -> Tglob (r, List.map subst l) | a -> a in subst t (* Simultaneous substitution of [[|Tvar 1; ... ; Tvar n|]] by [v] in a ML type. *) let type_subst_vect v t = let rec subst t = match t with | Tvar j -> v.(j-1) | Tmeta {contents=None} -> t | Tmeta {contents=Some u} -> subst u | Tarr (a,b) -> Tarr (subst a, subst b) | Tglob (r, l) -> Tglob (r, List.map subst l) | a -> a in subst t (*s From a type schema to a type. All [Tvar] become fresh [Tmeta]. *) let instantiation (nb,t) = type_subst_vect (Array.init nb new_meta) t (*s Occur-check of a free meta in a type *) let rec type_occurs alpha t = match t with | Tmeta {id=beta; contents=None} -> Int.equal alpha beta | Tmeta {contents=Some u} -> type_occurs alpha u | Tarr (t1, t2) -> type_occurs alpha t1 || type_occurs alpha t2 | Tglob (r,l) -> List.exists (type_occurs alpha) l | (Tdummy _ | Tvar _ | Tvar' _ | Taxiom | Tunknown) -> false (*s Most General Unificator *) let rec mgu = function | Tmeta m, Tmeta m' when Int.equal m.id m'.id -> () | Tmeta m, t | t, Tmeta m -> (match m.contents with | Some u -> mgu (u, t) | None when type_occurs m.id t -> raise Impossible | None -> m.contents <- Some t) | Tarr(a, b), Tarr(a', b') -> mgu (a, a'); mgu (b, b') | Tglob (r,l), Tglob (r',l') when GlobRef.equal r r' -> List.iter mgu (List.combine l l') | Tdummy _, Tdummy _ -> () | Tvar i, Tvar j when Int.equal i j -> () | Tvar' i, Tvar' j when Int.equal i j -> () | Tunknown, Tunknown -> () | Taxiom, Taxiom -> () | _ -> raise Impossible let skip_typing () = lang () == Scheme || is_extrcompute () let needs_magic p = if skip_typing () then false else try mgu p; false with Impossible -> true let put_magic_if b a = if b then MLmagic a else a let put_magic p a = if needs_magic p then MLmagic a else a let generalizable a = lang () != Ocaml || match a with | MLapp _ -> false | _ -> true (* TODO, this is just an approximation for the moment *) (*S ML type env. *) module Mlenv = struct let meta_cmp m m' = compare m.id m'.id module Metaset = Set.Make(struct type t = ml_meta let compare = meta_cmp end) (* Main MLenv type. [env] is the real environment, whereas [free] (tries to) record the free meta variables occurring in [env]. *) type t = { env : ml_schema list; mutable free : Metaset.t} (* Empty environment. *) let empty = { env = []; free = Metaset.empty } (* [get] returns a instantiated copy of the n-th most recently added type in the environment. *) let get mle n = assert (List.length mle.env >= n); instantiation (List.nth mle.env (n-1)) (* [find_free] finds the free meta in a type. *) let rec find_free set = function | Tmeta m when Option.is_empty m.contents -> Metaset.add m set | Tmeta {contents = Some t} -> find_free set t | Tarr (a,b) -> find_free (find_free set a) b | Tglob (_,l) -> List.fold_left find_free set l | _ -> set (* The [free] set of an environment can be outdate after some unifications. [clean_free] takes care of that. *) let clean_free mle = let rem = ref Metaset.empty and add = ref Metaset.empty in let clean m = match m.contents with | None -> () | Some u -> rem := Metaset.add m !rem; add := find_free !add u in Metaset.iter clean mle.free; mle.free <- Metaset.union (Metaset.diff mle.free !rem) !add (* From a type to a type schema. If a [Tmeta] is still uninstantiated and does appears in the [mle], then it becomes a [Tvar]. *) let generalization mle t = let c = ref 0 in let map = ref (Int.Map.empty : int Int.Map.t) in let add_new i = incr c; map := Int.Map.add i !c !map; !c in let rec meta2var t = match t with | Tmeta {contents=Some u} -> meta2var u | Tmeta ({id=i} as m) -> (try Tvar (Int.Map.find i !map) with Not_found -> if Metaset.mem m mle.free then t else Tvar (add_new i)) | Tarr (t1,t2) -> Tarr (meta2var t1, meta2var t2) | Tglob (r,l) -> Tglob (r, List.map meta2var l) | t -> t in !c, meta2var t (* Adding a type in an environment, after generalizing. *) let push_gen mle t = clean_free mle; { env = generalization mle t :: mle.env; free = mle.free } (* Adding a type with no [Tvar], hence no generalization needed. *) let push_type {env=e;free=f} t = { env = (0,t) :: e; free = find_free f t} (* Adding a type with no [Tvar] nor [Tmeta]. *) let push_std_type {env=e;free=f} t = { env = (0,t) :: e; free = f} end (*S Operations upon ML types (without meta). *) (*s Does a section path occur in a ML type ? *) let rec type_mem_kn kn = function | Tmeta {contents = Some t} -> type_mem_kn kn t | Tglob (r,l) -> occur_kn_in_ref kn r || List.exists (type_mem_kn kn) l | Tarr (a,b) -> (type_mem_kn kn a) || (type_mem_kn kn b) | _ -> false (*s Greatest variable occurring in [t]. *) let type_maxvar t = let rec parse n = function | Tmeta {contents = Some t} -> parse n t | Tvar i -> max i n | Tarr (a,b) -> parse (parse n a) b | Tglob (_,l) -> List.fold_left parse n l | _ -> n in parse 0 t (*s From [a -> b -> c] to [[a;b],c]. *) let rec type_decomp = function | Tmeta {contents = Some t} -> type_decomp t | Tarr (a,b) -> let l,h = type_decomp b in a::l, h | a -> [],a (*s The converse: From [[a;b],c] to [a -> b -> c]. *) let rec type_recomp (l,t) = match l with | [] -> t | a::l -> Tarr (a, type_recomp (l,t)) (*s Translating [Tvar] to [Tvar'] to avoid clash. *) let rec var2var' = function | Tmeta {contents = Some t} -> var2var' t | Tvar i -> Tvar' i | Tarr (a,b) -> Tarr (var2var' a, var2var' b) | Tglob (r,l) -> Tglob (r, List.map var2var' l) | a -> a type abbrev_map = GlobRef.t -> ml_type option (*s Delta-reduction of type constants everywhere in a ML type [t]. [env] is a function of type [ml_type_env]. *) let type_expand env t = let rec expand = function | Tmeta {contents = Some t} -> expand t | Tglob (r,l) -> (match env r with | Some mlt -> expand (type_subst_list l mlt) | None -> Tglob (r, List.map expand l)) | Tarr (a,b) -> Tarr (expand a, expand b) | a -> a in if Table.type_expand () then expand t else t let type_simpl = type_expand (fun _ -> None) (*s Generating a signature from a ML type. *) let type_to_sign env t = match type_expand env t with | Tdummy d when not (conservative_types ()) -> Kill d | _ -> Keep let type_to_signature env t = let rec f = function | Tmeta {contents = Some t} -> f t | Tarr (Tdummy d, b) when not (conservative_types ()) -> Kill d :: f b | Tarr (_, b) -> Keep :: f b | _ -> [] in f (type_expand env t) let isKill = function Kill _ -> true | _ -> false let isTdummy = function Tdummy _ -> true | _ -> false let isMLdummy = function MLdummy _ -> true | _ -> false let sign_of_id = function | Dummy -> Kill Kprop | (Id _ | Tmp _) -> Keep (* Classification of signatures *) type sign_kind = | EmptySig | NonLogicalSig (* at least a [Keep] *) | SafeLogicalSig (* only [Kill Ktype] *) | UnsafeLogicalSig (* No [Keep], not all [Kill Ktype] *) let rec sign_kind = function | [] -> EmptySig | Keep :: _ -> NonLogicalSig | Kill k :: s -> match k, sign_kind s with | _, NonLogicalSig -> NonLogicalSig | Ktype, (SafeLogicalSig | EmptySig) -> SafeLogicalSig | _, _ -> UnsafeLogicalSig (* Removing the final [Keep] in a signature *) let rec sign_no_final_keeps = function | [] -> [] | k :: s -> match k, sign_no_final_keeps s with | Keep, [] -> [] | k, l -> k::l (*s Removing [Tdummy] from the top level of a ML type. *) let type_expunge_from_sign env s t = let rec expunge s t = match s, t with | [], _ -> t | Keep :: s, Tarr(a,b) -> Tarr (a, expunge s b) | Kill _ :: s, Tarr(a,b) -> expunge s b | _, Tmeta {contents = Some t} -> expunge s t | _, Tglob (r,l) -> (match env r with | Some mlt -> expunge s (type_subst_list l mlt) | None -> assert false) | _ -> assert false in let t = expunge (sign_no_final_keeps s) t in if lang () != Haskell && sign_kind s == UnsafeLogicalSig then Tarr (Tdummy Kprop, t) else t let type_expunge env t = type_expunge_from_sign env (type_to_signature env t) t (*S Generic functions over ML ast terms. *) let mlapp f a = if List.is_empty a then f else MLapp (f,a) (** Equality *) let eq_ml_ident i1 i2 = match i1, i2 with | Dummy, Dummy -> true | Id id1, Id id2 -> Id.equal id1 id2 | Tmp id1, Tmp id2 -> Id.equal id1 id2 | Dummy, (Id _ | Tmp _) | Id _, (Dummy | Tmp _) | Tmp _, (Dummy | Id _) -> false let rec eq_ml_ast t1 t2 = match t1, t2 with | MLrel i1, MLrel i2 -> Int.equal i1 i2 | MLapp (f1, t1), MLapp (f2, t2) -> eq_ml_ast f1 f2 && List.equal eq_ml_ast t1 t2 | MLlam (na1, t1), MLlam (na2, t2) -> eq_ml_ident na1 na2 && eq_ml_ast t1 t2 | MLletin (na1, c1, t1), MLletin (na2, c2, t2) -> eq_ml_ident na1 na2 && eq_ml_ast c1 c2 && eq_ml_ast t1 t2 | MLglob gr1, MLglob gr2 -> GlobRef.equal gr1 gr2 | MLcons (t1, gr1, c1), MLcons (t2, gr2, c2) -> eq_ml_type t1 t2 && GlobRef.equal gr1 gr2 && List.equal eq_ml_ast c1 c2 | MLtuple t1, MLtuple t2 -> List.equal eq_ml_ast t1 t2 | MLcase (t1, c1, p1), MLcase (t2, c2, p2) -> eq_ml_type t1 t2 && eq_ml_ast c1 c2 && Array.equal eq_ml_branch p1 p2 | MLfix (i1, id1, t1), MLfix (i2, id2, t2) -> Int.equal i1 i2 && Array.equal Id.equal id1 id2 && Array.equal eq_ml_ast t1 t2 | MLexn e1, MLexn e2 -> String.equal e1 e2 | MLdummy k1, MLdummy k2 -> k1 == k2 | MLaxiom, MLaxiom -> true | MLmagic t1, MLmagic t2 -> eq_ml_ast t1 t2 | MLuint i1, MLuint i2 -> Uint63.equal i1 i2 | MLfloat f1, MLfloat f2 -> Float64.equal f1 f2 | _, _ -> false and eq_ml_pattern p1 p2 = match p1, p2 with | Pcons (gr1, p1), Pcons (gr2, p2) -> GlobRef.equal gr1 gr2 && List.equal eq_ml_pattern p1 p2 | Ptuple p1, Ptuple p2 -> List.equal eq_ml_pattern p1 p2 | Prel i1, Prel i2 -> Int.equal i1 i2 | Pwild, Pwild -> true | Pusual gr1, Pusual gr2 -> GlobRef.equal gr1 gr2 | _ -> false and eq_ml_branch (id1, p1, t1) (id2, p2, t2) = List.equal eq_ml_ident id1 id2 && eq_ml_pattern p1 p2 && eq_ml_ast t1 t2 (*s [ast_iter_rel f t] applies [f] on every [MLrel] in t. It takes care of the number of bingings crossed before reaching the [MLrel]. *) let ast_iter_rel f = let rec iter n = function | MLrel i -> f (i-n) | MLlam (_,a) -> iter (n+1) a | MLletin (_,a,b) -> iter n a; iter (n+1) b | MLcase (_,a,v) -> iter n a; Array.iter (fun (l,_,t) -> iter (n + (List.length l)) t) v | MLfix (_,ids,v) -> let k = Array.length ids in Array.iter (iter (n+k)) v | MLapp (a,l) -> iter n a; List.iter (iter n) l | MLcons (_,_,l) | MLtuple l -> List.iter (iter n) l | MLmagic a -> iter n a | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> () in iter 0 (*s Map over asts. *) let ast_map_branch f (c,ids,a) = (c,ids,f a) (* Warning: in [ast_map] we assume that [f] does not change the type of [MLcons] and of [MLcase] heads *) let ast_map f = function | MLlam (i,a) -> MLlam (i, f a) | MLletin (i,a,b) -> MLletin (i, f a, f b) | MLcase (typ,a,v) -> MLcase (typ,f a, Array.map (ast_map_branch f) v) | MLfix (i,ids,v) -> MLfix (i, ids, Array.map f v) | MLapp (a,l) -> MLapp (f a, List.map f l) | MLcons (typ,c,l) -> MLcons (typ,c, List.map f l) | MLtuple l -> MLtuple (List.map f l) | MLmagic a -> MLmagic (f a) | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ as a -> a (*s Map over asts, with binding depth as parameter. *) let ast_map_lift_branch f n (ids,p,a) = (ids,p, f (n+(List.length ids)) a) (* Same warning as for [ast_map]... *) let ast_map_lift f n = function | MLlam (i,a) -> MLlam (i, f (n+1) a) | MLletin (i,a,b) -> MLletin (i, f n a, f (n+1) b) | MLcase (typ,a,v) -> MLcase (typ,f n a,Array.map (ast_map_lift_branch f n) v) | MLfix (i,ids,v) -> let k = Array.length ids in MLfix (i,ids,Array.map (f (k+n)) v) | MLapp (a,l) -> MLapp (f n a, List.map (f n) l) | MLcons (typ,c,l) -> MLcons (typ,c, List.map (f n) l) | MLtuple l -> MLtuple (List.map (f n) l) | MLmagic a -> MLmagic (f n a) | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ as a -> a (*s Iter over asts. *) let ast_iter_branch f (c,ids,a) = f a let ast_iter f = function | MLlam (i,a) -> f a | MLletin (i,a,b) -> f a; f b | MLcase (_,a,v) -> f a; Array.iter (ast_iter_branch f) v | MLfix (i,ids,v) -> Array.iter f v | MLapp (a,l) -> f a; List.iter f l | MLcons (_,_,l) | MLtuple l -> List.iter f l | MLmagic a -> f a | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> () (*S Operations concerning De Bruijn indices. *) (*s [ast_occurs k t] returns [true] if [(Rel k)] occurs in [t]. *) let ast_occurs k t = try ast_iter_rel (fun i -> if Int.equal i k then raise Found) t; false with Found -> true (*s [occurs_itvl k k' t] returns [true] if there is a [(Rel i)] in [t] with [k<=i<=k'] *) let ast_occurs_itvl k k' t = try ast_iter_rel (fun i -> if (k <= i) && (i <= k') then raise Found) t; false with Found -> true (* Number of occurrences of [Rel 1] in [t], with special treatment of match: occurrences in different branches aren't added, but we rather use max. *) let nb_occur_match = let rec nb k = function | MLrel i -> if Int.equal i k then 1 else 0 | MLcase(_,a,v) -> (nb k a) + Array.fold_left (fun r (ids,_,a) -> max r (nb (k+(List.length ids)) a)) 0 v | MLletin (_,a,b) -> (nb k a) + (nb (k+1) b) | MLfix (_,ids,v) -> let k = k+(Array.length ids) in Array.fold_left (fun r a -> r+(nb k a)) 0 v | MLlam (_,a) -> nb (k+1) a | MLapp (a,l) -> List.fold_left (fun r a -> r+(nb k a)) (nb k a) l | MLcons (_,_,l) | MLtuple l -> List.fold_left (fun r a -> r+(nb k a)) 0 l | MLmagic a -> nb k a | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> 0 in nb 1 (* Replace unused variables by _ *) let dump_unused_vars a = let rec ren env a = match a with | MLrel i -> let () = (List.nth env (i-1)) := true in a | MLlam (id,b) -> let occ_id = ref false in let b' = ren (occ_id::env) b in if !occ_id then if b' == b then a else MLlam(id,b') else MLlam(Dummy,b') | MLletin (id,b,c) -> let occ_id = ref false in let b' = ren env b in let c' = ren (occ_id::env) c in if !occ_id then if b' == b && c' == c then a else MLletin(id,b',c') else (* 'let' without occurrence: shouldn't happen after simpl *) MLletin(Dummy,b',c') | MLcase (t,e,br) -> let e' = ren env e in let br' = Array.Smart.map (ren_branch env) br in if e' == e && br' == br then a else MLcase (t,e',br') | MLfix (i,ids,v) -> let env' = List.init (Array.length ids) (fun _ -> ref false) @ env in let v' = Array.Smart.map (ren env') v in if v' == v then a else MLfix (i,ids,v') | MLapp (b,l) -> let b' = ren env b and l' = List.Smart.map (ren env) l in if b' == b && l' == l then a else MLapp (b',l') | MLcons(t,r,l) -> let l' = List.Smart.map (ren env) l in if l' == l then a else MLcons (t,r,l') | MLtuple l -> let l' = List.Smart.map (ren env) l in if l' == l then a else MLtuple l' | MLmagic b -> let b' = ren env b in if b' == b then a else MLmagic b' | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> a and ren_branch env ((ids,p,b) as tr) = let occs = List.map (fun _ -> ref false) ids in let b' = ren (List.rev_append occs env) b in let ids' = List.map2 (fun id occ -> if !occ then id else Dummy) ids occs in if b' == b && List.equal eq_ml_ident ids ids' then tr else (ids',p,b') in ren [] a (*s Lifting on terms. [ast_lift k t] lifts the binding depth of [t] across [k] bindings. *) let ast_lift k t = let rec liftrec n = function | MLrel i as a -> if i-n < 1 then a else MLrel (i+k) | a -> ast_map_lift liftrec n a in if Int.equal k 0 then t else liftrec 0 t let ast_pop t = ast_lift (-1) t (*s [permut_rels k k' c] translates [Rel 1 ... Rel k] to [Rel (k'+1) ... Rel (k'+k)] and [Rel (k+1) ... Rel (k+k')] to [Rel 1 ... Rel k'] *) let permut_rels k k' = let rec permut n = function | MLrel i as a -> let i' = i-n in if i'<1 || i'>k+k' then a else if i'<=k then MLrel (i+k') else MLrel (i-k) | a -> ast_map_lift permut n a in permut 0 (*s Substitution. [ml_subst e t] substitutes [e] for [Rel 1] in [t]. Lifting (of one binder) is done at the same time. *) let ast_subst e = let rec subst n = function | MLrel i as a -> let i' = i-n in if Int.equal i' 1 then ast_lift n e else if i'<1 then a else MLrel (i-1) | a -> ast_map_lift subst n a in subst 0 (*s Generalized substitution. [gen_subst v d t] applies to [t] the substitution coded in the [v] array: [(Rel i)] becomes [v.(i-1)]. [d] is the correction applies to [Rel] greater than [Array.length v]. *) let gen_subst v d t = let rec subst n = function | MLrel i as a -> let i'= i-n in if i' < 1 then a else if i' <= Array.length v then match v.(i'-1) with | None -> assert false | Some u -> ast_lift n u else MLrel (i+d) | a -> ast_map_lift subst n a in subst 0 t (*S Operations concerning match patterns *) let is_basic_pattern = function | Prel _ | Pwild -> true | Pusual _ | Pcons _ | Ptuple _ -> false let has_deep_pattern br = let deep = function | Pcons (_,l) | Ptuple l -> not (List.for_all is_basic_pattern l) | Pusual _ | Prel _ | Pwild -> false in Array.exists (function (_,pat,_) -> deep pat) br let is_regular_match br = if Array.is_empty br then false (* empty match becomes MLexn *) else try let get_r (ids,pat,c) = match pat with | Pusual r -> r | Pcons (r,l) -> let is_rel i = function Prel j -> Int.equal i j | _ -> false in if not (List.for_all_i is_rel 1 (List.rev l)) then raise Impossible; r | _ -> raise Impossible in let ind = match get_r br.(0) with | GlobRef.ConstructRef (ind,_) -> ind | _ -> raise Impossible in let is_ref i tr = match get_r tr with | GlobRef.ConstructRef (ind', j) -> eq_ind ind ind' && Int.equal j (i + 1) | _ -> false in Array.for_all_i is_ref 0 br with Impossible -> false (*S Operations concerning lambdas. *) (*s [collect_lams MLlam(id1,...MLlam(idn,t)...)] returns [[idn;...;id1]] and the term [t]. *) let collect_lams = let rec collect acc = function | MLlam(id,t) -> collect (id::acc) t | x -> acc,x in collect [] (*s [collect_n_lams] does the same for a precise number of [MLlam]. *) let collect_n_lams = let rec collect acc n t = if Int.equal n 0 then acc,t else match t with | MLlam(id,t) -> collect (id::acc) (n-1) t | _ -> assert false in collect [] (*s [remove_n_lams] just removes some [MLlam]. *) let rec remove_n_lams n t = if Int.equal n 0 then t else match t with | MLlam(_,t) -> remove_n_lams (n-1) t | _ -> assert false (*s [nb_lams] gives the number of head [MLlam]. *) let rec nb_lams = function | MLlam(_,t) -> succ (nb_lams t) | _ -> 0 (*s [named_lams] does the converse of [collect_lams]. *) let rec named_lams ids a = match ids with | [] -> a | id :: ids -> named_lams ids (MLlam (id,a)) (*s The same for a specific identifier (resp. anonymous, dummy) *) let rec many_lams id a = function | 0 -> a | n -> many_lams id (MLlam (id,a)) (pred n) let anonym_tmp_lams a n = many_lams (Tmp anonymous_name) a n let dummy_lams a n = many_lams Dummy a n (*s mixed according to a signature. *) let rec anonym_or_dummy_lams a = function | [] -> a | Keep :: s -> MLlam(anonymous, anonym_or_dummy_lams a s) | Kill _ :: s -> MLlam(Dummy, anonym_or_dummy_lams a s) (*S Operations concerning eta. *) (*s The following function creates [MLrel n;...;MLrel 1] *) let rec eta_args n = if Int.equal n 0 then [] else (MLrel n)::(eta_args (pred n)) (*s Same, but filtered by a signature. *) let rec eta_args_sign n = function | [] -> [] | Keep :: s -> (MLrel n) :: (eta_args_sign (n-1) s) | Kill _ :: s -> eta_args_sign (n-1) s (*s This one tests [MLrel (n+k); ... ;MLrel (1+k)] *) let rec test_eta_args_lift k n = function | [] -> Int.equal n 0 | MLrel m :: q -> Int.equal (k+n) m && (test_eta_args_lift k (pred n) q) | _ -> false (*s Computes an eta-reduction. *) let eta_red e = let ids,t = collect_lams e in let n = List.length ids in if Int.equal n 0 then e else match t with | MLapp (f,a) -> let m = List.length a in let ids,body,args = if Int.equal m n then [], f, a else if m < n then List.skipn m ids, f, a else (* m > n *) let a1,a2 = List.chop (m-n) a in [], MLapp (f,a1), a2 in let p = List.length args in if test_eta_args_lift 0 p args && not (ast_occurs_itvl 1 p body) then named_lams ids (ast_lift (-p) body) else e | _ -> e (* Performs an eta-reduction when the core is atomic and value, or otherwise returns None *) let atomic_eta_red e = let ids,t = collect_lams e in let n = List.length ids in match t with | MLapp (f,a) when test_eta_args_lift 0 n a -> (match f with | MLrel k when k>n -> Some (MLrel (k-n)) | MLglob _ | MLdummy _ -> Some f | _ -> None) | _ -> None (*s Computes all head linear beta-reductions possible in [(t a)]. Non-linear head beta-redex become let-in. *) let rec linear_beta_red a t = match a,t with | [], _ -> t | a0::a, MLlam (id,t) -> (match nb_occur_match t with | 0 -> linear_beta_red a (ast_pop t) | 1 -> linear_beta_red a (ast_subst a0 t) | _ -> let a = List.map (ast_lift 1) a in MLletin (id, a0, linear_beta_red a t)) | _ -> MLapp (t, a) let rec tmp_head_lams = function | MLlam (id, t) -> MLlam (tmp_id id, tmp_head_lams t) | e -> e (*s Applies a substitution [s] of constants by their body, plus linear beta reductions at modified positions. Moreover, we mark some lambdas as suitable for later linear reduction (this helps the inlining of recursors). *) let rec ast_glob_subst s t = match t with | MLapp ((MLglob ((GlobRef.ConstRef kn) as refe)) as f, a) -> let a = List.map (fun e -> tmp_head_lams (ast_glob_subst s e)) a in (try linear_beta_red a (Refmap'.find refe s) with Not_found -> MLapp (f, a)) | MLglob ((GlobRef.ConstRef kn) as refe) -> (try Refmap'.find refe s with Not_found -> t) | _ -> ast_map (ast_glob_subst s) t (*S Auxiliary functions used in simplification of ML cases. *) (* Factorisation of some match branches into a common "x -> f x" branch may break types sometimes. Example: [type 'x a = A]. Then [let id = function A -> A] has type ['x a -> 'y a], which is incompatible with the type of [let id x = x]. We now check that the type arguments of the inductive are preserved by our transformation. TODO: this verification should be done someday modulo expansion of type definitions. *) (*s [branch_as_function b typ (l,p,c)] tries to see branch [c] as a function [f] applied to [MLcons(r,l)]. For that it transforms any [MLcons(r,l)] in [MLrel 1] and raises [Impossible] if any variable in [l] occurs outside such a [MLcons] *) let branch_as_fun typ (l,p,c) = let nargs = List.length l in let cons = match p with | Pusual r -> MLcons (typ, r, eta_args nargs) | Pcons (r,pl) -> let pat2rel = function Prel i -> MLrel i | _ -> raise Impossible in MLcons (typ, r, List.map pat2rel pl) | _ -> raise Impossible in let rec genrec n = function | MLrel i as c -> let i' = i-n in if i'<1 then c else if i'>nargs then MLrel (i-nargs+1) else raise Impossible | MLcons _ as cons' when eq_ml_ast cons' (ast_lift n cons) -> MLrel (n+1) | a -> ast_map_lift genrec n a in genrec 0 c (*s [branch_as_cst (l,p,c)] tries to see branch [c] as a constant independent from the pattern [MLcons(r,l)]. For that is raises [Impossible] if any variable in [l] occurs in [c], and otherwise returns [c] lifted to appear like a function with one arg (for uniformity with [branch_as_fun]). NB: [MLcons(r,l)] might occur nonetheless in [c], but only when [l] is empty, i.e. when [r] is a constant constructor *) let branch_as_cst (l,_,c) = let n = List.length l in if ast_occurs_itvl 1 n c then raise Impossible; ast_lift (1-n) c (* A branch [MLcons(r,l)->c] can be seen at the same time as a function branch and a constant branch, either because: - [MLcons(r,l)] doesn't occur in [c]. For example : "A -> B" - this constructor is constant (i.e. [l] is empty). For example "A -> A" When searching for the best factorisation below, we'll try both. *) (* The following structure allows recording which element occurred at what position, and then finally return the most frequent element and its positions. *) let census_add, census_max, census_clean = let h = ref [] in let clearf () = h := [] in let rec add k v = function | [] -> raise Not_found | (k', s) as p :: l -> if eq_ml_ast k k' then (k', Int.Set.add v s) :: l else p :: add k v l in let addf k i = try h := add k i !h with Not_found -> h := (k, Int.Set.singleton i) :: !h in let maxf () = let len = ref 0 and lst = ref Int.Set.empty and elm = ref MLaxiom in List.iter (fun (e, s) -> let n = Int.Set.cardinal s in if n > !len then begin len := n; lst := s; elm := e end) !h; (!elm,!lst) in (addf,maxf,clearf) (* [factor_branches] return the longest possible list of branches that have the same factorization, either as a function or as a constant. *) let is_opt_pat (_,p,_) = match p with | Prel _ | Pwild -> true | _ -> false let factor_branches o typ br = if Array.exists is_opt_pat br then None (* already optimized *) else begin census_clean (); for i = 0 to Array.length br - 1 do if o.opt_case_idr then (try census_add (branch_as_fun typ br.(i)) i with Impossible -> ()); if o.opt_case_cst then (try census_add (branch_as_cst br.(i)) i with Impossible -> ()); done; let br_factor, br_set = census_max () in census_clean (); let n = Int.Set.cardinal br_set in if Int.equal n 0 then None else if Array.length br >= 2 && n < 2 then None else Some (br_factor, br_set) end (*s If all branches are functions, try to permute the case and the functions. *) let rec merge_ids ids ids' = match ids,ids' with | [],l -> l | l,[] -> l | i::ids, i'::ids' -> (if i == Dummy then i' else i) :: (merge_ids ids ids') let is_exn = function MLexn _ -> true | _ -> false let permut_case_fun br acc = let nb = ref max_int in Array.iter (fun (_,_,t) -> let ids, c = collect_lams t in let n = List.length ids in if (n < !nb) && (not (is_exn c)) then nb := n) br; if Int.equal !nb max_int || Int.equal !nb 0 then ([],br) else begin let br = Array.copy br in let ids = ref [] in for i = 0 to Array.length br - 1 do let (l,p,t) = br.(i) in let local_nb = nb_lams t in if local_nb < !nb then (* t = MLexn ... *) br.(i) <- (l,p,remove_n_lams local_nb t) else begin let local_ids,t = collect_n_lams !nb t in ids := merge_ids !ids local_ids; br.(i) <- (l,p,permut_rels !nb (List.length l) t) end done; (!ids,br) end (*S Generalized iota-reduction. *) (* Definition of a generalized iota-redex: it's a [MLcase(e,br)] where the head [e] is a [MLcons] or made of [MLcase]'s with [MLcons] as leaf branches. A generalized iota-redex is transformed into beta-redexes. *) (* In [iota_red], we try to simplify a [MLcase(_,MLcons(typ,r,a),br)]. Argument [i] is the branch we consider, we should lift what comes from [br] by [lift] *) let rec iota_red i lift br ((typ,r,a) as cons) = if i >= Array.length br then raise Impossible; let (ids,p,c) = br.(i) in match p with | Pusual r' | Pcons (r',_) when not (GlobRef.equal r' r) -> iota_red (i+1) lift br cons | Pusual r' -> let c = named_lams (List.rev ids) c in let c = ast_lift lift c in MLapp (c,a) | Prel 1 when Int.equal (List.length ids) 1 -> let c = MLlam (List.hd ids, c) in let c = ast_lift lift c in MLapp(c,[MLcons(typ,r,a)]) | Pwild when List.is_empty ids -> ast_lift lift c | _ -> raise Impossible (* TODO: handle some more cases *) (* [iota_gen] is an extension of [iota_red] where we allow to traverse matches in the head of the first match *) let iota_gen br hd = let rec iota k = function | MLcons (typ,r,a) -> iota_red 0 k br (typ,r,a) | MLcase(typ,e,br') -> let new_br = Array.map (fun (i,p,c)->(i,p,iota (k+(List.length i)) c)) br' in MLcase(typ,e,new_br) | _ -> raise Impossible in iota 0 hd let is_atomic = function | MLrel _ | MLglob _ | MLexn _ | MLdummy _ -> true | _ -> false let is_imm_apply = function MLapp (MLrel 1, _) -> true | _ -> false (** Program creates a let-in named "program_branch_NN" for each branch of match. Unfolding them leads to more natural code (and more dummy removal) *) let is_program_branch = function | Tmp _ | Dummy -> false | Id id -> let s = Id.to_string id in try Scanf.sscanf s "program_branch_%d%!" (fun _ -> true) with Scanf.Scan_failure _ | End_of_file -> false let expand_linear_let o id e = o.opt_lin_let || is_tmp id || is_program_branch id || is_imm_apply e (*S The main simplification function. *) (* Some beta-iota reductions + simplifications. *) let rec unmagic = function MLmagic e -> unmagic e | e -> e let is_magic = function MLmagic _ -> true | _ -> false let magic_hd a = match a with | MLmagic _ :: _ -> a | e :: a -> MLmagic e :: a | [] -> assert false let rec simpl o = function | MLapp (f, []) -> simpl o f | MLapp (MLapp(f,a),a') -> simpl o (MLapp(f,a@a')) | MLapp (f, a) -> (* When the head of the application is magic, no need for magic on args *) let a = if is_magic f then List.map unmagic a else a in simpl_app o (List.map (simpl o) a) (simpl o f) | MLcase (typ,e,br) -> let br = Array.map (fun (l,p,t) -> (l,p,simpl o t)) br in simpl_case o typ br (simpl o e) | MLletin(Dummy,_,e) -> simpl o (ast_pop e) | MLletin(id,c,e) -> let e = simpl o e in if (is_atomic c) || (is_atomic e) || (let n = nb_occur_match e in (Int.equal n 0 || (Int.equal n 1 && expand_linear_let o id e))) then simpl o (ast_subst c e) else MLletin(id, simpl o c, e) | MLfix(i,ids,c) -> let n = Array.length ids in if ast_occurs_itvl 1 n c.(i) then MLfix (i, ids, Array.map (simpl o) c) else simpl o (ast_lift (-n) c.(i)) (* Dummy fixpoint *) | MLmagic(MLmagic _ as e) -> simpl o e | MLmagic(MLapp (f,l)) -> simpl o (MLapp (MLmagic f, l)) | MLmagic(MLletin(id,c,e)) -> simpl o (MLletin(id,c,MLmagic e)) | MLmagic(MLcase(typ,e,br)) -> let br' = Array.map (fun (ids,p,c) -> (ids,p,MLmagic c)) br in simpl o (MLcase(typ,e,br')) | MLmagic(MLdummy _ as e) when lang () == Haskell -> e | MLmagic(MLexn _ as e) -> e | MLlam _ as e -> (match atomic_eta_red e with | Some e' -> e' | None -> ast_map (simpl o) e) | a -> ast_map (simpl o) a (* invariant : list [a] of arguments is non-empty *) and simpl_app o a = function | MLlam (Dummy,t) -> simpl o (MLapp (ast_pop t, List.tl a)) | MLlam (id,t) -> (* Beta redex *) (match nb_occur_match t with | 0 -> simpl o (MLapp (ast_pop t, List.tl a)) | 1 when (is_tmp id || o.opt_lin_beta) -> simpl o (MLapp (ast_subst (List.hd a) t, List.tl a)) | _ -> let a' = List.map (ast_lift 1) (List.tl a) in simpl o (MLletin (id, List.hd a, MLapp (t, a')))) | MLmagic (MLlam (id,t)) -> (* When we've at least one argument, we permute the magic and the lambda, to simplify things a bit (see #2795). Alas, the 1st argument must also be magic then. *) simpl_app o (magic_hd a) (MLlam (id,MLmagic t)) | MLletin (id,e1,e2) when o.opt_let_app -> (* Application of a letin: we push arguments inside *) MLletin (id, e1, simpl o (MLapp (e2, List.map (ast_lift 1) a))) | MLcase (typ,e,br) when o.opt_case_app -> (* Application of a case: we push arguments inside *) let br' = Array.map (fun (l,p,t) -> let k = List.length l in let a' = List.map (ast_lift k) a in (l, p, simpl o (MLapp (t,a')))) br in simpl o (MLcase (typ,e,br')) | (MLdummy _ | MLexn _) as e -> e (* We just discard arguments in those cases. *) | f -> MLapp (f,a) (* Invariant : all empty matches should now be [MLexn] *) and simpl_case o typ br e = try (* Generalized iota-redex *) if not o.opt_case_iot then raise Impossible; simpl o (iota_gen br e) with Impossible -> (* Swap the case and the lam if possible *) let ids,br = if o.opt_case_fun then permut_case_fun br [] else [],br in let n = List.length ids in if not (Int.equal n 0) then simpl o (named_lams ids (MLcase (typ, ast_lift n e, br))) else (* Can we merge several branches as the same constant or function ? *) if lang() == Scheme || is_custom_match br then MLcase (typ, e, br) else match factor_branches o typ br with | Some (f,ints) when Int.equal (Int.Set.cardinal ints) (Array.length br) -> (* If all branches have been factorized, we remove the match *) simpl o (MLletin (Tmp anonymous_name, e, f)) | Some (f,ints) -> let last_br = if ast_occurs 1 f then ([Tmp anonymous_name], Prel 1, f) else ([], Pwild, ast_pop f) in let brl = Array.to_list br in let brl_opt = List.filteri (fun i _ -> not (Int.Set.mem i ints)) brl in let brl_opt = brl_opt @ [last_br] in MLcase (typ, e, Array.of_list brl_opt) | None -> MLcase (typ, e, br) (*S Local prop elimination. *) (* We try to eliminate as many [prop] as possible inside an [ml_ast]. *) (*s In a list, it selects only the elements corresponding to a [Keep] in the boolean list [l]. *) let rec select_via_bl l args = match l,args with | [],_ -> args | Keep::l,a::args -> a :: (select_via_bl l args) | Kill _::l,a::args -> select_via_bl l args | _ -> assert false (*s [kill_some_lams] removes some head lambdas according to the signature [bl]. This list is build on the identifier list model: outermost lambda is on the right. [Rels] corresponding to removed lambdas are not supposed to occur (except maybe in the case of Kimplicit), and the other [Rels] are made correct via a [gen_subst]. Output is not directly a [ml_ast], compose with [named_lams] if needed. *) let is_impl_kill = function Kill (Kimplicit _) -> true | _ -> false let kill_some_lams bl (ids,c) = let n = List.length bl in let n' = List.fold_left (fun n b -> if b == Keep then (n+1) else n) 0 bl in if Int.equal n n' then ids,c else if Int.equal n' 0 && not (List.exists is_impl_kill bl) then [],ast_lift (-n) c else begin let v = Array.make n None in let rec parse_ids i j = function | [] -> () | Keep :: l -> v.(i) <- Some (MLrel j); parse_ids (i+1) (j+1) l | Kill (Kimplicit _ as k) :: l -> v.(i) <- Some (MLdummy k); parse_ids (i+1) j l | Kill _ :: l -> parse_ids (i+1) j l in parse_ids 0 1 bl; select_via_bl bl ids, gen_subst v (n'-n) c end (*s [kill_dummy_lams] uses the last function to kill the lambdas corresponding to a [dummy_name]. It can raise [Impossible] if there is nothing to do, or if there is no lambda left at all. In addition, it now accepts a signature that may mention some implicits. *) let rec merge_implicits ids s = match ids, s with | [],_ -> [] | _,[] -> List.map sign_of_id ids | Dummy::ids, _::s -> Kill Kprop :: merge_implicits ids s | _::ids, (Kill (Kimplicit _) as k)::s -> k :: merge_implicits ids s | _::ids, _::s -> Keep :: merge_implicits ids s let kill_dummy_lams sign c = let ids,c = collect_lams c in let bl = merge_implicits ids (List.rev sign) in if not (List.memq Keep bl) then raise Impossible; let rec fst_kill n = function | [] -> raise Impossible | Kill _ :: bl -> n | Keep :: bl -> fst_kill (n+1) bl in let skip = max 0 ((fst_kill 0 bl) - 1) in let ids_skip, ids = List.chop skip ids in let _, bl = List.chop skip bl in let c = named_lams ids_skip c in let ids',c = kill_some_lams bl (ids,c) in (ids,bl), named_lams ids' c (*s [eta_expansion_sign] takes a function [fun idn ... id1 -> c] and a signature [s] and builds a eta-long version. *) (* For example, if [s = [Keep;Keep;Kill Prop;Keep]] then the output is : [fun idn ... id1 x x _ x -> (c' 4 3 __ 1)] with [c' = lift 4 c] *) let eta_expansion_sign s (ids,c) = let rec abs ids rels i = function | [] -> let a = List.rev_map (function MLrel x -> MLrel (i-x) | a -> a) rels in ids, MLapp (ast_lift (i-1) c, a) | Keep :: l -> abs (anonymous :: ids) (MLrel i :: rels) (i+1) l | Kill k :: l -> abs (Dummy :: ids) (MLdummy k :: rels) (i+1) l in abs ids [] 1 s (*s If [s = [b1; ... ; bn]] then [case_expunge] decomposes [e] in [n] lambdas (with eta-expansion if needed) and removes all dummy lambdas corresponding to [Kill _] in [s]. *) let case_expunge s e = let m = List.length s in let n = nb_lams e in let p = if m <= n then collect_n_lams m e else eta_expansion_sign (List.skipn n s) (collect_lams e) in kill_some_lams (List.rev s) p (*s [term_expunge] takes a function [fun idn ... id1 -> c] and a signature [s] and remove dummy lams. The difference with [case_expunge] is that we here leave one dummy lambda if all lambdas are logical dummy and the target language is strict. *) let term_expunge s (ids,c) = if List.is_empty s then c else let ids,c = kill_some_lams (List.rev s) (ids,c) in if List.is_empty ids && lang () != Haskell && sign_kind s == UnsafeLogicalSig then MLlam (Dummy, ast_lift 1 c) else named_lams ids c (*s [kill_dummy_args (ids,bl) r t] looks for occurrences of [MLrel r] in [t] and purge the args of [MLrel r] corresponding to a [Kill] in [bl]. It makes eta-expansion if needed. *) let kill_dummy_args (ids,bl) r t = let m = List.length ids in let sign = List.rev bl in let rec found n = function | MLrel r' when Int.equal r' (r + n) -> true | MLmagic e -> found n e | _ -> false in let rec killrec n = function | MLapp(e, a) when found n e -> let k = max 0 (m - (List.length a)) in let a = List.map (killrec n) a in let a = List.map (ast_lift k) a in let a = select_via_bl sign (a @ (eta_args k)) in named_lams (List.firstn k ids) (MLapp (ast_lift k e, a)) | e when found n e -> let a = select_via_bl sign (eta_args m) in named_lams ids (MLapp (ast_lift m e, a)) | e -> ast_map_lift killrec n e in killrec 0 t (*s The main function for local [dummy] elimination. *) let sign_of_args a = List.map (function MLdummy k -> Kill k | _ -> Keep) a let rec kill_dummy = function | MLfix(i,fi,c) -> (try let k,c = kill_dummy_fix i c [] in ast_subst (MLfix (i,fi,c)) (kill_dummy_args k 1 (MLrel 1)) with Impossible -> MLfix (i,fi,Array.map kill_dummy c)) | MLapp (MLfix (i,fi,c),a) -> let a = List.map kill_dummy a in (* Heuristics: if some arguments are implicit args, we try to eliminate the corresponding arguments of the fixpoint *) (try let k,c = kill_dummy_fix i c (sign_of_args a) in let fake = MLapp (MLrel 1, List.map (ast_lift 1) a) in let fake' = kill_dummy_args k 1 fake in ast_subst (MLfix (i,fi,c)) fake' with Impossible -> MLapp(MLfix(i,fi,Array.map kill_dummy c),a)) | MLletin(id, MLfix (i,fi,c),e) -> (try let k,c = kill_dummy_fix i c [] in let e = kill_dummy (kill_dummy_args k 1 e) in MLletin(id, MLfix(i,fi,c),e) with Impossible -> MLletin(id, MLfix(i,fi,Array.map kill_dummy c),kill_dummy e)) | MLletin(id,c,e) -> (try let k,c = kill_dummy_lams [] (kill_dummy_hd c) in let e = kill_dummy (kill_dummy_args k 1 e) in let c = kill_dummy c in if is_atomic c then ast_subst c e else MLletin (id, c, e) with Impossible -> MLletin(id,kill_dummy c,kill_dummy e)) | a -> ast_map kill_dummy a (* Similar function, but acting only on head lambdas and let-ins *) and kill_dummy_hd = function | MLlam(id,e) -> MLlam(id, kill_dummy_hd e) | MLletin(id,c,e) -> (try let k,c = kill_dummy_lams [] (kill_dummy_hd c) in let e = kill_dummy_hd (kill_dummy_args k 1 e) in let c = kill_dummy c in if is_atomic c then ast_subst c e else MLletin (id, c, e) with Impossible -> MLletin(id,kill_dummy c,kill_dummy_hd e)) | a -> a and kill_dummy_fix i c s = let n = Array.length c in let k,ci = kill_dummy_lams s (kill_dummy_hd c.(i)) in let c = Array.copy c in c.(i) <- ci; for j = 0 to (n-1) do c.(j) <- kill_dummy (kill_dummy_args k (n-i) c.(j)) done; k,c (*s Putting things together. *) let normalize a = let o = optims () in let rec norm a = let a' = if o.opt_kill_dum then kill_dummy (simpl o a) else simpl o a in if eq_ml_ast a a' then a else norm a' in norm a (*S Special treatment of fixpoint for pretty-printing purpose. *) let general_optimize_fix f ids n args m c = let v = Array.make n 0 in for i=0 to (n-1) do v.(i)<-i done; let aux i = function | MLrel j when v.(j-1)>=0 -> if ast_occurs (j+1) c then raise Impossible else v.(j-1)<-(-i-1) | _ -> raise Impossible in List.iteri aux args; let args_f = List.rev_map (fun i -> MLrel (i+m+1)) (Array.to_list v) in let new_f = anonym_tmp_lams (MLapp (MLrel (n+m+1),args_f)) m in let new_c = named_lams ids (normalize (MLapp ((ast_subst new_f c),args))) in MLfix(0,[|f|],[|new_c|]) let optimize_fix a = if not (optims()).opt_fix_fun then a else let ids,a' = collect_lams a in let n = List.length ids in if Int.equal n 0 then a else match a' with | MLfix(_,[|f|],[|c|]) -> let new_f = MLapp (MLrel (n+1),eta_args n) in let new_c = named_lams ids (normalize (ast_subst new_f c)) in MLfix(0,[|f|],[|new_c|]) | MLapp(a',args) -> let m = List.length args in (match a' with | MLfix(_,_,_) when (test_eta_args_lift 0 n args) && not (ast_occurs_itvl 1 m a') -> a' | MLfix(_,[|f|],[|c|]) -> (try general_optimize_fix f ids n args m c with Impossible -> a) | _ -> a) | _ -> a (*S Inlining. *) (* Utility functions used in the decision of inlining. *) let ml_size_branch size pv = Array.fold_left (fun a (_,_,t) -> a + size t) 0 pv let rec ml_size = function | MLapp(t,l) -> List.length l + ml_size t + ml_size_list l | MLlam(_,t) -> 1 + ml_size t | MLcons(_,_,l) | MLtuple l -> ml_size_list l | MLcase(_,t,pv) -> 1 + ml_size t + ml_size_branch ml_size pv | MLfix(_,_,f) -> ml_size_array f | MLletin (_,_,t) -> ml_size t | MLmagic t -> ml_size t | MLglob _ | MLrel _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> 0 and ml_size_list l = List.fold_left (fun a t -> a + ml_size t) 0 l and ml_size_array a = Array.fold_left (fun a t -> a + ml_size t) 0 a let is_fix = function MLfix _ -> true | _ -> false (*s Strictness *) (* A variable is strict if the evaluation of the whole term implies the evaluation of this variable. Non-strict variables can be found behind Match, for example. Expanding a term [t] is a good idea when it begins by at least one non-strict lambda, since the corresponding argument to [t] might be unevaluated in the expanded code. *) exception Toplevel let lift n l = List.map ((+) n) l let pop n l = List.map (fun x -> if x<=n then raise Toplevel else x-n) l (* This function returns a list of de Bruijn indices of non-strict variables, or raises [Toplevel] if it has an internal non-strict variable. In fact, not all variables are checked for strictness, only the ones which de Bruijn index is in the candidates list [cand]. The flag [add] controls the behaviour when going through a lambda: should we add the corresponding variable to the candidates? We use this flag to check only the external lambdas, those that will correspond to arguments. *) let rec non_stricts add cand = function | MLlam (id,t) -> let cand = lift 1 cand in let cand = if add then 1::cand else cand in pop 1 (non_stricts add cand t) | MLrel n -> List.filter (fun m -> not (Int.equal m n)) cand | MLapp (t,l)-> let cand = non_stricts false cand t in List.fold_left (non_stricts false) cand l | MLcons (_,_,l) -> List.fold_left (non_stricts false) cand l | MLletin (_,t1,t2) -> let cand = non_stricts false cand t1 in pop 1 (non_stricts add (lift 1 cand) t2) | MLfix (_,i,f)-> let n = Array.length i in let cand = lift n cand in let cand = Array.fold_left (non_stricts false) cand f in pop n cand | MLcase (_,t,v) -> (* The only interesting case: for a variable to be non-strict, *) (* it is sufficient that it appears non-strict in at least one branch, *) (* so we make an union (in fact a merge). *) let cand = non_stricts false cand t in Array.fold_left (fun c (i,_,t)-> let n = List.length i in let cand = lift n cand in let cand = pop n (non_stricts add cand t) in List.merge Int.compare cand c) [] v (* [merge] may duplicates some indices, but I don't mind. *) | MLmagic t -> non_stricts add cand t | _ -> cand (* The real test: we are looking for internal non-strict variables, so we start with no candidates, and the only positive answer is via the [Toplevel] exception. *) let is_not_strict t = try let _ = non_stricts true [] t in false with Toplevel -> true (*s Inlining decision *) (* [inline_test] answers the following question: If we could inline [t] (the user said nothing special), should we inline ? We expand small terms with at least one non-strict variable (i.e. a variable that may not be evaluated). Furthermore we don't expand fixpoints. Moreover, as mentioned by X. Leroy (bug #2241), inlining a constant from inside an opaque module might break types. To avoid that, we require below that both [r] and its body are globally visible. This isn't fully satisfactory, since [r] might not be visible (functor), and anyway it might be interesting to inline [r] at least inside its own structure. But to be safe, we adopt this restriction for the moment. *) open Declareops let inline_test r t = if not (auto_inline ()) then false else let c = match r with GlobRef.ConstRef c -> c | _ -> assert false in let has_body = try constant_has_body (Global.lookup_constant c) with Not_found -> false in has_body && (let t1 = eta_red t in let t2 = snd (collect_lams t1) in not (is_fix t2) && ml_size t < 12 && is_not_strict t) let con_of_string s = let d, id = Libnames.split_dirpath (dirpath_of_string s) in Constant.make2 (ModPath.MPfile d) (Label.of_id id) let manual_inline_set = List.fold_right (fun x -> Cset_env.add (con_of_string x)) [ "Coq.Init.Wf.well_founded_induction_type"; "Coq.Init.Wf.well_founded_induction"; "Coq.Init.Wf.Acc_iter"; "Coq.Init.Wf.Fix_F"; "Coq.Init.Wf.Fix"; "Coq.Init.Datatypes.andb"; "Coq.Init.Datatypes.orb"; "Coq.Init.Logic.eq_rec_r"; "Coq.Init.Logic.eq_rect_r"; "Coq.Init.Specif.proj1_sig"; ] Cset_env.empty let manual_inline = function | GlobRef.ConstRef c -> Cset_env.mem c manual_inline_set | _ -> false (* If the user doesn't say he wants to keep [t], we inline in two cases: \begin{itemize} \item the user explicitly requests it \item [expansion_test] answers that the inlining is a good idea, and we are free to act (AutoInline is set) \end{itemize} *) let inline r t = not (to_keep r) (* The user DOES want to keep it *) && not (is_inline_custom r) && (to_inline r (* The user DOES want to inline it *) || (lang () != Haskell && (is_projection r || is_recursor r || manual_inline r || inline_test r t))) coq-8.11.0/plugins/extraction/ExtrHaskellNatInteger.v0000644000175000017500000000071613612664533022531 0ustar treinentreinen(** Extraction of [nat] into Haskell's [Integer] *) Require Coq.extraction.Extraction. Require Import Arith. Require Import ExtrHaskellNatNum. (** * Disclaimer: trying to obtain efficient certified programs * by extracting [nat] into [Integer] isn't necessarily a good idea. * See comments in [ExtrOcamlNatInt.v]. *) Extract Inductive nat => "Prelude.Integer" [ "0" "Prelude.succ" ] "(\fO fS n -> if n Prelude.== 0 then fO () else fS (n Prelude.- 1))". coq-8.11.0/plugins/extraction/README0000644000175000017500000001130013612664533017002 0ustar treinentreinen Coq Extraction ============== What is it ? ------------ The extraction is a mechanism that produces functional code (Ocaml/Haskell/Scheme) out of any Coq terms (either programs or proofs). Who did it ? ------------ The current implementation (from version 7.0 up to now) has been done by P. Letouzey during his PhD, helped by J.C. Filliâtre and supervised by C. Paulin. An earlier implementation (versions 6.x) was due to B. Werner and C. Paulin. Where can we find more information ? ------------------------------------ - Coq Reference Manual includes a full chapter about extraction - P. Letouzey's PhD thesis [3] forms a complete document about both theory and implementation and test-cases of Coq-extraction - A more recent article [4] proposes a short overview of extraction - earlier documents [1] [2] may also be useful. Why a complete re-implementation ? ---------------------------------- Extraction code has been completely rewritten since version V6.3. 1) Principles The main goal of the new extraction is to handle any Coq term, even those upon sort Type, and to produce code that always compiles. Thus it will never answer something like "Not an ML type", but rather a dummy term like the ML unit. Translation between Coq and ML is based upon the following principles: - Terms of sort Prop don't have any computational meaning, so they are merged into one ML term "__". This part is done according to P. Letouzey's works [1] and [2]. This dummy constant "__" used to be implemented by the unit (), but we recently found that this constant might be applied in some cases. So "__" is now in Ocaml a fixpoint that forgets its arguments: let __ = let rec f _ = Obj.repr f in Obj.repr f - Terms that are type schemes (i.e. something of type ( : )( : )...s with s a sort ) don't have any ML counterpart at the term level, since they are types transformers. In fact they do not have any computational meaning either. So we also merge them into that dummy term "__". - A Coq term gives a ML term or a ML type depending of its type: type schemes will (try to) give ML types, and all other terms give ML terms. And the rest of the translation is (almost) straightforward: an inductive gives an inductive, etc... This gives ML code that have no special reason to typecheck, due to the incompatibilities between Coq and ML typing systems. In fact most of the time everything goes right. We now verify during extraction that the produced code is typecheckable, and if it is not we insert unsafe type casting at critical points in the code, with either "Obj.magic" in Ocaml or "unsafeCoerce" in Haskell. 2) Differences with previous extraction (V6.3 and before) 2.a) The pros The ability to extract every Coq term, as explain in the previous paragraph. The ability to extract from a file an ML module (cf Extraction Library in the documentation) You can have a taste of extraction directly at the toplevel by using the "Extraction " or the "Recursive Extraction ". This toplevel extraction was already there in V6.3, but was printing Fw terms. It now prints in the language of your choice: Ocaml, Haskell or Scheme. The optimization done on extracted code has been ported between V6.3 and V7 and enhanced, and in particular the mechanism of automatic expansion. 2.b) The cons The presence of some parasite "__" as dummy arguments in functions. This denotes the rests of a proof part. The previous extraction was able to remove them totally. The current implementation removes a good deal of them, but not all. This problem is due to extraction upon Type. For example, let's take this pathological term: (if b then Set else Prop) : Type The only way to know if this is an Set (to keep) or a Prop (to remove) is to compute the boolean b, and we do not want to do that during extraction. There is no more "ML import" feature. You can compensate by using Axioms, and then "Extract Constant ..." [1]: Exécution de termes de preuves: une nouvelle méthode d'extraction pour le Calcul des Constructions Inductives, Pierre Letouzey, DEA thesis, 2000, http://www.pps.jussieu.fr/~letouzey/download/rapport_dea.ps.gz [2]: A New Extraction for Coq, Pierre Letouzey, Types 2002 Post-Workshop Proceedings. http://www.pps.jussieu.fr/~letouzey/download/extraction2002.ps.gz [3]: Programmation fonctionnelle certifiée: l'extraction de programmes dans l'assistant Coq. Pierre Letouzey, PhD thesis, 2004. http://www.pps.jussieu.fr/~letouzey/download/these_letouzey.ps.gz http://www.pps.jussieu.fr/~letouzey/download/these_letouzey_English.ps.gz [4]: Coq Extraction, An overview. Pierre Letouzey. CiE2008. http://www.pps.jussieu.fr/~letouzey/download/letouzey_extr_cie08.pdf coq-8.11.0/plugins/extraction/ExtrHaskellNatInt.v0000644000175000017500000000070713612664533021666 0ustar treinentreinen(** Extraction of [nat] into Haskell's [Int] *) Require Coq.extraction.Extraction. Require Import Arith. Require Import ExtrHaskellNatNum. (** * Disclaimer: trying to obtain efficient certified programs * by extracting [nat] into [Int] is definitively *not* a good idea. * See comments in [ExtrOcamlNatInt.v]. *) Extract Inductive nat => "Prelude.Int" [ "0" "Prelude.succ" ] "(\fO fS n -> if n Prelude.== 0 then fO () else fS (n Prelude.- 1))". coq-8.11.0/plugins/extraction/scheme.ml0000644000175000017500000002011513612664533017724 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Id.Set.add (Id.of_string s)) [ "define"; "let"; "lambda"; "lambdas"; "match"; "apply"; "car"; "cdr"; "error"; "delay"; "force"; "_"; "__"] Id.Set.empty let pp_comment s = str";; "++h 0 s++fnl () let pp_header_comment = function | None -> mt () | Some com -> pp_comment com ++ fnl () ++ fnl () let preamble _ comment _ usf = pp_header_comment comment ++ str ";; This extracted scheme code relies on some additional macros\n" ++ str ";; available at http://www.pps.univ-paris-diderot.fr/~letouzey/scheme\n" ++ str "(load \"macros_extr.scm\")\n\n" ++ (if usf.mldummy then str "(define __ (lambda (_) __))\n\n" else mt ()) let pr_id id = str @@ String.map (fun c -> if c == '\'' then '~' else c) (Id.to_string id) let paren = pp_par true let pp_abst st = function | [] -> assert false | [id] -> paren (str "lambda " ++ paren (pr_id id) ++ spc () ++ st) | l -> paren (str "lambdas " ++ paren (prlist_with_sep spc pr_id l) ++ spc () ++ st) let pp_apply st _ = function | [] -> st | [a] -> hov 2 (paren (st ++ spc () ++ a)) | args -> hov 2 (paren (str "@ " ++ st ++ (prlist_strict (fun x -> spc () ++ x) args))) (*s The pretty-printer for Scheme syntax *) let pp_global k r = str (Common.pp_global k r) (*s Pretty-printing of expressions. *) let rec pp_expr env args = let apply st = pp_apply st true args in function | MLrel n -> let id = get_db_name n env in apply (pr_id id) | MLapp (f,args') -> let stl = List.map (pp_expr env []) args' in pp_expr env (stl @ args) f | MLlam _ as a -> let fl,a' = collect_lams a in let fl,env' = push_vars (List.map id_of_mlid fl) env in apply (pp_abst (pp_expr env' [] a') (List.rev fl)) | MLletin (id,a1,a2) -> let i,env' = push_vars [id_of_mlid id] env in apply (hv 0 (hov 2 (paren (str "let " ++ paren (paren (pr_id (List.hd i) ++ spc () ++ pp_expr env [] a1)) ++ spc () ++ hov 0 (pp_expr env' [] a2))))) | MLglob r -> apply (pp_global Term r) | MLcons (_,r,args') -> assert (List.is_empty args); let st = str "`" ++ paren (pp_global Cons r ++ (if List.is_empty args' then mt () else spc ()) ++ prlist_with_sep spc (pp_cons_args env) args') in if is_coinductive r then paren (str "delay " ++ st) else st | MLtuple _ -> user_err Pp.(str "Cannot handle tuples in Scheme yet.") | MLcase (_,_,pv) when not (is_regular_match pv) -> user_err Pp.(str "Cannot handle general patterns in Scheme yet.") | MLcase (_,t,pv) when is_custom_match pv -> let mkfun (ids,_,e) = if not (List.is_empty ids) then named_lams (List.rev ids) e else dummy_lams (ast_lift 1 e) 1 in apply (paren (hov 2 (str (find_custom_match pv) ++ fnl () ++ prvect (fun tr -> pp_expr env [] (mkfun tr) ++ fnl ()) pv ++ pp_expr env [] t))) | MLcase (typ,t, pv) -> let e = if not (is_coinductive_type typ) then pp_expr env [] t else paren (str "force" ++ spc () ++ pp_expr env [] t) in apply (v 3 (paren (str "match " ++ e ++ fnl () ++ pp_pat env pv))) | MLfix (i,ids,defs) -> let ids',env' = push_vars (List.rev (Array.to_list ids)) env in pp_fix env' i (Array.of_list (List.rev ids'),defs) args | MLexn s -> (* An [MLexn] may be applied, but I don't really care. *) paren (str "error" ++ spc () ++ qs s) | MLdummy _ -> str "__" (* An [MLdummy] may be applied, but I don't really care. *) | MLmagic a -> pp_expr env args a | MLaxiom -> paren (str "error \"AXIOM TO BE REALIZED\"") | MLuint _ -> paren (str "Prelude.error \"EXTRACTION OF UINT NOT IMPLEMENTED\"") | MLfloat _ -> paren (str "Prelude.error \"EXTRACTION OF FLOAT NOT IMPLEMENTED\"") and pp_cons_args env = function | MLcons (_,r,args) when is_coinductive r -> paren (pp_global Cons r ++ (if List.is_empty args then mt () else spc ()) ++ prlist_with_sep spc (pp_cons_args env) args) | e -> str "," ++ pp_expr env [] e and pp_one_pat env (ids,p,t) = let r = match p with | Pusual r -> r | Pcons (r,l) -> r (* cf. the check [is_regular_match] above *) | _ -> assert false in let ids,env' = push_vars (List.rev_map id_of_mlid ids) env in let args = if List.is_empty ids then mt () else (str " " ++ prlist_with_sep spc pr_id (List.rev ids)) in (pp_global Cons r ++ args), (pp_expr env' [] t) and pp_pat env pv = prvect_with_sep fnl (fun x -> let s1,s2 = pp_one_pat env x in hov 2 (str "((" ++ s1 ++ str ")" ++ spc () ++ s2 ++ str ")")) pv (*s names of the functions ([ids]) are already pushed in [env], and passed here just for convenience. *) and pp_fix env j (ids,bl) args = paren (str "letrec " ++ (v 0 (paren (prvect_with_sep fnl (fun (fi,ti) -> paren ((pr_id fi) ++ spc () ++ (pp_expr env [] ti))) (Array.map2 (fun id b -> (id,b)) ids bl)) ++ fnl () ++ hov 2 (pp_apply (pr_id (ids.(j))) true args)))) (*s Pretty-printing of a declaration. *) let pp_decl = function | Dind _ -> mt () | Dtype _ -> mt () | Dfix (rv, defs,_) -> let names = Array.map (fun r -> if is_inline_custom r then mt () else pp_global Term r) rv in prvecti (fun i r -> let void = is_inline_custom r || (not (is_custom r) && match defs.(i) with MLexn "UNUSED" -> true | _ -> false) in if void then mt () else hov 2 (paren (str "define " ++ names.(i) ++ spc () ++ (if is_custom r then str (find_custom r) else pp_expr (empty_env ()) [] defs.(i))) ++ fnl ()) ++ fnl ()) rv | Dterm (r, a, _) -> if is_inline_custom r then mt () else hov 2 (paren (str "define " ++ pp_global Term r ++ spc () ++ (if is_custom r then str (find_custom r) else pp_expr (empty_env ()) [] a))) ++ fnl2 () let rec pp_structure_elem = function | (l,SEdecl d) -> pp_decl d | (l,SEmodule m) -> pp_module_expr m.ml_mod_expr | (l,SEmodtype m) -> mt () (* for the moment we simply discard module type *) and pp_module_expr = function | MEstruct (mp,sel) -> prlist_strict pp_structure_elem sel | MEfunctor _ -> mt () (* for the moment we simply discard unapplied functors *) | MEident _ | MEapply _ -> assert false (* should be expanded in extract_env *) let pp_struct = let pp_sel (mp,sel) = push_visible mp []; let p = prlist_strict pp_structure_elem sel in pop_visible (); p in prlist_strict pp_sel let scheme_descr = { keywords = keywords; file_suffix = ".scm"; file_naming = file_of_modfile; preamble = preamble; pp_struct = pp_struct; sig_suffix = None; sig_preamble = (fun _ _ _ _ -> mt ()); pp_sig = (fun _ -> mt ()); pp_decl = pp_decl; } coq-8.11.0/plugins/extraction/ExtrOcamlNatInt.v0000644000175000017500000000701713612664533021337 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int [ "0" "Pervasives.succ" ] "(fun fO fS n -> if n=0 then fO () else fS (n-1))". (** Efficient (but uncertified) versions for usual [nat] functions *) Extract Constant plus => "(+)". Extract Constant pred => "fun n -> Pervasives.max 0 (n-1)". Extract Constant minus => "fun n m -> Pervasives.max 0 (n-m)". Extract Constant mult => "( * )". Extract Inlined Constant max => "Pervasives.max". Extract Inlined Constant min => "Pervasives.min". (*Extract Inlined Constant nat_beq => "(=)".*) Extract Inlined Constant EqNat.beq_nat => "(=)". Extract Inlined Constant EqNat.eq_nat_decide => "(=)". Extract Inlined Constant Peano_dec.eq_nat_dec => "(=)". Extract Constant Nat.compare => "fun n m -> if n=m then Eq else if n "(<=)". Extract Inlined Constant Compare_dec.le_lt_dec => "(<=)". Extract Inlined Constant Compare_dec.lt_dec => "(<)". Extract Constant Compare_dec.lt_eq_lt_dec => "fun n m -> if n>m then None else Some (n "fun n -> n mod 2 = 0". Extract Constant Div2.div2 => "fun n -> n/2". Extract Inductive Euclid.diveucl => "(int * int)" [ "" ]. Extract Constant Euclid.eucl_dev => "fun n m -> (m/n, m mod n)". Extract Constant Euclid.quotient => "fun n m -> m/n". Extract Constant Euclid.modulo => "fun n m -> m mod n". (* Definition test n m (H:m>0) := let (q,r,_,_) := eucl_dev m H n in nat_compare n (q*m+r). Recursive Extraction test fact. *) coq-8.11.0/plugins/extraction/modutil.mli0000644000175000017500000000346713612664533020321 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool) -> ml_structure -> bool val struct_type_search : (ml_type -> bool) -> ml_structure -> bool type do_ref = GlobRef.t -> unit val type_iter_references : do_ref -> ml_type -> unit val ast_iter_references : do_ref -> do_ref -> do_ref -> ml_ast -> unit val decl_iter_references : do_ref -> do_ref -> do_ref -> ml_decl -> unit val spec_iter_references : do_ref -> do_ref -> do_ref -> ml_spec -> unit val signature_of_structure : ml_structure -> ml_signature val mtyp_of_mexpr : ml_module_expr -> ml_module_type val msid_of_mt : ml_module_type -> ModPath.t val get_decl_in_structure : GlobRef.t -> ml_structure -> ml_decl (* Some transformations of ML terms. [optimize_struct] simplify all beta redexes (when the argument does not occur, it is just thrown away; when it occurs exactly once it is substituted; otherwise a let-in redex is created for clarity) and iota redexes, plus some other optimizations. The first argument is the list of objects we want to appear. *) val optimize_struct : GlobRef.t list * ModPath.t list -> ml_structure -> ml_structure coq-8.11.0/plugins/extraction/miniml.ml0000644000175000017500000001434513612664533017755 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* string; (* the second argument is a comment to add to the preamble *) preamble : Id.t -> Pp.t option -> ModPath.t list -> unsafe_needs -> Pp.t; pp_struct : ml_structure -> Pp.t; (* Concerning a possible interface file *) sig_suffix : string option; (* the second argument is a comment to add to the preamble *) sig_preamble : Id.t -> Pp.t option -> ModPath.t list -> unsafe_needs -> Pp.t; pp_sig : ml_signature -> Pp.t; (* for an isolated declaration print *) pp_decl : ml_decl -> Pp.t; } coq-8.11.0/plugins/extraction/ExtrOcamlNatBigInt.v0000644000175000017500000000542213612664533021757 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* "Big.big_int" [ "Big.zero" "Big.succ" ] "Big.nat_case". (** Efficient (but uncertified) versions for usual [nat] functions *) Extract Constant plus => "Big.add". Extract Constant mult => "Big.mult". Extract Constant pred => "fun n -> Big.max Big.zero (Big.pred n)". Extract Constant minus => "fun n m -> Big.max Big.zero (Big.sub n m)". Extract Constant max => "Big.max". Extract Constant min => "Big.min". (*Extract Constant nat_beq => "Big.eq".*) Extract Constant EqNat.beq_nat => "Big.eq". Extract Constant EqNat.eq_nat_decide => "Big.eq". Extract Constant Peano_dec.eq_nat_dec => "Big.eq". Extract Constant Nat.compare => "Big.compare_case Eq Lt Gt". Extract Constant Compare_dec.leb => "Big.le". Extract Constant Compare_dec.le_lt_dec => "Big.le". Extract Constant Compare_dec.lt_eq_lt_dec => "Big.compare_case (Some false) (Some true) None". Extract Constant Even.even_odd_dec => "fun n -> Big.sign (Big.mod n Big.two) = 0". Extract Constant Div2.div2 => "fun n -> Big.div n Big.two". Extract Inductive Euclid.diveucl => "(Big.big_int * Big.big_int)" [""]. Extract Constant Euclid.eucl_dev => "fun n m -> Big.quomod m n". Extract Constant Euclid.quotient => "fun n m -> Big.div m n". Extract Constant Euclid.modulo => "fun n m -> Big.modulo m n". (* Require Import Euclid. Definition test n m (H:m>0) := let (q,r,_,_) := eucl_dev m H n in nat_compare n (q*m+r). Extraction "/tmp/test.ml" test fact pred minus max min Div2.div2. *) coq-8.11.0/plugins/extraction/scheme.mli0000644000175000017500000000131513612664533020076 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* char [ "(* If this appears, you're using Ascii internals. Please don't *) (fun (b0,b1,b2,b3,b4,b5,b6,b7) -> let f b i = if b then 1 lsl i else 0 in Char.chr (f b0 0 + f b1 1 + f b2 2 + f b3 3 + f b4 4 + f b5 5 + f b6 6 + f b7 7))" ] "(* If this appears, you're using Ascii internals. Please don't *) (fun f c -> let n = Char.code c in let h i = (n land (1 lsl i)) <> 0 in f (h 0) (h 1) (h 2) (h 3) (h 4) (h 5) (h 6) (h 7))". Extract Constant zero => "'\000'". Extract Constant one => "'\001'". Extract Constant shift => "fun b c -> Char.chr (((Char.code c) lsl 1) land 255 + if b then 1 else 0)". Extract Inlined Constant ascii_dec => "(=)". Extract Inlined Constant Ascii.eqb => "(=)". Extract Inductive string => "char list" [ "[]" "(::)" ]. (* python -c 'print(" ".join(r""" "%s" """.strip() % (r"'"'\''"'" if chr(i) == "'"'"'" else repr(""" "" """.strip()) if chr(i) == """ " """.strip() else repr(chr(i))) for i in range(256)))' # " to satisfy Coq's comment parser *) Extract Inductive byte => char ["'\x00'" "'\x01'" "'\x02'" "'\x03'" "'\x04'" "'\x05'" "'\x06'" "'\x07'" "'\x08'" "'\t'" "'\n'" "'\x0b'" "'\x0c'" "'\r'" "'\x0e'" "'\x0f'" "'\x10'" "'\x11'" "'\x12'" "'\x13'" "'\x14'" "'\x15'" "'\x16'" "'\x17'" "'\x18'" "'\x19'" "'\x1a'" "'\x1b'" "'\x1c'" "'\x1d'" "'\x1e'" "'\x1f'" "' '" "'!'" "'""'" "'#'" "'$'" "'%'" "'&'" "'\''" "'('" "')'" "'*'" "'+'" "','" "'-'" "'.'" "'/'" "'0'" "'1'" "'2'" "'3'" "'4'" "'5'" "'6'" "'7'" "'8'" "'9'" "':'" "';'" "'<'" "'='" "'>'" "'?'" "'@'" "'A'" "'B'" "'C'" "'D'" "'E'" "'F'" "'G'" "'H'" "'I'" "'J'" "'K'" "'L'" "'M'" "'N'" "'O'" "'P'" "'Q'" "'R'" "'S'" "'T'" "'U'" "'V'" "'W'" "'X'" "'Y'" "'Z'" "'['" "'\\'" "']'" "'^'" "'_'" "'`'" "'a'" "'b'" "'c'" "'d'" "'e'" "'f'" "'g'" "'h'" "'i'" "'j'" "'k'" "'l'" "'m'" "'n'" "'o'" "'p'" "'q'" "'r'" "'s'" "'t'" "'u'" "'v'" "'w'" "'x'" "'y'" "'z'" "'{'" "'|'" "'}'" "'~'" "'\x7f'" "'\x80'" "'\x81'" "'\x82'" "'\x83'" "'\x84'" "'\x85'" "'\x86'" "'\x87'" "'\x88'" "'\x89'" "'\x8a'" "'\x8b'" "'\x8c'" "'\x8d'" "'\x8e'" "'\x8f'" "'\x90'" "'\x91'" "'\x92'" "'\x93'" "'\x94'" "'\x95'" "'\x96'" "'\x97'" "'\x98'" "'\x99'" "'\x9a'" "'\x9b'" "'\x9c'" "'\x9d'" "'\x9e'" "'\x9f'" "'\xa0'" "'\xa1'" "'\xa2'" "'\xa3'" "'\xa4'" "'\xa5'" "'\xa6'" "'\xa7'" "'\xa8'" "'\xa9'" "'\xaa'" "'\xab'" "'\xac'" "'\xad'" "'\xae'" "'\xaf'" "'\xb0'" "'\xb1'" "'\xb2'" "'\xb3'" "'\xb4'" "'\xb5'" "'\xb6'" "'\xb7'" "'\xb8'" "'\xb9'" "'\xba'" "'\xbb'" "'\xbc'" "'\xbd'" "'\xbe'" "'\xbf'" "'\xc0'" "'\xc1'" "'\xc2'" "'\xc3'" "'\xc4'" "'\xc5'" "'\xc6'" "'\xc7'" "'\xc8'" "'\xc9'" "'\xca'" "'\xcb'" "'\xcc'" "'\xcd'" "'\xce'" "'\xcf'" "'\xd0'" "'\xd1'" "'\xd2'" "'\xd3'" "'\xd4'" "'\xd5'" "'\xd6'" "'\xd7'" "'\xd8'" "'\xd9'" "'\xda'" "'\xdb'" "'\xdc'" "'\xdd'" "'\xde'" "'\xdf'" "'\xe0'" "'\xe1'" "'\xe2'" "'\xe3'" "'\xe4'" "'\xe5'" "'\xe6'" "'\xe7'" "'\xe8'" "'\xe9'" "'\xea'" "'\xeb'" "'\xec'" "'\xed'" "'\xee'" "'\xef'" "'\xf0'" "'\xf1'" "'\xf2'" "'\xf3'" "'\xf4'" "'\xf5'" "'\xf6'" "'\xf7'" "'\xf8'" "'\xf9'" "'\xfa'" "'\xfb'" "'\xfc'" "'\xfd'" "'\xfe'" "'\xff'"]. Extract Inlined Constant Byte.eqb => "(=)". Extract Inlined Constant Byte.byte_eq_dec => "(=)". Extract Inlined Constant Ascii.ascii_of_byte => "(fun x -> x)". Extract Inlined Constant Ascii.byte_of_ascii => "(fun x -> x)". (* Definition test := "ceci est un test"%string. Definition test2 := List.map (option_map Byte.to_nat) (List.map Byte.of_nat (List.seq 0 256)). Definition test3 := List.map ascii_of_nat (List.seq 0 256). Recursive Extraction test Ascii.zero Ascii.one test2 test3 byte_rect. *) coq-8.11.0/plugins/extraction/plugin_base.dune0000644000175000017500000000022213612664533021270 0ustar treinentreinen(library (name extraction_plugin) (public_name coq.plugins.extraction) (synopsis "Coq's extraction plugin") (libraries num coq.plugins.ltac)) coq-8.11.0/plugins/extraction/extraction_plugin.mlpack0000644000175000017500000000014113612664533023052 0ustar treinentreinenMiniml Table Mlutil Modutil Extraction Common Ocaml Haskell Scheme Json Extract_env G_extraction coq-8.11.0/plugins/syntax/0000755000175000017500000000000013612664533015275 5ustar treinentreinencoq-8.11.0/plugins/syntax/float_syntax_plugin.mlpack0000644000175000017500000000001513612664533022553 0ustar treinentreinenFloat_syntax coq-8.11.0/plugins/syntax/r_syntax.ml0000644000175000017500000001752113612664533017504 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* GlobRef.equal r gr | _ -> false let positive_modpath = MPfile (make_dir binnums) let positive_kn = MutInd.make2 positive_modpath (Label.make "positive") let path_of_xI = ((positive_kn,0),1) let path_of_xO = ((positive_kn,0),2) let path_of_xH = ((positive_kn,0),3) let glob_xI = GlobRef.ConstructRef path_of_xI let glob_xO = GlobRef.ConstructRef path_of_xO let glob_xH = GlobRef.ConstructRef path_of_xH let pos_of_bignat ?loc x = let ref_xI = DAst.make @@ GRef (glob_xI, None) in let ref_xH = DAst.make @@ GRef (glob_xH, None) in let ref_xO = DAst.make @@ GRef (glob_xO, None) in let rec pos_of x = match div2_with_rest x with | (q,false) -> DAst.make @@ GApp (ref_xO,[pos_of q]) | (q,true) when not (Bigint.equal q zero) -> DAst.make @@ GApp (ref_xI,[pos_of q]) | (q,true) -> ref_xH in pos_of x (**********************************************************************) (* Printing positive via scopes *) (**********************************************************************) let rec bignat_of_pos c = match DAst.get c with | GApp (r, [a]) when is_gr r glob_xO -> mult_2(bignat_of_pos a) | GApp (r, [a]) when is_gr r glob_xI -> add_1(mult_2(bignat_of_pos a)) | GRef (a, _) when GlobRef.equal a glob_xH -> Bigint.one | _ -> raise Non_closed_number (**********************************************************************) (* Parsing Z via scopes *) (**********************************************************************) let z_kn = MutInd.make2 positive_modpath (Label.make "Z") let path_of_ZERO = ((z_kn,0),1) let path_of_POS = ((z_kn,0),2) let path_of_NEG = ((z_kn,0),3) let glob_ZERO = GlobRef.ConstructRef path_of_ZERO let glob_POS = GlobRef.ConstructRef path_of_POS let glob_NEG = GlobRef.ConstructRef path_of_NEG let z_of_int ?loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in DAst.make @@ GApp(DAst.make @@ GRef (sgn,None), [pos_of_bignat ?loc n]) else DAst.make @@ GRef (glob_ZERO, None) (**********************************************************************) (* Printing Z via scopes *) (**********************************************************************) let bigint_of_z c = match DAst.get c with | GApp (r,[a]) when is_gr r glob_POS -> bignat_of_pos a | GApp (r,[a]) when is_gr r glob_NEG -> Bigint.neg (bignat_of_pos a) | GRef (a, _) when GlobRef.equal a glob_ZERO -> Bigint.zero | _ -> raise Non_closed_number (**********************************************************************) (* Parsing R via scopes *) (**********************************************************************) let rdefinitions = ["Coq";"Reals";"Rdefinitions"] let r_modpath = MPfile (make_dir rdefinitions) let r_base_modpath = MPdot (r_modpath, Label.make "RbaseSymbolsImpl") let r_path = make_path ["Coq";"Reals";"Rdefinitions";"RbaseSymbolsImpl"] "R" let glob_IZR = GlobRef.ConstRef (Constant.make2 r_modpath @@ Label.make "IZR") let glob_Rmult = GlobRef.ConstRef (Constant.make2 r_base_modpath @@ Label.make "Rmult") let glob_Rdiv = GlobRef.ConstRef (Constant.make2 r_modpath @@ Label.make "Rdiv") let binintdef = ["Coq";"ZArith";"BinIntDef"] let z_modpath = MPdot (MPfile (make_dir binintdef), Label.make "Z") let glob_pow_pos = GlobRef.ConstRef (Constant.make2 z_modpath @@ Label.make "pow_pos") let r_of_rawnum ?loc (sign,n) = let n, f, e = NumTok.(n.int, n.frac, n.exp) in let izr z = DAst.make @@ GApp (DAst.make @@ GRef(glob_IZR,None), [z]) in let rmult r r' = DAst.make @@ GApp (DAst.make @@ GRef(glob_Rmult,None), [r; r']) in let rdiv r r' = DAst.make @@ GApp (DAst.make @@ GRef(glob_Rdiv,None), [r; r']) in let pow10 e = let ten = z_of_int ?loc (Bigint.of_int 10) in let e = pos_of_bignat e in DAst.make @@ GApp (DAst.make @@ GRef(glob_pow_pos,None), [ten; e]) in let n = let n = Bigint.of_string (n ^ f) in let n = match sign with SPlus -> n | SMinus -> Bigint.(neg n) in izr (z_of_int ?loc n) in let e = let e = if e = "" then Bigint.zero else match e.[1] with | '+' -> Bigint.of_string (String.sub e 2 (String.length e - 2)) | '-' -> Bigint.(neg (of_string (String.sub e 2 (String.length e - 2)))) | _ -> Bigint.of_string (String.sub e 1 (String.length e - 1)) in Bigint.(sub e (of_int (String.length f))) in if Bigint.is_strictly_pos e then rmult n (izr (pow10 e)) else if Bigint.is_strictly_neg e then rdiv n (izr (pow10 (neg e))) else n (* e = 0 *) (**********************************************************************) (* Printing R via scopes *) (**********************************************************************) let rawnum_of_r c = match DAst.get c with | GApp (r, [a]) when is_gr r glob_IZR -> let n = bigint_of_z a in let s, n = if is_strictly_neg n then SMinus, neg n else SPlus, n in s, NumTok.int (to_string n) | GApp (md, [l; r]) when is_gr md glob_Rmult || is_gr md glob_Rdiv -> begin match DAst.get l, DAst.get r with | GApp (i, [l]), GApp (i', [r]) when is_gr i glob_IZR && is_gr i' glob_IZR -> begin match DAst.get r with | GApp (p, [t; e]) when is_gr p glob_pow_pos -> let t = bigint_of_z t in if not (Bigint.(equal t (of_int 10))) then raise Non_closed_number else let i = bigint_of_z l in let e = bignat_of_pos e in let s, i = if is_pos_or_zero i then SPlus, i else SMinus, neg i in let i = Bigint.to_string i in let se = if is_gr md glob_Rdiv then "-" else "" in let e = se ^ Bigint.to_string e in s, { NumTok.int = i; frac = ""; exp = e } | _ -> raise Non_closed_number end | _ -> raise Non_closed_number end | _ -> raise Non_closed_number let uninterp_r (AnyGlobConstr p) = try Some (rawnum_of_r p) with Non_closed_number -> None open Notation let at_declare_ml_module f x = Mltop.declare_cache_obj (fun () -> f x) __coq_plugin_name let r_scope = "R_scope" let _ = register_rawnumeral_interpretation r_scope (r_of_rawnum,uninterp_r); at_declare_ml_module enable_prim_token_interpretation { pt_local = false; pt_scope = r_scope; pt_interp_info = Uid r_scope; pt_required = (r_path,["Coq";"Reals";"Rdefinitions"]); pt_refs = [glob_IZR; glob_Rmult; glob_Rdiv]; pt_in_match = false } coq-8.11.0/plugins/syntax/string_notation_plugin.mlpack0000644000175000017500000000003113612664533023257 0ustar treinentreinenString_notation G_string coq-8.11.0/plugins/syntax/g_string.mlg0000644000175000017500000000203713612664533017614 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { vernac_string_notation (Locality.make_module_locality locality) ty f g (Id.to_string sc) } END coq-8.11.0/plugins/syntax/int63_syntax_plugin.mlpack0000644000175000017500000000001513612664533022411 0ustar treinentreinenInt63_syntax coq-8.11.0/plugins/syntax/float_syntax.ml0000644000175000017500000000342513612664533020346 0ustar treinentreinen(************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* "" | SMinus -> "-") in DAst.make ?loc (GFloat (Float64.of_string (sign ^ NumTok.to_string n))) (* Pretty printing is already handled in constrextern.ml *) let uninterp_float _ = None (* Actually declares the interpreter for float *) open Notation let at_declare_ml_module f x = Mltop.declare_cache_obj (fun () -> f x) __coq_plugin_name let float_module = ["Coq"; "Floats"; "PrimFloat"] let float_path = make_path float_module "float" let float_scope = "float_scope" let _ = register_rawnumeral_interpretation float_scope (interp_float,uninterp_float); at_declare_ml_module enable_prim_token_interpretation { pt_local = false; pt_scope = float_scope; pt_interp_info = Uid float_scope; pt_required = (float_path,float_module); pt_refs = []; pt_in_match = false } coq-8.11.0/plugins/syntax/r_syntax_plugin.mlpack0000644000175000017500000000001113612664533021703 0ustar treinentreinenR_syntax coq-8.11.0/plugins/syntax/numeral.ml0000644000175000017500000001560613612664533017302 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* strbrk "The 'abstract after' directive has no effect when " ++ strbrk "the parsing function (" ++ Nametab.pr_global_env (Termops.vars_of_env (Global.env ())) f ++ strbrk ") targets an " ++ strbrk "option type.") let get_constructors ind = let mib,oib = Global.lookup_inductive ind in let mc = oib.Declarations.mind_consnames in Array.to_list (Array.mapi (fun j c -> GlobRef.ConstructRef (ind, j + 1)) mc) let qualid_of_ref n = n |> Coqlib.lib_ref |> Nametab.shortest_qualid_of_global Id.Set.empty let q_option () = qualid_of_ref "core.option.type" let unsafe_locate_ind q = match Nametab.locate q with | GlobRef.IndRef i -> i | _ -> raise Not_found let locate_z () = let zn = "num.Z.type" in let pn = "num.pos.type" in if Coqlib.has_ref zn && Coqlib.has_ref pn then let q_z = qualid_of_ref zn in let q_pos = qualid_of_ref pn in Some ({ z_ty = unsafe_locate_ind q_z; pos_ty = unsafe_locate_ind q_pos; }, mkRefC q_z) else None let locate_decimal () = let int = "num.int.type" in let uint = "num.uint.type" in let dec = "num.decimal.type" in if Coqlib.has_ref int && Coqlib.has_ref uint && Coqlib.has_ref dec then let q_int = qualid_of_ref int in let q_uint = qualid_of_ref uint in let q_dec = qualid_of_ref dec in let int_ty = { int = unsafe_locate_ind q_int; uint = unsafe_locate_ind q_uint; } in let dec_ty = { int = int_ty; decimal = unsafe_locate_ind q_dec; } in Some (int_ty, mkRefC q_int, mkRefC q_uint, dec_ty, mkRefC q_dec) else None let locate_int63 () = let int63n = "num.int63.type" in if Coqlib.has_ref int63n then let q_int63 = qualid_of_ref int63n in Some (mkRefC q_int63) else None let has_type env sigma f ty = let c = mkCastC (mkRefC f, Glob_term.CastConv ty) in try let _ = Constrintern.interp_constr env sigma c in true with Pretype_errors.PretypeError _ -> false let type_error_to f ty = CErrors.user_err (pr_qualid f ++ str " should go from Decimal.int to " ++ pr_qualid ty ++ str " or (option " ++ pr_qualid ty ++ str ")." ++ fnl () ++ str "Instead of Decimal.int, the types Decimal.uint or Z or Int63.int or Decimal.decimal could be used (you may need to require BinNums or Decimal or Int63 first).") let type_error_of g ty = CErrors.user_err (pr_qualid g ++ str " should go from " ++ pr_qualid ty ++ str " to Decimal.int or (option Decimal.int)." ++ fnl () ++ str "Instead of Decimal.int, the types Decimal.uint or Z or Int63.int or Decimal.decimal could be used (you may need to require BinNums or Decimal or Int63 first).") let vernac_numeral_notation local ty f g scope opts = let env = Global.env () in let sigma = Evd.from_env env in let dec_ty = locate_decimal () in let z_pos_ty = locate_z () in let int63_ty = locate_int63 () in let tyc = Smartlocate.global_inductive_with_alias ty in let to_ty = Smartlocate.global_with_alias f in let of_ty = Smartlocate.global_with_alias g in let cty = mkRefC ty in let app x y = mkAppC (x,[y]) in let arrow x y = mkProdC ([CAst.make Anonymous],Default Glob_term.Explicit, x, y) in let opt r = app (mkRefC (q_option ())) r in let constructors = get_constructors tyc in (* Check the type of f *) let to_kind = match dec_ty with | Some (int_ty, cint, _, _, _) when has_type env sigma f (arrow cint cty) -> Int int_ty, Direct | Some (int_ty, cint, _, _, _) when has_type env sigma f (arrow cint (opt cty)) -> Int int_ty, Option | Some (int_ty, _, cuint, _, _) when has_type env sigma f (arrow cuint cty) -> UInt int_ty.uint, Direct | Some (int_ty, _, cuint, _, _) when has_type env sigma f (arrow cuint (opt cty)) -> UInt int_ty.uint, Option | Some (_, _, _, dec_ty, cdec) when has_type env sigma f (arrow cdec cty) -> Decimal dec_ty, Direct | Some (_, _, _, dec_ty, cdec) when has_type env sigma f (arrow cdec (opt cty)) -> Decimal dec_ty, Option | _ -> match z_pos_ty with | Some (z_pos_ty, cZ) when has_type env sigma f (arrow cZ cty) -> Z z_pos_ty, Direct | Some (z_pos_ty, cZ) when has_type env sigma f (arrow cZ (opt cty)) -> Z z_pos_ty, Option | _ -> match int63_ty with | Some cint63 when has_type env sigma f (arrow cint63 cty) -> Int63, Direct | Some cint63 when has_type env sigma f (arrow cint63 (opt cty)) -> Int63, Option | _ -> type_error_to f ty in (* Check the type of g *) let of_kind = match dec_ty with | Some (int_ty, cint, _, _, _) when has_type env sigma g (arrow cty cint) -> Int int_ty, Direct | Some (int_ty, cint, _, _, _) when has_type env sigma g (arrow cty (opt cint)) -> Int int_ty, Option | Some (int_ty, _, cuint, _, _) when has_type env sigma g (arrow cty cuint) -> UInt int_ty.uint, Direct | Some (int_ty, _, cuint, _, _) when has_type env sigma g (arrow cty (opt cuint)) -> UInt int_ty.uint, Option | Some (_, _, _, dec_ty, cdec) when has_type env sigma g (arrow cty cdec) -> Decimal dec_ty, Direct | Some (_, _, _, dec_ty, cdec) when has_type env sigma g (arrow cty (opt cdec)) -> Decimal dec_ty, Option | _ -> match z_pos_ty with | Some (z_pos_ty, cZ) when has_type env sigma g (arrow cty cZ) -> Z z_pos_ty, Direct | Some (z_pos_ty, cZ) when has_type env sigma g (arrow cty (opt cZ)) -> Z z_pos_ty, Option | _ -> match int63_ty with | Some cint63 when has_type env sigma g (arrow cty cint63) -> Int63, Direct | Some cint63 when has_type env sigma g (arrow cty (opt cint63)) -> Int63, Option | _ -> type_error_of g ty in let o = { to_kind; to_ty; of_kind; of_ty; ty_name = ty; warning = opts } in (match opts, to_kind with | Abstract _, (_, Option) -> warn_abstract_large_num_no_op o.to_ty | _ -> ()); let i = { pt_local = local; pt_scope = scope; pt_interp_info = NumeralNotation o; pt_required = Nametab.path_of_global (GlobRef.IndRef tyc),[]; pt_refs = constructors; pt_in_match = true } in enable_prim_token_interpretation i coq-8.11.0/plugins/syntax/r_syntax.mli0000644000175000017500000000124313612664533017647 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* qualid -> qualid -> qualid -> Notation_term.scope_name -> numnot_option -> unit coq-8.11.0/plugins/syntax/string_notation.mli0000644000175000017500000000160413612664533021222 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* qualid -> qualid -> qualid -> Notation_term.scope_name -> unit coq-8.11.0/plugins/syntax/string_notation.ml0000644000175000017500000000717513612664533021062 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* GlobRef.ConstructRef (ind, j + 1)) mc) let qualid_of_ref n = n |> Coqlib.lib_ref |> Nametab.shortest_qualid_of_global Id.Set.empty let q_option () = qualid_of_ref "core.option.type" let q_list () = qualid_of_ref "core.list.type" let q_byte () = qualid_of_ref "core.byte.type" let has_type env sigma f ty = let c = mkCastC (mkRefC f, Glob_term.CastConv ty) in try let _ = Constrintern.interp_constr env sigma c in true with Pretype_errors.PretypeError _ -> false let type_error_to f ty = CErrors.user_err (pr_qualid f ++ str " should go from Byte.byte or (list Byte.byte) to " ++ pr_qualid ty ++ str " or (option " ++ pr_qualid ty ++ str ").") let type_error_of g ty = CErrors.user_err (pr_qualid g ++ str " should go from " ++ pr_qualid ty ++ str " to Byte.byte or (option Byte.byte) or (list Byte.byte) or (option (list Byte.byte)).") let vernac_string_notation local ty f g scope = let env = Global.env () in let sigma = Evd.from_env env in let app x y = mkAppC (x,[y]) in let cref q = mkRefC q in let cbyte = cref (q_byte ()) in let clist = cref (q_list ()) in let coption = cref (q_option ()) in let opt r = app coption r in let clist_byte = app clist cbyte in let tyc = Smartlocate.global_inductive_with_alias ty in let to_ty = Smartlocate.global_with_alias f in let of_ty = Smartlocate.global_with_alias g in let cty = cref ty in let arrow x y = mkProdC ([CAst.make Anonymous],Default Glob_term.Explicit, x, y) in let constructors = get_constructors tyc in (* Check the type of f *) let to_kind = if has_type env sigma f (arrow clist_byte cty) then ListByte, Direct else if has_type env sigma f (arrow clist_byte (opt cty)) then ListByte, Option else if has_type env sigma f (arrow cbyte cty) then Byte, Direct else if has_type env sigma f (arrow cbyte (opt cty)) then Byte, Option else type_error_to f ty in (* Check the type of g *) let of_kind = if has_type env sigma g (arrow cty clist_byte) then ListByte, Direct else if has_type env sigma g (arrow cty (opt clist_byte)) then ListByte, Option else if has_type env sigma g (arrow cty cbyte) then Byte, Direct else if has_type env sigma g (arrow cty (opt cbyte)) then Byte, Option else type_error_of g ty in let o = { to_kind = to_kind; to_ty = to_ty; of_kind = of_kind; of_ty = of_ty; ty_name = ty; warning = () } in let i = { pt_local = local; pt_scope = scope; pt_interp_info = StringNotation o; pt_required = Nametab.path_of_global (GlobRef.IndRef tyc),[]; pt_refs = constructors; pt_in_match = true } in enable_prim_token_interpretation i coq-8.11.0/plugins/syntax/numeral_notation_plugin.mlpack0000644000175000017500000000002213612664533023414 0ustar treinentreinenNumeral G_numeral coq-8.11.0/plugins/syntax/g_numeral.mlg0000644000175000017500000000272713612664533017757 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* mt () | Warning n -> str "(warning after " ++ str n ++ str ")" | Abstract n -> str "(abstract after " ++ str n ++ str ")" } VERNAC ARGUMENT EXTEND numnotoption PRINTED BY { pr_numnot_option } | [ ] -> { Nop } | [ "(" "warning" "after" bigint(waft) ")" ] -> { Warning waft } | [ "(" "abstract" "after" bigint(n) ")" ] -> { Abstract n } END VERNAC COMMAND EXTEND NumeralNotation CLASSIFIED AS SIDEFF | #[ locality = Attributes.locality; ] [ "Numeral" "Notation" reference(ty) reference(f) reference(g) ":" ident(sc) numnotoption(o) ] -> { vernac_numeral_notation (Locality.make_module_locality locality) ty f g (Id.to_string sc) o } END coq-8.11.0/plugins/syntax/plugin_base.dune0000644000175000017500000000152513612664533020445 0ustar treinentreinen(library (name numeral_notation_plugin) (public_name coq.plugins.numeral_notation) (synopsis "Coq numeral notation plugin") (modules g_numeral numeral) (libraries coq.vernac)) (library (name string_notation_plugin) (public_name coq.plugins.string_notation) (synopsis "Coq string notation plugin") (modules g_string string_notation) (libraries coq.vernac)) (library (name r_syntax_plugin) (public_name coq.plugins.r_syntax) (synopsis "Coq syntax plugin: reals") (modules r_syntax) (libraries coq.vernac)) (library (name int63_syntax_plugin) (public_name coq.plugins.int63_syntax) (synopsis "Coq syntax plugin: int63") (modules int63_syntax) (libraries coq.vernac)) (library (name float_syntax_plugin) (public_name coq.plugins.float_syntax) (synopsis "Coq syntax plugin: float") (modules float_syntax) (libraries coq.vernac)) coq-8.11.0/plugins/syntax/int63_syntax.ml0000644000175000017500000000411213612664533020176 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* f x) __coq_plugin_name (* Actually declares the interpreter for int63 *) let _ = let open Notation in at_declare_ml_module (fun () -> let id_int63 = Nametab.locate q_id_int63 in let o = { to_kind = Int63, Direct; to_ty = id_int63; of_kind = Int63, Direct; of_ty = id_int63; ty_name = q_int63; warning = Nop } in enable_prim_token_interpretation { pt_local = false; pt_scope = int63_scope; pt_interp_info = NumeralNotation o; pt_required = (int63_path, int63_module); pt_refs = []; pt_in_match = false }) () coq-8.11.0/plugins/ssrmatching/0000755000175000017500000000000013612664533016271 5ustar treinentreinencoq-8.11.0/plugins/ssrmatching/ssrmatching.mli0000644000175000017500000002523013612664533021320 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Pp.t (** The type of rewrite patterns, the patterns of the [rewrite] tactic. These patterns also include patterns that identify all the subterms of a context (i.e. "in" prefix) *) type rpattern val pr_rpattern : rpattern -> Pp.t (** Pattern interpretation and matching *) exception NoMatch exception NoProgress (** AST for [rpattern] (and consequently [cpattern]) *) type ('ident, 'term) ssrpattern = | T of 'term | In_T of 'term | X_In_T of 'ident * 'term | In_X_In_T of 'ident * 'term | E_In_X_In_T of 'term * 'ident * 'term | E_As_X_In_T of 'term * 'ident * 'term type pattern = evar_map * (constr, constr) ssrpattern val pp_pattern : env -> pattern -> Pp.t (** Extracts the redex and applies to it the substitution part of the pattern. @raise Anomaly if called on [In_T] or [In_X_In_T] *) val redex_of_pattern : ?resolve_typeclasses:bool -> env -> pattern -> constr Evd.in_evar_universe_context (** [interp_rpattern ise gl rpat] "internalizes" and "interprets" [rpat] in the current [Ltac] interpretation signature [ise] and tactic input [gl]*) val interp_rpattern : goal sigma -> rpattern -> pattern (** [interp_cpattern ise gl cpat ty] "internalizes" and "interprets" [cpat] in the current [Ltac] interpretation signature [ise] and tactic input [gl]. [ty] is an optional type for the redex of [cpat] *) val interp_cpattern : goal sigma -> cpattern -> (glob_constr_and_expr * Geninterp.interp_sign) option -> pattern (** The set of occurrences to be matched. The boolean is set to true * to signal the complement of this set (i.e. \{-1 3\}) *) type occ = (bool * int list) option (** [subst e p t i]. [i] is the number of binders traversed so far, [p] the term from the pattern, [t] the matched one *) type subst = env -> constr -> constr -> int -> constr (** [eval_pattern b env sigma t pat occ subst] maps [t] calling [subst] on every [occ] occurrence of [pat]. The [int] argument is the number of binders traversed. If [pat] is [None] then then subst is called on [t]. [t] must live in [env] and [sigma], [pat] must have been interpreted in (an extension of) [sigma]. @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) @return [t] where all [occ] occurrences of [pat] have been mapped using [subst] *) val eval_pattern : ?raise_NoMatch:bool -> env -> evar_map -> constr -> pattern option -> occ -> subst -> constr (** [fill_occ_pattern b env sigma t pat occ h] is a simplified version of [eval_pattern]. It replaces all [occ] occurrences of [pat] in [t] with Rel [h]. [t] must live in [env] and [sigma], [pat] must have been interpreted in (an extension of) [sigma]. @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) @return the instance of the redex of [pat] that was matched and [t] transformed as described above. *) val fill_occ_pattern : ?raise_NoMatch:bool -> env -> evar_map -> constr -> pattern -> occ -> int -> constr Evd.in_evar_universe_context * constr (** *************************** Low level APIs ****************************** *) (* The primitive matching facility. It matches of a term with holes, like the T pattern above, and calls a continuation on its occurrences. *) type ssrdir = L2R | R2L val pr_dir_side : ssrdir -> Pp.t (** a pattern for a term with wildcards *) type tpattern (** [mk_tpattern env sigma0 sigma_p ok p_origin dir t] compiles a term [t] living in [env] [sigma] (an extension of [sigma0]) intro a [tpattern]. The [tpattern] can hold a (proof) term [p] and a diction [dir]. The [ok] callback is used to filter occurrences. @return the compiled [tpattern] and its [evar_map] @raise UserEerror is the pattern is a wildcard *) val mk_tpattern : ?p_origin:ssrdir * constr -> env -> evar_map -> evar_map * constr -> (constr -> evar_map -> bool) -> ssrdir -> constr -> evar_map * tpattern (** [findP env t i k] is a stateful function that finds the next occurrence of a tpattern and calls the callback [k] to map the subterm matched. The [int] argument passed to [k] is the number of binders traversed so far plus the initial value [i]. @return [t] where the subterms identified by the selected occurrences of the patter have been mapped using [k] @raise NoMatch if the raise_NoMatch flag given to [mk_tpattern_matcher] is [true] and if the pattern did not match @raise UserEerror if the raise_NoMatch flag given to [mk_tpattern_matcher] is [false] and if the pattern did not match *) type find_P = env -> constr -> int -> k:subst -> constr (** [conclude ()] asserts that all mentioned occurrences have been visited. @return the instance of the pattern, the evarmap after the pattern instantiation, the proof term and the ssrdit stored in the tpattern @raise UserEerror if too many occurrences were specified *) type conclude = unit -> constr * ssrdir * (evar_map * UState.t * constr) (** [mk_tpattern_matcher b o sigma0 occ sigma_tplist] creates a pair a function [find_P] and [conclude] with the behaviour explained above. The flag [b] (default [false]) changes the error reporting behaviour of [find_P] if none of the [tpattern] matches. The argument [o] can be passed to tune the [UserError] eventually raised (useful if the pattern is coming from the LHS/RHS of an equation) *) val mk_tpattern_matcher : ?all_instances:bool -> ?raise_NoMatch:bool -> ?upats_origin:ssrdir * constr -> evar_map -> occ -> evar_map * tpattern list -> find_P * conclude (** Example of [mk_tpattern_matcher] to implement [rewrite \{occ\}\[in t\]rules]. It first matches "in t" (called [pat]), then in all matched subterms it matches the LHS of the rules using [find_R]. [concl0] is the initial goal, [concl] will be the goal where some terms are replaced by a De Bruijn index. The [rw_progress] extra check selects only occurrences that are not rewritten to themselves (e.g. an occurrence "x + x" rewritten with the commutativity law of addition is skipped) {[ let find_R, conclude = match pat with | Some (_, In_T _) -> let aux (sigma, pats) (d, r, lhs, rhs) = let sigma, pat = mk_tpattern env0 sigma0 (sigma, r) (rw_progress rhs) d lhs in sigma, pats @ [pat] in let rpats = List.fold_left aux (r_sigma, []) rules in let find_R, end_R = mk_tpattern_matcher sigma0 occ rpats in find_R ~k:(fun _ _ h -> mkRel h), fun cl -> let rdx, d, r = end_R () in (d,r),rdx | _ -> ... in let concl = eval_pattern env0 sigma0 concl0 pat occ find_R in let (d, r), rdx = conclude concl in ]} *) (* convenience shortcut: [pf_fill_occ_term gl occ (sigma,t)] returns * the conclusion of [gl] where [occ] occurrences of [t] have been replaced * by [Rel 1] and the instance of [t] *) val pf_fill_occ_term : goal sigma -> occ -> evar_map * EConstr.t -> EConstr.t * EConstr.t (* It may be handy to inject a simple term into the first form of cpattern *) val cpattern_of_term : char * glob_constr_and_expr -> Geninterp.interp_sign -> cpattern (** Helpers to make stateful closures. Example: a [find_P] function may be called many times, but the pattern instantiation phase is performed only the first time. The corresponding [conclude] has to return the instantiated pattern redex. Since it is up to [find_P] to raise [NoMatch] if the pattern has no instance, [conclude] considers it an anomaly if the pattern did not match *) (** [do_once r f] calls [f] and updates the ref only once *) val do_once : 'a option ref -> (unit -> 'a) -> unit (** [assert_done r] return the content of r. @raise Anomaly is r is [None] *) val assert_done : 'a option ref -> 'a (** Very low level APIs. these are calls to evarconv's [the_conv_x] followed by [solve_unif_constraints_with_heuristics]. In case of failure they raise [NoMatch] *) val unify_HO : env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map val pf_unify_HO : goal sigma -> EConstr.constr -> EConstr.constr -> goal sigma (** Some more low level functions needed to implement the full SSR language on top of the former APIs *) val tag_of_cpattern : cpattern -> char val loc_of_cpattern : cpattern -> Loc.t option val id_of_pattern : pattern -> Names.Id.t option val is_wildcard : cpattern -> bool val cpattern_of_id : Names.Id.t -> cpattern val pr_constr_pat : env -> evar_map -> constr -> Pp.t val pr_econstr_pat : env -> evar_map -> econstr -> Pp.t val pf_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma val pf_unsafe_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma (* One can also "Set SsrMatchingDebug" from a .v *) val debug : bool -> unit val ssrinstancesof : cpattern -> Tacmach.tactic (** Functions used for grammar extensions. Do not use. *) module Internal : sig val wit_rpatternty : (rpattern, rpattern, rpattern) Genarg.genarg_type val glob_rpattern : Genintern.glob_sign -> rpattern -> rpattern val subst_rpattern : Mod_subst.substitution -> rpattern -> rpattern val interp_rpattern : Geninterp.interp_sign -> Goal.goal Evd.sigma -> rpattern -> Evd.evar_map * rpattern val pr_rpattern : rpattern -> Pp.t val mk_rpattern : (cpattern, cpattern) ssrpattern -> rpattern val mk_lterm : Constrexpr.constr_expr -> Geninterp.interp_sign option -> cpattern val mk_term : char -> Constrexpr.constr_expr -> Geninterp.interp_sign option -> cpattern val glob_cpattern : Genintern.glob_sign -> cpattern -> cpattern val subst_ssrterm : Mod_subst.substitution -> cpattern -> cpattern val interp_ssrterm : Geninterp.interp_sign -> Goal.goal Evd.sigma -> cpattern -> Evd.evar_map * cpattern val pr_ssrterm : cpattern -> Pp.t end (* eof *) coq-8.11.0/plugins/ssrmatching/ssrmatching.ml0000644000175000017500000016002613612664533021152 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ()) let ssr_pp s = Feedback.msg_debug (str"SSR: "++Lazy.force s) let _ = try ignore(Sys.getenv "SSRMATCHINGDEBUG"); pp_ref := ssr_pp with Not_found -> () let debug b = if b then pp_ref := ssr_pp else pp_ref := fun _ -> () let _ = Goptions.declare_bool_option { Goptions.optname = "ssrmatching debugging"; Goptions.optkey = ["Debug";"SsrMatching"]; Goptions.optdepr = false; Goptions.optread = (fun _ -> !pp_ref == ssr_pp); Goptions.optwrite = debug } let pp s = !pp_ref s (** Utils *)(* {{{ *****************************************************************) let env_size env = List.length (Environ.named_context env) let safeDestApp c = match kind c with App (f, a) -> f, a | _ -> c, [| |] (* Toplevel constr must be globalized twice ! *) let glob_constr ist genv sigma t = match t, ist with | (_, Some ce), Some ist -> let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.lfun Id.Set.empty in let ltacvars = { Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in Constrintern.intern_gen WithoutTypeConstraint ~ltacvars:ltacvars genv sigma ce | (rc, None), _ -> rc | (_, Some _), None -> CErrors.anomaly Pp.(str"glob_constr: term with no ist") (* Term printing utilities functions for deciding bracketing. *) let pr_paren prx x = hov 1 (str "(" ++ prx x ++ str ")") (* String lexing utilities *) let skip_wschars s = let rec loop i = match s.[i] with '\n'..' ' -> loop (i + 1) | _ -> i in loop (* We also guard characters that might interfere with the ssreflect *) (* tactic syntax. *) let guard_term ch1 s i = match s.[i] with | '(' -> false | '{' | '/' | '=' -> true | _ -> ch1 = '(' (* The call 'guard s i' should return true if the contents of s *) (* starting at i need bracketing to avoid ambiguities. *) let pr_guarded guard prc c = let s = Pp.string_of_ppcmds (prc c) ^ "$" in if guard s (skip_wschars s 0) then pr_paren prc c else prc c (* More sensible names for constr printers *) let prl_glob_constr c = pr_lglob_constr_env (Global.env ()) c let pr_glob_constr c = pr_glob_constr_env (Global.env ()) c let prl_constr_expr = pr_lconstr_expr let pr_constr_expr = pr_constr_expr let prl_glob_constr_and_expr env sigma = function | _, Some c -> prl_constr_expr env sigma c | c, None -> prl_glob_constr c let pr_glob_constr_and_expr env sigma = function | _, Some c -> pr_constr_expr env sigma c | c, None -> pr_glob_constr c let pr_term (k, c, _) = let env = Global.env () in let sigma = Evd.from_env env in pr_guarded (guard_term k) (pr_glob_constr_and_expr env sigma) c let prl_term (k, c, _) = let env = Global.env () in let sigma = Evd.from_env env in pr_guarded (guard_term k) (prl_glob_constr_and_expr env sigma) c (** Adding a new uninterpreted generic argument type *) let add_genarg tag pr = let wit = Genarg.make0 tag in let tag = Geninterp.Val.create tag in let glob ist x = (ist, x) in let subst _ x = x in let interp ist x = Ftactic.return (Geninterp.Val.Dyn (tag, x)) in let gen_pr env sigma _ _ _ = pr env sigma in let () = Genintern.register_intern0 wit glob in let () = Genintern.register_subst0 wit subst in let () = Geninterp.register_interp0 wit interp in let () = Geninterp.register_val0 wit (Some (Geninterp.Val.Base tag)) in Pptactic.declare_extra_genarg_pprule wit gen_pr gen_pr gen_pr; wit (** Constructors for cast type *) let dC t = CastConv t (** Constructors for constr_expr *) let isCVar = function { CAst.v = CRef (qid,_) } -> qualid_is_ident qid | _ -> false let destCVar = function | { CAst.v = CRef (qid,_) } when qualid_is_ident qid -> qualid_basename qid | _ -> CErrors.anomaly (str"not a CRef.") let isGLambda c = match DAst.get c with GLambda (Name _, _, _, _) -> true | _ -> false let destGLambda c = match DAst.get c with GLambda (Name id, _, _, c) -> (id, c) | _ -> CErrors.anomaly (str "not a GLambda") let isGHole c = match DAst.get c with GHole _ -> true | _ -> false let mkCHole ~loc = CAst.make ?loc @@ CHole (None, IntroAnonymous, None) let mkCLambda ?loc name ty t = CAst.make ?loc @@ CLambdaN ([CLocalAssum([CAst.make ?loc name], Default Explicit, ty)], t) let mkCLetIn ?loc name bo t = CAst.make ?loc @@ CLetIn ((CAst.make ?loc name), bo, None, t) let mkCCast ?loc t ty = CAst.make ?loc @@ CCast (t, dC ty) (** Constructors for rawconstr *) let mkRHole = DAst.make @@ GHole (InternalHole, IntroAnonymous, None) let mkRApp f args = if args = [] then f else DAst.make @@ GApp (f, args) let mkRCast rc rt = DAst.make @@ GCast (rc, dC rt) let mkRLambda n s t = DAst.make @@ GLambda (n, Explicit, s, t) (* ssrterm conbinators *) let combineCG t1 t2 f g = let mk_ist i1 i2 = match i1, i2 with | None, Some i -> Some i | Some i, None -> Some i | None, None -> None | Some i, Some j when i == j -> Some i | _ -> CErrors.anomaly (Pp.str "combineCG: different ist") in match t1, t2 with | (x, (t1, None), i1), (_, (t2, None), i2) -> x, (g t1 t2, None), mk_ist i1 i2 | (x, (_, Some t1), i1), (_, (_, Some t2), i2) -> x, (mkRHole, Some (f t1 t2)), mk_ist i1 i2 | _, (_, (_, None), _) -> CErrors.anomaly (str"have: mixed C-G constr.") | _ -> CErrors.anomaly (str"have: mixed G-C constr.") let loc_ofCG = function | (_, (s, None), _) -> Glob_ops.loc_of_glob_constr s | (_, (_, Some s), _) -> Constrexpr_ops.constr_loc s let mk_term k c ist = k, (mkRHole, Some c), ist let mk_lterm = mk_term ' ' let pf_type_of gl t = let sigma, ty = pf_type_of gl t in re_sig (sig_it gl) sigma, ty let nf_evar sigma c = EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr c)) (* }}} *) exception NoProgress (** Unification procedures. *) (* To enforce the rigidity of the rooted match we always split *) (* top applications, so the unification procedures operate on *) (* arrays of patterns and terms. *) (* We perform three kinds of unification: *) (* EQ: exact conversion check *) (* FO: first-order unification of evars, without conversion *) (* HO: higher-order unification with conversion *) (* The subterm unification strategy is to find the first FO *) (* match, if possible, and the first HO match otherwise, then *) (* compute all the occurrences that are EQ matches for the *) (* relevant subterm. *) (* Additional twists: *) (* - If FO/HO fails then we attempt to fill evars using *) (* typeclasses before raising an outright error. We also *) (* fill typeclasses even after a successful match, since *) (* beta-reduction and canonical instances may leave *) (* undefined evars. *) (* - We do postchecks to rule out matches that are not *) (* closed or that assign to a global evar; these can be *) (* disabled for rewrite or dependent family matches. *) (* - We do a full FO scan before turning to HO, as the FO *) (* comparison can be much faster than the HO one. *) let unif_EQ env sigma p c = let evars = existential_opt_value0 sigma, Evd.universes sigma in try let _ = Reduction.conv env p ~evars c in true with _ -> false let unif_EQ_args env sigma pa a = let n = Array.length pa in let rec loop i = (i = n) || unif_EQ env sigma pa.(i) a.(i) && loop (i + 1) in loop 0 let unif_HO env ise p c = try Evarconv.unify_delay env ise p c with Evarconv.UnableToUnify(ise, err) -> raise Pretype_errors.(PretypeError(env,ise,CannotUnify(p,c,Some err))) let unif_HO_args env ise0 pa i ca = let n = Array.length pa in let rec loop ise j = if j = n then ise else loop (unif_HO env ise (EConstr.of_constr pa.(j)) (EConstr.of_constr ca.(i + j))) (j + 1) in loop ise0 0 (* FO unification should boil down to calling w_unify with no_delta, but *) (* alas things are not so simple: w_unify does partial type-checking, *) (* which breaks down when the no-delta flag is on (as the Coq type system *) (* requires full convertibility. The workaround here is to convert all *) (* evars into metas, since 8.2 does not TC metas. This means some lossage *) (* for HO evars, though hopefully Miller patterns can pick up some of *) (* those cases, and HO matching will mop up the rest. *) let flags_FO env = let oracle = Environ.oracle env in let ts = Conv_oracle.get_transp_state oracle in let flags = { (Unification.default_no_delta_unify_flags ts).Unification.core_unify_flags with Unification.modulo_conv_on_closed_terms = None; Unification.modulo_eta = true; Unification.modulo_betaiota = true; Unification.modulo_delta_types = ts } in { Unification.core_unify_flags = flags; Unification.merge_unify_flags = flags; Unification.subterm_unify_flags = flags; Unification.allow_K_in_toplevel_higher_order_unification = false; Unification.resolve_evars = (Unification.default_no_delta_unify_flags ts).Unification.resolve_evars } let unif_FO env ise p c = Unification.w_unify env ise Reduction.CONV ~flags:(flags_FO env) (EConstr.of_constr p) (EConstr.of_constr c) (* Perform evar substitution in main term and prune substitution. *) let nf_open_term sigma0 ise c = let c = EConstr.Unsafe.to_constr c in let s = ise and s' = ref sigma0 in let rec nf c' = match kind c' with | Evar ex -> begin try nf (existential_value0 s ex) with _ -> let k, a = ex in let a' = Array.map nf a in if not (Evd.mem !s' k) then s' := Evd.add !s' k (Evarutil.nf_evar_info s (Evd.find s k)); mkEvar (k, a') end | _ -> map nf c' in let copy_def k evi () = if evar_body evi != Evd.Evar_empty then () else match Evd.evar_body (Evd.find s k) with | Evar_defined c' -> let c' = EConstr.of_constr (nf (EConstr.Unsafe.to_constr c')) in s' := Evd.define k c' !s' | _ -> () in let c' = nf c in let _ = Evd.fold copy_def sigma0 () in !s', Evd.evar_universe_context s, EConstr.of_constr c' let unif_end ?(solve_TC=true) env sigma0 ise0 pt ok = let ise = Evarconv.solve_unif_constraints_with_heuristics env ise0 in let tcs = Evd.get_typeclass_evars ise in let s, uc, t = nf_open_term sigma0 ise pt in let ise1 = create_evar_defs s in let ise1 = Evd.set_typeclass_evars ise1 (Evar.Set.filter (fun ev -> Evd.is_undefined ise1 ev) tcs) in let ise1 = Evd.set_universe_context ise1 uc in let ise2 = if solve_TC then Typeclasses.resolve_typeclasses ~fail:true env ise1 else ise1 in if not (ok ise) then raise NoProgress else if ise2 == ise1 then (s, uc, t) else let s, uc', t = nf_open_term sigma0 ise2 t in s, UState.union uc uc', t let unify_HO env sigma0 t1 t2 = let sigma = unif_HO env sigma0 t1 t2 in let sigma, uc, _ = unif_end ~solve_TC:false env sigma0 sigma t2 (fun _ -> true) in Evd.set_universe_context sigma uc let pf_unify_HO gl t1 t2 = let env, sigma0, si = pf_env gl, project gl, sig_it gl in let sigma = unify_HO env sigma0 t1 t2 in re_sig si sigma (* This is what the definition of iter_constr should be... *) let iter_constr_LR f c = match kind c with | Evar (k, a) -> Array.iter f a | Cast (cc, _, t) -> f cc; f t | Prod (_, t, b) | Lambda (_, t, b) -> f t; f b | LetIn (_, v, t, b) -> f v; f t; f b | App (cf, a) -> f cf; Array.iter f a | Case (_, p, v, b) -> f v; f p; Array.iter f b | Fix (_, (_, t, b)) | CoFix (_, (_, t, b)) -> for i = 0 to Array.length t - 1 do f t.(i); f b.(i) done | Proj(_,a) -> f a | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> () (* The comparison used to determine which subterms matches is KEYED *) (* CONVERSION. This looks for convertible terms that either have the same *) (* same head constant as pat if pat is an application (after beta-iota), *) (* or start with the same constr constructor (esp. for LetIn); this is *) (* disregarded if the head term is let x := ... in x, and casts are always *) (* ignored and removed). *) (* Record projections get special treatment: in addition to the projection *) (* constant itself, ssreflect also recognizes head constants of canonical *) (* projections. *) exception NoMatch type ssrdir = L2R | R2L let pr_dir_side = function L2R -> str "LHS" | R2L -> str "RHS" let inv_dir = function L2R -> R2L | R2L -> L2R type pattern_class = | KpatFixed | KpatConst | KpatEvar of Evar.t | KpatLet | KpatLam | KpatRigid | KpatFlex | KpatProj of Constant.t type tpattern = { up_k : pattern_class; up_FO : constr; up_f : constr; up_a : constr array; up_t : constr; (* equation proof term or matched term *) up_dir : ssrdir; (* direction of the rule *) up_ok : constr -> evar_map -> bool; (* progress test for rewrite *) } let all_ok _ _ = true let proj_nparams c = try 1 + Recordops.find_projection_nparams (GlobRef.ConstRef c) with _ -> 0 let isRigid c = match kind c with | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true | _ -> false let hole_var = mkVar (Id.of_string "_") let pr_constr_pat env sigma c0 = let rec wipe_evar c = if isEvar c then hole_var else map wipe_evar c in pr_constr_env env sigma (wipe_evar c0) let ehole_var = EConstr.mkVar (Id.of_string "_") let pr_econstr_pat env sigma c0 = let rec wipe_evar c = let open EConstr in if isEvar sigma c then ehole_var else map sigma wipe_evar c in pr_econstr_env env sigma (wipe_evar c0) (* Turn (new) evars into metas *) let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 = let ise = ref ise0 in let sigma = ref ise0 in let nenv = env_size env + if hack then 1 else 0 in let rec put c = match kind c with | Evar (k, a as ex) -> begin try put (existential_value0 !sigma ex) with NotInstantiatedEvar -> if Evd.mem sigma0 k then map put c else let evi = Evd.find !sigma k in let dc = List.firstn (max 0 (Array.length a - nenv)) (evar_filtered_context evi) in let abs_dc (d, c) = function | Context.Named.Declaration.LocalDef (x, b, t) -> d, mkNamedLetIn x (put b) (put t) c | Context.Named.Declaration.LocalAssum (x, t) -> mkVar x.binder_name :: d, mkNamedProd x (put t) c in let a, t = Context.Named.fold_inside abs_dc ~init:([], (put @@ EConstr.Unsafe.to_constr evi.evar_concl)) (EConstr.Unsafe.to_named_context dc) in let m = Evarutil.new_meta () in ise := meta_declare m (EConstr.of_constr t) !ise; sigma := Evd.define k (EConstr.of_constr (applistc (mkMeta m) a)) !sigma; put (existential_value0 !sigma ex) end | _ -> map put c in let c1 = put c0 in !ise, c1 (* Compile a match pattern from a term; t is the term to fill. *) (* p_origin can be passed to obtain a better error message *) let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p = let k, f, a = let f, a = Reductionops.whd_betaiota_stack ise (EConstr.of_constr p) in let f = EConstr.Unsafe.to_constr f in let a = List.map EConstr.Unsafe.to_constr a in match kind f with | Const (p,_) -> let np = proj_nparams p in if np = 0 || np > List.length a then KpatConst, f, a else let a1, a2 = List.chop np a in KpatProj p, (applistc f a1), a2 | Proj (p,arg) -> KpatProj (Projection.constant p), f, a | Var _ | Ind _ | Construct _ -> KpatFixed, f, a | Evar (k, _) -> if Evd.mem sigma0 k then KpatEvar k, f, a else if a <> [] then KpatFlex, f, a else (match p_origin with None -> CErrors.user_err Pp.(str "indeterminate pattern") | Some (dir, rule) -> errorstrm (str "indeterminate " ++ pr_dir_side dir ++ str " in " ++ pr_constr_pat env ise rule)) | LetIn (_, v, _, b) -> if b <> mkRel 1 then KpatLet, f, a else KpatFlex, v, a | Lambda _ -> KpatLam, f, a | _ -> KpatRigid, f, a in let aa = Array.of_list a in let ise', p' = evars_for_FO ~hack env sigma0 ise (mkApp (f, aa)) in ise', { up_k = k; up_FO = p'; up_f = f; up_a = aa; up_ok = ok; up_dir = dir; up_t = t} (* Specialize a pattern after a successful match: assign a precise head *) (* kind and arity for Proj and Flex patterns. *) let ungen_upat lhs (sigma, uc, t) u = let f, a = safeDestApp lhs in let k = match kind f with | Var _ | Ind _ | Construct _ -> KpatFixed | Const _ -> KpatConst | Evar (k, _) -> if is_defined sigma k then raise NoMatch else KpatEvar k | LetIn _ -> KpatLet | Lambda _ -> KpatLam | _ -> KpatRigid in sigma, uc, {u with up_k = k; up_FO = lhs; up_f = f; up_a = a; up_t = t} let nb_cs_proj_args pc f u = let na k = List.length (snd (lookup_canonical_conversion (GlobRef.ConstRef pc, k))).o_TCOMPS in let nargs_of_proj t = match kind t with | App(_,args) -> Array.length args | Proj _ -> 0 (* if splay_app calls expand_projection, this has to be the number of arguments including the projected *) | _ -> assert false in try match kind f with | Prod _ -> na Prod_cs | Sort s -> na (Sort_cs (Sorts.family s)) | Const (c',_) when Constant.equal c' pc -> nargs_of_proj u.up_f | Proj (c',_) when Constant.equal (Projection.constant c') pc -> nargs_of_proj u.up_f | Var _ | Ind _ | Construct _ | Const _ -> na (Const_cs (global_of_constr f)) | _ -> -1 with Not_found -> -1 let isEvar_k k f = match kind f with Evar (k', _) -> k = k' | _ -> false let nb_args c = match kind c with App (_, a) -> Array.length a | _ -> 0 let mkSubArg i a = if i = Array.length a then a else Array.sub a 0 i let mkSubApp f i a = if i = 0 then f else mkApp (f, mkSubArg i a) let splay_app ise = let rec loop c a = match kind c with | App (f, a') -> loop f (Array.append a' a) | Cast (c', _, _) -> loop c' a | Evar ex -> (try loop (existential_value0 ise ex) a with _ -> c, a) | _ -> c, a in fun c -> match kind c with | App (f, a) -> loop f a | Cast _ | Evar _ -> loop c [| |] | _ -> c, [| |] let filter_upat i0 f n u fpats = let na = Array.length u.up_a in if n < na then fpats else let np = match u.up_k with | KpatConst when eq_constr_nounivs u.up_f f -> na | KpatFixed when eq_constr_nounivs u.up_f f -> na | KpatEvar k when isEvar_k k f -> na | KpatLet when isLetIn f -> na | KpatLam when isLambda f -> na | KpatRigid when isRigid f -> na | KpatFlex -> na | KpatProj pc -> let np = na + nb_cs_proj_args pc f u in if n < np then -1 else np | _ -> -1 in if np < na then fpats else let () = if !i0 < np then i0 := n in (u, np) :: fpats let eq_prim_proj c t = match kind t with | Proj(p,_) -> Constant.equal (Projection.constant p) c | _ -> false let filter_upat_FO i0 f n u fpats = let np = nb_args u.up_FO in if n < np then fpats else let ok = match u.up_k with | KpatConst -> eq_constr_nounivs u.up_f f | KpatFixed -> eq_constr_nounivs u.up_f f | KpatEvar k -> isEvar_k k f | KpatLet -> isLetIn f | KpatLam -> isLambda f | KpatRigid -> isRigid f | KpatProj pc -> equal f (mkConst pc) || eq_prim_proj pc f | KpatFlex -> i0 := n; true in if ok then begin if !i0 < np then i0 := np; (u, np) :: fpats end else fpats exception FoundUnif of (evar_map * UState.t * tpattern) (* Note: we don't update env as we descend into the term, as the primitive *) (* unification procedure always rejects subterms with bound variables. *) let dont_impact_evars_in sigma0 cl = let evs_in_cl = Evd.evars_of_term sigma0 cl in fun sigma -> Evar.Set.for_all (fun k -> try let _ = Evd.find_undefined sigma k in true with Not_found -> false) evs_in_cl (* We are forced to duplicate code between the FO/HO matching because we *) (* have to work around several kludges in unify.ml: *) (* - w_unify drops into second-order unification when the pattern is an *) (* application whose head is a meta. *) (* - w_unify tries to unify types without subsumption when the pattern *) (* head is an evar or meta (e.g., it fails on ?1 = nat when ?1 : Type). *) (* - w_unify expands let-in (zeta conversion) eagerly, whereas we want to *) (* match a head let rigidly. *) let match_upats_FO upats env sigma0 ise orig_c = let dont_impact_evars = dont_impact_evars_in sigma0 (EConstr.of_constr orig_c) in let rec loop c = let f, a = splay_app ise c in let i0 = ref (-1) in let fpats = List.fold_right (filter_upat_FO i0 f (Array.length a)) upats [] in while !i0 >= 0 do let i = !i0 in i0 := -1; let c' = mkSubApp f i a in let one_match (u, np) = let skip = if i <= np then i < np else if u.up_k == KpatFlex then begin i0 := i - 1; false end else begin if !i0 < np then i0 := np; true end in if skip || not (closed0 c') then () else try let _ = match u.up_k with | KpatFlex -> let kludge v = mkLambda (make_annot Anonymous Sorts.Relevant, mkProp, v) in unif_FO env ise (kludge u.up_FO) (kludge c') | KpatLet -> let kludge vla = let vl, a = safeDestApp vla in let x, v, t, b = destLetIn vl in mkApp (mkLambda (x, t, b), Array.cons v a) in unif_FO env ise (kludge u.up_FO) (kludge c') | _ -> unif_FO env ise u.up_FO c' in let ise' = (* Unify again using HO to assign evars *) let p = mkApp (u.up_f, u.up_a) in try unif_HO env ise (EConstr.of_constr p) (EConstr.of_constr c') with e when CErrors.noncritical e -> raise NoMatch in let lhs = mkSubApp f i a in let pt' = unif_end env sigma0 ise' (EConstr.of_constr u.up_t) (u.up_ok lhs) in let pt' = pi1 pt', pi2 pt', EConstr.Unsafe.to_constr (pi3 pt') in raise (FoundUnif (ungen_upat lhs pt' u)) with FoundUnif (s,_,_) as sig_u when dont_impact_evars s -> raise sig_u | Not_found -> CErrors.anomaly (str"incomplete ise in match_upats_FO.") | e when CErrors.noncritical e -> () in List.iter one_match fpats done; iter_constr_LR loop f; Array.iter loop a in try loop orig_c with Invalid_argument _ -> CErrors.anomaly (str"IN FO.") let match_upats_HO ~on_instance upats env sigma0 ise c = let dont_impact_evars = dont_impact_evars_in sigma0 (EConstr.of_constr c) in let it_did_match = ref false in let failed_because_of_TC = ref false in let rec aux upats env sigma0 ise c = let f, a = splay_app ise c in let i0 = ref (-1) in let fpats = List.fold_right (filter_upat i0 f (Array.length a)) upats [] in while !i0 >= 0 do let i = !i0 in i0 := -1; let one_match (u, np) = let skip = if i <= np then i < np else if u.up_k == KpatFlex then begin i0 := i - 1; false end else begin if !i0 < np then i0 := np; true end in if skip then () else try let ise' = match u.up_k with | KpatFixed | KpatConst -> ise | KpatEvar _ -> let _, pka = destEvar u.up_f and _, ka = destEvar f in unif_HO_args env ise pka 0 ka | KpatLet -> let x, v, t, b = destLetIn f in let _, pv, _, pb = destLetIn u.up_f in let ise' = unif_HO env ise (EConstr.of_constr pv) (EConstr.of_constr v) in unif_HO (Environ.push_rel (Context.Rel.Declaration.LocalAssum(x, t)) env) ise' (EConstr.of_constr pb) (EConstr.of_constr b) | KpatFlex | KpatProj _ -> unif_HO env ise (EConstr.of_constr u.up_f) (EConstr.of_constr(mkSubApp f (i - Array.length u.up_a) a)) | _ -> unif_HO env ise (EConstr.of_constr u.up_f) (EConstr.of_constr f) in let ise'' = unif_HO_args env ise' u.up_a (i - Array.length u.up_a) a in let lhs = mkSubApp f i a in let pt' = unif_end env sigma0 ise'' (EConstr.of_constr u.up_t) (u.up_ok lhs) in let pt' = pi1 pt', pi2 pt', EConstr.Unsafe.to_constr (pi3 pt') in on_instance (ungen_upat lhs pt' u) with FoundUnif (s,_,_) as sig_u when dont_impact_evars s -> raise sig_u | NoProgress -> it_did_match := true | Pretype_errors.PretypeError (_,_,Pretype_errors.UnsatisfiableConstraints _) -> failed_because_of_TC:=true | e when CErrors.noncritical e -> () in List.iter one_match fpats done; iter_constr_LR (aux upats env sigma0 ise) f; Array.iter (aux upats env sigma0 ise) a in aux upats env sigma0 ise c; if !it_did_match then raise NoProgress; !failed_because_of_TC let fixed_upat evd = function | {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false | {up_t = t} -> not (occur_existential evd (EConstr.of_constr t)) (** FIXME *) let do_once r f = match !r with Some _ -> () | None -> r := Some (f ()) let assert_done r = match !r with Some x -> x | None -> CErrors.anomaly (str"do_once never called.") let assert_done_multires r = match !r with | None -> CErrors.anomaly (str"do_once never called.") | Some (e, n, xs) -> r := Some (e, n+1,xs); try List.nth xs n with Failure _ -> raise NoMatch type subst = Environ.env -> constr -> constr -> int -> constr type find_P = Environ.env -> constr -> int -> k:subst -> constr type conclude = unit -> constr * ssrdir * (Evd.evar_map * UState.t * constr) (* upats_origin makes a better error message only *) let mk_tpattern_matcher ?(all_instances=false) ?(raise_NoMatch=false) ?upats_origin sigma0 occ (ise, upats) = let nocc = ref 0 and skip_occ = ref false in let use_occ, occ_list = match occ with | Some (true, ol) -> ol = [], ol | Some (false, ol) -> ol <> [], ol | None -> false, [] in let max_occ = List.fold_right max occ_list 0 in let subst_occ = let occ_set = Array.make max_occ (not use_occ) in let _ = List.iter (fun i -> occ_set.(i - 1) <- use_occ) occ_list in let _ = if max_occ = 0 then skip_occ := use_occ in fun () -> incr nocc; if !nocc = max_occ then skip_occ := use_occ; if !nocc <= max_occ then occ_set.(!nocc - 1) else not use_occ in let upat_that_matched = ref None in let match_EQ env sigma u = match u.up_k with | KpatLet -> let x, pv, t, pb = destLetIn u.up_f in let env' = Environ.push_rel (Context.Rel.Declaration.LocalAssum(x, t)) env in let match_let f = match kind f with | LetIn (_, v, _, b) -> unif_EQ env sigma pv v && unif_EQ env' sigma pb b | _ -> false in match_let | KpatFixed -> eq_constr_nounivs u.up_f | KpatConst -> eq_constr_nounivs u.up_f | KpatLam -> fun c -> (match kind c with | Lambda _ -> unif_EQ env sigma u.up_f c | _ -> false) | _ -> unif_EQ env sigma u.up_f in let p2t p = mkApp(p.up_f,p.up_a) in let source env = match upats_origin, upats with | None, [p] -> (if fixed_upat ise p then str"term " else str"partial term ") ++ pr_constr_pat env ise (p2t p) ++ spc() | Some (dir,rule), [p] -> str"The " ++ pr_dir_side dir ++ str" of " ++ pr_constr_pat env ise rule ++ fnl() ++ ws 4 ++ pr_constr_pat env ise (p2t p) ++ fnl() | Some (dir,rule), _ -> str"The " ++ pr_dir_side dir ++ str" of " ++ pr_constr_pat env ise rule ++ spc() | _, [] | None, _::_::_ -> CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin.") in let on_instance, instances = let instances = ref [] in (fun x -> if all_instances then instances := !instances @ [x] else raise (FoundUnif x)), (fun () -> !instances) in let rec uniquize = function | [] -> [] | (sigma,_,{ up_f = f; up_a = a; up_t = t } as x) :: xs -> let t = nf_evar sigma t in let f = nf_evar sigma f in let a = Array.map (nf_evar sigma) a in let neq (sigma1,_,{ up_f = f1; up_a = a1; up_t = t1 }) = let t1 = nf_evar sigma1 t1 in let f1 = nf_evar sigma1 f1 in let a1 = Array.map (nf_evar sigma1) a1 in not (equal t t1 && equal f f1 && CArray.for_all2 equal a a1) in x :: uniquize (List.filter neq xs) in ((fun env c h ~k -> do_once upat_that_matched (fun () -> let failed_because_of_TC = ref false in try if not all_instances then match_upats_FO upats env sigma0 ise c; failed_because_of_TC:=match_upats_HO ~on_instance upats env sigma0 ise c; raise NoMatch with FoundUnif sigma_u -> env,0,[sigma_u] | (NoMatch|NoProgress) when all_instances && instances () <> [] -> env, 0, uniquize (instances ()) | NoMatch when (not raise_NoMatch) -> if !failed_because_of_TC then errorstrm (source env ++ strbrk"matches but type classes inference fails") else errorstrm (source env ++ str "does not match any subterm of the goal") | NoProgress when (not raise_NoMatch) -> let dir = match upats_origin with Some (d,_) -> d | _ -> CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin.") in errorstrm (str"all matches of "++ source env ++ str"are equal to the " ++ pr_dir_side (inv_dir dir)) | NoProgress -> raise NoMatch); let sigma, _, ({up_f = pf; up_a = pa} as u) = if all_instances then assert_done_multires upat_that_matched else List.hd (pi3(assert_done upat_that_matched)) in (* pp(lazy(str"sigma@tmatch=" ++ pr_evar_map None sigma)); *) if !skip_occ then ((*ignore(k env u.up_t 0);*) c) else let match_EQ = match_EQ env sigma u in let pn = Array.length pa in let rec subst_loop (env,h as acc) c' = if !skip_occ then c' else let f, a = splay_app sigma c' in if Array.length a >= pn && match_EQ f && unif_EQ_args env sigma pa a then let a1, a2 = Array.chop (Array.length pa) a in let fa1 = mkApp (f, a1) in let f' = if subst_occ () then k env u.up_t fa1 h else fa1 in mkApp (f', Array.map_left (subst_loop acc) a2) else (* TASSI: clear letin values to avoid unfolding *) let inc_h rd (env,h') = let ctx_item = match rd with | Context.Rel.Declaration.LocalAssum _ as x -> x | Context.Rel.Declaration.LocalDef (x,_,y) -> Context.Rel.Declaration.LocalAssum(x,y) in EConstr.push_rel ctx_item env, h' + 1 in let self acc c = EConstr.of_constr (subst_loop acc (EConstr.Unsafe.to_constr c)) in let f = EConstr.of_constr f in let f' = map_constr_with_binders_left_to_right sigma inc_h self acc f in let f' = EConstr.Unsafe.to_constr f' in mkApp (f', Array.map_left (subst_loop acc) a) in subst_loop (env,h) c) : find_P), ((fun () -> let env, (sigma, uc, ({up_f = pf; up_a = pa} as u)) = match !upat_that_matched with | Some (env,_,x) -> env,List.hd x | None when raise_NoMatch -> raise NoMatch | None -> CErrors.anomaly (str"companion function never called.") in let p' = mkApp (pf, pa) in if max_occ <= !nocc then p', u.up_dir, (sigma, uc, u.up_t) else errorstrm (str"Only " ++ int !nocc ++ str" < " ++ int max_occ ++ str(String.plural !nocc " occurrence") ++ match upats_origin with | None -> str" of" ++ spc() ++ pr_constr_pat env sigma p' | Some (dir,rule) -> str" of the " ++ pr_dir_side dir ++ fnl() ++ ws 4 ++ pr_constr_pat env sigma p' ++ fnl () ++ str"of " ++ pr_constr_pat env sigma rule)) : conclude) type ('ident, 'term) ssrpattern = | T of 'term | In_T of 'term | X_In_T of 'ident * 'term | In_X_In_T of 'ident * 'term | E_In_X_In_T of 'term * 'ident * 'term | E_As_X_In_T of 'term * 'ident * 'term let pr_pattern = function | T t -> prl_term t | In_T t -> str "in " ++ prl_term t | X_In_T (x,t) -> prl_term x ++ str " in " ++ prl_term t | In_X_In_T (x,t) -> str "in " ++ prl_term x ++ str " in " ++ prl_term t | E_In_X_In_T (e,x,t) -> prl_term e ++ str " in " ++ prl_term x ++ str " in " ++ prl_term t | E_As_X_In_T (e,x,t) -> prl_term e ++ str " as " ++ prl_term x ++ str " in " ++ prl_term t let pr_pattern_w_ids = function | T t -> prl_term t | In_T t -> str "in " ++ prl_term t | X_In_T (x,t) -> pr_id x ++ str " in " ++ prl_term t | In_X_In_T (x,t) -> str "in " ++ pr_id x ++ str " in " ++ prl_term t | E_In_X_In_T (e,x,t) -> prl_term e ++ str " in " ++ pr_id x ++ str " in " ++ prl_term t | E_As_X_In_T (e,x,t) -> prl_term e ++ str " as " ++ pr_id x ++ str " in " ++ prl_term t let pr_pattern_aux pr_constr = function | T t -> pr_constr t | In_T t -> str "in " ++ pr_constr t | X_In_T (x,t) -> pr_constr x ++ str " in " ++ pr_constr t | In_X_In_T (x,t) -> str "in " ++ pr_constr x ++ str " in " ++ pr_constr t | E_In_X_In_T (e,x,t) -> pr_constr e ++ str " in " ++ pr_constr x ++ str " in " ++ pr_constr t | E_As_X_In_T (e,x,t) -> pr_constr e ++ str " as " ++ pr_constr x ++ str " in " ++ pr_constr t let pp_pattern env (sigma, p) = pr_pattern_aux (fun t -> pr_econstr_pat env sigma (pi3 (nf_open_term sigma sigma (EConstr.of_constr t)))) p let pr_cpattern = pr_term let wit_rpatternty = add_genarg "rpatternty" (fun env sigma -> pr_pattern) let glob_ssrterm gs = function | k, (_, Some c), None -> let x = Tacintern.intern_constr gs c in k, (fst x, Some c), None | ct -> ct (* This piece of code asserts the following notations are reserved *) (* Reserved Notation "( a 'in' b )" (at level 0). *) (* Reserved Notation "( a 'as' b )" (at level 0). *) (* Reserved Notation "( a 'in' b 'in' c )" (at level 0). *) (* Reserved Notation "( a 'as' b 'in' c )" (at level 0). *) let glob_cpattern gs p = pp(lazy(str"globbing pattern: " ++ pr_term p)); let glob x = pi2 (glob_ssrterm gs (mk_lterm x None)) in let encode k s l = let name = Name (Id.of_string ("_ssrpat_" ^ s)) in k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None), None in let bind_in t1 t2 = let mkCHole = mkCHole ~loc:None in let n = Name (destCVar t1) in fst (glob (mkCCast mkCHole (mkCLambda n mkCHole t2))) in let check_var t2 = if not (isCVar t2) then loc_error (constr_loc t2) "Only identifiers are allowed here" in match p with | _, (_, None), _ as x -> x | k, (v, Some t), _ as orig -> if k = 'x' then glob_ssrterm gs ('(', (v, Some t), None) else match t.CAst.v with | CNotation((InConstrEntrySomeLevel,"( _ in _ )"), ([t1; t2], [], [], [])) -> (try match glob t1, glob t2 with | (r1, None), (r2, None) -> encode k "In" [r1;r2] | (r1, Some _), (r2, Some _) when isCVar t1 -> encode k "In" [r1; r2; bind_in t1 t2] | (r1, Some _), (r2, Some _) -> encode k "In" [r1; r2] | _ -> CErrors.anomaly (str"where are we?.") with _ when isCVar t1 -> encode k "In" [bind_in t1 t2]) | CNotation((InConstrEntrySomeLevel,"( _ in _ in _ )"), ([t1; t2; t3], [], [], [])) -> check_var t2; encode k "In" [fst (glob t1); bind_in t2 t3] | CNotation((InConstrEntrySomeLevel,"( _ as _ )"), ([t1; t2], [], [], [])) -> encode k "As" [fst (glob t1); fst (glob t2)] | CNotation((InConstrEntrySomeLevel,"( _ as _ in _ )"), ([t1; t2; t3], [], [], [])) -> check_var t2; encode k "As" [fst (glob t1); bind_in t2 t3] | _ -> glob_ssrterm gs orig ;; let glob_rpattern s p = match p with | T t -> T (glob_cpattern s t) | In_T t -> In_T (glob_ssrterm s t) | X_In_T(x,t) -> X_In_T (x,glob_ssrterm s t) | In_X_In_T(x,t) -> In_X_In_T (x,glob_ssrterm s t) | E_In_X_In_T(e,x,t) -> E_In_X_In_T (glob_ssrterm s e,x,glob_ssrterm s t) | E_As_X_In_T(e,x,t) -> E_As_X_In_T (glob_ssrterm s e,x,glob_ssrterm s t) let subst_ssrterm s (k, c, ist) = k, Tacsubst.subst_glob_constr_and_expr s c, ist let subst_rpattern s = function | T t -> T (subst_ssrterm s t) | In_T t -> In_T (subst_ssrterm s t) | X_In_T(x,t) -> X_In_T (x,subst_ssrterm s t) | In_X_In_T(x,t) -> In_X_In_T (x,subst_ssrterm s t) | E_In_X_In_T(e,x,t) -> E_In_X_In_T (subst_ssrterm s e,x,subst_ssrterm s t) | E_As_X_In_T(e,x,t) -> E_As_X_In_T (subst_ssrterm s e,x,subst_ssrterm s t) let interp_ssrterm ist (k,t,_) = k, t, Some ist let interp_rpattern s = function | T t -> T (interp_ssrterm s t) | In_T t -> In_T (interp_ssrterm s t) | X_In_T(x,t) -> X_In_T (interp_ssrterm s x,interp_ssrterm s t) | In_X_In_T(x,t) -> In_X_In_T (interp_ssrterm s x,interp_ssrterm s t) | E_In_X_In_T(e,x,t) -> E_In_X_In_T (interp_ssrterm s e,interp_ssrterm s x,interp_ssrterm s t) | E_As_X_In_T(e,x,t) -> E_As_X_In_T (interp_ssrterm s e,interp_ssrterm s x,interp_ssrterm s t) let interp_rpattern0 ist gl t = Tacmach.project gl, interp_rpattern ist t type cpattern = char * Genintern.glob_constr_and_expr * Geninterp.interp_sign option let tag_of_cpattern = pi1 let loc_of_cpattern = loc_ofCG let cpattern_of_term (c, t) ist = c, t, Some ist type occ = (bool * int list) option type rpattern = (cpattern, cpattern) ssrpattern type pattern = Evd.evar_map * (constr, constr) ssrpattern let id_of_cpattern (_, (c1, c2), _) = let open CAst in match DAst.get c1, c2 with | _, Some { v = CRef (qid, _) } when qualid_is_ident qid -> Some (qualid_basename qid) | _, Some { v = CAppExpl ((_, qid, _), []) } when qualid_is_ident qid -> Some (qualid_basename qid) | GRef (GlobRef.VarRef x, _), None -> Some x | _ -> None let id_of_Cterm t = match id_of_cpattern t with | Some x -> x | None -> loc_error (loc_of_cpattern t) "Only identifiers are allowed here" let of_ftactic ftac gl = let r = ref None in let tac = Ftactic.run ftac (fun ans -> r := Some ans; Proofview.tclUNIT ()) in let tac = Proofview.V82.of_tactic tac in let { sigma = sigma } = tac gl in let ans = match !r with | None -> assert false (* If the tactic failed we should not reach this point *) | Some ans -> ans in (sigma, ans) let interp_wit wit ist gl x = let globarg = in_gen (glbwit wit) x in let arg = interp_genarg ist globarg in let (sigma, arg) = of_ftactic arg gl in sigma, Value.cast (topwit wit) arg let interp_open_constr ist gl gc = interp_wit wit_open_constr ist gl gc let pf_intern_term gl (_, c, ist) = glob_constr ist (pf_env gl) (project gl) c let interp_ssrterm ist gl t = Tacmach.project gl, interp_ssrterm ist t let interp_term gl = function | (_, c, Some ist) -> on_snd EConstr.Unsafe.to_constr (interp_open_constr ist gl c) | _ -> errorstrm (str"interpreting a term with no ist") let thin id sigma goal = let ids = Id.Set.singleton id in let env = Goal.V82.env sigma goal in let cl = Goal.V82.concl sigma goal in let sigma = Evd.clear_metas sigma in let ans = try Some (Evarutil.clear_hyps_in_evi env sigma (Environ.named_context_val env) cl ids) with Evarutil.ClearDependencyError _ -> None in match ans with | None -> sigma | Some (sigma, hyps, concl) -> let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl in let sigma = Goal.V82.partial_solution_to env sigma goal gl ev in sigma (* let pr_ist { lfun= lfun } = prlist_with_sep spc (fun (id, Geninterp.Val.Dyn(ty,_)) -> pr_id id ++ str":" ++ Geninterp.Val.pr ty) (Id.Map.bindings lfun) *) let interp_pattern ?wit_ssrpatternarg gl red redty = pp(lazy(str"interpreting: " ++ pr_pattern red)); let xInT x y = X_In_T(x,y) and inXInT x y = In_X_In_T(x,y) in let inT x = In_T x and eInXInT e x t = E_In_X_In_T(e,x,t) in let eAsXInT e x t = E_As_X_In_T(e,x,t) in let mkG ?(k=' ') x ist = k,(x,None), ist in let ist_of (_,_,ist) = ist in let decode (_,_,ist as t) ?reccall f g = try match DAst.get (pf_intern_term gl t) with | GCast(t,CastConv c) when isGHole t && isGLambda c-> let (x, c) = destGLambda c in f x (' ',(c,None),ist) | GVar id when Option.has_some ist && let ist = Option.get ist in Id.Map.mem id ist.lfun && not(Option.is_empty reccall) && not(Option.is_empty wit_ssrpatternarg) -> let v = Id.Map.find id (Option.get ist).lfun in Option.get reccall (Value.cast (topwit (Option.get wit_ssrpatternarg)) v) | it -> g t with e when CErrors.noncritical e -> g t in let decodeG ist t f g = decode (mkG t ist) f g in let bad_enc id _ = CErrors.anomaly (str"bad encoding for pattern "++str id++str".") in let cleanup_XinE h x rp sigma = let h_k = match kind h with Evar (k,_) -> k | _ -> assert false in let to_clean, update = (* handle rename if x is already used *) let ctx = pf_hyps gl in let len = Context.Named.length ctx in let name = ref None in try ignore(Context.Named.lookup x ctx); (name, fun k -> if !name = None then let nctx = Evd.evar_context (Evd.find sigma k) in let nlen = Context.Named.length nctx in if nlen > len then begin name := Some (Context.Named.Declaration.get_id (List.nth nctx (nlen - len - 1))) end) with Not_found -> ref (Some x), fun _ -> () in let sigma0 = project gl in let new_evars = let rec aux acc t = match kind t with | Evar (k,_) -> if k = h_k || List.mem k acc || Evd.mem sigma0 k then acc else (update k; k::acc) | _ -> CoqConstr.fold aux acc t in aux [] (nf_evar sigma rp) in let sigma = List.fold_left (fun sigma e -> if Evd.is_defined sigma e then sigma else (* clear may be recursive *) if Option.is_empty !to_clean then sigma else let name = Option.get !to_clean in pp(lazy(pr_id name)); thin name sigma e) sigma new_evars in sigma in let red = let rec decode_red = function | T(k,(t,None),ist) -> begin match DAst.get t with | GCast (c,CastConv t) when isGHole c && let (id, t) = destGLambda t in let id = Id.to_string id in let len = String.length id in (len > 8 && String.sub id 0 8 = "_ssrpat_") -> let (id, t) = destGLambda t in let id = Id.to_string id in let len = String.length id in (match String.sub id 8 (len - 8), DAst.get t with | "In", GApp( _, [t]) -> decodeG ist t xInT (fun x -> T x) | "In", GApp( _, [e; t]) -> decodeG ist t (eInXInT (mkG e ist)) (bad_enc id) | "In", GApp( _, [e; t; e_in_t]) -> decodeG ist t (eInXInT (mkG e ist)) (fun _ -> decodeG ist e_in_t xInT (fun _ -> assert false)) | "As", GApp(_, [e; t]) -> decodeG ist t (eAsXInT (mkG e ist)) (bad_enc id) | _ -> bad_enc id ()) | _ -> decode ~reccall:decode_red (mkG ~k t ist) xInT (fun x -> T x) end | T t -> decode ~reccall:decode_red t xInT (fun x -> T x) | In_T t -> decode t inXInT inT | X_In_T (e,t) -> decode t (eInXInT e) (fun x -> xInT (id_of_Cterm e) x) | In_X_In_T (e,t) -> inXInT (id_of_Cterm e) t | E_In_X_In_T (e,x,rp) -> eInXInT e (id_of_Cterm x) rp | E_As_X_In_T (e,x,rp) -> eAsXInT e (id_of_Cterm x) rp in decode_red red in pp(lazy(str"decoded as: " ++ pr_pattern_w_ids red)); let red = match redty with | None -> red | Some (ty, ist) -> let ty = ' ', ty, Some ist in match red with | T t -> T (combineCG t ty (mkCCast ?loc:(loc_ofCG t)) mkRCast) | X_In_T (x,t) -> let gty = pf_intern_term gl ty in E_As_X_In_T (mkG (mkRCast mkRHole gty) (ist_of ty), x, t) | E_In_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term gl ty) (ist_of ty) in E_In_X_In_T (combineCG e ty (mkCCast ?loc:(loc_ofCG t)) mkRCast, x, t) | E_As_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term gl ty) (ist_of ty) in E_As_X_In_T (combineCG e ty (mkCCast ?loc:(loc_ofCG t)) mkRCast, x, t) | red -> red in pp(lazy(str"typed as: " ++ pr_pattern_w_ids red)); let mkXLetIn ?loc x (a,(g,c),ist) = match c with | Some b -> a,(g,Some (mkCLetIn ?loc x (mkCHole ~loc) b)), ist | None -> a,(DAst.make ?loc @@ GLetIn (x, DAst.make ?loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None), ist in match red with | T t -> let sigma, t = interp_term gl t in sigma, T t | In_T t -> let sigma, t = interp_term gl t in sigma, In_T t | X_In_T (x, rp) | In_X_In_T (x, rp) -> let mk x p = match red with X_In_T _ -> X_In_T(x,p) | _ -> In_X_In_T(x,p) in let rp = mkXLetIn (Name x) rp in let sigma, rp = interp_term gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in let rp = subst1 h (nf_evar sigma rp) in sigma, mk h rp | E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) -> let mk e x p = match red with E_In_X_In_T _ ->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p) in let rp = mkXLetIn (Name x) rp in let sigma, rp = interp_term gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in let rp = subst1 h (nf_evar sigma rp) in let sigma, e = interp_term (re_sig (sig_it gl) sigma) e in sigma, mk e h rp ;; let interp_cpattern gl red redty = interp_pattern gl (T red) redty;; let interp_rpattern ~wit_ssrpatternarg gl red = interp_pattern ~wit_ssrpatternarg gl red None;; let id_of_pattern = function | _, T t -> (match kind t with Var id -> Some id | _ -> None) | _ -> None (* The full occurrence set *) let noindex = Some(false,[]) (* calls do_subst on every sub-term identified by (pattern,occ) *) let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let fs sigma x = nf_evar sigma x in let pop_evar sigma e p = let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in let e_body = match e_body with Evar_defined c -> EConstr.Unsafe.to_constr c | _ -> errorstrm (str "Matching the pattern " ++ pr_constr_env env0 sigma0 p ++ str " did not instantiate ?" ++ int (Evar.repr e) ++ spc () ++ str "Does the variable bound by the \"in\" construct occur "++ str "in the pattern?") in let sigma = Evd.add (Evd.remove sigma e) e {e_def with Evd.evar_body = Evar_empty} in sigma, e_body in let ex_value hole = match kind hole with Evar (e,_) -> e | _ -> assert false in let mk_upat_for ?hack env sigma0 (sigma, t) ?(p=t) ok = let sigma,pat= mk_tpattern ?hack env sigma0 (sigma,p) ok L2R (fs sigma t) in sigma, [pat] in match pattern with | None -> do_subst env0 concl0 concl0 1, UState.empty | Some (sigma, (T rp | In_T rp)) -> let rp = fs sigma rp in let ise = create_evar_defs sigma in let occ = match pattern with Some (_, T _) -> occ | _ -> noindex in let rp = mk_upat_for env0 sigma0 (ise, rp) all_ok in let find_T, end_T = mk_tpattern_matcher ?raise_NoMatch sigma0 occ rp in let concl = find_T env0 concl0 1 ~k:do_subst in let _, _, (_, us, _) = end_T () in concl, us | Some (sigma, (X_In_T (hole, p) | In_X_In_T (hole, p))) -> let p = fs sigma p in let occ = match pattern with Some (_, X_In_T _) -> occ | _ -> noindex in let ex = ex_value hole in let rp = mk_upat_for ~hack:true env0 sigma0 (sigma, p) all_ok in let find_T, end_T = mk_tpattern_matcher sigma0 noindex rp in (* we start from sigma, so hole is considered a rigid head *) let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher ?raise_NoMatch sigma occ holep in let concl = find_T env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h ~k:(fun env _ -> do_subst env e_body))) in let _ = end_X () in let _, _, (_, us, _) = end_T () in concl, us | Some (sigma, E_In_X_In_T (e, hole, p)) -> let p, e = fs sigma p, fs sigma e in let ex = ex_value hole in let rp = mk_upat_for ~hack:true env0 sigma0 (sigma, p) all_ok in let find_T, end_T = mk_tpattern_matcher sigma0 noindex rp in let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher sigma noindex holep in let re = mk_upat_for env0 sigma0 (sigma, e) all_ok in let find_E, end_E = mk_tpattern_matcher ?raise_NoMatch sigma0 occ re in let concl = find_T env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h ~k:(fun env c _ h -> find_E env e_body h ~k:do_subst))) in let _,_,(_,us,_) = end_E () in let _ = end_X () in let _ = end_T () in concl, us | Some (sigma, E_As_X_In_T (e, hole, p)) -> let p, e = fs sigma p, fs sigma e in let ex = ex_value hole in let rp = let e_sigma = unify_HO env0 sigma (EConstr.of_constr hole) (EConstr.of_constr e) in e_sigma, fs e_sigma p in let rp = mk_upat_for ~hack:true env0 sigma0 rp all_ok in let find_TE, end_TE = mk_tpattern_matcher sigma0 noindex rp in let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher sigma occ holep in let concl = find_TE env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h ~k:(fun env c _ h -> let e_sigma = unify_HO env sigma (EConstr.of_constr e_body) (EConstr.of_constr e) in let e_body = fs e_sigma e in do_subst env e_body e_body h))) in let _ = end_X () in let _,_,(_,us,_) = end_TE () in concl, us ;; let redex_of_pattern ?(resolve_typeclasses=false) env (sigma, p) = let e = match p with | In_T _ | In_X_In_T _ -> CErrors.anomaly (str"pattern without redex.") | T e | X_In_T (e, _) | E_As_X_In_T (e, _, _) | E_In_X_In_T (e, _, _) -> e in let sigma = if not resolve_typeclasses then sigma else Typeclasses.resolve_typeclasses ~fail:false env sigma in nf_evar sigma e, Evd.evar_universe_context sigma let fill_occ_pattern ?raise_NoMatch env sigma cl pat occ h = let do_make_rel, occ = if occ = Some(true,[]) then false, Some(false,[1]) else true, occ in let find_R, conclude = let r = ref None in (fun env c _ h' -> do_once r (fun () -> c); if do_make_rel then mkRel (h'+h-1) else c), (fun _ -> if !r = None then fst(redex_of_pattern env pat) else assert_done r) in let cl, us = eval_pattern ?raise_NoMatch env sigma cl (Some pat) occ find_R in let e = conclude cl in (e, us), cl ;; (* clenup interface for external use *) let mk_tpattern ?p_origin env sigma0 sigma_t f dir c = mk_tpattern ?p_origin env sigma0 sigma_t f dir c ;; let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = fst (eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst) ;; let pf_fill_occ env concl occ sigma0 p (sigma, t) ok h = let p = EConstr.Unsafe.to_constr p in let concl = EConstr.Unsafe.to_constr concl in let ise = create_evar_defs sigma in let ise, u = mk_tpattern env sigma0 (ise,EConstr.Unsafe.to_constr t) ok L2R p in let find_U, end_U = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in let concl = find_U env concl h ~k:(fun _ _ _ -> mkRel) in let rdx, _, (sigma, uc, p) = end_U () in sigma, uc, EConstr.of_constr p, EConstr.of_constr concl, EConstr.of_constr rdx let fill_occ_term env cl occ sigma0 (sigma, t) = try let sigma',uc,t',cl,_= pf_fill_occ env cl occ sigma0 t (sigma, t) all_ok 1 in if sigma' != sigma0 then CErrors.user_err Pp.(str "matching impacts evars") else cl, (Evd.merge_universe_context sigma' uc, t') with NoMatch -> try let sigma', uc, t' = unif_end env sigma0 (create_evar_defs sigma) t (fun _ -> true) in if sigma' != sigma0 then raise NoMatch else cl, (Evd.merge_universe_context sigma' uc, t') with _ -> errorstrm (str "partial term " ++ pr_econstr_pat env sigma t ++ str " does not match any subterm of the goal") let pf_fill_occ_term gl occ t = let sigma0 = project gl and env = pf_env gl and concl = pf_concl gl in let cl,(_,t) = fill_occ_term env concl occ sigma0 t in cl, t let cpattern_of_id id = ' ', (DAst.make @@ GRef (GlobRef.VarRef id, None), None), Some Geninterp.({ lfun = Id.Map.empty; poly = false; extra = Tacinterp.TacStore.empty }) let is_wildcard ((_, (l, r), _) : cpattern) : bool = match DAst.get l, r with | _, Some { CAst.v = CHole _ } | GHole _, None -> true | _ -> false (* "ssrpattern" *) let pr_rpattern = pr_pattern let pf_merge_uc uc gl = re_sig (sig_it gl) (Evd.merge_universe_context (project gl) uc) let pf_unsafe_merge_uc uc gl = re_sig (sig_it gl) (Evd.set_universe_context (project gl) uc) (** All the pattern types reuse the same dynamic toplevel tag *) let wit_ssrpatternarg = wit_rpatternty let interp_rpattern = interp_rpattern ~wit_ssrpatternarg let ssrpatterntac _ist arg gl = let pat = interp_rpattern gl arg in let sigma0 = project gl in let concl0 = pf_concl gl in let concl0 = EConstr.Unsafe.to_constr concl0 in let (t, uc), concl_x = fill_occ_pattern (pf_env gl) sigma0 concl0 pat noindex 1 in let t = EConstr.of_constr t in let concl_x = EConstr.of_constr concl_x in let gl, tty = pf_type_of gl t in let concl = EConstr.mkLetIn (make_annot (Name (Id.of_string "selected")) Sorts.Relevant, t, tty, concl_x) in Proofview.V82.of_tactic (convert_concl ~check:true concl DEFAULTcast) gl (* Register "ssrpattern" tactic *) let () = let mltac _ ist = let arg = let v = Id.Map.find (Names.Id.of_string "pattern") ist.lfun in Value.cast (topwit wit_ssrpatternarg) v in Proofview.V82.tactic (ssrpatterntac ist arg) in let name = { mltac_plugin = "ssrmatching_plugin"; mltac_tactic = "ssrpattern"; } in let () = Tacenv.register_ml_tactic name [|mltac|] in let tac = TacFun ([Name (Id.of_string "pattern")], TacML (CAst.make ({ mltac_name = name; mltac_index = 0 }, []))) in let obj () = Tacenv.register_ltac true false (Id.of_string "ssrpattern") tac in Mltop.declare_cache_obj obj "ssrmatching_plugin" let ssrinstancesof arg gl = let ok rhs lhs ise = true in (* not (equal lhs (Evarutil.nf_evar ise rhs)) in *) let env, sigma, concl = pf_env gl, project gl, pf_concl gl in let concl = EConstr.Unsafe.to_constr concl in let sigma0, cpat = interp_cpattern gl arg None in let pat = match cpat with T x -> x | _ -> errorstrm (str"Not supported") in let etpat, tpat = mk_tpattern env sigma (sigma0,pat) (ok pat) L2R pat in let find, conclude = mk_tpattern_matcher ~all_instances:true ~raise_NoMatch:true sigma None (etpat,[tpat]) in let print env p c _ = ppnl (hov 1 (str"instance:" ++ spc() ++ pr_constr_env (pf_env gl) (gl.sigma) p ++ spc() ++ str "matches:" ++ spc() ++ pr_constr_env (pf_env gl) (gl.sigma) c)); c in ppnl (str"BEGIN INSTANCES"); try while true do ignore(find env concl 1 ~k:print) done; raise NoMatch with NoMatch -> ppnl (str"END INSTANCES"); tclIDTAC gl module Internal = struct let wit_rpatternty = wit_rpatternty let glob_rpattern = glob_rpattern let subst_rpattern = subst_rpattern let interp_rpattern = interp_rpattern0 let pr_rpattern = pr_rpattern let mk_rpattern x = x let mk_lterm = mk_lterm let mk_term = mk_term let glob_cpattern = glob_cpattern let subst_ssrterm = subst_ssrterm let interp_ssrterm = interp_ssrterm let pr_ssrterm = pr_term end (* vim: set filetype=ocaml foldmethod=marker: *) coq-8.11.0/plugins/ssrmatching/ssrmatching.v0000644000175000017500000000324213612664533021003 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t) : ssrpatternscope. (* Some shortcuts for recurrent "X in t" parts. *) Notation RHS := (X in _ = X)%pattern. Notation LHS := (X in X = _)%pattern. End SsrMatchingSyntax. Export SsrMatchingSyntax. Tactic Notation "ssrpattern" ssrpatternarg(p) := ssrpattern p . coq-8.11.0/plugins/ssrmatching/g_ssrmatching.mlg0000644000175000017500000000740613612664533021631 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { mk_rpattern (T (mk_lterm c None)) } | [ "in" lconstr(c) ] -> { mk_rpattern (In_T (mk_lterm c None)) } | [ lconstr(x) "in" lconstr(c) ] -> { mk_rpattern (X_In_T (mk_lterm x None, mk_lterm c None)) } | [ "in" lconstr(x) "in" lconstr(c) ] -> { mk_rpattern (In_X_In_T (mk_lterm x None, mk_lterm c None)) } | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] -> { mk_rpattern (E_In_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) } | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] -> { mk_rpattern (E_As_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) } END { let pr_ssrterm _ _ _ = pr_ssrterm } ARGUMENT EXTEND cpattern PRINTED BY { pr_ssrterm } INTERPRETED BY { interp_ssrterm } GLOBALIZED BY { glob_cpattern } SUBSTITUTED BY { subst_ssrterm } RAW_PRINTED BY { pr_ssrterm } GLOB_PRINTED BY { pr_ssrterm } | [ "Qed" constr(c) ] -> { mk_lterm c None } END { let input_ssrtermkind _ strm = match Util.stream_nth 0 strm with | Tok.KEYWORD "(" -> '(' | Tok.KEYWORD "@" -> '@' | _ -> ' ' let ssrtermkind = Pcoq.Entry.of_parser "ssrtermkind" input_ssrtermkind } GRAMMAR EXTEND Gram GLOBAL: cpattern; cpattern: [[ k = ssrtermkind; c = constr -> { let pattern = mk_term k c None in if loc_of_cpattern pattern <> Some loc && k = '(' then mk_term 'x' c None else pattern } ]]; END ARGUMENT EXTEND lcpattern TYPED AS cpattern PRINTED BY { pr_ssrterm } INTERPRETED BY { interp_ssrterm } GLOBALIZED BY { glob_cpattern } SUBSTITUTED BY { subst_ssrterm } RAW_PRINTED BY { pr_ssrterm } GLOB_PRINTED BY { pr_ssrterm } | [ "Qed" lconstr(c) ] -> { mk_lterm c None } END GRAMMAR EXTEND Gram GLOBAL: lcpattern; lcpattern: [[ k = ssrtermkind; c = lconstr -> { let pattern = mk_term k c None in if loc_of_cpattern pattern <> Some loc && k = '(' then mk_term 'x' c None else pattern } ]]; END ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY { pr_rpattern } | [ rpattern(pat) ] -> { pat } END TACTIC EXTEND ssrinstoftpat | [ "ssrinstancesoftpat" cpattern(arg) ] -> { Proofview.V82.tactic (ssrinstancesof arg) } END { (* We wipe out all the keywords generated by the grammar rules we defined. *) (* The user is supposed to Require Import ssreflect or Require ssreflect *) (* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) (* consequence the extended ssreflect grammar. *) let () = CLexer.set_keyword_state frozen_lexer ;; } coq-8.11.0/plugins/ssrmatching/ssrmatching_plugin.mlpack0000644000175000017500000000003213612664533023355 0ustar treinentreinenSsrmatching G_ssrmatching coq-8.11.0/plugins/ssrmatching/plugin_base.dune0000644000175000017500000000021713612664533021436 0ustar treinentreinen(library (name ssrmatching_plugin) (public_name coq.plugins.ssrmatching) (synopsis "Coq ssrmatching plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/setoid_ring/0000755000175000017500000000000013612664533016255 5ustar treinentreinencoq-8.11.0/plugins/setoid_ring/RealField.v0000644000175000017500000001033613612664533020276 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 0. unfold Rgt. induction x; simpl; intros. apply Rlt_trans with (1 + 0). rewrite Rplus_comm. apply Rlt_n_Sn. apply Rplus_lt_compat_l. rewrite <- (Rmul_0_l Rset Rext RTheory 2). rewrite Rmult_comm. apply Rmult_lt_compat_l. apply Rlt_0_2. trivial. rewrite <- (Rmul_0_l Rset Rext RTheory 2). rewrite Rmult_comm. apply Rmult_lt_compat_l. apply Rlt_0_2. trivial. replace 1 with (0 + 1). apply Rlt_n_Sn. apply Rplus_0_l. Qed. Lemma Rgen_phiPOS_not_0 : forall x, InitialRing.gen_phiPOS1 1 Rplus Rmult x <> 0. red; intros. specialize (Rgen_phiPOS x). rewrite H; intro. apply (Rlt_asym 0 0); trivial. Qed. Lemma Zeq_bool_complete : forall x y, InitialRing.gen_phiZ 0%R 1%R Rplus Rmult Ropp x = InitialRing.gen_phiZ 0%R 1%R Rplus Rmult Ropp y -> Zeq_bool x y = true. Proof gen_phiZ_complete Rset Rext Rfield Rgen_phiPOS_not_0. Lemma Rdef_pow_add : forall (x:R) (n m:nat), pow x (n + m) = pow x n * pow x m. Proof. intros x n; elim n; simpl; auto with real. intros n0 H' m; rewrite H'; auto with real. Qed. Lemma R_power_theory : power_theory 1%R Rmult (@eq R) N.to_nat pow. Proof. constructor. destruct n. reflexivity. simpl. induction p. - rewrite Pos2Nat.inj_xI. simpl. now rewrite plus_0_r, Rdef_pow_add, IHp. - rewrite Pos2Nat.inj_xO. simpl. now rewrite plus_0_r, Rdef_pow_add, IHp. - simpl. rewrite Rmult_comm;apply Rmult_1_l. Qed. Ltac Rpow_tac t := match isnatcst t with | false => constr:(InitialRing.NotConstant) | _ => constr:(N.of_nat t) end. Ltac IZR_tac t := match t with | R0 => constr:(0%Z) | R1 => constr:(1%Z) | IZR (Z.pow_pos 10 ?p) => match isPcst p with | true => constr:(Z.pow_pos 10 p) | _ => constr:(InitialRing.NotConstant) end | IZR ?u => match isZcst u with | true => u | _ => constr:(InitialRing.NotConstant) end | _ => constr:(InitialRing.NotConstant) end. Add Field RField : Rfield (completeness Zeq_bool_complete, constants [IZR_tac], power_tac R_power_theory [Rpow_tac]). coq-8.11.0/plugins/setoid_ring/Ring_tac.v0000644000175000017500000003652613612664533020206 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* _ => R | _ => fail 1000 "Equality has no relation type" end. Ltac Get_goal := match goal with [|- ?G] => G end. (********************************************************************) (* Tacticals to build reflexive tactics *) Ltac OnEquation req := match goal with | |- req ?lhs ?rhs => (fun f => f lhs rhs) | _ => (fun _ => fail "Goal is not an equation (of expected equality)") end. Ltac OnEquationHyp req h := match type of h with | req ?lhs ?rhs => fun f => f lhs rhs | _ => (fun _ => fail "Hypothesis is not an equation (of expected equality)") end. (* Note: auxiliary subgoals in reverse order *) Ltac OnMainSubgoal H ty := match ty with | _ -> ?ty' => let subtac := OnMainSubgoal H ty' in fun kont => lapply H; [clear H; intro H; subtac kont | idtac] | _ => (fun kont => kont()) end. (* A generic pattern to have reflexive tactics do some computation: lemmas of the form [forall x', x=x' -> P(x')] are understood as: compute the normal form of x, instantiate x' with it, prove hypothesis x=x' with vm_compute and reflexivity, and pass the instantiated lemma to the continuation. *) Ltac ProveLemmaHyp lemma := match type of lemma with forall x', ?x = x' -> _ => (fun kont => let x' := fresh "res" in let H := fresh "res_eq" in compute_assertion H x' x; let lemma' := constr:(lemma x' H) in kont lemma'; (clear H||idtac"ProveLemmaHyp: cleanup failed"); subst x') | _ => (fun _ => fail "ProveLemmaHyp: lemma not of the expected form") end. Ltac ProveLemmaHyps lemma := match type of lemma with forall x', ?x = x' -> _ => (fun kont => let x' := fresh "res" in let H := fresh "res_eq" in compute_assertion H x' x; let lemma' := constr:(lemma x' H) in ProveLemmaHyps lemma' kont; (clear H||idtac"ProveLemmaHyps: cleanup failed"); subst x') | _ => (fun kont => kont lemma) end. (* Ltac ProveLemmaHyps lemma := (* expects a continuation *) let try_step := ProveLemmaHyp lemma in (fun kont => try_step ltac:(fun lemma' => ProveLemmaHyps lemma' kont) || kont lemma). *) Ltac ApplyLemmaThen lemma expr kont := let lem := constr:(lemma expr) in ProveLemmaHyp lem ltac:(fun lem' => let Heq := fresh "thm" in assert (Heq:=lem'); OnMainSubgoal Heq ltac:(type of Heq) ltac:(fun _ => kont Heq); (clear Heq||idtac"ApplyLemmaThen: cleanup failed")). (* Ltac ApplyLemmaThenAndCont lemma expr tac CONT_tac cont_arg := let pe := match type of (lemma expr) with forall pe', ?pe = pe' -> _ => pe | _ => fail 1 "ApplyLemmaThenAndCont: cannot find norm expression" end in let pe' := fresh "expr_nf" in let nf_pe := fresh "pe_eq" in compute_assertion nf_pe pe' pe; let Heq := fresh "thm" in (assert (Heq:=lemma pe pe' H) || fail "anomaly: failed to apply lemma"); clear nf_pe; OnMainSubgoal Heq ltac:(type of Heq) ltac:(try tac Heq; clear Heq pe';CONT_tac cont_arg)). *) Ltac ApplyLemmaThenAndCont lemma expr tac CONT_tac := ApplyLemmaThen lemma expr ltac:(fun lemma' => try tac lemma'; CONT_tac()). (* General scheme of reflexive tactics using of correctness lemma that involves normalisation of one expression - [FV_tac term fv] is a tactic that adds the atomic expressions of [term] into [fv] - [SYN_tac term fv] reifies [term] given the list of atomic expressions - [LEMMA_tac fv kont] computes the correctness lemma and passes it to continuation kont - [MAIN_tac H] process H which is the conclusion of the correctness lemma instantiated with each reified term - [fv] is the initial value of atomic expressions (to be completed by the reification of the terms - [terms] the list (a constr of type list) of terms to reify and process. *) Ltac ReflexiveRewriteTactic FV_tac SYN_tac LEMMA_tac MAIN_tac fv terms := (* extend the atom list *) let fv := list_fold_left FV_tac fv terms in let RW_tac lemma := let fcons term CONT_tac := let expr := SYN_tac term fv in let main H := match type of H with | (?req _ ?rhs) => change (req term rhs) in H end; MAIN_tac H in (ApplyLemmaThenAndCont lemma expr main CONT_tac) in (* rewrite steps *) lazy_list_fold_right fcons ltac:(fun _=>idtac) terms in LEMMA_tac fv RW_tac. (********************************************************) Ltac FV_hypo_tac mkFV req lH := let R := relation_carrier req in let FV_hypo_l_tac h := match h with @mkhypo (req ?pe _) _ => mkFV pe end in let FV_hypo_r_tac h := match h with @mkhypo (req _ ?pe) _ => mkFV pe end in let fv := list_fold_right FV_hypo_l_tac (@nil R) lH in list_fold_right FV_hypo_r_tac fv lH. Ltac mkHyp_tac C req Reify lH := let mkHyp h res := match h with | @mkhypo (req ?r1 ?r2) _ => let pe1 := Reify r1 in let pe2 := Reify r2 in constr:(cons (pe1,pe2) res) | _ => fail 1 "hypothesis is not a ring equality" end in list_fold_right mkHyp (@nil (PExpr C * PExpr C)) lH. Ltac proofHyp_tac lH := let get_proof h := match h with | @mkhypo _ ?p => p end in let rec bh l := match l with | nil => constr:(I) | cons ?h nil => get_proof h | cons ?h ?tl => let l := get_proof h in let r := bh tl in constr:(conj l r) end in bh lH. Ltac get_MonPol lemma := match type of lemma with | context [(mk_monpol_list ?cO ?cI ?cadd ?cmul ?csub ?copp ?cdiv ?ceqb _)] => constr:(mk_monpol_list cO cI cadd cmul csub copp cdiv ceqb) | _ => fail 1 "ring/field anomaly: bad correctness lemma (get_MonPol)" end. (********************************************************) (* Building the atom list of a ring expression *) (* We do not assume that Cst recognizes the rO and rI terms as constants, as *) (* the tactic could be used to discriminate occurrences of an opaque *) (* constant phi, with (phi 0) not convertible to 0 for instance *) Ltac FV Cst CstPow rO rI add mul sub opp pow t fv := let rec TFV t fv := let f := match Cst t with | NotConstant => match t with | rO => fun _ => fv | rI => fun _ => fv | (add ?t1 ?t2) => fun _ => TFV t2 ltac:(TFV t1 fv) | (mul ?t1 ?t2) => fun _ => TFV t2 ltac:(TFV t1 fv) | (sub ?t1 ?t2) => fun _ => TFV t2 ltac:(TFV t1 fv) | (opp ?t1) => fun _ => TFV t1 fv | (pow ?t1 ?n) => match CstPow n with | InitialRing.NotConstant => fun _ => AddFvTail t fv | _ => fun _ => TFV t1 fv end | _ => fun _ => AddFvTail t fv end | _ => fun _ => fv end in f() in TFV t fv. (* syntaxification of ring expressions *) (* We do not assume that Cst recognizes the rO and rI terms as constants, as *) (* the tactic could be used to discriminate occurrences of an opaque *) (* constant phi, with (phi 0) not convertible to 0 for instance *) Ltac mkPolexpr C Cst CstPow rO rI radd rmul rsub ropp rpow t fv := let rec mkP t := let f := match Cst t with | InitialRing.NotConstant => match t with | rO => fun _ => constr:(@PEO C) | rI => fun _ => constr:(@PEI C) | (radd ?t1 ?t2) => fun _ => let e1 := mkP t1 in let e2 := mkP t2 in constr:(@PEadd C e1 e2) | (rmul ?t1 ?t2) => fun _ => let e1 := mkP t1 in let e2 := mkP t2 in constr:(@PEmul C e1 e2) | (rsub ?t1 ?t2) => fun _ => let e1 := mkP t1 in let e2 := mkP t2 in constr:(@PEsub C e1 e2) | (ropp ?t1) => fun _ => let e1 := mkP t1 in constr:(@PEopp C e1) | (rpow ?t1 ?n) => match CstPow n with | InitialRing.NotConstant => fun _ => let p := Find_at t fv in constr:(PEX C p) | ?c => fun _ => let e1 := mkP t1 in constr:(@PEpow C e1 c) end | _ => fun _ => let p := Find_at t fv in constr:(PEX C p) end | ?c => fun _ => constr:(@PEc C c) end in f () in mkP t. (* packaging the ring structure *) Ltac PackRing F req sth ext morph arth cst_tac pow_tac lemma1 lemma2 pre post := let RNG := match type of lemma1 with | context [@PEeval ?R ?r0 ?r1 ?add ?mul ?sub ?opp ?C ?phi ?Cpow ?powphi ?pow _ _] => (fun proj => proj cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2) | _ => fail 1 "field anomaly: bad correctness lemma (parse)" end in F RNG. Ltac get_Carrier RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => R). Ltac get_Eq RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => req). Ltac get_Pre RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => pre). Ltac get_Post RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => post). Ltac get_NormLemma RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => lemma1). Ltac get_SimplifyLemma RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => lemma2). Ltac get_RingFV RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => FV cst_tac pow_tac r0 r1 add mul sub opp pow). Ltac get_RingMeta RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => mkPolexpr C cst_tac pow_tac r0 r1 add mul sub opp pow). Ltac get_RingHypTac RNG := RNG ltac:(fun cst_tac pow_tac pre post R req r0 r1 add mul sub opp C Cpow powphi pow lemma1 lemma2 => let mkPol := mkPolexpr C cst_tac pow_tac r0 r1 add mul sub opp pow in fun fv lH => mkHyp_tac C req ltac:(fun t => mkPol t fv) lH). (* ring tactics *) Definition ring_subst_niter := (10*10*10)%nat. Ltac Ring RNG lemma lH := let req := get_Eq RNG in OnEquation req ltac:(fun lhs rhs => let mkFV := get_RingFV RNG in let mkPol := get_RingMeta RNG in let mkHyp := get_RingHypTac RNG in let fv := FV_hypo_tac mkFV ltac:(get_Eq RNG) lH in let fv := mkFV lhs fv in let fv := mkFV rhs fv in check_fv fv; let pe1 := mkPol lhs fv in let pe2 := mkPol rhs fv in let lpe := mkHyp fv lH in let vlpe := fresh "hyp_list" in let vfv := fresh "fv_list" in pose (vlpe := lpe); pose (vfv := fv); (apply (lemma vfv vlpe pe1 pe2) || fail "typing error while applying ring"); [ ((let prh := proofHyp_tac lH in exact prh) || idtac "can not automatically prove hypothesis :"; [> idtac " maybe a left member of a hypothesis is not a monomial"..]) | vm_compute; (exact (eq_refl true) || fail "not a valid ring equation")]). Ltac Ring_norm_gen f RNG lemma lH rl := let mkFV := get_RingFV RNG in let mkPol := get_RingMeta RNG in let mkHyp := get_RingHypTac RNG in let mk_monpol := get_MonPol lemma in let fv := FV_hypo_tac mkFV ltac:(get_Eq RNG) lH in let lemma_tac fv kont := let lpe := mkHyp fv lH in let vlpe := fresh "list_hyp" in let vlmp := fresh "list_hyp_norm" in let vlmp_eq := fresh "list_hyp_norm_eq" in let prh := proofHyp_tac lH in pose (vlpe := lpe); compute_assertion vlmp_eq vlmp (mk_monpol vlpe); let H := fresh "ring_lemma" in (assert (H := lemma vlpe fv prh vlmp vlmp_eq) || fail "type error when build the rewriting lemma"); clear vlmp_eq; kont H; (clear H||idtac"Ring_norm_gen: cleanup failed"); subst vlpe vlmp in let simpl_ring H := (protect_fv "ring" in H; f H) in ReflexiveRewriteTactic mkFV mkPol lemma_tac simpl_ring fv rl. Ltac Ring_gen RNG lH rl := let lemma := get_NormLemma RNG in get_Pre RNG (); Ring RNG (lemma ring_subst_niter) lH. Tactic Notation (at level 0) "ring" := let G := Get_goal in ring_lookup (PackRing Ring_gen) [] G. Tactic Notation (at level 0) "ring" "[" constr_list(lH) "]" := let G := Get_goal in ring_lookup (PackRing Ring_gen) [lH] G. (* Simplification *) Ltac Ring_simplify_gen f RNG lH rl := let lemma := get_SimplifyLemma RNG in let l := fresh "to_rewrite" in pose (l:= rl); generalize (eq_refl l); unfold l at 2; get_Pre RNG (); let rl := match goal with | [|- l = ?RL -> _ ] => RL | _ => fail 1 "ring_simplify anomaly: bad goal after pre" end in let Heq := fresh "Heq" in intros Heq;clear Heq l; Ring_norm_gen f RNG (lemma ring_subst_niter) lH rl; get_Post RNG (). Ltac Ring_simplify := Ring_simplify_gen ltac:(fun H => rewrite H). Tactic Notation (at level 0) "ring_simplify" constr_list(rl) := let G := Get_goal in ring_lookup (PackRing Ring_simplify) [] rl G. Tactic Notation (at level 0) "ring_simplify" "[" constr_list(lH) "]" constr_list(rl) := let G := Get_goal in ring_lookup (PackRing Ring_simplify) [lH] rl G. Tactic Notation "ring_simplify" constr_list(rl) "in" hyp(H):= let G := Get_goal in let t := type of H in let g := fresh "goal" in set (g:= G); generalize H; ring_lookup (PackRing Ring_simplify) [] rl t; (* Correction of bug 1859: we want to leave H at its initial position this is obtained by adding a copy of H (H'), move it just after H, remove H and finally rename H into H' *) let H' := fresh "H" in intro H'; move H' after H; clear H;rename H' into H; unfold g;clear g. Tactic Notation "ring_simplify" "["constr_list(lH)"]" constr_list(rl) "in" hyp(H):= let G := Get_goal in let t := type of H in let g := fresh "goal" in set (g:= G); generalize H; ring_lookup (PackRing Ring_simplify) [lH] rl t; (* Correction of bug 1859: we want to leave H at its initial position this is obtained by adding a copy of H (H'), move it just after H, remove H and finally rename H into H' *) let H' := fresh "H" in intro H'; move H' after H; clear H;rename H' into H; unfold g;clear g. coq-8.11.0/plugins/setoid_ring/Ncring.v0000644000175000017500000002263613612664533017675 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* T->T} {mul:T->T->T} {sub:T->T->T} {opp:T->T} {ring_eq:T->T->Prop}. Instance zero_notation(T:Type)`{Ring_ops T}:Zero T:= ring0. Instance one_notation(T:Type)`{Ring_ops T}:One T:= ring1. Instance add_notation(T:Type)`{Ring_ops T}:Addition T:= add. Instance mul_notation(T:Type)`{Ring_ops T}:@Multiplication T T:= mul. Instance sub_notation(T:Type)`{Ring_ops T}:Subtraction T:= sub. Instance opp_notation(T:Type)`{Ring_ops T}:Opposite T:= opp. Instance eq_notation(T:Type)`{Ring_ops T}:@Equality T:= ring_eq. Class Ring `{Ro:Ring_ops}:={ ring_setoid: Equivalence _==_; ring_plus_comp: Proper (_==_ ==> _==_ ==>_==_) _+_; ring_mult_comp: Proper (_==_ ==> _==_ ==>_==_) _*_; ring_sub_comp: Proper (_==_ ==> _==_ ==>_==_) _-_; ring_opp_comp: Proper (_==_==>_==_) -_; ring_add_0_l : forall x, 0 + x == x; ring_add_comm : forall x y, x + y == y + x; ring_add_assoc : forall x y z, x + (y + z) == (x + y) + z; ring_mul_1_l : forall x, 1 * x == x; ring_mul_1_r : forall x, x * 1 == x; ring_mul_assoc : forall x y z, x * (y * z) == (x * y) * z; ring_distr_l : forall x y z, (x + y) * z == x * z + y * z; ring_distr_r : forall x y z, z * ( x + y) == z * x + z * y; ring_sub_def : forall x y, x - y == x + -y; ring_opp_def : forall x, x + -x == 0 }. (* inutile! je sais plus pourquoi j'ai mis ca... Instance ring_Ring_ops(R:Type)`{Ring R} :@Ring_ops R 0 1 addition multiplication subtraction opposite equality. *) Existing Instance ring_setoid. Existing Instance ring_plus_comp. Existing Instance ring_mult_comp. Existing Instance ring_sub_comp. Existing Instance ring_opp_comp. Section Ring_power. Context {R:Type}`{Ring R}. Fixpoint pow_pos (x:R) (i:positive) {struct i}: R := match i with | xH => x | xO i => let p := pow_pos x i in p * p | xI i => let p := pow_pos x i in x * (p * p) end. Definition pow_N (x:R) (p:N) := match p with | N0 => 1 | Npos p => pow_pos x p end. End Ring_power. Definition ZN(x:Z):= match x with Z0 => N0 |Zpos p | Zneg p => Npos p end. Instance power_ring {R:Type}`{Ring R} : Power:= {power x y := pow_N x (ZN y)}. (** Interpretation morphisms definition*) Class Ring_morphism (C R:Type)`{Cr:Ring C} `{Rr:Ring R}`{Rh:Bracket C R}:= { ring_morphism0 : [0] == 0; ring_morphism1 : [1] == 1; ring_morphism_add : forall x y, [x + y] == [x] + [y]; ring_morphism_sub : forall x y, [x - y] == [x] - [y]; ring_morphism_mul : forall x y, [x * y] == [x] * [y]; ring_morphism_opp : forall x, [-x] == -[x]; ring_morphism_eq : forall x y, x == y -> [x] == [y]}. Section Ring. Context {R:Type}`{Rr:Ring R}. (* Powers *) Lemma pow_pos_comm : forall x j, x * pow_pos x j == pow_pos x j * x. Proof. induction j; simpl. rewrite <- ring_mul_assoc. rewrite <- ring_mul_assoc. rewrite <- IHj. rewrite (ring_mul_assoc (pow_pos x j) x (pow_pos x j)). rewrite <- IHj. rewrite <- ring_mul_assoc. reflexivity. rewrite <- ring_mul_assoc. rewrite <- IHj. rewrite ring_mul_assoc. rewrite IHj. rewrite <- ring_mul_assoc. rewrite IHj. reflexivity. reflexivity. Qed. Lemma pow_pos_succ : forall x j, pow_pos x (Pos.succ j) == x * pow_pos x j. Proof. induction j; simpl. rewrite IHj. rewrite <- (ring_mul_assoc x (pow_pos x j) (x * pow_pos x j)). rewrite (ring_mul_assoc (pow_pos x j) x (pow_pos x j)). rewrite <- pow_pos_comm. rewrite <- ring_mul_assoc. reflexivity. reflexivity. reflexivity. Qed. Lemma pow_pos_add : forall x i j, pow_pos x (i + j) == pow_pos x i * pow_pos x j. Proof. intro x;induction i;intros. rewrite Pos.xI_succ_xO;rewrite <- Pos.add_1_r. rewrite <- Pos.add_diag;repeat rewrite <- Pos.add_assoc. repeat rewrite IHi. rewrite Pos.add_comm;rewrite Pos.add_1_r; rewrite pow_pos_succ. simpl;repeat rewrite ring_mul_assoc. reflexivity. rewrite <- Pos.add_diag;repeat rewrite <- Pos.add_assoc. repeat rewrite IHi. rewrite ring_mul_assoc. reflexivity. rewrite Pos.add_comm;rewrite Pos.add_1_r;rewrite pow_pos_succ. simpl. reflexivity. Qed. Definition id_phi_N (x:N) : N := x. Lemma pow_N_pow_N : forall x n, pow_N x (id_phi_N n) == pow_N x n. Proof. intros; reflexivity. Qed. (** Identity is a morphism *) (* Instance IDmorph : Ring_morphism _ _ _ (fun x => x). Proof. apply (Build_Ring_morphism H6 H6 (fun x => x));intros; try reflexivity. trivial. Qed. *) (** rings are almost rings*) Lemma ring_mul_0_l : forall x, 0 * x == 0. Proof. intro x. setoid_replace (0*x) with ((0+1)*x + -x). rewrite ring_add_0_l. rewrite ring_mul_1_l . rewrite ring_opp_def . fold zero. reflexivity. rewrite ring_distr_l . rewrite ring_mul_1_l . rewrite <- ring_add_assoc ; rewrite ring_opp_def . rewrite ring_add_comm ; rewrite ring_add_0_l ;reflexivity. Qed. Lemma ring_mul_0_r : forall x, x * 0 == 0. Proof. intro x; setoid_replace (x*0) with (x*(0+1) + -x). rewrite ring_add_0_l ; rewrite ring_mul_1_r . rewrite ring_opp_def ; fold zero; reflexivity. rewrite ring_distr_r ;rewrite ring_mul_1_r . rewrite <- ring_add_assoc ; rewrite ring_opp_def . rewrite ring_add_comm ; rewrite ring_add_0_l ;reflexivity. Qed. Lemma ring_opp_mul_l : forall x y, -(x * y) == -x * y. Proof. intros x y;rewrite <- (ring_add_0_l (- x * y)). rewrite ring_add_comm . rewrite <- (ring_opp_def (x*y)). rewrite ring_add_assoc . rewrite <- ring_distr_l. rewrite (ring_add_comm (-x));rewrite ring_opp_def . rewrite ring_mul_0_l;rewrite ring_add_0_l ;reflexivity. Qed. Lemma ring_opp_mul_r : forall x y, -(x * y) == x * -y. Proof. intros x y;rewrite <- (ring_add_0_l (x * - y)). rewrite ring_add_comm . rewrite <- (ring_opp_def (x*y)). rewrite ring_add_assoc . rewrite <- ring_distr_r . rewrite (ring_add_comm (-y));rewrite ring_opp_def . rewrite ring_mul_0_r;rewrite ring_add_0_l ;reflexivity. Qed. Lemma ring_opp_add : forall x y, -(x + y) == -x + -y. Proof. intros x y;rewrite <- (ring_add_0_l (-(x+y))). rewrite <- (ring_opp_def x). rewrite <- (ring_add_0_l (x + - x + - (x + y))). rewrite <- (ring_opp_def y). rewrite (ring_add_comm x). rewrite (ring_add_comm y). rewrite <- (ring_add_assoc (-y)). rewrite <- (ring_add_assoc (- x)). rewrite (ring_add_assoc y). rewrite (ring_add_comm y). rewrite <- (ring_add_assoc (- x)). rewrite (ring_add_assoc y). rewrite (ring_add_comm y);rewrite ring_opp_def . rewrite (ring_add_comm (-x) 0);rewrite ring_add_0_l . rewrite ring_add_comm; reflexivity. Qed. Lemma ring_opp_opp : forall x, - -x == x. Proof. intros x; rewrite <- (ring_add_0_l (- -x)). rewrite <- (ring_opp_def x). rewrite <- ring_add_assoc ; rewrite ring_opp_def . rewrite (ring_add_comm x); rewrite ring_add_0_l . reflexivity. Qed. Lemma ring_sub_ext : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 - y1 == x2 - y2. Proof. intros. setoid_replace (x1 - y1) with (x1 + -y1). setoid_replace (x2 - y2) with (x2 + -y2). rewrite H;rewrite H0;reflexivity. rewrite ring_sub_def. reflexivity. rewrite ring_sub_def. reflexivity. Qed. Ltac mrewrite := repeat first [ rewrite ring_add_0_l | rewrite <- (ring_add_comm 0) | rewrite ring_mul_1_l | rewrite ring_mul_0_l | rewrite ring_distr_l | reflexivity ]. Lemma ring_add_0_r : forall x, (x + 0) == x. Proof. intros; mrewrite. Qed. Lemma ring_add_assoc1 : forall x y z, (x + y) + z == (y + z) + x. Proof. intros;rewrite <- (ring_add_assoc x). rewrite (ring_add_comm x);reflexivity. Qed. Lemma ring_add_assoc2 : forall x y z, (y + x) + z == (y + z) + x. Proof. intros; repeat rewrite <- ring_add_assoc. rewrite (ring_add_comm x); reflexivity. Qed. Lemma ring_opp_zero : -0 == 0. Proof. rewrite <- (ring_mul_0_r 0). rewrite ring_opp_mul_l. repeat rewrite ring_mul_0_r. reflexivity. Qed. End Ring. (** Some simplification tactics*) Ltac gen_reflexivity := reflexivity. Ltac gen_rewrite := repeat first [ reflexivity | progress rewrite ring_opp_zero | rewrite ring_add_0_l | rewrite ring_add_0_r | rewrite ring_mul_1_l | rewrite ring_mul_1_r | rewrite ring_mul_0_l | rewrite ring_mul_0_r | rewrite ring_distr_l | rewrite ring_distr_r | rewrite ring_add_assoc | rewrite ring_mul_assoc | progress rewrite ring_opp_add | progress rewrite ring_sub_def | progress rewrite <- ring_opp_mul_l | progress rewrite <- ring_opp_mul_r ]. Ltac gen_add_push x := repeat (match goal with | |- context [(?y + x) + ?z] => progress rewrite (ring_add_assoc2 x y z) | |- context [(x + ?y) + ?z] => progress rewrite (ring_add_assoc1 x y z) end). coq-8.11.0/plugins/setoid_ring/newring.ml0000644000175000017500000011172413612664533020266 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* EConstr.t -> GlobRef.t -> (Int.t -> protect_flag) option let global_head_of_constr sigma c = let f, args = decompose_app sigma c in try fst (Termops.global_of_constr sigma f) with Not_found -> CErrors.anomaly (str "global_head_of_constr.") let global_of_constr_nofail c = try global_of_constr c with Not_found -> GlobRef.VarRef (Id.of_string "dummy") let rec mk_clos_but f_map n t = let (f, args) = Constr.decompose_appvect t in match f_map (global_of_constr_nofail f) with | Some tag -> let map i t = tag_arg f_map n (tag i) t in if Array.is_empty args then map (-1) f else mk_red (FApp (map (-1) f, Array.mapi map args)) | None -> mk_atom t and tag_arg f_map n tag c = match tag with | Eval -> mk_clos (Esubst.subs_id n) c | Prot -> mk_atom c | Rec -> mk_clos_but f_map n c let interp_map l t = try Some(List.assoc_f GlobRef.equal t l) with Not_found -> None let protect_maps : protection String.Map.t ref = ref String.Map.empty let add_map s m = protect_maps := String.Map.add s m !protect_maps let lookup_map map = try String.Map.find map !protect_maps with Not_found -> CErrors.user_err ~hdr:"lookup_map" (str"map "++qs map++str"not found") let protect_red map env sigma c0 = let evars ev = Evarutil.safe_evar_value sigma ev in let c = EConstr.Unsafe.to_constr c0 in let tab = create_tab () in let infos = create_clos_infos ~evars all env in let map = lookup_map map sigma c0 in let rec eval n c = match Constr.kind c with | Prod (na, t, u) -> Constr.mkProd (na, eval n t, eval (n + 1) u) | _ -> kl infos tab (mk_clos_but map n c) in EConstr.of_constr (eval 0 c) let protect_tac map = Tactics.reduct_option ~check:false (protect_red map,DEFAULTcast) None let protect_tac_in map id = Tactics.reduct_option ~check:false (protect_red map,DEFAULTcast) (Some(id, Locus.InHyp)) (****************************************************************************) let rec closed_under sigma cset t = try let (gr, _) = Termops.global_of_constr sigma t in GlobRef.Set_env.mem gr cset with Not_found -> match EConstr.kind sigma t with | Cast(c,_,_) -> closed_under sigma cset c | App(f,l) -> closed_under sigma cset f && Array.for_all (closed_under sigma cset) l | _ -> false let closed_term args _ = match args with | [t; l] -> let t = Option.get (Value.to_constr t) in let l = List.map (fun c -> Value.cast (Genarg.topwit Stdarg.wit_ref) c) (Option.get (Value.to_list l)) in Proofview.tclEVARMAP >>= fun sigma -> let cs = List.fold_right GlobRef.Set_env.add l GlobRef.Set_env.empty in if closed_under sigma cs t then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (mt()) | _ -> assert false let closed_term_ast = let tacname = { mltac_plugin = "newring_plugin"; mltac_tactic = "closed_term"; } in let () = Tacenv.register_ml_tactic tacname [|closed_term|] in let tacname = { mltac_name = tacname; mltac_index = 0; } in fun l -> let l = List.map (fun gr -> ArgArg(Loc.tag gr)) l in TacFun([Name(Id.of_string"t")], TacML(CAst.make (tacname, [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (DAst.make @@ GVar(Id.of_string"t"),None)); TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)]))) (* let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term" *) (****************************************************************************) let ic c = let env = Global.env() in let sigma = Evd.from_env env in let c, uctx = Constrintern.interp_constr env sigma c in (Evd.from_ctx uctx, c) let ic_unsafe c = (*FIXME remove *) let env = Global.env() in let sigma = Evd.from_env env in fst (Constrintern.interp_constr env sigma c) let decl_constant na univs c = let open Constr in let vars = CVars.universes_of_constr c in let univs = UState.restrict_universe_context ~lbound:(Global.universes_lbound ()) univs vars in let () = Declare.declare_universe_context ~poly:false univs in let types = (Typeops.infer (Global.env ()) c).uj_type in let univs = Monomorphic_entry Univ.ContextSet.empty in mkConst(declare_constant ~name:(Id.of_string na) ~kind:Decls.(IsProof Lemma) (DefinitionEntry (definition_entry ~opaque:true ~types ~univs c))) (* Calling a global tactic *) let ltac_call tac (args:glob_tactic_arg list) = TacArg(CAst.make @@ TacCall (CAst.make (ArgArg(Loc.tag @@ Lazy.force tac),args))) let dummy_goal env sigma = let (gl,_,sigma) = Goal.V82.mk_goal sigma (named_context_val env) EConstr.mkProp in {Evd.it = gl; Evd.sigma = sigma} let constr_of evd v = match Value.to_constr v with | Some c -> EConstr.to_constr evd c | None -> failwith "Ring.exec_tactic: anomaly" let tactic_res = ref [||] let get_res = let open Tacexpr in let name = { mltac_plugin = "newring_plugin"; mltac_tactic = "get_res"; } in let entry = { mltac_name = name; mltac_index = 0 } in let tac args ist = let n = Tacinterp.Value.cast (Genarg.topwit Stdarg.wit_int) (List.hd args) in let init i = Id.Map.find (Id.of_string ("x" ^ string_of_int i)) ist.lfun in tactic_res := Array.init n init; Proofview.tclUNIT () in Tacenv.register_ml_tactic name [| tac |]; entry let exec_tactic env evd n f args = let fold arg (i, vars, lfun) = let id = Id.of_string ("x" ^ string_of_int i) in let x = Reference (ArgVar CAst.(make id)) in (succ i, x :: vars, Id.Map.add id (Value.of_constr arg) lfun) in let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in let ist = { (Tacinterp.default_ist ()) with Tacinterp.lfun = lfun; } in (* Build the getter *) let lid = List.init n (fun i -> Id.of_string("x"^string_of_int i)) in let n = Genarg.in_gen (Genarg.glbwit Stdarg.wit_int) n in let get_res = TacML (CAst.make (get_res, [TacGeneric n])) in let getter = Tacexp (TacFun (List.map (fun n -> Name n) lid, get_res)) in (* Evaluate the whole result *) let gl = dummy_goal env evd in let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in let evd = Evd.minimize_universes (Refiner.project gls) in let nf c = constr_of evd c in Array.map nf !tactic_res, Evd.universe_context_set evd let gen_constant n = lazy (EConstr.of_constr (UnivGen.constr_of_monomorphic_global (Coqlib.lib_ref n))) let gen_reference n = lazy (Coqlib.lib_ref n) let coq_mk_Setoid = gen_constant "plugins.setoid_ring.Build_Setoid_Theory" let coq_None = gen_reference "core.option.None" let coq_Some = gen_reference "core.option.Some" let coq_eq = gen_constant "core.eq.type" let coq_cons = gen_reference "core.list.cons" let coq_nil = gen_reference "core.list.nil" let lapp f args = mkApp(Lazy.force f,args) let plapp evdref f args = let evd, fc = Evarutil.new_global !evdref (Lazy.force f) in evdref := evd; mkApp(fc,args) let dest_rel0 sigma t = match EConstr.kind sigma t with | App(f,args) when Array.length args >= 2 -> let rel = mkApp(f,Array.sub args 0 (Array.length args - 2)) in if closed0 sigma rel then (rel,args.(Array.length args - 2),args.(Array.length args - 1)) else error "ring: cannot find relation (not closed)" | _ -> error "ring: cannot find relation" let rec dest_rel sigma t = match EConstr.kind sigma t with | Prod(_,_,c) -> dest_rel sigma c | _ -> dest_rel0 sigma t (****************************************************************************) (* Library linking *) let plugin_dir = "setoid_ring" let cdir = ["Coq";plugin_dir] let plugin_modules = List.map (fun d -> cdir@d) [["Ring_theory"];["Ring_polynom"]; ["Ring_tac"];["InitialRing"]; ["Field_tac"]; ["Field_theory"] ] let my_constant c = lazy (EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.gen_reference_in_modules "Ring" plugin_modules c)) [@@ocaml.warning "-3"] let my_reference c = lazy (Coqlib.gen_reference_in_modules "Ring" plugin_modules c) [@@ocaml.warning "-3"] let znew_ring_path = DirPath.make (List.map Id.of_string ["InitialRing";plugin_dir;"Coq"]) let zltac s = lazy(KerName.make (ModPath.MPfile znew_ring_path) (Label.make s)) let mk_cst l s = lazy (Coqlib.coq_reference "newring" l s) [@@ocaml.warning "-3"] let pol_cst s = mk_cst [plugin_dir;"Ring_polynom"] s (* Ring theory *) (* almost_ring defs *) let coq_almost_ring_theory = my_constant "almost_ring_theory" (* setoid and morphism utilities *) let coq_eq_setoid = my_reference "Eqsth" let coq_eq_morph = my_reference "Eq_ext" let coq_eq_smorph = my_reference "Eq_s_ext" (* ring -> almost_ring utilities *) let coq_ring_theory = my_constant "ring_theory" let coq_mk_reqe = my_constant "mk_reqe" (* semi_ring -> almost_ring utilities *) let coq_semi_ring_theory = my_constant "semi_ring_theory" let coq_mk_seqe = my_constant "mk_seqe" let coq_abstract = my_constant"Abstract" let coq_comp = my_constant"Computational" let coq_morph = my_constant"Morphism" (* power function *) let ltac_inv_morph_nothing = zltac"inv_morph_nothing" (* hypothesis *) let coq_mkhypo = my_reference "mkhypo" let coq_hypo = my_reference "hypo" (* Equality: do not evaluate but make recursive call on both sides *) let map_with_eq arg_map sigma c = let (req,_,_) = dest_rel sigma c in interp_map ((global_head_of_constr sigma req,(function -1->Prot|_->Rec)):: List.map (fun (c,map) -> (Lazy.force c,map)) arg_map) let map_without_eq arg_map _ _ = interp_map (List.map (fun (c,map) -> (Lazy.force c,map)) arg_map) let _ = add_map "ring" (map_with_eq [coq_cons,(function -1->Eval|2->Rec|_->Prot); coq_nil, (function -1->Eval|_ -> Prot); my_reference "IDphi", (function _->Eval); my_reference "gen_phiZ", (function _->Eval); (* Pphi_dev: evaluate polynomial and coef operations, protect ring operations and make recursive call on the var map *) pol_cst "Pphi_dev", (function -1|8|9|10|12|14->Eval|11|13->Rec|_->Prot); pol_cst "Pphi_pow", (function -1|8|9|10|13|15|17->Eval|11|16->Rec|_->Prot); (* PEeval: evaluate polynomial, protect ring operations and make recursive call on the var map *) pol_cst "PEeval", (function -1|10|13->Eval|8|12->Rec|_->Prot)]) (****************************************************************************) (* Ring database *) module Cmap = Map.Make(Constr) let from_carrier = Summary.ref Cmap.empty ~name:"ring-tac-carrier-table" let print_rings () = Feedback.msg_notice (strbrk "The following ring structures have been declared:"); Cmap.iter (fun _carrier ring -> let env = Global.env () in let sigma = Evd.from_env env in Feedback.msg_notice (hov 2 (Ppconstr.pr_id ring.ring_name ++ spc() ++ str"with carrier "++ pr_constr_env env sigma ring.ring_carrier++spc()++ str"and equivalence relation "++ pr_constr_env env sigma ring.ring_req)) ) !from_carrier let ring_for_carrier r = Cmap.find r !from_carrier let find_ring_structure env sigma l = match l with | t::cl' -> let ty = Retyping.get_type_of env sigma t in let check c = let ty' = Retyping.get_type_of env sigma c in if not (Reductionops.is_conv env sigma ty ty') then CErrors.user_err ~hdr:"ring" (str"arguments of ring_simplify do not have all the same type") in List.iter check cl'; (try ring_for_carrier (EConstr.to_constr sigma ty) with Not_found -> CErrors.user_err ~hdr:"ring" (str"cannot find a declared ring structure over"++ spc() ++ str"\"" ++ pr_econstr_env env sigma ty ++ str"\"")) | [] -> assert false let add_entry (sp,_kn) e = from_carrier := Cmap.add e.ring_carrier e !from_carrier let subst_th (subst,th) = let c' = subst_mps subst th.ring_carrier in let eq' = subst_mps subst th.ring_req in let set' = subst_mps subst th.ring_setoid in let ext' = subst_mps subst th.ring_ext in let morph' = subst_mps subst th.ring_morph in let th' = subst_mps subst th.ring_th in let thm1' = subst_mps subst th.ring_lemma1 in let thm2' = subst_mps subst th.ring_lemma2 in let tac'= Tacsubst.subst_tactic subst th.ring_cst_tac in let pow_tac'= Tacsubst.subst_tactic subst th.ring_pow_tac in let pretac'= Tacsubst.subst_tactic subst th.ring_pre_tac in let posttac'= Tacsubst.subst_tactic subst th.ring_post_tac in if c' == th.ring_carrier && eq' == th.ring_req && Constr.equal set' th.ring_setoid && ext' == th.ring_ext && morph' == th.ring_morph && th' == th.ring_th && thm1' == th.ring_lemma1 && thm2' == th.ring_lemma2 && tac' == th.ring_cst_tac && pow_tac' == th.ring_pow_tac && pretac' == th.ring_pre_tac && posttac' == th.ring_post_tac then th else { ring_name = th.ring_name; ring_carrier = c'; ring_req = eq'; ring_setoid = set'; ring_ext = ext'; ring_morph = morph'; ring_th = th'; ring_cst_tac = tac'; ring_pow_tac = pow_tac'; ring_lemma1 = thm1'; ring_lemma2 = thm2'; ring_pre_tac = pretac'; ring_post_tac = posttac' } let theory_to_obj : ring_info -> obj = let cache_th (name,th) = add_entry name th in declare_object @@ global_object_nodischarge "tactic-new-ring-theory" ~cache:cache_th ~subst:(Some subst_th) let setoid_of_relation env evd a r = try let evm = !evd in let evm, refl = Rewrite.get_reflexive_proof env evm a r in let evm, sym = Rewrite.get_symmetric_proof env evm a r in let evm, trans = Rewrite.get_transitive_proof env evm a r in evd := evm; lapp coq_mk_Setoid [|a ; r ; refl; sym; trans |] with Not_found -> error "cannot find setoid relation" let op_morph r add mul opp req m1 m2 m3 = lapp coq_mk_reqe [| r; add; mul; opp; req; m1; m2; m3 |] let op_smorph r add mul req m1 m2 = lapp coq_mk_seqe [| r; add; mul; req; m1; m2 |] let ring_equality env evd (r,add,mul,opp,req) = match EConstr.kind !evd req with | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) -> let setoid = plapp evd coq_eq_setoid [|r|] in let op_morph = match opp with Some opp -> plapp evd coq_eq_morph [|r;add;mul;opp|] | None -> plapp evd coq_eq_smorph [|r;add;mul|] in let sigma = !evd in let sigma, setoid = Typing.solve_evars env sigma setoid in let sigma, op_morph = Typing.solve_evars env sigma op_morph in evd := sigma; (setoid,op_morph) | _ -> let setoid = setoid_of_relation (Global.env ()) evd r req in let signature = [Some (r,Some req);Some (r,Some req)],Some(r,Some req) in let add_m, add_m_lem = try Rewrite.default_morphism signature add with Not_found -> error "ring addition should be declared as a morphism" in let mul_m, mul_m_lem = try Rewrite.default_morphism signature mul with Not_found -> error "ring multiplication should be declared as a morphism" in let op_morph = match opp with | Some opp -> (let opp_m,opp_m_lem = try Rewrite.default_morphism ([Some(r,Some req)],Some(r,Some req)) opp with Not_found -> error "ring opposite should be declared as a morphism" in let op_morph = op_morph r add mul opp req add_m_lem mul_m_lem opp_m_lem in Flags.if_verbose Feedback.msg_info (str"Using setoid \""++ pr_econstr_env env !evd req++str"\""++spc()++ str"and morphisms \""++pr_econstr_env env !evd add_m_lem ++ str"\","++spc()++ str"\""++pr_econstr_env env !evd mul_m_lem++ str"\""++spc()++str"and \""++pr_econstr_env env !evd opp_m_lem++ str"\""); op_morph) | None -> (Flags.if_verbose Feedback.msg_info (str"Using setoid \""++pr_econstr_env env !evd req ++str"\"" ++ spc() ++ str"and morphisms \""++pr_econstr_env env !evd add_m_lem ++ str"\""++spc()++str"and \""++ pr_econstr_env env !evd mul_m_lem++str"\""); op_smorph r add mul req add_m_lem mul_m_lem) in (setoid,op_morph) let build_setoid_params env evd r add mul opp req eqth = match eqth with Some th -> th | None -> ring_equality env evd (r,add,mul,opp,req) let dest_ring env sigma th_spec = let th_typ = Retyping.get_type_of env sigma th_spec in match EConstr.kind sigma th_typ with App(f,[|r;zero;one;add;mul;sub;opp;req|]) when eq_constr_nounivs sigma f (Lazy.force coq_almost_ring_theory) -> (None,r,zero,one,add,mul,Some sub,Some opp,req) | App(f,[|r;zero;one;add;mul;req|]) when eq_constr_nounivs sigma f (Lazy.force coq_semi_ring_theory) -> (Some true,r,zero,one,add,mul,None,None,req) | App(f,[|r;zero;one;add;mul;sub;opp;req|]) when eq_constr_nounivs sigma f (Lazy.force coq_ring_theory) -> (Some false,r,zero,one,add,mul,Some sub,Some opp,req) | _ -> error "bad ring structure" let reflect_coeff rkind = (* We build an ill-typed terms on purpose... *) match rkind with Abstract -> Lazy.force coq_abstract | Computational c -> lapp coq_comp [|c|] | Morphism m -> lapp coq_morph [|m|] let interp_cst_tac env sigma rk kind (zero,one,add,mul,opp) cst_tac = match cst_tac with Some (CstTac t) -> Tacintern.glob_tactic t | Some (Closed lc) -> closed_term_ast (List.map Smartlocate.global_with_alias lc) | None -> let t = ArgArg(Loc.tag @@ Lazy.force ltac_inv_morph_nothing) in TacArg(CAst.make (TacCall(CAst.make (t,[])))) let make_hyp env evd c = let t = Retyping.get_type_of env !evd c in plapp evd coq_mkhypo [|t;c|] let make_hyp_list env evdref lH = let evd, carrier = Evarutil.new_global !evdref (Lazy.force coq_hypo) in evdref := evd; let l = List.fold_right (fun c l -> plapp evdref coq_cons [|carrier; (make_hyp env evdref c); l|]) lH (plapp evdref coq_nil [|carrier|]) in let sigma, l' = Typing.solve_evars env !evdref l in evdref := sigma; let l' = EConstr.Unsafe.to_constr l' in Evarutil.nf_evars_universes !evdref l' let interp_power env evdref pow = let evd, carrier = Evarutil.new_global !evdref (Lazy.force coq_hypo) in evdref := evd; match pow with | None -> let t = ArgArg(Loc.tag (Lazy.force ltac_inv_morph_nothing)) in (TacArg(CAst.make (TacCall(CAst.make (t,[])))), plapp evdref coq_None [|carrier|]) | Some (tac, spec) -> let tac = match tac with | CstTac t -> Tacintern.glob_tactic t | Closed lc -> closed_term_ast (List.map Smartlocate.global_with_alias lc) in let spec = make_hyp env evdref (ic_unsafe spec) in (tac, plapp evdref coq_Some [|carrier; spec|]) let interp_sign env evdref sign = let evd, carrier = Evarutil.new_global !evdref (Lazy.force coq_hypo) in evdref := evd; match sign with | None -> plapp evdref coq_None [|carrier|] | Some spec -> let spec = make_hyp env evdref (ic_unsafe spec) in plapp evdref coq_Some [|carrier;spec|] (* Same remark on ill-typed terms ... *) let interp_div env evdref div = let evd, carrier = Evarutil.new_global !evdref (Lazy.force coq_hypo) in evdref := evd; match div with | None -> plapp evdref coq_None [|carrier|] | Some spec -> let spec = make_hyp env evdref (ic_unsafe spec) in plapp evdref coq_Some [|carrier;spec|] (* Same remark on ill-typed terms ... *) let add_theory0 name (sigma, rth) eqth morphth cst_tac (pre,post) power sign div = check_required_library (cdir@["Ring_base"]); let env = Global.env() in let (kind,r,zero,one,add,mul,sub,opp,req) = dest_ring env sigma rth in let evd = ref sigma in let (sth,ext) = build_setoid_params env evd r add mul opp req eqth in let (pow_tac, pspec) = interp_power env evd power in let sspec = interp_sign env evd sign in let dspec = interp_div env evd div in let rk = reflect_coeff morphth in let params,ctx = exec_tactic env !evd 5 (zltac "ring_lemmas") [sth;ext;rth;pspec;sspec;dspec;rk] in let lemma1 = params.(3) in let lemma2 = params.(4) in let lemma1 = decl_constant (Id.to_string name^"_ring_lemma1") ctx lemma1 in let lemma2 = decl_constant (Id.to_string name^"_ring_lemma2") ctx lemma2 in let cst_tac = interp_cst_tac env sigma morphth kind (zero,one,add,mul,opp) cst_tac in let pretac = match pre with Some t -> Tacintern.glob_tactic t | _ -> TacId [] in let posttac = match post with Some t -> Tacintern.glob_tactic t | _ -> TacId [] in let r = EConstr.to_constr sigma r in let req = EConstr.to_constr sigma req in let sth = EConstr.to_constr sigma sth in let _ = Lib.add_leaf name (theory_to_obj { ring_name = name; ring_carrier = r; ring_req = req; ring_setoid = sth; ring_ext = params.(1); ring_morph = params.(2); ring_th = params.(0); ring_cst_tac = cst_tac; ring_pow_tac = pow_tac; ring_lemma1 = lemma1; ring_lemma2 = lemma2; ring_pre_tac = pretac; ring_post_tac = posttac }) in () let ic_coeff_spec = function | Computational t -> Computational (ic_unsafe t) | Morphism t -> Morphism (ic_unsafe t) | Abstract -> Abstract let set_once s r v = if Option.is_empty !r then r := Some v else error (s^" cannot be set twice") let process_ring_mods l = let kind = ref None in let set = ref None in let cst_tac = ref None in let pre = ref None in let post = ref None in let sign = ref None in let power = ref None in let div = ref None in List.iter(function Ring_kind k -> set_once "ring kind" kind (ic_coeff_spec k) | Const_tac t -> set_once "tactic recognizing constants" cst_tac t | Pre_tac t -> set_once "preprocess tactic" pre t | Post_tac t -> set_once "postprocess tactic" post t | Setoid(sth,ext) -> set_once "setoid" set (ic_unsafe sth,ic_unsafe ext) | Pow_spec(t,spec) -> set_once "power" power (t,spec) | Sign_spec t -> set_once "sign" sign t | Div_spec t -> set_once "div" div t) l; let k = match !kind with Some k -> k | None -> Abstract in (k, !set, !cst_tac, !pre, !post, !power, !sign, !div) let add_theory id rth l = let (sigma, rth) = ic rth in let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in add_theory0 id (sigma, rth) set k cst (pre,post) power sign div (*****************************************************************************) (* The tactics consist then only in a lookup in the ring database and call the appropriate ltac. *) let make_args_list sigma rl t = match rl with | [] -> let (_,t1,t2) = dest_rel0 sigma t in [t1;t2] | _ -> rl let make_term_list env evd carrier rl = let l = List.fold_right (fun x l -> plapp evd coq_cons [|carrier;x;l|]) rl (plapp evd coq_nil [|carrier|]) in let sigma, l = Typing.solve_evars env !evd l in evd := sigma; l let carg c = Tacinterp.Value.of_constr (EConstr.of_constr c) let tacarg expr = Tacinterp.Value.of_closure (Tacinterp.default_ist ()) expr let ltac_ring_structure e = let req = carg e.ring_req in let sth = carg e.ring_setoid in let ext = carg e.ring_ext in let morph = carg e.ring_morph in let th = carg e.ring_th in let cst_tac = tacarg e.ring_cst_tac in let pow_tac = tacarg e.ring_pow_tac in let lemma1 = carg e.ring_lemma1 in let lemma2 = carg e.ring_lemma2 in let pretac = tacarg (TacFun([Anonymous],e.ring_pre_tac)) in let posttac = tacarg (TacFun([Anonymous],e.ring_post_tac)) in [req;sth;ext;morph;th;cst_tac;pow_tac; lemma1;lemma2;pretac;posttac] let ring_lookup (f : Value.t) lH rl t = Proofview.Goal.enter begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in try (* find_ring_strucure can raise an exception *) let rl = make_args_list sigma rl t in let evdref = ref sigma in let e = find_ring_structure env sigma rl in let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.ring_carrier) rl) in let lH = carg (make_hyp_list env evdref lH) in let ring = ltac_ring_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Value.apply f (ring@[lH;rl])) with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e end (***********************************************************************) let new_field_path = DirPath.make (List.map Id.of_string ["Field_tac";plugin_dir;"Coq"]) let field_ltac s = lazy(KerName.make (ModPath.MPfile new_field_path) (Label.make s)) let _ = add_map "field" (map_with_eq [coq_cons,(function -1->Eval|2->Rec|_->Prot); coq_nil, (function -1->Eval|_ -> Prot); my_reference "IDphi", (function _->Eval); my_reference "gen_phiZ", (function _->Eval); (* display_linear: evaluate polynomials and coef operations, protect field operations and make recursive call on the var map *) my_reference "display_linear", (function -1|9|10|11|13|15|16->Eval|12|14->Rec|_->Prot); my_reference "display_pow_linear", (function -1|9|10|11|14|16|18|19->Eval|12|17->Rec|_->Prot); (* Pphi_dev: evaluate polynomial and coef operations, protect ring operations and make recursive call on the var map *) pol_cst "Pphi_dev", (function -1|8|9|10|12|14->Eval|11|13->Rec|_->Prot); pol_cst "Pphi_pow", (function -1|8|9|10|13|15|17->Eval|11|16->Rec|_->Prot); (* PEeval: evaluate polynomial, protect ring operations and make recursive call on the var map *) pol_cst "PEeval", (function -1|10|13->Eval|8|12->Rec|_->Prot); (* FEeval: evaluate polynomial, protect field operations and make recursive call on the var map *) my_reference "FEeval", (function -1|12|15->Eval|10|14->Rec|_->Prot)]);; let _ = add_map "field_cond" (map_without_eq [coq_cons,(function -1->Eval|2->Rec|_->Prot); coq_nil, (function -1->Eval|_ -> Prot); my_reference "IDphi", (function _->Eval); my_reference "gen_phiZ", (function _->Eval); (* PCond: evaluate denum list, protect ring operations and make recursive call on the var map *) my_reference "PCond", (function -1|11|14->Eval|9|13->Rec|_->Prot)]);; let _ = Redexpr.declare_reduction "simpl_field_expr" (protect_red "field") let afield_theory = my_reference "almost_field_theory" let field_theory = my_reference "field_theory" let sfield_theory = my_reference "semi_field_theory" let af_ar = my_reference"AF_AR" let f_r = my_reference"F_R" let sf_sr = my_reference"SF_SR" let dest_field env evd th_spec = let open Termops in let th_typ = Retyping.get_type_of env !evd th_spec in match EConstr.kind !evd th_typ with | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|]) when is_global !evd (Lazy.force afield_theory) f -> let rth = plapp evd af_ar [|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in (None,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth) | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|]) when is_global !evd (Lazy.force field_theory) f -> let rth = plapp evd f_r [|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in (Some false,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth) | App(f,[|r;zero;one;add;mul;div;inv;req|]) when is_global !evd (Lazy.force sfield_theory) f -> let rth = plapp evd sf_sr [|r;zero;one;add;mul;div;inv;req;th_spec|] in (Some true,r,zero,one,add,mul,None,None,div,inv,req,rth) | _ -> error "bad field structure" let field_from_carrier = Summary.ref Cmap.empty ~name:"field-tac-carrier-table" let print_fields () = Feedback.msg_notice (strbrk "The following field structures have been declared:"); Cmap.iter (fun _carrier fi -> let env = Global.env () in let sigma = Evd.from_env env in Feedback.msg_notice (hov 2 (Id.print fi.field_name ++ spc() ++ str"with carrier "++ pr_constr_env env sigma fi.field_carrier++spc()++ str"and equivalence relation "++ pr_constr_env env sigma fi.field_req)) ) !field_from_carrier let field_for_carrier r = Cmap.find r !field_from_carrier let find_field_structure env sigma l = check_required_library (cdir@["Field_tac"]); match l with | t::cl' -> let ty = Retyping.get_type_of env sigma t in let check c = let ty' = Retyping.get_type_of env sigma c in if not (Reductionops.is_conv env sigma ty ty') then CErrors.user_err ~hdr:"field" (str"arguments of field_simplify do not have all the same type") in List.iter check cl'; (try field_for_carrier (EConstr.to_constr sigma ty) with Not_found -> CErrors.user_err ~hdr:"field" (str"cannot find a declared field structure over"++ spc()++str"\""++pr_econstr_env env sigma ty++str"\"")) | [] -> assert false let add_field_entry (sp,_kn) e = field_from_carrier := Cmap.add e.field_carrier e !field_from_carrier let subst_th (subst,th) = let c' = subst_mps subst th.field_carrier in let eq' = subst_mps subst th.field_req in let thm1' = subst_mps subst th.field_ok in let thm2' = subst_mps subst th.field_simpl_eq_ok in let thm3' = subst_mps subst th.field_simpl_ok in let thm4' = subst_mps subst th.field_simpl_eq_in_ok in let thm5' = subst_mps subst th.field_cond in let tac'= Tacsubst.subst_tactic subst th.field_cst_tac in let pow_tac' = Tacsubst.subst_tactic subst th.field_pow_tac in let pretac'= Tacsubst.subst_tactic subst th.field_pre_tac in let posttac'= Tacsubst.subst_tactic subst th.field_post_tac in if c' == th.field_carrier && eq' == th.field_req && thm1' == th.field_ok && thm2' == th.field_simpl_eq_ok && thm3' == th.field_simpl_ok && thm4' == th.field_simpl_eq_in_ok && thm5' == th.field_cond && tac' == th.field_cst_tac && pow_tac' == th.field_pow_tac && pretac' == th.field_pre_tac && posttac' == th.field_post_tac then th else { field_name = th.field_name; field_carrier = c'; field_req = eq'; field_cst_tac = tac'; field_pow_tac = pow_tac'; field_ok = thm1'; field_simpl_eq_ok = thm2'; field_simpl_ok = thm3'; field_simpl_eq_in_ok = thm4'; field_cond = thm5'; field_pre_tac = pretac'; field_post_tac = posttac' } let ftheory_to_obj : field_info -> obj = let cache_th (name,th) = add_field_entry name th in declare_object @@ global_object_nodischarge "tactic-new-field-theory" ~cache:cache_th ~subst:(Some subst_th) let field_equality evd r inv req = match EConstr.kind !evd req with | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) -> let c = UnivGen.constr_of_monomorphic_global Coqlib.(lib_ref "core.eq.congr") in let c = EConstr.of_constr c in mkApp(c,[|r;r;inv|]) | _ -> let _setoid = setoid_of_relation (Global.env ()) evd r req in let signature = [Some (r,Some req)],Some(r,Some req) in let inv_m, inv_m_lem = try Rewrite.default_morphism signature inv with Not_found -> error "field inverse should be declared as a morphism" in inv_m_lem let add_field_theory0 name fth eqth morphth cst_tac inj (pre,post) power sign odiv = let open Constr in check_required_library (cdir@["Field_tac"]); let (sigma,fth) = ic fth in let env = Global.env() in let evd = ref sigma in let (kind,r,zero,one,add,mul,sub,opp,div,inv,req,rth) = dest_field env evd fth in let (sth,ext) = build_setoid_params env evd r add mul opp req eqth in let eqth = Some(sth,ext) in let _ = add_theory0 name (!evd,rth) eqth morphth cst_tac (None,None) power sign odiv in let (pow_tac, pspec) = interp_power env evd power in let sspec = interp_sign env evd sign in let dspec = interp_div env evd odiv in let inv_m = field_equality evd r inv req in let rk = reflect_coeff morphth in let params,ctx = exec_tactic env !evd 9 (field_ltac"field_lemmas") [sth;ext;inv_m;fth;pspec;sspec;dspec;rk] in let lemma1 = params.(3) in let lemma2 = params.(4) in let lemma3 = params.(5) in let lemma4 = params.(6) in let cond_lemma = match inj with | Some thm -> mkApp(params.(8),[|EConstr.to_constr sigma thm|]) | None -> params.(7) in let lemma1 = decl_constant (Id.to_string name^"_field_lemma1") ctx lemma1 in let lemma2 = decl_constant (Id.to_string name^"_field_lemma2") ctx lemma2 in let lemma3 = decl_constant (Id.to_string name^"_field_lemma3") ctx lemma3 in let lemma4 = decl_constant (Id.to_string name^"_field_lemma4") ctx lemma4 in let cond_lemma = decl_constant (Id.to_string name^"_lemma5") ctx cond_lemma in let cst_tac = interp_cst_tac env sigma morphth kind (zero,one,add,mul,opp) cst_tac in let pretac = match pre with Some t -> Tacintern.glob_tactic t | _ -> TacId [] in let posttac = match post with Some t -> Tacintern.glob_tactic t | _ -> TacId [] in let r = EConstr.to_constr sigma r in let req = EConstr.to_constr sigma req in let _ = Lib.add_leaf name (ftheory_to_obj { field_name = name; field_carrier = r; field_req = req; field_cst_tac = cst_tac; field_pow_tac = pow_tac; field_ok = lemma1; field_simpl_eq_ok = lemma2; field_simpl_ok = lemma3; field_simpl_eq_in_ok = lemma4; field_cond = cond_lemma; field_pre_tac = pretac; field_post_tac = posttac }) in () let process_field_mods l = let kind = ref None in let set = ref None in let cst_tac = ref None in let pre = ref None in let post = ref None in let inj = ref None in let sign = ref None in let power = ref None in let div = ref None in List.iter(function Ring_mod(Ring_kind k) -> set_once "field kind" kind (ic_coeff_spec k) | Ring_mod(Const_tac t) -> set_once "tactic recognizing constants" cst_tac t | Ring_mod(Pre_tac t) -> set_once "preprocess tactic" pre t | Ring_mod(Post_tac t) -> set_once "postprocess tactic" post t | Ring_mod(Setoid(sth,ext)) -> set_once "setoid" set (ic_unsafe sth,ic_unsafe ext) | Ring_mod(Pow_spec(t,spec)) -> set_once "power" power (t,spec) | Ring_mod(Sign_spec t) -> set_once "sign" sign t | Ring_mod(Div_spec t) -> set_once "div" div t | Inject i -> set_once "infinite property" inj (ic_unsafe i)) l; let k = match !kind with Some k -> k | None -> Abstract in (k, !set, !inj, !cst_tac, !pre, !post, !power, !sign, !div) let add_field_theory id t mods = let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods mods in add_field_theory0 id t set k cst_tac inj (pre,post) power sign div let ltac_field_structure e = let req = carg e.field_req in let cst_tac = tacarg e.field_cst_tac in let pow_tac = tacarg e.field_pow_tac in let field_ok = carg e.field_ok in let field_simpl_ok = carg e.field_simpl_ok in let field_simpl_eq_ok = carg e.field_simpl_eq_ok in let field_simpl_eq_in_ok = carg e.field_simpl_eq_in_ok in let cond_ok = carg e.field_cond in let pretac = tacarg (TacFun([Anonymous],e.field_pre_tac)) in let posttac = tacarg (TacFun([Anonymous],e.field_post_tac)) in [req;cst_tac;pow_tac;field_ok;field_simpl_ok;field_simpl_eq_ok; field_simpl_eq_in_ok;cond_ok;pretac;posttac] let field_lookup (f : Value.t) lH rl t = Proofview.Goal.enter begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in try let rl = make_args_list sigma rl t in let evdref = ref sigma in let e = find_field_structure env sigma rl in let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.field_carrier) rl) in let lH = carg (make_hyp_list env evdref lH) in let field = ltac_field_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Value.apply f (field@[lH;rl])) with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e end coq-8.11.0/plugins/setoid_ring/ArithRing.v0000644000175000017500000000435313612664533020340 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr:(N.of_nat t) | _ => constr:(InitialRing.NotConstant) end. Ltac Ss_to_add f acc := match f with | S ?f1 => Ss_to_add f1 (S acc) | _ => constr:((acc + f)%nat) end. (* For internal use only *) Local Definition protected_to_nat := N.to_nat. Ltac natprering := match goal with |- context C [S ?p] => match p with O => fail 1 (* avoid replacing 1 with 1+0 ! *) | p => match isnatcst p with | true => fail 1 | false => let v := Ss_to_add p (S 0) in fold v; natprering end end | _ => change N.to_nat with protected_to_nat end. Ltac natpostring := match goal with | |- context [N.to_nat ?x] => let v := eval cbv in (N.to_nat x) in change (N.to_nat x) with v; natpostring | _ => change protected_to_nat with N.to_nat end. Add Ring natr : natSRth (morphism nat_morph_N, constants [natcst], preprocess [natprering], postprocess [natpostring]). coq-8.11.0/plugins/setoid_ring/newring_ast.mli0000644000175000017500000000432213612664533021301 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* x == 0 \/ y == 0; integral_domain_one_zero: not (1 == 0)}. Section integral_domain. Context {R:Type}`{Rid:Integral_domain R}. Lemma integral_domain_minus_one_zero: ~ - (1:R) == 0. red;intro. apply integral_domain_one_zero. assert (0 == - (0:R)). cring. rewrite H0. rewrite <- H. cring. Qed. Definition pow (r : R) (n : nat) := Ring_theory.pow_N 1 mul r (N.of_nat n). Lemma pow_not_zero: forall p n, pow p n == 0 -> p == 0. induction n. unfold pow; simpl. intros. absurd (1 == 0). simpl. apply integral_domain_one_zero. trivial. setoid_replace (pow p (S n)) with (p * (pow p n)). intros. case (integral_domain_product p (pow p n) H). trivial. trivial. unfold pow; simpl. clear IHn. induction n; simpl; try cring. rewrite Ring_theory.pow_pos_succ. cring. apply ring_setoid. apply ring_mult_comp. apply ring_mul_assoc. Qed. Lemma Rintegral_domain_pow: forall c p r, ~c == 0 -> c * (pow p r) == ring0 -> p == ring0. intros. case (integral_domain_product c (pow p r) H0). intros; absurd (c == ring0); auto. intros. apply pow_not_zero with r. trivial. Qed. End integral_domain. coq-8.11.0/plugins/setoid_ring/g_newring.mlg0000644000175000017500000001153013612664533020735 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { protect_tac_in map id } | [ "protect_fv" string(map) ] -> { protect_tac map } END { open Pptactic open Ppconstr let pr_ring_mod env sigma = function | Ring_kind (Computational eq_test) -> str "decidable" ++ pr_arg (pr_constr_expr env sigma) eq_test | Ring_kind Abstract -> str "abstract" | Ring_kind (Morphism morph) -> str "morphism" ++ pr_arg (pr_constr_expr env sigma) morph | Const_tac (CstTac cst_tac) -> str "constants" ++ spc () ++ str "[" ++ pr_raw_tactic env sigma cst_tac ++ str "]" | Const_tac (Closed l) -> str "closed" ++ spc () ++ str "[" ++ prlist_with_sep spc pr_qualid l ++ str "]" | Pre_tac t -> str "preprocess" ++ spc () ++ str "[" ++ pr_raw_tactic env sigma t ++ str "]" | Post_tac t -> str "postprocess" ++ spc () ++ str "[" ++ pr_raw_tactic env sigma t ++ str "]" | Setoid(sth,ext) -> str "setoid" ++ pr_arg (pr_constr_expr env sigma) sth ++ pr_arg (pr_constr_expr env sigma) ext | Pow_spec(Closed l,spec) -> str "power_tac" ++ pr_arg (pr_constr_expr env sigma) spec ++ spc () ++ str "[" ++ prlist_with_sep spc pr_qualid l ++ str "]" | Pow_spec(CstTac cst_tac,spec) -> str "power_tac" ++ pr_arg (pr_constr_expr env sigma) spec ++ spc () ++ str "[" ++ pr_raw_tactic env sigma cst_tac ++ str "]" | Sign_spec t -> str "sign" ++ pr_arg (pr_constr_expr env sigma) t | Div_spec t -> str "div" ++ pr_arg (pr_constr_expr env sigma) t } VERNAC ARGUMENT EXTEND ring_mod PRINTED BY { pr_ring_mod env sigma } | [ "decidable" constr(eq_test) ] -> { Ring_kind(Computational eq_test) } | [ "abstract" ] -> { Ring_kind Abstract } | [ "morphism" constr(morph) ] -> { Ring_kind(Morphism morph) } | [ "constants" "[" tactic(cst_tac) "]" ] -> { Const_tac(CstTac cst_tac) } | [ "closed" "[" ne_global_list(l) "]" ] -> { Const_tac(Closed l) } | [ "preprocess" "[" tactic(pre) "]" ] -> { Pre_tac pre } | [ "postprocess" "[" tactic(post) "]" ] -> { Post_tac post } | [ "setoid" constr(sth) constr(ext) ] -> { Setoid(sth,ext) } | [ "sign" constr(sign_spec) ] -> { Sign_spec sign_spec } | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] -> { Pow_spec (Closed l, pow_spec) } | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] -> { Pow_spec (CstTac cst_tac, pow_spec) } | [ "div" constr(div_spec) ] -> { Div_spec div_spec } END { let pr_ring_mods env sigma l = surround (prlist_with_sep pr_comma (pr_ring_mod env sigma) l) } VERNAC ARGUMENT EXTEND ring_mods PRINTED BY { pr_ring_mods env sigma } | [ "(" ne_ring_mod_list_sep(mods, ",") ")" ] -> { mods } END VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] -> { add_theory id t (Option.default [] l) } | [ "Print" "Rings" ] => { Vernacextend.classify_as_query } -> { print_rings () } END TACTIC EXTEND ring_lookup | [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] -> { let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t } END { let pr_field_mod env sigma = function | Ring_mod m -> pr_ring_mod env sigma m | Inject inj -> str "completeness" ++ pr_arg (pr_constr_expr env sigma) inj } VERNAC ARGUMENT EXTEND field_mod PRINTED BY { pr_field_mod env sigma } | [ ring_mod(m) ] -> { Ring_mod m } | [ "completeness" constr(inj) ] -> { Inject inj } END { let pr_field_mods env sigma l = surround (prlist_with_sep pr_comma (pr_field_mod env sigma) l) } VERNAC ARGUMENT EXTEND field_mods PRINTED BY { pr_field_mods env sigma } | [ "(" ne_field_mod_list_sep(mods, ",") ")" ] -> { mods } END VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF | [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] -> { let l = match l with None -> [] | Some l -> l in add_field_theory id t l } | [ "Print" "Fields" ] => {Vernacextend.classify_as_query} -> { print_fields () } END TACTIC EXTEND field_lookup | [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] -> { let (t,l) = List.sep_last lt in field_lookup f lH l t } END coq-8.11.0/plugins/setoid_ring/Ncring_polynom.v0000644000175000017500000004141713612664533021450 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* : non commutative polynomials on a commutative ring A *) Set Implicit Arguments. Require Import Setoid. Require Import BinList. Require Import BinPos. Require Import BinNat. Require Import BinInt. Require Export Ring_polynom. (* n'utilise que PExpr *) Require Export Ncring. Section MakeRingPol. Context (C R:Type) `{Rh:Ring_morphism C R}. Variable phiCR_comm: forall (c:C)(x:R), x * [c] == [c] * x. Ltac rsimpl := repeat (gen_rewrite || rewrite phiCR_comm). Ltac add_push := gen_add_push . (* Definition of non commutative multivariable polynomials with coefficients in C : *) Inductive Pol : Type := | Pc : C -> Pol | PX : Pol -> positive -> positive -> Pol -> Pol. (* PX P i n Q represents P * X_i^n + Q *) Definition cO:C . exact ring0. Defined. Definition cI:C . exact ring1. Defined. Definition P0 := Pc 0. Definition P1 := Pc 1. Variable Ceqb:C->C->bool. #[universes(template)] Class Equalityb (A : Type):= {equalityb : A -> A -> bool}. Notation "x =? y" := (equalityb x y) (at level 70, no associativity). Variable Ceqb_eq: forall x y:C, Ceqb x y = true -> (x == y). Instance equalityb_coef : Equalityb C := {equalityb x y := Ceqb x y}. Fixpoint Peq (P P' : Pol) {struct P'} : bool := match P, P' with | Pc c, Pc c' => c =? c' | PX P i n Q, PX P' i' n' Q' => match Pos.compare i i', Pos.compare n n' with | Eq, Eq => if Peq P P' then Peq Q Q' else false | _,_ => false end | _, _ => false end. Instance equalityb_pol : Equalityb Pol := {equalityb x y := Peq x y}. (* Q a ses variables de queue < i *) Definition mkPX P i n Q := match P with | Pc c => if c =? 0 then Q else PX P i n Q | PX P' i' n' Q' => match Pos.compare i i' with | Eq => if Q' =? P0 then PX P' i (n + n') Q else PX P i n Q | _ => PX P i n Q end end. Definition mkXi i n := PX P1 i n P0. Definition mkX i := mkXi i 1. (** Opposite of addition *) Fixpoint Popp (P:Pol) : Pol := match P with | Pc c => Pc (- c) | PX P i n Q => PX (Popp P) i n (Popp Q) end. Notation "-- P" := (Popp P)(at level 30). (** Addition et subtraction *) Fixpoint PaddCl (c:C)(P:Pol) {struct P} : Pol := match P with | Pc c1 => Pc (c + c1) | PX P i n Q => PX P i n (PaddCl c Q) end. (* Q quelconque *) Section PaddX. Variable Padd:Pol->Pol->Pol. Variable P:Pol. (* Xi^n * P + Q les variables de tete de Q ne sont pas forcement < i mais Q est normalisé : variables de tete decroissantes *) Fixpoint PaddX (i n:positive)(Q:Pol){struct Q}:= match Q with | Pc c => mkPX P i n Q | PX P' i' n' Q' => match Pos.compare i i' with | (* i > i' *) Gt => mkPX P i n Q | (* i < i' *) Lt => mkPX P' i' n' (PaddX i n Q') | (* i = i' *) Eq => match Z.pos_sub n n' with | (* n > n' *) Zpos k => mkPX (PaddX i k P') i' n' Q' | (* n = n' *) Z0 => mkPX (Padd P P') i n Q' | (* n < n' *) Zneg k => mkPX (Padd P (mkPX P' i k P0)) i n Q' end end end. End PaddX. Fixpoint Padd (P1 P2: Pol) {struct P1} : Pol := match P1 with | Pc c => PaddCl c P2 | PX P' i' n' Q' => PaddX Padd P' i' n' (Padd Q' P2) end. Notation "P ++ P'" := (Padd P P'). Definition Psub(P P':Pol):= P ++ (--P'). Notation "P -- P'" := (Psub P P')(at level 50). (** Multiplication *) Fixpoint PmulC_aux (P:Pol) (c:C) {struct P} : Pol := match P with | Pc c' => Pc (c' * c) | PX P i n Q => mkPX (PmulC_aux P c) i n (PmulC_aux Q c) end. Definition PmulC P c := if c =? 0 then P0 else if c =? 1 then P else PmulC_aux P c. Fixpoint Pmul (P1 P2 : Pol) {struct P2} : Pol := match P2 with | Pc c => PmulC P1 c | PX P i n Q => PaddX Padd (Pmul P1 P) i n (Pmul P1 Q) end. Notation "P ** P'" := (Pmul P P')(at level 40). Definition Psquare (P:Pol) : Pol := P ** P. (** Evaluation of a polynomial towards R *) Fixpoint Pphi(l:list R) (P:Pol) {struct P} : R := match P with | Pc c => [c] | PX P i n Q => let x := nth 0 i l in let xn := pow_pos x n in (Pphi l P) * xn + (Pphi l Q) end. Reserved Notation "P @ l " (at level 10, no associativity). Notation "P @ l " := (Pphi l P). (** Proofs *) Ltac destr_pos_sub H := match goal with |- context [Z.pos_sub ?x ?y] => assert (H := Z.pos_sub_discr x y); destruct (Z.pos_sub x y) end. Lemma Peq_ok : forall P P', (P =? P') = true -> forall l, P@l == P'@ l. Proof. induction P;destruct P';simpl;intros ;try easy. - now apply ring_morphism_eq, Ceqb_eq. - specialize (IHP1 P'1). specialize (IHP2 P'2). simpl in IHP1, IHP2. destruct (Pos.compare_spec p p1); try discriminate; destruct (Pos.compare_spec p0 p2); try discriminate. destruct (Peq P2 P'1); try discriminate. subst; now rewrite IHP1, IHP2. Qed. Lemma Pphi0 : forall l, P0@l == 0. Proof. intros;simpl. rewrite ring_morphism0. reflexivity. Qed. Lemma Pphi1 : forall l, P1@l == 1. Proof. intros;simpl; rewrite ring_morphism1. reflexivity. Qed. Lemma mkPX_ok : forall l P i n Q, (mkPX P i n Q)@l == P@l * (pow_pos (nth 0 i l) n) + Q@l. Proof. intros l P i n Q;unfold mkPX. destruct P;try (simpl;reflexivity). assert (Hh := ring_morphism_eq c 0). simpl; case_eq (Ceqb c 0);simpl;try reflexivity. intros. rewrite Hh. rewrite ring_morphism0. rsimpl. apply Ceqb_eq. trivial. destruct (Pos.compare_spec i p). assert (Hh := @Peq_ok P3 P0). case_eq (P3=? P0). intro. simpl. rewrite Hh. rewrite Pphi0. rsimpl. rewrite Pos.add_comm. rewrite pow_pos_add;rsimpl. subst;trivial. reflexivity. trivial. intros. simpl. reflexivity. simpl. reflexivity. simpl. reflexivity. Qed. Ltac Esimpl := repeat (progress ( match goal with | |- context [?P@?l] => match P with | P0 => rewrite (Pphi0 l) | P1 => rewrite (Pphi1 l) | (mkPX ?P ?i ?n ?Q) => rewrite (mkPX_ok l P i n Q) end | |- context [[?c]] => match c with | 0 => rewrite ring_morphism0 | 1 => rewrite ring_morphism1 | ?x + ?y => rewrite ring_morphism_add | ?x * ?y => rewrite ring_morphism_mul | ?x - ?y => rewrite ring_morphism_sub | - ?x => rewrite ring_morphism_opp end end)); simpl; rsimpl. Lemma PaddCl_ok : forall c P l, (PaddCl c P)@l == [c] + P@l . Proof. induction P; simpl; intros; Esimpl; try reflexivity. rewrite IHP2. rsimpl. rewrite (ring_add_comm (P2 @ l * pow_pos (nth 0 p l) p0) [c]). reflexivity. Qed. Lemma PmulC_aux_ok : forall c P l, (PmulC_aux P c)@l == P@l * [c]. Proof. induction P;simpl;intros. rewrite ring_morphism_mul. try reflexivity. simpl. Esimpl. rewrite IHP1;rewrite IHP2;rsimpl. Qed. Lemma PmulC_ok : forall c P l, (PmulC P c)@l == P@l * [c]. Proof. intros c P l; unfold PmulC. assert (Hh:= ring_morphism_eq c 0);case_eq (c =? 0). intros. rewrite Hh;Esimpl. apply Ceqb_eq;trivial. assert (H1h:= ring_morphism_eq c 1);case_eq (c =? 1);intros. rewrite H1h;Esimpl. apply Ceqb_eq;trivial. apply PmulC_aux_ok. Qed. Lemma Popp_ok : forall P l, (--P)@l == - P@l. Proof. induction P;simpl;intros. Esimpl. rewrite IHP1;rewrite IHP2;rsimpl. Qed. Ltac Esimpl2 := Esimpl; repeat (progress ( match goal with | |- context [(PaddCl ?c ?P)@?l] => rewrite (PaddCl_ok c P l) | |- context [(PmulC ?P ?c)@?l] => rewrite (PmulC_ok c P l) | |- context [(--?P)@?l] => rewrite (Popp_ok P l) end)); Esimpl. Lemma PaddXPX: forall P i n Q, PaddX Padd P i n Q = match Q with | Pc c => mkPX P i n Q | PX P' i' n' Q' => match Pos.compare i i' with | (* i > i' *) Gt => mkPX P i n Q | (* i < i' *) Lt => mkPX P' i' n' (PaddX Padd P i n Q') | (* i = i' *) Eq => match Z.pos_sub n n' with | (* n > n' *) Zpos k => mkPX (PaddX Padd P i k P') i' n' Q' | (* n = n' *) Z0 => mkPX (Padd P P') i n Q' | (* n < n' *) Zneg k => mkPX (Padd P (mkPX P' i k P0)) i n Q' end end end. induction Q; reflexivity. Qed. Lemma PaddX_ok2 : forall P2, (forall P l, (P2 ++ P) @ l == P2 @ l + P @ l) /\ (forall P k n l, (PaddX Padd P2 k n P) @ l == P2 @ l * pow_pos (nth 0 k l) n + P @ l). induction P2;simpl;intros. split. intros. apply PaddCl_ok. induction P. unfold PaddX. intros. rewrite mkPX_ok. simpl. rsimpl. intros. simpl. destruct (Pos.compare_spec k p) as [Hh|Hh|Hh]. destr_pos_sub H1h. Esimpl2. rewrite Hh; trivial. rewrite H1h. reflexivity. simpl. rewrite mkPX_ok. rewrite IHP1. Esimpl2. rewrite Pos.add_comm in H1h. rewrite H1h. rewrite pow_pos_add. Esimpl2. rewrite Hh; trivial. reflexivity. rewrite mkPX_ok. rewrite PaddCl_ok. Esimpl2. rewrite Pos.add_comm in H1h. rewrite H1h. Esimpl2. rewrite pow_pos_add. Esimpl2. rewrite Hh; trivial. reflexivity. rewrite mkPX_ok. rewrite IHP2. Esimpl2. rewrite (ring_add_comm (P2 @ l * pow_pos (nth 0 p l) p0) ([c] * pow_pos (nth 0 k l) n)). reflexivity. assert (H1h := ring_morphism_eq c 0);case_eq (Ceqb c 0); intros; simpl. rewrite H1h;trivial. Esimpl2. apply Ceqb_eq; trivial. reflexivity. decompose [and] IHP2_1. decompose [and] IHP2_2. clear IHP2_1 IHP2_2. split. intros. rewrite H0. rewrite H1. Esimpl2. induction P. unfold PaddX. intros. rewrite mkPX_ok. simpl. reflexivity. intros. rewrite PaddXPX. destruct (Pos.compare_spec k p1) as [H3h|H3h|H3h]. destr_pos_sub H4h. rewrite mkPX_ok. simpl. rewrite H0. rewrite H1. Esimpl2. rewrite H4h. rewrite H3h;trivial. reflexivity. rewrite mkPX_ok. rewrite IHP1. Esimpl2. rewrite H3h;trivial. rewrite Pos.add_comm in H4h. rewrite H4h. rewrite pow_pos_add. Esimpl2. rewrite mkPX_ok. simpl. rewrite H0. rewrite H1. rewrite mkPX_ok. Esimpl2. rewrite H3h;trivial. rewrite Pos.add_comm in H4h. rewrite H4h. rewrite pow_pos_add. Esimpl2. rewrite mkPX_ok. simpl. rewrite IHP2. Esimpl2. gen_add_push (P2 @ l * pow_pos (nth 0 p1 l) p2). try reflexivity. rewrite mkPX_ok. simpl. reflexivity. Qed. Lemma Padd_ok : forall P Q l, (P ++ Q) @ l == P @ l + Q @ l. intro P. elim (PaddX_ok2 P); auto. Qed. Lemma PaddX_ok : forall P2 P k n l, (PaddX Padd P2 k n P) @ l == P2 @ l * pow_pos (nth 0 k l) n + P @ l. intro P2. elim (PaddX_ok2 P2); auto. Qed. Lemma Psub_ok : forall P' P l, (P -- P')@l == P@l - P'@l. unfold Psub. intros. rewrite Padd_ok. rewrite Popp_ok. rsimpl. Qed. Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l. induction P'; simpl; intros. rewrite PmulC_ok. reflexivity. rewrite PaddX_ok. rewrite IHP'1. rewrite IHP'2. Esimpl2. Qed. Lemma Psquare_ok : forall P l, (Psquare P)@l == P@l * P@l. Proof. intros. unfold Psquare. apply Pmul_ok. Qed. (** Definition of polynomial expressions *) (* Inductive PExpr : Type := | PEc : C -> PExpr | PEX : positive -> PExpr | PEadd : PExpr -> PExpr -> PExpr | PEsub : PExpr -> PExpr -> PExpr | PEmul : PExpr -> PExpr -> PExpr | PEopp : PExpr -> PExpr | PEpow : PExpr -> N -> PExpr. *) (** Specification of the power function *) Section POWER. Variable Cpow : Set. Variable Cp_phi : N -> Cpow. Variable rpow : R -> Cpow -> R. Record power_theory : Prop := mkpow_th { rpow_pow_N : forall r n, (rpow r (Cp_phi n))== (pow_N r n) }. End POWER. Variable Cpow : Set. Variable Cp_phi : N -> Cpow. Variable rpow : R -> Cpow -> R. Variable pow_th : power_theory Cp_phi rpow. (** evaluation of polynomial expressions towards R *) Fixpoint PEeval (l:list R) (pe:PExpr C) {struct pe} : R := match pe with | PEO => 0 | PEI => 1 | PEc c => [c] | PEX _ j => nth 0 j l | PEadd pe1 pe2 => (PEeval l pe1) + (PEeval l pe2) | PEsub pe1 pe2 => (PEeval l pe1) - (PEeval l pe2) | PEmul pe1 pe2 => (PEeval l pe1) * (PEeval l pe2) | PEopp pe1 => - (PEeval l pe1) | PEpow pe1 n => rpow (PEeval l pe1) (Cp_phi n) end. Strategy expand [PEeval]. Definition mk_X j := mkX j. (** Correctness proofs *) Lemma mkX_ok : forall p l, nth 0 p l == (mk_X p) @ l. Proof. destruct p;simpl;intros;Esimpl;trivial. Qed. Ltac Esimpl3 := repeat match goal with | |- context [(?P1 ++ ?P2)@?l] => rewrite (Padd_ok P1 P2 l) | |- context [(?P1 -- ?P2)@?l] => rewrite (Psub_ok P1 P2 l) end;try Esimpl2;try reflexivity;try apply ring_add_comm. (* Power using the chinise algorithm *) Section POWER2. Variable subst_l : Pol -> Pol. Fixpoint Ppow_pos (res P:Pol) (p:positive){struct p} : Pol := match p with | xH => subst_l (Pmul P res) | xO p => Ppow_pos (Ppow_pos res P p) P p | xI p => subst_l (Pmul P (Ppow_pos (Ppow_pos res P p) P p)) end. Definition Ppow_N P n := match n with | N0 => P1 | Npos p => Ppow_pos P1 P p end. Fixpoint pow_pos_gen (R:Type)(m:R->R->R)(x:R) (i:positive) {struct i}: R := match i with | xH => x | xO i => let p := pow_pos_gen m x i in m p p | xI i => let p := pow_pos_gen m x i in m x (m p p) end. Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) -> forall res P p, (Ppow_pos res P p)@l == (pow_pos_gen Pmul P p)@l * res@l. Proof. intros l subst_l_ok res P p. generalize res;clear res. induction p;simpl;intros. try rewrite subst_l_ok. repeat rewrite Pmul_ok. repeat rewrite IHp. rsimpl. repeat rewrite Pmul_ok. repeat rewrite IHp. rsimpl. try rewrite subst_l_ok. repeat rewrite Pmul_ok. reflexivity. Qed. Definition pow_N_gen (R:Type)(x1:R)(m:R->R->R)(x:R) (p:N) := match p with | N0 => x1 | Npos p => pow_pos_gen m x p end. Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) -> forall P n, (Ppow_N P n)@l == (pow_N_gen P1 Pmul P n)@l. Proof. destruct n;simpl. reflexivity. rewrite Ppow_pos_ok; trivial. Esimpl. Qed. End POWER2. (** Normalization and rewriting *) Section NORM_SUBST_REC. Let subst_l (P:Pol) := P. Let Pmul_subst P1 P2 := subst_l (Pmul P1 P2). Let Ppow_subst := Ppow_N subst_l. Fixpoint norm_aux (pe:PExpr C) : Pol := match pe with | PEO => Pc cO | PEI => Pc cI | PEc c => Pc c | PEX _ j => mk_X j | PEadd pe1 (PEopp pe2) => Psub (norm_aux pe1) (norm_aux pe2) | PEadd pe1 pe2 => Padd (norm_aux pe1) (norm_aux pe2) | PEsub pe1 pe2 => Psub (norm_aux pe1) (norm_aux pe2) | PEmul pe1 pe2 => Pmul (norm_aux pe1) (norm_aux pe2) | PEopp pe1 => Popp (norm_aux pe1) | PEpow pe1 n => Ppow_N (fun p => p) (norm_aux pe1) n end. Definition norm_subst pe := subst_l (norm_aux pe). Lemma norm_aux_spec : forall l pe, PEeval l pe == (norm_aux pe)@l. Proof. intros. induction pe. - now simpl; rewrite <- ring_morphism0. - now simpl; rewrite <- ring_morphism1. - Esimpl3. - Esimpl3. - simpl. rewrite IHpe1;rewrite IHpe2. destruct pe2; Esimpl3. unfold Psub. destruct pe1; destruct pe2; rewrite Padd_ok; rewrite Popp_ok; reflexivity. - simpl. unfold Psub. rewrite IHpe1;rewrite IHpe2. now destruct pe1; [destruct pe2; rewrite Padd_ok; rewrite Popp_ok; Esimpl3 | Esimpl3..]. - simpl. rewrite IHpe1;rewrite IHpe2. rewrite Pmul_ok. reflexivity. - now simpl; rewrite IHpe; Esimpl3. - simpl. rewrite Ppow_N_ok; (intros;try reflexivity). rewrite rpow_pow_N; [| now apply pow_th]. induction n;simpl; [now Esimpl3|]. induction p; simpl; trivial. + try rewrite IHp;try rewrite IHpe; repeat rewrite Pms_ok; repeat rewrite Pmul_ok;reflexivity. + rewrite Pmul_ok. try rewrite IHp;try rewrite IHpe; repeat rewrite Pms_ok; repeat rewrite Pmul_ok;reflexivity. Qed. Lemma norm_subst_spec : forall l pe, PEeval l pe == (norm_subst pe)@l. Proof. intros;unfold norm_subst. unfold subst_l. apply norm_aux_spec. Qed. End NORM_SUBST_REC. Fixpoint interp_PElist (l:list R) (lpe:list (PExpr C * PExpr C)) {struct lpe} : Prop := match lpe with | nil => True | (me,pe)::lpe => match lpe with | nil => PEeval l me == PEeval l pe | _ => PEeval l me == PEeval l pe /\ interp_PElist l lpe end end. Lemma norm_subst_ok : forall l pe, PEeval l pe == (norm_subst pe)@l. Proof. intros;apply norm_subst_spec. Qed. Lemma ring_correct : forall l pe1 pe2, (norm_subst pe1 =? norm_subst pe2) = true -> PEeval l pe1 == PEeval l pe2. Proof. simpl;intros. do 2 (rewrite (norm_subst_ok l);trivial). apply Peq_ok;trivial. Qed. End MakeRingPol. coq-8.11.0/plugins/setoid_ring/newring_ast.ml0000644000175000017500000000432213612664533021130 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* b) (eq(A:=bool)). split; simpl. destruct x; reflexivity. destruct x; destruct y; reflexivity. destruct x; destruct y; destruct z; reflexivity. reflexivity. destruct x; destruct y; reflexivity. destruct x; destruct y; reflexivity. destruct x; destruct y; destruct z; reflexivity. reflexivity. destruct x; reflexivity. Qed. Definition bool_eq (b1 b2:bool) := if b1 then b2 else negb b2. Lemma bool_eq_ok : forall b1 b2, bool_eq b1 b2 = true -> b1 = b2. destruct b1; destruct b2; auto. Qed. Ltac bool_cst t := let t := eval hnf in t in match t with true => constr:(true) | false => constr:(false) | _ => constr:(NotConstant) end. Add Ring bool_ring : BoolTheory (decidable bool_eq_ok, constants [bool_cst]). coq-8.11.0/plugins/setoid_ring/Rings_Z.v0000644000175000017500000000175413612664533020026 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 0%Z. omega. Qed. Instance Zdi : (Integral_domain (Rcr:=Zcri)). constructor. exact Zmult_integral. exact Z_one_zero. Defined. coq-8.11.0/plugins/setoid_ring/BinList.v0000644000175000017500000000446613612664533020022 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* tl l | xO p => jump p (jump p l) | xI p => jump p (jump p (tl l)) end. Fixpoint nth (p:positive) (l:list A) {struct p} : A:= match p with | xH => hd default l | xO p => nth p (jump p l) | xI p => nth p (jump p (tl l)) end. Lemma jump_tl : forall j l, tl (jump j l) = jump j (tl l). Proof. induction j;simpl;intros; now rewrite ?IHj. Qed. Lemma jump_succ : forall j l, jump (Pos.succ j) l = jump 1 (jump j l). Proof. induction j;simpl;intros. - rewrite !IHj; simpl; now rewrite !jump_tl. - now rewrite !jump_tl. - trivial. Qed. Lemma jump_add : forall i j l, jump (i + j) l = jump i (jump j l). Proof. induction i using Pos.peano_ind; intros. - now rewrite Pos.add_1_l, jump_succ. - now rewrite Pos.add_succ_l, !jump_succ, IHi. Qed. Lemma jump_pred_double : forall i l, jump (Pos.pred_double i) (tl l) = jump i (jump i l). Proof. induction i;intros;simpl. - now rewrite !jump_tl. - now rewrite IHi, <- 2 jump_tl, IHi. - trivial. Qed. Lemma nth_jump : forall p l, nth p (tl l) = hd default (jump p l). Proof. induction p;simpl;intros. - now rewrite <-jump_tl, IHp. - now rewrite <-jump_tl, IHp. - trivial. Qed. Lemma nth_pred_double : forall p l, nth (Pos.pred_double p) (tl l) = nth p (jump p l). Proof. induction p;simpl;intros. - now rewrite !jump_tl. - now rewrite jump_pred_double, <- !jump_tl, IHp. - trivial. Qed. End MakeBinList. coq-8.11.0/plugins/setoid_ring/Field_theory.v0000644000175000017500000015433113612664533021070 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* R->R) (ropp : R->R). Variable (rdiv : R->R->R) (rinv : R->R). Variable req : R -> R -> Prop. Notation "0" := rO : R_scope. Notation "1" := rI : R_scope. Infix "+" := radd : R_scope. Infix "-" := rsub : R_scope. Infix "*" := rmul : R_scope. Infix "/" := rdiv : R_scope. Notation "- x" := (ropp x) : R_scope. Notation "/ x" := (rinv x) : R_scope. Infix "==" := req (at level 70, no associativity) : R_scope. (* Equality properties *) Variable Rsth : Equivalence req. Variable Reqe : ring_eq_ext radd rmul ropp req. Variable SRinv_ext : forall p q, p == q -> / p == / q. (* Field properties *) Record almost_field_theory : Prop := mk_afield { AF_AR : almost_ring_theory rO rI radd rmul rsub ropp req; AF_1_neq_0 : ~ 1 == 0; AFdiv_def : forall p q, p / q == p * / q; AFinv_l : forall p, ~ p == 0 -> / p * p == 1 }. Section AlmostField. Variable AFth : almost_field_theory. Let ARth := (AF_AR AFth). Let rI_neq_rO := (AF_1_neq_0 AFth). Let rdiv_def := (AFdiv_def AFth). Let rinv_l := (AFinv_l AFth). Add Morphism radd with signature (req ==> req ==> req) as radd_ext. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext. Proof. exact (Ropp_ext Reqe). Qed. Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed. Add Morphism rinv with signature (req ==> req) as rinv_ext. Proof. exact SRinv_ext. Qed. Let eq_trans := Setoid.Seq_trans _ _ Rsth. Let eq_sym := Setoid.Seq_sym _ _ Rsth. Let eq_refl := Setoid.Seq_refl _ _ Rsth. Let radd_0_l := ARadd_0_l ARth. Let radd_comm := ARadd_comm ARth. Let radd_assoc := ARadd_assoc ARth. Let rmul_1_l := ARmul_1_l ARth. Let rmul_0_l := ARmul_0_l ARth. Let rmul_comm := ARmul_comm ARth. Let rmul_assoc := ARmul_assoc ARth. Let rdistr_l := ARdistr_l ARth. Let ropp_mul_l := ARopp_mul_l ARth. Let ropp_add := ARopp_add ARth. Let rsub_def := ARsub_def ARth. Let radd_0_r := ARadd_0_r Rsth ARth. Let rmul_0_r := ARmul_0_r Rsth ARth. Let rmul_1_r := ARmul_1_r Rsth ARth. Let ropp_0 := ARopp_zero Rsth Reqe ARth. Let rdistr_r := ARdistr_r Rsth Reqe ARth. (* Coefficients : C *) Variable C: Type. Declare Scope C_scope. Bind Scope C_scope with C. Delimit Scope C_scope with coef. Variable (cO cI: C) (cadd cmul csub : C->C->C) (copp : C->C). Variable ceqb : C->C->bool. Variable phi : C -> R. Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi. Notation "0" := cO : C_scope. Notation "1" := cI : C_scope. Infix "+" := cadd : C_scope. Infix "-" := csub : C_scope. Infix "*" := cmul : C_scope. Notation "- x" := (copp x) : C_scope. Infix "=?" := ceqb : C_scope. Notation "[ x ]" := (phi x) (at level 0). Let phi_0 := (morph0 CRmorph). Let phi_1 := (morph1 CRmorph). Lemma ceqb_spec c c' : BoolSpec ([c] == [c']) True (c =? c')%coef. Proof. generalize ((morph_eq CRmorph) c c'). destruct (c =? c')%coef; auto. Qed. (* Power coefficients : Cpow *) Variable Cpow : Type. Variable Cp_phi : N -> Cpow. Variable rpow : R -> Cpow -> R. Variable pow_th : power_theory rI rmul req Cp_phi rpow. (* sign function *) Variable get_sign : C -> option C. Variable get_sign_spec : sign_theory copp ceqb get_sign. Variable cdiv:C -> C -> C*C. Variable cdiv_th : div_theory req cadd cmul phi cdiv. Let rpow_pow := (rpow_pow_N pow_th). (* Polynomial expressions : (PExpr C) *) Declare Scope PE_scope. Bind Scope PE_scope with PExpr. Delimit Scope PE_scope with poly. Notation NPEeval := (PEeval rO rI radd rmul rsub ropp phi Cp_phi rpow). Notation "P @ l" := (NPEeval l P) (at level 10, no associativity). Arguments PEc _ _%coef. Notation "0" := (PEc 0) : PE_scope. Notation "1" := (PEc 1) : PE_scope. Infix "+" := PEadd : PE_scope. Infix "-" := PEsub : PE_scope. Infix "*" := PEmul : PE_scope. Notation "- e" := (PEopp e) : PE_scope. Infix "^" := PEpow : PE_scope. Definition NPEequiv e e' := forall l, e@l == e'@l. Infix "===" := NPEequiv (at level 70, no associativity) : PE_scope. Instance NPEequiv_eq : Equivalence NPEequiv. Proof. split; red; unfold NPEequiv; intros; [reflexivity|symmetry|etransitivity]; eauto. Qed. Instance NPEeval_ext : Proper (eq ==> NPEequiv ==> req) NPEeval. Proof. intros l l' <- e e' He. now rewrite (He l). Qed. Notation Nnorm := (norm_subst cO cI cadd cmul csub copp ceqb cdiv). Notation NPphi_dev := (Pphi_dev rO rI radd rmul rsub ropp cO cI ceqb phi get_sign). Notation NPphi_pow := (Pphi_pow rO rI radd rmul rsub ropp cO cI ceqb phi Cp_phi rpow get_sign). (* add abstract semi-ring to help with some proofs *) Add Ring Rring : (ARth_SRth ARth). (* additional ring properties *) Lemma rsub_0_l r : 0 - r == - r. Proof. rewrite rsub_def; ring. Qed. Lemma rsub_0_r r : r - 0 == r. Proof. rewrite rsub_def, ropp_0; ring. Qed. (*************************************************************************** Properties of division ***************************************************************************) Theorem rdiv_simpl p q : ~ q == 0 -> q * (p / q) == p. Proof. intros. rewrite rdiv_def. transitivity (/ q * q * p); [ ring | ]. now rewrite rinv_l. Qed. Instance rdiv_ext: Proper (req ==> req ==> req) rdiv. Proof. intros p1 p2 Ep q1 q2 Eq. now rewrite !rdiv_def, Ep, Eq. Qed. Lemma rmul_reg_l p q1 q2 : ~ p == 0 -> p * q1 == p * q2 -> q1 == q2. Proof. intros H EQ. assert (H' : p * (q1 / p) == p * (q2 / p)). { now rewrite !rdiv_def, !rmul_assoc, EQ. } now rewrite !rdiv_simpl in H'. Qed. Theorem field_is_integral_domain r1 r2 : ~ r1 == 0 -> ~ r2 == 0 -> ~ r1 * r2 == 0. Proof. intros H1 H2. contradict H2. transitivity (/r1 * r1 * r2). - now rewrite rinv_l. - now rewrite <- rmul_assoc, H2. Qed. Theorem ropp_neq_0 r : ~ -(1) == 0 -> ~ r == 0 -> ~ -r == 0. Proof. intros. setoid_replace (- r) with (- (1) * r). - apply field_is_integral_domain; trivial. - now rewrite <- ropp_mul_l, rmul_1_l. Qed. Theorem rdiv_r_r r : ~ r == 0 -> r / r == 1. Proof. intros. rewrite rdiv_def, rmul_comm. now apply rinv_l. Qed. Theorem rdiv1 r : r == r / 1. Proof. transitivity (1 * (r / 1)). - symmetry; apply rdiv_simpl. apply rI_neq_rO. - apply rmul_1_l. Qed. Theorem rdiv2 a b c d : ~ b == 0 -> ~ d == 0 -> a / b + c / d == (a * d + c * b) / (b * d). Proof. intros H H0. assert (~ b * d == 0) by now apply field_is_integral_domain. apply rmul_reg_l with (b * d); trivial. rewrite rdiv_simpl; trivial. rewrite rdistr_r. apply radd_ext. - now rewrite <- rmul_assoc, (rmul_comm d), rmul_assoc, rdiv_simpl. - now rewrite (rmul_comm c), <- rmul_assoc, rdiv_simpl. Qed. Theorem rdiv2b a b c d e : ~ (b*e) == 0 -> ~ (d*e) == 0 -> a / (b*e) + c / (d*e) == (a * d + c * b) / (b * (d * e)). Proof. intros H H0. assert (~ b == 0) by (contradict H; rewrite H; ring). assert (~ e == 0) by (contradict H; rewrite H; ring). assert (~ d == 0) by (contradict H0; rewrite H0; ring). assert (~ b * (d * e) == 0) by (repeat apply field_is_integral_domain; trivial). apply rmul_reg_l with (b * (d * e)); trivial. rewrite rdiv_simpl; trivial. rewrite rdistr_r. apply radd_ext. - transitivity ((b * e) * (a / (b * e)) * d); [ ring | now rewrite rdiv_simpl ]. - transitivity ((d * e) * (c / (d * e)) * b); [ ring | now rewrite rdiv_simpl ]. Qed. Theorem rdiv5 a b : - (a / b) == - a / b. Proof. now rewrite !rdiv_def, ropp_mul_l. Qed. Theorem rdiv3b a b c d e : ~ (b * e) == 0 -> ~ (d * e) == 0 -> a / (b*e) - c / (d*e) == (a * d - c * b) / (b * (d * e)). Proof. intros H H0. rewrite !rsub_def, rdiv5, ropp_mul_l. now apply rdiv2b. Qed. Theorem rdiv6 a b : ~ a == 0 -> ~ b == 0 -> / (a / b) == b / a. Proof. intros H H0. assert (Hk : ~ a / b == 0). { contradict H. transitivity (b * (a / b)). - now rewrite rdiv_simpl. - rewrite H. apply rmul_0_r. } apply rmul_reg_l with (a / b); trivial. rewrite (rmul_comm (a / b)), rinv_l; trivial. rewrite !rdiv_def. transitivity (/ a * a * (/ b * b)); [ | ring ]. now rewrite !rinv_l, rmul_1_l. Qed. Theorem rdiv4 a b c d : ~ b == 0 -> ~ d == 0 -> (a / b) * (c / d) == (a * c) / (b * d). Proof. intros H H0. assert (~ b * d == 0) by now apply field_is_integral_domain. apply rmul_reg_l with (b * d); trivial. rewrite rdiv_simpl; trivial. transitivity (b * (a / b) * (d * (c / d))); [ ring | ]. rewrite !rdiv_simpl; trivial. Qed. Theorem rdiv4b a b c d e f : ~ b * e == 0 -> ~ d * f == 0 -> ((a * f) / (b * e)) * ((c * e) / (d * f)) == (a * c) / (b * d). Proof. intros H H0. assert (~ b == 0) by (contradict H; rewrite H; ring). assert (~ e == 0) by (contradict H; rewrite H; ring). assert (~ d == 0) by (contradict H0; rewrite H0; ring). assert (~ f == 0) by (contradict H0; rewrite H0; ring). assert (~ b*d == 0) by now apply field_is_integral_domain. assert (~ e*f == 0) by now apply field_is_integral_domain. rewrite rdiv4; trivial. transitivity ((e * f) * (a * c) / ((e * f) * (b * d))). - apply rdiv_ext; ring. - rewrite <- rdiv4, rdiv_r_r; trivial. Qed. Theorem rdiv7 a b c d : ~ b == 0 -> ~ c == 0 -> ~ d == 0 -> (a / b) / (c / d) == (a * d) / (b * c). Proof. intros. rewrite (rdiv_def (a / b)). rewrite rdiv6; trivial. apply rdiv4; trivial. Qed. Theorem rdiv7b a b c d e f : ~ b * f == 0 -> ~ c * e == 0 -> ~ d * f == 0 -> ((a * e) / (b * f)) / ((c * e) / (d * f)) == (a * d) / (b * c). Proof. intros Hbf Hce Hdf. assert (~ c==0) by (contradict Hce; rewrite Hce; ring). assert (~ e==0) by (contradict Hce; rewrite Hce; ring). assert (~ b==0) by (contradict Hbf; rewrite Hbf; ring). assert (~ f==0) by (contradict Hbf; rewrite Hbf; ring). assert (~ b*c==0) by now apply field_is_integral_domain. assert (~ e*f==0) by now apply field_is_integral_domain. rewrite rdiv7; trivial. transitivity ((e * f) * (a * d) / ((e * f) * (b * c))). - apply rdiv_ext; ring. - now rewrite <- rdiv4, rdiv_r_r. Qed. Theorem rinv_nz a : ~ a == 0 -> ~ /a == 0. Proof. intros H H0. apply rI_neq_rO. rewrite <- (rdiv_r_r H), rdiv_def, H0. apply rmul_0_r. Qed. Theorem rdiv8 a b : ~ b == 0 -> a == 0 -> a / b == 0. Proof. intros H H0. now rewrite rdiv_def, H0, rmul_0_l. Qed. Theorem cross_product_eq a b c d : ~ b == 0 -> ~ d == 0 -> a * d == c * b -> a / b == c / d. Proof. intros. transitivity (a / b * (d / d)). - now rewrite rdiv_r_r, rmul_1_r. - now rewrite rdiv4, H1, (rmul_comm b d), <- rdiv4, rdiv_r_r. Qed. (* Results about [pow_pos] and [pow_N] *) Instance pow_ext : Proper (req ==> eq ==> req) (pow_pos rmul). Proof. intros x y H p p' <-. induction p as [p IH| p IH|];simpl; trivial; now rewrite !IH, ?H. Qed. Instance pow_N_ext : Proper (req ==> eq ==> req) (pow_N rI rmul). Proof. intros x y H n n' <-. destruct n; simpl; trivial. now apply pow_ext. Qed. Lemma pow_pos_0 p : pow_pos rmul 0 p == 0. Proof. induction p;simpl;trivial; now rewrite !IHp. Qed. Lemma pow_pos_1 p : pow_pos rmul 1 p == 1. Proof. induction p;simpl;trivial; ring [IHp]. Qed. Lemma pow_pos_cst c p : pow_pos rmul [c] p == [pow_pos cmul c p]. Proof. induction p;simpl;trivial; now rewrite !(morph_mul CRmorph), !IHp. Qed. Lemma pow_pos_mul_l x y p : pow_pos rmul (x * y) p == pow_pos rmul x p * pow_pos rmul y p. Proof. induction p;simpl;trivial; ring [IHp]. Qed. Lemma pow_pos_add_r x p1 p2 : pow_pos rmul x (p1+p2) == pow_pos rmul x p1 * pow_pos rmul x p2. Proof. exact (Ring_theory.pow_pos_add Rsth rmul_ext rmul_assoc x p1 p2). Qed. Lemma pow_pos_mul_r x p1 p2 : pow_pos rmul x (p1*p2) == pow_pos rmul (pow_pos rmul x p1) p2. Proof. induction p1;simpl;intros; rewrite ?pow_pos_mul_l, ?pow_pos_add_r; simpl; trivial; ring [IHp1]. Qed. Lemma pow_pos_nz x p : ~x==0 -> ~pow_pos rmul x p == 0. Proof. intros Hx. induction p;simpl;trivial; repeat (apply field_is_integral_domain; trivial). Qed. Lemma pow_pos_div a b p : ~ b == 0 -> pow_pos rmul (a / b) p == pow_pos rmul a p / pow_pos rmul b p. Proof. intros. induction p; simpl; trivial. - rewrite IHp. assert (nz := pow_pos_nz p H). rewrite !rdiv4; trivial. apply field_is_integral_domain; trivial. - rewrite IHp. assert (nz := pow_pos_nz p H). rewrite !rdiv4; trivial. Qed. (* === is a morphism *) Instance PEadd_ext : Proper (NPEequiv ==> NPEequiv ==> NPEequiv) (@PEadd C). Proof. intros ? ? E ? ? E' l. simpl. now rewrite E, E'. Qed. Instance PEsub_ext : Proper (NPEequiv ==> NPEequiv ==> NPEequiv) (@PEsub C). Proof. intros ? ? E ? ? E' l. simpl. now rewrite E, E'. Qed. Instance PEmul_ext : Proper (NPEequiv ==> NPEequiv ==> NPEequiv) (@PEmul C). Proof. intros ? ? E ? ? E' l. simpl. now rewrite E, E'. Qed. Instance PEopp_ext : Proper (NPEequiv ==> NPEequiv) (@PEopp C). Proof. intros ? ? E l. simpl. now rewrite E. Qed. Instance PEpow_ext : Proper (NPEequiv ==> eq ==> NPEequiv) (@PEpow C). Proof. intros ? ? E ? ? <- l. simpl. rewrite !rpow_pow. apply pow_N_ext; trivial. Qed. Lemma PE_1_l (e : PExpr C) : (1 * e === e)%poly. Proof. intros l. simpl. rewrite phi_1. apply rmul_1_l. Qed. Lemma PE_1_r (e : PExpr C) : (e * 1 === e)%poly. Proof. intros l. simpl. rewrite phi_1. apply rmul_1_r. Qed. Lemma PEpow_0_r (e : PExpr C) : (e ^ 0 === 1)%poly. Proof. intros l. simpl. now rewrite !rpow_pow. Qed. Lemma PEpow_1_r (e : PExpr C) : (e ^ 1 === e)%poly. Proof. intros l. simpl. now rewrite !rpow_pow. Qed. Lemma PEpow_1_l n : (1 ^ n === 1)%poly. Proof. intros l. simpl. rewrite rpow_pow. destruct n; simpl. - now rewrite phi_1. - now rewrite phi_1, pow_pos_1. Qed. Lemma PEpow_add_r (e : PExpr C) n n' : (e ^ (n+n') === e ^ n * e ^ n')%poly. Proof. intros l. simpl. rewrite !rpow_pow. destruct n; simpl. - rewrite rmul_1_l. trivial. - destruct n'; simpl. + rewrite rmul_1_r. trivial. + apply pow_pos_add_r. Qed. Lemma PEpow_mul_l (e e' : PExpr C) n : ((e * e') ^ n === e ^ n * e' ^ n)%poly. Proof. intros l. simpl. rewrite !rpow_pow. destruct n; simpl; trivial. - symmetry; apply rmul_1_l. - apply pow_pos_mul_l. Qed. Lemma PEpow_mul_r (e : PExpr C) n n' : (e ^ (n * n') === (e ^ n) ^ n')%poly. Proof. intros l. simpl. rewrite !rpow_pow. destruct n, n'; simpl; trivial. - now rewrite pow_pos_1. - apply pow_pos_mul_r. Qed. Lemma PEpow_nz l e n : ~ e @ l == 0 -> ~ (e^n) @ l == 0. Proof. intros. simpl. rewrite rpow_pow. destruct n; simpl. - apply rI_neq_rO. - now apply pow_pos_nz. Qed. (*************************************************************************** Some equality test ***************************************************************************) Local Notation "a &&& b" := (if a then b else false) (at level 40, left associativity). (* equality test *) Fixpoint PExpr_eq (e e' : PExpr C) {struct e} : bool := match e, e' with | PEc c, PEc c' => ceqb c c' | PEX _ p, PEX _ p' => Pos.eqb p p' | e1 + e2, e1' + e2' => PExpr_eq e1 e1' &&& PExpr_eq e2 e2' | e1 - e2, e1' - e2' => PExpr_eq e1 e1' &&& PExpr_eq e2 e2' | e1 * e2, e1' * e2' => PExpr_eq e1 e1' &&& PExpr_eq e2 e2' | - e, - e' => PExpr_eq e e' | e ^ n, e' ^ n' => N.eqb n n' &&& PExpr_eq e e' | _, _ => false end%poly. Lemma if_true (a b : bool) : a &&& b = true -> a = true /\ b = true. Proof. destruct a, b; split; trivial. Qed. Theorem PExpr_eq_semi_ok e e' : PExpr_eq e e' = true -> (e === e')%poly. Proof. revert e'; induction e; destruct e'; simpl; try discriminate. - intros H l. now apply (morph_eq CRmorph). - case Pos.eqb_spec; intros; now subst. - intros H; destruct (if_true _ _ H). now rewrite IHe1, IHe2. - intros H; destruct (if_true _ _ H). now rewrite IHe1, IHe2. - intros H; destruct (if_true _ _ H). now rewrite IHe1, IHe2. - intros H. now rewrite IHe. - intros H. destruct (if_true _ _ H). apply N.eqb_eq in H0. now rewrite IHe, H0. Qed. Lemma PExpr_eq_spec e e' : BoolSpec (e === e')%poly True (PExpr_eq e e'). Proof. assert (H := PExpr_eq_semi_ok e e'). destruct PExpr_eq; constructor; intros; trivial. now apply H. Qed. (** Smart constructors for polynomial expression, with reduction of constants *) Definition NPEadd e1 e2 := match e1, e2 with | PEc c1, PEc c2 => PEc (c1 + c2) | PEc c, _ => if (c =? 0)%coef then e2 else e1 + e2 | _, PEc c => if (c =? 0)%coef then e1 else e1 + e2 (* Peut t'on factoriser ici ??? *) | _, _ => (e1 + e2) end%poly. Infix "++" := NPEadd (at level 60, right associativity). Theorem NPEadd_ok e1 e2 : (e1 ++ e2 === e1 + e2)%poly. Proof. intros l. destruct e1, e2; simpl; try reflexivity; try (case ceqb_spec); try intro H; try rewrite H; simpl; try apply eq_refl; try (ring [phi_0]). apply (morph_add CRmorph). Qed. Definition NPEsub e1 e2 := match e1, e2 with | PEc c1, PEc c2 => PEc (c1 - c2) | PEc c, _ => if (c =? 0)%coef then - e2 else e1 - e2 | _, PEc c => if (c =? 0)%coef then e1 else e1 - e2 (* Peut-on factoriser ici *) | _, _ => e1 - e2 end%poly. Infix "--" := NPEsub (at level 50, left associativity). Theorem NPEsub_ok e1 e2: (e1 -- e2 === e1 - e2)%poly. Proof. intros l. destruct e1, e2; simpl; try reflexivity; try case ceqb_spec; try intro H; try rewrite H; simpl; try rewrite phi_0; try reflexivity; try (symmetry; apply rsub_0_l); try (symmetry; apply rsub_0_r). apply (morph_sub CRmorph). Qed. Definition NPEopp e1 := match e1 with PEc c1 => PEc (- c1) | _ => - e1 end%poly. Theorem NPEopp_ok e : (NPEopp e === -e)%poly. Proof. intros l. destruct e; simpl; trivial. apply (morph_opp CRmorph). Qed. Definition NPEpow x n := match n with | N0 => 1 | Npos p => if (p =? 1)%positive then x else match x with | PEc c => if (c =? 1)%coef then 1 else if (c =? 0)%coef then 0 else PEc (pow_pos cmul c p) | _ => x ^ n end end%poly. Infix "^^" := NPEpow (at level 35, right associativity). Theorem NPEpow_ok e n : (e ^^ n === e ^ n)%poly. Proof. intros l. unfold NPEpow; destruct n. - simpl; now rewrite rpow_pow. - case Pos.eqb_spec; [intro; subst | intros _]. + simpl. now rewrite rpow_pow. + destruct e;simpl;trivial. repeat case ceqb_spec; intros; rewrite ?rpow_pow, ?H; simpl. * now rewrite phi_1, pow_pos_1. * now rewrite phi_0, pow_pos_0. * now rewrite pow_pos_cst. Qed. Fixpoint NPEmul (x y : PExpr C) {struct x} : PExpr C := match x, y with | PEc c1, PEc c2 => PEc (c1 * c2) | PEc c, _ => if (c =? 1)%coef then y else if (c =? 0)%coef then 0 else x * y | _, PEc c => if (c =? 1)%coef then x else if (c =? 0)%coef then 0 else x * y | e1 ^ n1, e2 ^ n2 => if (n1 =? n2)%N then (NPEmul e1 e2)^^n1 else x * y | _, _ => x * y end%poly. Infix "**" := NPEmul (at level 40, left associativity). Theorem NPEmul_ok e1 e2 : (e1 ** e2 === e1 * e2)%poly. Proof. intros l. revert e2; induction e1;destruct e2; simpl;try reflexivity; repeat (case ceqb_spec; intro H; try rewrite H; clear H); simpl; try reflexivity; try ring [phi_0 phi_1]. apply (morph_mul CRmorph). case N.eqb_spec; [intros <- | reflexivity]. rewrite NPEpow_ok. simpl. rewrite !rpow_pow. rewrite IHe1. destruct n; simpl; [ ring | apply pow_pos_mul_l ]. Qed. (* simplification *) Fixpoint PEsimp (e : PExpr C) : PExpr C := match e with | e1 + e2 => (PEsimp e1) ++ (PEsimp e2) | e1 * e2 => (PEsimp e1) ** (PEsimp e2) | e1 - e2 => (PEsimp e1) -- (PEsimp e2) | - e1 => NPEopp (PEsimp e1) | e1 ^ n1 => (PEsimp e1) ^^ n1 | _ => e end%poly. Theorem PEsimp_ok e : (PEsimp e === e)%poly. Proof. induction e; simpl. - reflexivity. - reflexivity. - intro l; trivial. - intro l; trivial. - rewrite NPEadd_ok. now f_equiv. - rewrite NPEsub_ok. now f_equiv. - rewrite NPEmul_ok. now f_equiv. - rewrite NPEopp_ok. now f_equiv. - rewrite NPEpow_ok. now f_equiv. Qed. (**************************************************************************** Datastructure ***************************************************************************) (* The input: syntax of a field expression *) Inductive FExpr : Type := | FEO : FExpr | FEI : FExpr | FEc: C -> FExpr | FEX: positive -> FExpr | FEadd: FExpr -> FExpr -> FExpr | FEsub: FExpr -> FExpr -> FExpr | FEmul: FExpr -> FExpr -> FExpr | FEopp: FExpr -> FExpr | FEinv: FExpr -> FExpr | FEdiv: FExpr -> FExpr -> FExpr | FEpow: FExpr -> N -> FExpr . Fixpoint FEeval (l : list R) (pe : FExpr) {struct pe} : R := match pe with | FEO => rO | FEI => rI | FEc c => phi c | FEX x => BinList.nth 0 x l | FEadd x y => FEeval l x + FEeval l y | FEsub x y => FEeval l x - FEeval l y | FEmul x y => FEeval l x * FEeval l y | FEopp x => - FEeval l x | FEinv x => / FEeval l x | FEdiv x y => FEeval l x / FEeval l y | FEpow x n => rpow (FEeval l x) (Cp_phi n) end. Strategy expand [FEeval]. (* The result of the normalisation *) Record linear : Type := mk_linear { num : PExpr C; denum : PExpr C; condition : list (PExpr C) }. (*************************************************************************** Semantics and properties of side condition ***************************************************************************) Fixpoint PCond (l : list R) (le : list (PExpr C)) {struct le} : Prop := match le with | nil => True | e1 :: nil => ~ req (e1 @ l) rO | e1 :: l1 => ~ req (e1 @ l) rO /\ PCond l l1 end. Theorem PCond_cons l a l1 : PCond l (a :: l1) <-> ~ a @ l == 0 /\ PCond l l1. Proof. destruct l1. - simpl. split; [split|destruct 1]; trivial. - reflexivity. Qed. Theorem PCond_cons_inv_l l a l1 : PCond l (a::l1) -> ~ a @ l == 0. Proof. rewrite PCond_cons. now destruct 1. Qed. Theorem PCond_cons_inv_r l a l1 : PCond l (a :: l1) -> PCond l l1. Proof. rewrite PCond_cons. now destruct 1. Qed. Theorem PCond_app l l1 l2 : PCond l (l1 ++ l2) <-> PCond l l1 /\ PCond l l2. Proof. induction l1. - simpl. split; [split|destruct 1]; trivial. - simpl app. rewrite !PCond_cons, IHl1. symmetry; apply and_assoc. Qed. (* An unsatisfiable condition: issued when a division by zero is detected *) Definition absurd_PCond := cons 0%poly nil. Lemma absurd_PCond_bottom : forall l, ~ PCond l absurd_PCond. Proof. unfold absurd_PCond; simpl. red; intros. apply H. apply phi_0. Qed. (*************************************************************************** Normalisation ***************************************************************************) Definition default_isIn e1 p1 e2 p2 := if PExpr_eq e1 e2 then match Z.pos_sub p1 p2 with | Zpos p => Some (Npos p, 1%poly) | Z0 => Some (N0, 1%poly) | Zneg p => Some (N0, e2 ^^ Npos p) end else None. Fixpoint isIn e1 p1 e2 p2 {struct e2}: option (N * PExpr C) := match e2 with | e3 * e4 => match isIn e1 p1 e3 p2 with | Some (N0, e5) => Some (N0, e5 ** (e4 ^^ Npos p2)) | Some (Npos p, e5) => match isIn e1 p e4 p2 with | Some (n, e6) => Some (n, e5 ** e6) | None => Some (Npos p, e5 ** (e4 ^^ Npos p2)) end | None => match isIn e1 p1 e4 p2 with | Some (n, e5) => Some (n, (e3 ^^ Npos p2) ** e5) | None => None end end | e3 ^ N0 => None | e3 ^ Npos p3 => isIn e1 p1 e3 (Pos.mul p3 p2) | _ => default_isIn e1 p1 e2 p2 end%poly. Definition ZtoN z := match z with Zpos p => Npos p | _ => N0 end. Definition NtoZ n := match n with Npos p => Zpos p | _ => Z0 end. Lemma Z_pos_sub_gt p q : (p > q)%positive -> Z.pos_sub p q = Zpos (p - q). Proof. intros; now apply Z.pos_sub_gt, Pos.gt_lt. Qed. Ltac simpl_pos_sub := rewrite ?Z_pos_sub_gt in * by assumption. Lemma default_isIn_ok e1 e2 p1 p2 : match default_isIn e1 p1 e2 p2 with | Some(n, e3) => let n' := ZtoN (Zpos p1 - NtoZ n) in (e2 ^ N.pos p2 === e1 ^ n' * e3)%poly /\ (Zpos p1 > NtoZ n)%Z | _ => True end. Proof. unfold default_isIn. case PExpr_eq_spec; trivial. intros EQ. rewrite Z.pos_sub_spec. case Pos.compare_spec;intros H; split; try reflexivity. - simpl. now rewrite PE_1_r, H, EQ. - rewrite NPEpow_ok, EQ, <- PEpow_add_r. f_equiv. simpl. f_equiv. now rewrite Pos.add_comm, Pos.sub_add. - simpl. rewrite PE_1_r, EQ. f_equiv. rewrite Z.pos_sub_gt by now apply Pos.sub_decr. simpl. f_equiv. rewrite Pos.sub_sub_distr, Pos.add_comm; trivial. rewrite Pos.add_sub; trivial. apply Pos.sub_decr; trivial. - simpl. now apply Z.lt_gt, Pos.sub_decr. Qed. Ltac npe_simpl := rewrite ?NPEmul_ok, ?NPEpow_ok, ?PEpow_mul_l. Ltac npe_ring := intro l; simpl; ring. Theorem isIn_ok e1 p1 e2 p2 : match isIn e1 p1 e2 p2 with | Some(n, e3) => let n' := ZtoN (Zpos p1 - NtoZ n) in (e2 ^ N.pos p2 === e1 ^ n' * e3)%poly /\ (Zpos p1 > NtoZ n)%Z | _ => True end. Proof. Opaque NPEpow. revert p1 p2. induction e2; intros p1 p2; try refine (default_isIn_ok e1 _ p1 p2); simpl isIn. - specialize (IHe2_1 p1 p2). destruct isIn as [([|p],e)|]. + split; [|reflexivity]. clear IHe2_2. destruct IHe2_1 as (IH,_). npe_simpl. rewrite IH. npe_ring. + specialize (IHe2_2 p p2). destruct isIn as [([|p'],e')|]. * destruct IHe2_1 as (IH1,GT1). destruct IHe2_2 as (IH2,GT2). split; [|simpl; apply Zgt_trans with (Z.pos p); trivial]. npe_simpl. rewrite IH1, IH2. simpl. simpl_pos_sub. simpl. replace (N.pos p1) with (N.pos p + N.pos (p1 - p))%N. rewrite PEpow_add_r; npe_ring. { simpl. f_equal. rewrite Pos.add_comm, Pos.sub_add. trivial. now apply Pos.gt_lt. } * destruct IHe2_1 as (IH1,GT1). destruct IHe2_2 as (IH2,GT2). assert (Z.pos p1 > Z.pos p')%Z by (now apply Zgt_trans with (Zpos p)). split; [|simpl; trivial]. npe_simpl. rewrite IH1, IH2. simpl. simpl_pos_sub. simpl. replace (N.pos (p1 - p')) with (N.pos (p1 - p) + N.pos (p - p'))%N. rewrite PEpow_add_r; npe_ring. { simpl. f_equal. rewrite Pos.add_sub_assoc, Pos.sub_add; trivial. now apply Pos.gt_lt. now apply Pos.gt_lt. } * destruct IHe2_1 as (IH,GT). split; trivial. npe_simpl. rewrite IH. npe_ring. + specialize (IHe2_2 p1 p2). destruct isIn as [(n,e)|]; trivial. destruct IHe2_2 as (IH,GT). split; trivial. set (d := ZtoN (Z.pos p1 - NtoZ n)) in *; clearbody d. npe_simpl. rewrite IH. npe_ring. - destruct n; trivial. specialize (IHe2 p1 (p * p2)%positive). destruct isIn as [(n,e)|]; trivial. destruct IHe2 as (IH,GT). split; trivial. set (d := ZtoN (Z.pos p1 - NtoZ n)) in *; clearbody d. now rewrite <- PEpow_mul_r. Qed. Record rsplit : Type := mk_rsplit { rsplit_left : PExpr C; rsplit_common : PExpr C; rsplit_right : PExpr C}. (* Stupid name clash *) Notation left := rsplit_left. Notation right := rsplit_right. Notation common := rsplit_common. Fixpoint split_aux e1 p e2 {struct e1}: rsplit := match e1 with | e3 * e4 => let r1 := split_aux e3 p e2 in let r2 := split_aux e4 p (right r1) in mk_rsplit (left r1 ** left r2) (common r1 ** common r2) (right r2) | e3 ^ N0 => mk_rsplit 1 1 e2 | e3 ^ Npos p3 => split_aux e3 (Pos.mul p3 p) e2 | _ => match isIn e1 p e2 1 with | Some (N0,e3) => mk_rsplit 1 (e1 ^^ Npos p) e3 | Some (Npos q, e3) => mk_rsplit (e1 ^^ Npos q) (e1 ^^ Npos (p - q)) e3 | None => mk_rsplit (e1 ^^ Npos p) 1 e2 end end%poly. Lemma split_aux_ok1 e1 p e2 : (let res := match isIn e1 p e2 1 with | Some (N0,e3) => mk_rsplit 1 (e1 ^^ Npos p) e3 | Some (Npos q, e3) => mk_rsplit (e1 ^^ Npos q) (e1 ^^ Npos (p - q)) e3 | None => mk_rsplit (e1 ^^ Npos p) 1 e2 end in e1 ^ Npos p === left res * common res /\ e2 === right res * common res)%poly. Proof. Opaque NPEpow NPEmul. intros. unfold res;clear res; generalize (isIn_ok e1 p e2 xH). destruct (isIn e1 p e2 1) as [([|p'],e')|]; simpl. - intros (H1,H2); split; npe_simpl. + now rewrite PE_1_l. + rewrite PEpow_1_r in H1. rewrite H1. npe_ring. - intros (H1,H2); split; npe_simpl. + rewrite <- PEpow_add_r. f_equiv. simpl. f_equal. rewrite Pos.add_comm, Pos.sub_add; trivial. now apply Z.gt_lt in H2. + rewrite PEpow_1_r in H1. rewrite H1. simpl_pos_sub. simpl. npe_ring. - intros _; split; npe_simpl; now rewrite PE_1_r. Qed. Theorem split_aux_ok: forall e1 p e2, (e1 ^ Npos p === left (split_aux e1 p e2) * common (split_aux e1 p e2) /\ e2 === right (split_aux e1 p e2) * common (split_aux e1 p e2))%poly. Proof. induction e1;intros k e2; try refine (split_aux_ok1 _ k e2);simpl. destruct (IHe1_1 k e2) as (H1,H2). destruct (IHe1_2 k (right (split_aux e1_1 k e2))) as (H3,H4). clear IHe1_1 IHe1_2. - npe_simpl; split. * rewrite H1, H3. npe_ring. * rewrite H2 at 1. rewrite H4 at 1. npe_ring. - destruct n; simpl. + rewrite PEpow_0_r, PEpow_1_l, !PE_1_r. now split. + rewrite <- PEpow_mul_r. simpl. apply IHe1. Qed. Definition split e1 e2 := split_aux e1 xH e2. Theorem split_ok_l e1 e2 : (e1 === left (split e1 e2) * common (split e1 e2))%poly. Proof. destruct (split_aux_ok e1 xH e2) as (H,_). now rewrite <- H, PEpow_1_r. Qed. Theorem split_ok_r e1 e2 : (e2 === right (split e1 e2) * common (split e1 e2))%poly. Proof. destruct (split_aux_ok e1 xH e2) as (_,H). trivial. Qed. Lemma split_nz_l l e1 e2 : ~ e1 @ l == 0 -> ~ left (split e1 e2) @ l == 0. Proof. intros H. contradict H. rewrite (split_ok_l e1 e2); simpl. now rewrite H, rmul_0_l. Qed. Lemma split_nz_r l e1 e2 : ~ e2 @ l == 0 -> ~ right (split e1 e2) @ l == 0. Proof. intros H. contradict H. rewrite (split_ok_r e1 e2); simpl. now rewrite H, rmul_0_l. Qed. Fixpoint Fnorm (e : FExpr) : linear := match e with | FEO => mk_linear 0 1 nil | FEI => mk_linear 1 1 nil | FEc c => mk_linear (PEc c) 1 nil | FEX x => mk_linear (PEX C x) 1 nil | FEadd e1 e2 => let x := Fnorm e1 in let y := Fnorm e2 in let s := split (denum x) (denum y) in mk_linear ((num x ** right s) ++ (num y ** left s)) (left s ** (right s ** common s)) (condition x ++ condition y)%list | FEsub e1 e2 => let x := Fnorm e1 in let y := Fnorm e2 in let s := split (denum x) (denum y) in mk_linear ((num x ** right s) -- (num y ** left s)) (left s ** (right s ** common s)) (condition x ++ condition y)%list | FEmul e1 e2 => let x := Fnorm e1 in let y := Fnorm e2 in let s1 := split (num x) (denum y) in let s2 := split (num y) (denum x) in mk_linear (left s1 ** left s2) (right s2 ** right s1) (condition x ++ condition y)%list | FEopp e1 => let x := Fnorm e1 in mk_linear (NPEopp (num x)) (denum x) (condition x) | FEinv e1 => let x := Fnorm e1 in mk_linear (denum x) (num x) (num x :: condition x) | FEdiv e1 e2 => let x := Fnorm e1 in let y := Fnorm e2 in let s1 := split (num x) (num y) in let s2 := split (denum x) (denum y) in mk_linear (left s1 ** right s2) (left s2 ** right s1) (num y :: condition x ++ condition y)%list | FEpow e1 n => let x := Fnorm e1 in mk_linear ((num x)^^n) ((denum x)^^n) (condition x) end. (* Example *) (* Eval compute in (Fnorm (FEdiv (FEc cI) (FEadd (FEinv (FEX xH%positive)) (FEinv (FEX (xO xH)%positive))))). *) Theorem Pcond_Fnorm l e : PCond l (condition (Fnorm e)) -> ~ (denum (Fnorm e))@l == 0. Proof. induction e; simpl condition; rewrite ?PCond_cons, ?PCond_app; simpl denum; intros (Hc1,Hc2) || intros Hc; rewrite ?NPEmul_ok. - simpl. rewrite phi_1; exact rI_neq_rO. - simpl. rewrite phi_1; exact rI_neq_rO. - simpl; intros. rewrite phi_1; exact rI_neq_rO. - simpl; intros. rewrite phi_1; exact rI_neq_rO. - rewrite <- split_ok_r. simpl. apply field_is_integral_domain. + apply split_nz_l, IHe1, Hc1. + apply IHe2, Hc2. - rewrite <- split_ok_r. simpl. apply field_is_integral_domain. + apply split_nz_l, IHe1, Hc1. + apply IHe2, Hc2. - simpl. apply field_is_integral_domain. + apply split_nz_r, IHe1, Hc1. + apply split_nz_r, IHe2, Hc2. - now apply IHe. - trivial. - destruct Hc2 as (Hc2,_). simpl. apply field_is_integral_domain. + apply split_nz_l, IHe1, Hc2. + apply split_nz_r, Hc1. - rewrite NPEpow_ok. apply PEpow_nz, IHe, Hc. Qed. (*************************************************************************** Main theorem ***************************************************************************) Ltac uneval := repeat match goal with | |- context [ ?x @ ?l * ?y @ ?l ] => change (x@l * y@l) with ((x*y)@l) | |- context [ ?x @ ?l + ?y @ ?l ] => change (x@l + y@l) with ((x+y)@l) end. Theorem Fnorm_FEeval_PEeval l fe: PCond l (condition (Fnorm fe)) -> FEeval l fe == (num (Fnorm fe)) @ l / (denum (Fnorm fe)) @ l. Proof. induction fe; simpl condition; rewrite ?PCond_cons, ?PCond_app; simpl; intros (Hc1,Hc2) || intros Hc; try (specialize (IHfe1 Hc1);apply Pcond_Fnorm in Hc1); try (specialize (IHfe2 Hc2);apply Pcond_Fnorm in Hc2); try set (F1 := Fnorm fe1) in *; try set (F2 := Fnorm fe2) in *. - now rewrite phi_1, phi_0, rdiv_def. - now rewrite phi_1; apply rdiv1. - rewrite phi_1; apply rdiv1. - rewrite phi_1; apply rdiv1. - rewrite NPEadd_ok, !NPEmul_ok. simpl. rewrite <- rdiv2b; uneval; rewrite <- ?split_ok_l, <- ?split_ok_r; trivial. now f_equiv. - rewrite NPEsub_ok, !NPEmul_ok. simpl. rewrite <- rdiv3b; uneval; rewrite <- ?split_ok_l, <- ?split_ok_r; trivial. now f_equiv. - rewrite !NPEmul_ok. simpl. rewrite IHfe1, IHfe2. rewrite (split_ok_l (num F1) (denum F2) l), (split_ok_r (num F1) (denum F2) l), (split_ok_l (num F2) (denum F1) l), (split_ok_r (num F2) (denum F1) l) in *. apply rdiv4b; trivial. - rewrite NPEopp_ok; simpl; rewrite (IHfe Hc); apply rdiv5. - rewrite (IHfe Hc2); apply rdiv6; trivial; apply Pcond_Fnorm; trivial. - destruct Hc2 as (Hc2,Hc3). rewrite !NPEmul_ok. simpl. assert (U1 := split_ok_l (num F1) (num F2) l). assert (U2 := split_ok_r (num F1) (num F2) l). assert (U3 := split_ok_l (denum F1) (denum F2) l). assert (U4 := split_ok_r (denum F1) (denum F2) l). rewrite (IHfe1 Hc2), (IHfe2 Hc3), U1, U2, U3, U4. simpl in U2, U3, U4. apply rdiv7b; rewrite <- ?U2, <- ?U3, <- ?U4; try apply Pcond_Fnorm; trivial. - rewrite !NPEpow_ok. simpl. rewrite !rpow_pow, (IHfe Hc). destruct n; simpl. + apply rdiv1. + apply pow_pos_div. apply Pcond_Fnorm; trivial. Qed. Theorem Fnorm_crossproduct l fe1 fe2 : let nfe1 := Fnorm fe1 in let nfe2 := Fnorm fe2 in (num nfe1 * denum nfe2) @ l == (num nfe2 * denum nfe1) @ l -> PCond l (condition nfe1 ++ condition nfe2) -> FEeval l fe1 == FEeval l fe2. Proof. simpl. rewrite PCond_app. intros Hcrossprod (Hc1,Hc2). rewrite !Fnorm_FEeval_PEeval; trivial. apply cross_product_eq; trivial; apply Pcond_Fnorm; trivial. Qed. (* Correctness lemmas of reflexive tactics *) Notation Ninterp_PElist := (interp_PElist rO rI radd rmul rsub ropp req phi Cp_phi rpow). Notation Nmk_monpol_list := (mk_monpol_list cO cI cadd cmul csub copp ceqb cdiv). Theorem Fnorm_ok: forall n l lpe fe, Ninterp_PElist l lpe -> Peq ceqb (Nnorm n (Nmk_monpol_list lpe) (num (Fnorm fe))) (Pc cO) = true -> PCond l (condition (Fnorm fe)) -> FEeval l fe == 0. Proof. intros n l lpe fe Hlpe H H1. rewrite (Fnorm_FEeval_PEeval l fe H1). apply rdiv8. apply Pcond_Fnorm; trivial. transitivity (0@l); trivial. rewrite (norm_subst_ok Rsth Reqe ARth CRmorph pow_th cdiv_th n l lpe); trivial. change (0 @ l) with (Pphi 0 radd rmul phi l (Pc cO)). apply (Peq_ok Rsth Reqe CRmorph); trivial. Qed. Notation ring_rw_correct := (ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec). Notation ring_rw_pow_correct := (ring_rw_pow_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec). Notation ring_correct := (ring_correct Rsth Reqe ARth CRmorph pow_th cdiv_th). (* simplify a field expression into a fraction *) Definition display_linear l num den := let lnum := NPphi_dev l num in match den with | Pc c => if ceqb c cI then lnum else lnum / NPphi_dev l den | _ => lnum / NPphi_dev l den end. Definition display_pow_linear l num den := let lnum := NPphi_pow l num in match den with | Pc c => if ceqb c cI then lnum else lnum / NPphi_pow l den | _ => lnum / NPphi_pow l den end. Theorem Field_rw_correct n lpe l : Ninterp_PElist l lpe -> forall lmp, Nmk_monpol_list lpe = lmp -> forall fe nfe, Fnorm fe = nfe -> PCond l (condition nfe) -> FEeval l fe == display_linear l (Nnorm n lmp (num nfe)) (Nnorm n lmp (denum nfe)). Proof. intros Hlpe lmp lmp_eq fe nfe eq_nfe H; subst nfe lmp. rewrite (Fnorm_FEeval_PEeval _ _ H). unfold display_linear. destruct (Nnorm _ _ _) as [c | | ] eqn: HN; try ( apply rdiv_ext; eapply ring_rw_correct; eauto). destruct (ceqb_spec c cI). set (nnum := NPphi_dev _ _). apply eq_trans with (nnum / NPphi_dev l (Pc c)). apply rdiv_ext; eapply ring_rw_correct; eauto. rewrite Pphi_dev_ok; try eassumption. now simpl; rewrite H0, phi_1, <- rdiv1. apply rdiv_ext; eapply ring_rw_correct; eauto. Qed. Theorem Field_rw_pow_correct n lpe l : Ninterp_PElist l lpe -> forall lmp, Nmk_monpol_list lpe = lmp -> forall fe nfe, Fnorm fe = nfe -> PCond l (condition nfe) -> FEeval l fe == display_pow_linear l (Nnorm n lmp (num nfe)) (Nnorm n lmp (denum nfe)). Proof. intros Hlpe lmp lmp_eq fe nfe eq_nfe H; subst nfe lmp. rewrite (Fnorm_FEeval_PEeval _ _ H). unfold display_pow_linear. destruct (Nnorm _ _ _) as [c | | ] eqn: HN; try ( apply rdiv_ext; eapply ring_rw_pow_correct; eauto). destruct (ceqb_spec c cI). set (nnum := NPphi_pow _ _). apply eq_trans with (nnum / NPphi_pow l (Pc c)). apply rdiv_ext; eapply ring_rw_pow_correct; eauto. rewrite Pphi_pow_ok; try eassumption. now simpl; rewrite H0, phi_1, <- rdiv1. apply rdiv_ext; eapply ring_rw_pow_correct; eauto. Qed. Theorem Field_correct n l lpe fe1 fe2 : Ninterp_PElist l lpe -> forall lmp, Nmk_monpol_list lpe = lmp -> forall nfe1, Fnorm fe1 = nfe1 -> forall nfe2, Fnorm fe2 = nfe2 -> Peq ceqb (Nnorm n lmp (num nfe1 * denum nfe2)) (Nnorm n lmp (num nfe2 * denum nfe1)) = true -> PCond l (condition nfe1 ++ condition nfe2) -> FEeval l fe1 == FEeval l fe2. Proof. intros Hlpe lmp eq_lmp nfe1 eq1 nfe2 eq2 Hnorm Hcond; subst nfe1 nfe2 lmp. apply Fnorm_crossproduct; trivial. eapply ring_correct; eauto. Qed. (* simplify a field equation : generate the crossproduct and simplify polynomials *) (** This allows rewriting modulo the simplification of PEeval on PMul *) Declare Equivalent Keys PEeval rmul. Theorem Field_simplify_eq_correct : forall n l lpe fe1 fe2, Ninterp_PElist l lpe -> forall lmp, Nmk_monpol_list lpe = lmp -> forall nfe1, Fnorm fe1 = nfe1 -> forall nfe2, Fnorm fe2 = nfe2 -> forall den, split (denum nfe1) (denum nfe2) = den -> NPphi_dev l (Nnorm n lmp (num nfe1 * right den)) == NPphi_dev l (Nnorm n lmp (num nfe2 * left den)) -> PCond l (condition nfe1 ++ condition nfe2) -> FEeval l fe1 == FEeval l fe2. Proof. intros n l lpe fe1 fe2 Hlpe lmp Hlmp nfe1 eq1 nfe2 eq2 den eq3 Hcrossprod Hcond. apply Fnorm_crossproduct; rewrite ?eq1, ?eq2; trivial. simpl. rewrite (split_ok_l (denum nfe1) (denum nfe2) l), eq3. rewrite (split_ok_r (denum nfe1) (denum nfe2) l), eq3. simpl. rewrite !rmul_assoc. apply rmul_ext; trivial. rewrite (ring_rw_correct n lpe l Hlpe Logic.eq_refl (num nfe1 * right den) Logic.eq_refl), (ring_rw_correct n lpe l Hlpe Logic.eq_refl (num nfe2 * left den) Logic.eq_refl). rewrite Hlmp. apply Hcrossprod. Qed. Theorem Field_simplify_eq_pow_correct : forall n l lpe fe1 fe2, Ninterp_PElist l lpe -> forall lmp, Nmk_monpol_list lpe = lmp -> forall nfe1, Fnorm fe1 = nfe1 -> forall nfe2, Fnorm fe2 = nfe2 -> forall den, split (denum nfe1) (denum nfe2) = den -> NPphi_pow l (Nnorm n lmp (num nfe1 * right den)) == NPphi_pow l (Nnorm n lmp (num nfe2 * left den)) -> PCond l (condition nfe1 ++ condition nfe2) -> FEeval l fe1 == FEeval l fe2. Proof. intros n l lpe fe1 fe2 Hlpe lmp Hlmp nfe1 eq1 nfe2 eq2 den eq3 Hcrossprod Hcond. apply Fnorm_crossproduct; rewrite ?eq1, ?eq2; trivial. simpl. rewrite (split_ok_l (denum nfe1) (denum nfe2) l), eq3. rewrite (split_ok_r (denum nfe1) (denum nfe2) l), eq3. simpl. rewrite !rmul_assoc. apply rmul_ext; trivial. rewrite (ring_rw_pow_correct n lpe l Hlpe Logic.eq_refl (num nfe1 * right den) Logic.eq_refl), (ring_rw_pow_correct n lpe l Hlpe Logic.eq_refl (num nfe2 * left den) Logic.eq_refl). rewrite Hlmp. apply Hcrossprod. Qed. Theorem Field_simplify_aux_ok l fe1 fe2 den : FEeval l fe1 == FEeval l fe2 -> split (denum (Fnorm fe1)) (denum (Fnorm fe2)) = den -> PCond l (condition (Fnorm fe1) ++ condition (Fnorm fe2)) -> (num (Fnorm fe1) * right den) @ l == (num (Fnorm fe2) * left den) @ l. Proof. rewrite PCond_app; intros Hfe Hden (Hc1,Hc2); simpl. assert (Hc1' := Pcond_Fnorm _ _ Hc1). assert (Hc2' := Pcond_Fnorm _ _ Hc2). set (N1 := num (Fnorm fe1)) in *. set (N2 := num (Fnorm fe2)) in *. set (D1 := denum (Fnorm fe1)) in *. set (D2 := denum (Fnorm fe2)) in *. assert (~ (common den) @ l == 0). { intro H. apply Hc1'. rewrite (split_ok_l D1 D2 l). rewrite Hden. simpl. ring [H]. } apply (@rmul_reg_l ((common den) @ l)); trivial. rewrite !(rmul_comm ((common den) @ l)), <- !rmul_assoc. change (N1@l * (right den * common den) @ l == N2@l * (left den * common den) @ l). rewrite <- Hden, <- split_ok_l, <- split_ok_r. apply (@rmul_reg_l (/ D2@l)). { apply rinv_nz; trivial. } rewrite (rmul_comm (/ D2 @ l)), <- !rmul_assoc. rewrite <- rdiv_def, rdiv_r_r, rmul_1_r by trivial. apply (@rmul_reg_l (/ (D1@l))). { apply rinv_nz; trivial. } rewrite !(rmul_comm (/ D1@l)), <- !rmul_assoc. rewrite <- !rdiv_def, rdiv_r_r, rmul_1_r by trivial. rewrite (rmul_comm (/ D2@l)), <- rdiv_def. unfold N1,N2,D1,D2; rewrite <- !Fnorm_FEeval_PEeval; trivial. Qed. Theorem Field_simplify_eq_pow_in_correct : forall n l lpe fe1 fe2, Ninterp_PElist l lpe -> forall lmp, Nmk_monpol_list lpe = lmp -> forall nfe1, Fnorm fe1 = nfe1 -> forall nfe2, Fnorm fe2 = nfe2 -> forall den, split (denum nfe1) (denum nfe2) = den -> forall np1, Nnorm n lmp (num nfe1 * right den) = np1 -> forall np2, Nnorm n lmp (num nfe2 * left den) = np2 -> FEeval l fe1 == FEeval l fe2 -> PCond l (condition nfe1 ++ condition nfe2) -> NPphi_pow l np1 == NPphi_pow l np2. Proof. intros. subst nfe1 nfe2 lmp np1 np2. rewrite !(Pphi_pow_ok Rsth Reqe ARth CRmorph pow_th get_sign_spec). repeat (rewrite <- (norm_subst_ok Rsth Reqe ARth CRmorph pow_th);trivial). simpl. apply Field_simplify_aux_ok; trivial. Qed. Theorem Field_simplify_eq_in_correct : forall n l lpe fe1 fe2, Ninterp_PElist l lpe -> forall lmp, Nmk_monpol_list lpe = lmp -> forall nfe1, Fnorm fe1 = nfe1 -> forall nfe2, Fnorm fe2 = nfe2 -> forall den, split (denum nfe1) (denum nfe2) = den -> forall np1, Nnorm n lmp (num nfe1 * right den) = np1 -> forall np2, Nnorm n lmp (num nfe2 * left den) = np2 -> FEeval l fe1 == FEeval l fe2 -> PCond l (condition nfe1 ++ condition nfe2) -> NPphi_dev l np1 == NPphi_dev l np2. Proof. intros. subst nfe1 nfe2 lmp np1 np2. rewrite !(Pphi_dev_ok Rsth Reqe ARth CRmorph get_sign_spec). repeat (rewrite <- (norm_subst_ok Rsth Reqe ARth CRmorph pow_th);trivial). apply Field_simplify_aux_ok; trivial. Qed. Section Fcons_impl. Variable Fcons : PExpr C -> list (PExpr C) -> list (PExpr C). Hypothesis PCond_fcons_inv : forall l a l1, PCond l (Fcons a l1) -> ~ a @ l == 0 /\ PCond l l1. Fixpoint Fapp (l m:list (PExpr C)) {struct l} : list (PExpr C) := match l with | nil => m | cons a l1 => Fcons a (Fapp l1 m) end. Lemma fcons_ok : forall l l1, (forall lock, lock = PCond l -> lock (Fapp l1 nil)) -> PCond l l1. Proof. intros l l1 h1; assert (H := h1 (PCond l) (refl_equal _));clear h1. induction l1; simpl; intros. trivial. elim PCond_fcons_inv with (1 := H); intros. destruct l1; trivial. split; trivial. apply IHl1; trivial. Qed. End Fcons_impl. Section Fcons_simpl. (* Some general simpifications of the condition: eliminate duplicates, split multiplications *) Fixpoint Fcons (e:PExpr C) (l:list (PExpr C)) {struct l} : list (PExpr C) := match l with nil => cons e nil | cons a l1 => if PExpr_eq e a then l else cons a (Fcons e l1) end. Theorem PFcons_fcons_inv: forall l a l1, PCond l (Fcons a l1) -> ~ a @ l == 0 /\ PCond l l1. Proof. induction l1 as [|e l1]; simpl Fcons. - simpl; now split. - case PExpr_eq_spec; intros H; rewrite !PCond_cons; intros (H1,H2); repeat split; trivial. + now rewrite H. + now apply IHl1. + now apply IHl1. Qed. (* equality of normal forms rather than syntactic equality *) Fixpoint Fcons0 (e:PExpr C) (l:list (PExpr C)) {struct l} : list (PExpr C) := match l with nil => cons e nil | cons a l1 => if Peq ceqb (Nnorm O nil e) (Nnorm O nil a) then l else cons a (Fcons0 e l1) end. Theorem PFcons0_fcons_inv: forall l a l1, PCond l (Fcons0 a l1) -> ~ a @ l == 0 /\ PCond l l1. Proof. induction l1 as [|e l1]; simpl Fcons0. - simpl; now split. - generalize (ring_correct O l nil a e). lazy zeta; simpl Peq. case Peq; intros H; rewrite !PCond_cons; intros (H1,H2); repeat split; trivial. + now rewrite H. + now apply IHl1. + now apply IHl1. Qed. (* split factorized denominators *) Fixpoint Fcons00 (e:PExpr C) (l:list (PExpr C)) {struct e} : list (PExpr C) := match e with PEmul e1 e2 => Fcons00 e1 (Fcons00 e2 l) | PEpow e1 _ => Fcons00 e1 l | _ => Fcons0 e l end. Theorem PFcons00_fcons_inv: forall l a l1, PCond l (Fcons00 a l1) -> ~ a @ l == 0 /\ PCond l l1. Proof. intros l a; elim a; try (intros; apply PFcons0_fcons_inv; trivial; fail). - intros p H p0 H0 l1 H1. simpl in H1. destruct (H _ H1) as (H2,H3). destruct (H0 _ H3) as (H4,H5). split; trivial. simpl. apply field_is_integral_domain; trivial. - intros. destruct (H _ H0). split; trivial. apply PEpow_nz; trivial. Qed. Definition Pcond_simpl_gen := fcons_ok _ PFcons00_fcons_inv. (* Specific case when the equality test of coefs is complete w.r.t. the field equality: non-zero coefs can be eliminated, and opposite can be simplified (if -1 <> 0) *) Hypothesis ceqb_complete : forall c1 c2, [c1] == [c2] -> ceqb c1 c2 = true. Lemma ceqb_spec' c1 c2 : Bool.reflect ([c1] == [c2]) (ceqb c1 c2). Proof. assert (H := morph_eq CRmorph c1 c2). assert (H' := @ceqb_complete c1 c2). destruct (ceqb c1 c2); constructor. - now apply H. - intro E. specialize (H' E). discriminate. Qed. Fixpoint Fcons1 (e:PExpr C) (l:list (PExpr C)) {struct e} : list (PExpr C) := match e with | PEmul e1 e2 => Fcons1 e1 (Fcons1 e2 l) | PEpow e _ => Fcons1 e l | PEopp e => if (-(1) =? 0)%coef then absurd_PCond else Fcons1 e l | PEc c => if (c =? 0)%coef then absurd_PCond else l | _ => Fcons0 e l end. Theorem PFcons1_fcons_inv: forall l a l1, PCond l (Fcons1 a l1) -> ~ a @ l == 0 /\ PCond l l1. Proof. intros l a; elim a; try (intros; apply PFcons0_fcons_inv; trivial; fail). - simpl; intros c l1. case ceqb_spec'; intros H H0. + elim (@absurd_PCond_bottom l H0). + split; trivial. rewrite <- phi_0; trivial. - intros p H p0 H0 l1 H1. simpl in H1. destruct (H _ H1) as (H2,H3). destruct (H0 _ H3) as (H4,H5). split; trivial. simpl. apply field_is_integral_domain; trivial. - simpl; intros p H l1. case ceqb_spec'; intros H0 H1. + elim (@absurd_PCond_bottom l H1). + destruct (H _ H1). split; trivial. apply ropp_neq_0; trivial. rewrite (morph_opp CRmorph), phi_0, phi_1 in H0. trivial. - intros. destruct (H _ H0);split;trivial. apply PEpow_nz; trivial. Qed. Definition Fcons2 e l := Fcons1 (PEsimp e) l. Theorem PFcons2_fcons_inv: forall l a l1, PCond l (Fcons2 a l1) -> ~ a @ l == 0 /\ PCond l l1. Proof. unfold Fcons2; intros l a l1 H; split; case (PFcons1_fcons_inv l (PEsimp a) l1); trivial. intros H1 H2 H3; case H1. transitivity (a@l); trivial. apply PEsimp_ok. Qed. Definition Pcond_simpl_complete := fcons_ok _ PFcons2_fcons_inv. End Fcons_simpl. End AlmostField. Section FieldAndSemiField. Record field_theory : Prop := mk_field { F_R : ring_theory rO rI radd rmul rsub ropp req; F_1_neq_0 : ~ 1 == 0; Fdiv_def : forall p q, p / q == p * / q; Finv_l : forall p, ~ p == 0 -> / p * p == 1 }. Definition F2AF f := mk_afield (Rth_ARth Rsth Reqe (F_R f)) (F_1_neq_0 f) (Fdiv_def f) (Finv_l f). Record semi_field_theory : Prop := mk_sfield { SF_SR : semi_ring_theory rO rI radd rmul req; SF_1_neq_0 : ~ 1 == 0; SFdiv_def : forall p q, p / q == p * / q; SFinv_l : forall p, ~ p == 0 -> / p * p == 1 }. End FieldAndSemiField. End MakeFieldPol. Definition SF2AF R (rO rI:R) radd rmul rdiv rinv req Rsth (sf:semi_field_theory rO rI radd rmul rdiv rinv req) := mk_afield _ _ (SRth_ARth Rsth (SF_SR sf)) (SF_1_neq_0 sf) (SFdiv_def sf) (SFinv_l sf). Section Complete. Variable R : Type. Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R). Variable (rdiv : R -> R -> R) (rinv : R -> R). Variable req : R -> R -> Prop. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y). Notation "x - y " := (rsub x y). Notation "- x" := (ropp x). Notation "x / y " := (rdiv x y). Notation "/ x" := (rinv x). Notation "x == y" := (req x y) (at level 70, no associativity). Variable Rsth : Setoid_Theory R req. Add Parametric Relation : R req reflexivity proved by (@Equivalence_Reflexive _ _ Rsth) symmetry proved by (@Equivalence_Symmetric _ _ Rsth) transitivity proved by (@Equivalence_Transitive _ _ Rsth) as R_setoid3. Variable Reqe : ring_eq_ext radd rmul ropp req. Add Morphism radd with signature (req ==> req ==> req) as radd_ext3. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext3. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext3. Proof. exact (Ropp_ext Reqe). Qed. Section AlmostField. Variable AFth : almost_field_theory rO rI radd rmul rsub ropp rdiv rinv req. Let ARth := (AF_AR AFth). Let rI_neq_rO := (AF_1_neq_0 AFth). Let rdiv_def := (AFdiv_def AFth). Let rinv_l := (AFinv_l AFth). Hypothesis S_inj : forall x y, 1+x==1+y -> x==y. Hypothesis gen_phiPOS_not_0 : forall p, ~ gen_phiPOS1 rI radd rmul p == 0. Lemma add_inj_r p x y : gen_phiPOS1 rI radd rmul p + x == gen_phiPOS1 rI radd rmul p + y -> x==y. Proof. elim p using Pos.peano_ind; simpl; intros. apply S_inj; trivial. apply H. apply S_inj. rewrite !(ARadd_assoc ARth). rewrite <- (ARgen_phiPOS_Psucc Rsth Reqe ARth); trivial. Qed. Lemma gen_phiPOS_inj x y : gen_phiPOS rI radd rmul x == gen_phiPOS rI radd rmul y -> x = y. Proof. rewrite <- !(same_gen Rsth Reqe ARth). case (Pos.compare_spec x y). intros. trivial. intros. elim gen_phiPOS_not_0 with (y - x)%positive. apply add_inj_r with x. symmetry. rewrite (ARadd_0_r Rsth ARth). rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth). now rewrite Pos.add_comm, Pos.sub_add. intros. elim gen_phiPOS_not_0 with (x - y)%positive. apply add_inj_r with y. rewrite (ARadd_0_r Rsth ARth). rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth). now rewrite Pos.add_comm, Pos.sub_add. Qed. Lemma gen_phiN_inj x y : gen_phiN rO rI radd rmul x == gen_phiN rO rI radd rmul y -> x = y. Proof. destruct x; destruct y; simpl; intros; trivial. elim gen_phiPOS_not_0 with p. symmetry . rewrite (same_gen Rsth Reqe ARth); trivial. elim gen_phiPOS_not_0 with p. rewrite (same_gen Rsth Reqe ARth); trivial. rewrite gen_phiPOS_inj with (1 := H); trivial. Qed. Lemma gen_phiN_complete x y : gen_phiN rO rI radd rmul x == gen_phiN rO rI radd rmul y -> N.eqb x y = true. Proof. intros. now apply N.eqb_eq, gen_phiN_inj. Qed. End AlmostField. Section Field. Variable Fth : field_theory rO rI radd rmul rsub ropp rdiv rinv req. Let Rth := (F_R Fth). Let rI_neq_rO := (F_1_neq_0 Fth). Let rdiv_def := (Fdiv_def Fth). Let rinv_l := (Finv_l Fth). Let AFth := F2AF Rsth Reqe Fth. Let ARth := Rth_ARth Rsth Reqe Rth. Lemma ring_S_inj x y : 1+x==1+y -> x==y. Proof. intros. rewrite <- (ARadd_0_l ARth x), <- (ARadd_0_l ARth y). rewrite <- (Ropp_def Rth 1), (ARadd_comm ARth 1). rewrite <- !(ARadd_assoc ARth). now apply (Radd_ext Reqe). Qed. Hypothesis gen_phiPOS_not_0 : forall p, ~ gen_phiPOS1 rI radd rmul p == 0. Let gen_phiPOS_inject := gen_phiPOS_inj AFth ring_S_inj gen_phiPOS_not_0. Lemma gen_phiPOS_discr_sgn x y : ~ gen_phiPOS rI radd rmul x == - gen_phiPOS rI radd rmul y. Proof. red; intros. apply gen_phiPOS_not_0 with (y + x)%positive. rewrite (ARgen_phiPOS_add Rsth Reqe ARth). transitivity (gen_phiPOS1 1 radd rmul y + - gen_phiPOS1 1 radd rmul y). apply (Radd_ext Reqe); trivial. reflexivity. rewrite (same_gen Rsth Reqe ARth). rewrite (same_gen Rsth Reqe ARth). trivial. apply (Ropp_def Rth). Qed. Lemma gen_phiZ_inj x y : gen_phiZ rO rI radd rmul ropp x == gen_phiZ rO rI radd rmul ropp y -> x = y. Proof. destruct x; destruct y; simpl; intros. trivial. elim gen_phiPOS_not_0 with p. rewrite (same_gen Rsth Reqe ARth). symmetry ; trivial. elim gen_phiPOS_not_0 with p. rewrite (same_gen Rsth Reqe ARth). rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p)). rewrite <- H. apply (ARopp_zero Rsth Reqe ARth). elim gen_phiPOS_not_0 with p. rewrite (same_gen Rsth Reqe ARth). trivial. rewrite gen_phiPOS_inject with (1 := H); trivial. elim gen_phiPOS_discr_sgn with (1 := H). elim gen_phiPOS_not_0 with p. rewrite (same_gen Rsth Reqe ARth). rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p)). rewrite H. apply (ARopp_zero Rsth Reqe ARth). elim gen_phiPOS_discr_sgn with p0 p. symmetry ; trivial. replace p0 with p; trivial. apply gen_phiPOS_inject. rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p)). rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p0)). rewrite H; trivial. reflexivity. Qed. Lemma gen_phiZ_complete x y : gen_phiZ rO rI radd rmul ropp x == gen_phiZ rO rI radd rmul ropp y -> Zeq_bool x y = true. Proof. intros. replace y with x. unfold Zeq_bool. rewrite Z.compare_refl; trivial. apply gen_phiZ_inj; trivial. Qed. End Field. End Complete. Arguments FEO {C}. Arguments FEI {C}. coq-8.11.0/plugins/setoid_ring/Field.v0000644000175000017500000000133313612664533017467 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* idtac | ?e1::?e2::_ => match goal with |- (?op ?u1 ?u2) => change (op (@Ring_polynom.PEeval _ zero one _+_ _*_ _-_ -_ Z Ncring_initial.gen_phiZ N (fun n:N => n) (@Ring_theory.pow_N _ 1 multiplication) lvar e1) (@Ring_polynom.PEeval _ zero one _+_ _*_ _-_ -_ Z Ncring_initial.gen_phiZ N (fun n:N => n) (@Ring_theory.pow_N _ 1 multiplication) lvar e2)) end end. Section cring. Context {R:Type}`{Rr:Cring R}. Lemma cring_eq_ext: ring_eq_ext _+_ _*_ -_ _==_. Proof. intros. apply mk_reqe; solve_proper. Defined. Lemma cring_almost_ring_theory: almost_ring_theory (R:=R) zero one _+_ _*_ _-_ -_ _==_. intros. apply mk_art ;intros. rewrite ring_add_0_l; reflexivity. rewrite ring_add_comm; reflexivity. rewrite ring_add_assoc; reflexivity. rewrite ring_mul_1_l; reflexivity. apply ring_mul_0_l. rewrite cring_mul_comm; reflexivity. rewrite ring_mul_assoc; reflexivity. rewrite ring_distr_l; reflexivity. rewrite ring_opp_mul_l; reflexivity. apply ring_opp_add. rewrite ring_sub_def ; reflexivity. Defined. Lemma cring_morph: ring_morph zero one _+_ _*_ _-_ -_ _==_ 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool Ncring_initial.gen_phiZ. intros. apply mkmorph ; intros; simpl; try reflexivity. rewrite Ncring_initial.gen_phiZ_add; reflexivity. rewrite ring_sub_def. unfold Z.sub. rewrite Ncring_initial.gen_phiZ_add. rewrite Ncring_initial.gen_phiZ_opp; reflexivity. rewrite Ncring_initial.gen_phiZ_mul; reflexivity. rewrite Ncring_initial.gen_phiZ_opp; reflexivity. rewrite (Zeqb_ok x y H). reflexivity. Defined. Lemma cring_power_theory : @Ring_theory.power_theory R one _*_ _==_ N (fun n:N => n) (@Ring_theory.pow_N _ 1 multiplication). intros; apply Ring_theory.mkpow_th. reflexivity. Defined. Lemma cring_div_theory: div_theory _==_ Z.add Z.mul Ncring_initial.gen_phiZ Z.quotrem. intros. apply InitialRing.Ztriv_div_th. unfold Setoid_Theory. simpl. apply ring_setoid. Defined. End cring. Ltac cring_gen := match goal with |- ?g => let lterm := lterm_goal g in match eval red in (list_reifyl (lterm:=lterm)) with | (?fv, ?lexpr) => (*idtac "variables:";idtac fv; idtac "terms:"; idtac lterm; idtac "reifications:"; idtac lexpr; *) reify_goal fv lexpr lterm; match goal with |- ?g => generalize (@Ring_polynom.ring_correct _ 0 1 _+_ _*_ _-_ -_ _==_ ring_setoid cring_eq_ext cring_almost_ring_theory Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool Ncring_initial.gen_phiZ cring_morph N (fun n:N => n) (@Ring_theory.pow_N _ 1 multiplication) cring_power_theory Z.quotrem cring_div_theory O fv nil); let rc := fresh "rc"in intro rc; apply rc end end end. Ltac cring_compute:= vm_compute; reflexivity. Ltac cring:= intros; cring_gen; cring_compute. Instance Zcri: (Cring (Rr:=Zr)). red. exact Z.mul_comm. Defined. (* Cring_simplify *) Ltac cring_simplify_aux lterm fv lexpr hyp := match lterm with | ?t0::?lterm => match lexpr with | ?e::?le => let t := constr:(@Ring_polynom.norm_subst Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool Z.quotrem O nil e) in let te := constr:(@Ring_polynom.Pphi_dev _ 0 1 _+_ _*_ _-_ -_ Z 0%Z 1%Z Zeq_bool Ncring_initial.gen_phiZ get_signZ fv t) in let eq1 := fresh "ring" in let nft := eval vm_compute in t in let t':= fresh "t" in pose (t' := nft); assert (eq1 : t = t'); [vm_cast_no_check (eq_refl t')| let eq2 := fresh "ring" in assert (eq2:(@Ring_polynom.PEeval _ zero _+_ _*_ _-_ -_ Z Ncring_initial.gen_phiZ N (fun n:N => n) (@Ring_theory.pow_N _ 1 multiplication) fv e) == te); [let eq3 := fresh "ring" in generalize (@ring_rw_correct _ 0 1 _+_ _*_ _-_ -_ _==_ ring_setoid cring_eq_ext cring_almost_ring_theory Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool Ncring_initial.gen_phiZ cring_morph N (fun n:N => n) (@Ring_theory.pow_N _ 1 multiplication) cring_power_theory Z.quotrem cring_div_theory get_signZ get_signZ_th O nil fv I nil (eq_refl nil) ); intro eq3; apply eq3; reflexivity| match hyp with | 1%nat => rewrite eq2 | ?H => try rewrite eq2 in H end]; let P:= fresh "P" in match hyp with | 1%nat => rewrite eq1; pattern (@Ring_polynom.Pphi_dev _ 0 1 _+_ _*_ _-_ -_ Z 0%Z 1%Z Zeq_bool Ncring_initial.gen_phiZ get_signZ fv t'); match goal with |- (?p ?t) => set (P:=p) end; unfold t' in *; clear t' eq1 eq2; unfold Pphi_dev, Pphi_avoid; simpl; repeat (unfold mkmult1, mkmultm1, mkmult_c_pos, mkmult_c, mkadd_mult, mkmult_c_pos, mkmult_pow, mkadd_mult, mkpow;simpl) | ?H => rewrite eq1 in H; pattern (@Ring_polynom.Pphi_dev _ 0 1 _+_ _*_ _-_ -_ Z 0%Z 1%Z Zeq_bool Ncring_initial.gen_phiZ get_signZ fv t') in H; match type of H with | (?p ?t) => set (P:=p) in H end; unfold t' in *; clear t' eq1 eq2; unfold Pphi_dev, Pphi_avoid in H; simpl in H; repeat (unfold mkmult1, mkmultm1, mkmult_c_pos, mkmult_c, mkadd_mult, mkmult_c_pos, mkmult_pow, mkadd_mult, mkpow in H;simpl in H) end; unfold P in *; clear P ]; cring_simplify_aux lterm fv le hyp | nil => idtac end | nil => idtac end. Ltac set_variables fv := match fv with | nil => idtac | ?t::?fv => let v := fresh "X" in set (v:=t) in *; set_variables fv end. Ltac deset n:= match n with | 0%nat => idtac | S ?n1 => match goal with | h:= ?v : ?t |- ?g => unfold h in *; clear h; deset n1 end end. (* a est soit un terme de l'anneau, soit une liste de termes. J'ai pas réussi à un décomposer les Vlists obtenues avec ne_constr_list dans Tactic Notation *) Ltac cring_simplify_gen a hyp := let lterm := match a with | _::_ => a | _ => constr:(a::nil) end in match eval red in (list_reifyl (lterm:=lterm)) with | (?fv, ?lexpr) => idtac lterm; idtac fv; idtac lexpr; let n := eval compute in (length fv) in idtac n; let lt:=fresh "lt" in set (lt:= lterm); let lv:=fresh "fv" in set (lv:= fv); (* les termes de fv sont remplacés par des variables pour pouvoir utiliser simpl ensuite sans risquer des simplifications indésirables *) set_variables fv; let lterm1 := eval unfold lt in lt in let lv1 := eval unfold lv in lv in idtac lterm1; idtac lv1; cring_simplify_aux lterm1 lv1 lexpr hyp; clear lt lv; (* on remet les termes de fv *) deset n end. Tactic Notation "cring_simplify" constr(lterm):= cring_simplify_gen lterm 1%nat. Tactic Notation "cring_simplify" constr(lterm) "in" ident(H):= cring_simplify_gen lterm H. coq-8.11.0/plugins/setoid_ring/Ncring_tac.v0000644000175000017500000002302713612664533020517 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr:(t1::t2::nil) | ?t1 = ?t2 => constr:(t1::t2::nil) | (_ ?t1 ?t2) => constr:(t1::t2::nil) end. Lemma Zeqb_ok: forall x y : Z, Zeq_bool x y = true -> x == y. intros x y H. rewrite (Zeq_bool_eq x y H). reflexivity. Qed. Ltac reify_goal lvar lexpr lterm:= (*idtac lvar; idtac lexpr; idtac lterm;*) match lexpr with nil => idtac | ?e1::?e2::_ => match goal with |- (?op ?u1 ?u2) => change (op (@PEeval Z _ _ _ _ _ _ _ _ _ (@gen_phiZ _ _ _ _ _ _ _ _ _) N (fun n:N => n) (@pow_N _ _ _ _ _ _ _ _ _) lvar e1) (@PEeval Z _ _ _ _ _ _ _ _ _ (@gen_phiZ _ _ _ _ _ _ _ _ _) N (fun n:N => n) (@pow_N _ _ _ _ _ _ _ _ _) lvar e2)) end end. Lemma comm: forall (R:Type)`{Ring R}(c : Z) (x : R), x * (gen_phiZ c) == (gen_phiZ c) * x. induction c. intros. simpl. gen_rewrite. simpl. intros. rewrite <- same_gen. induction p. simpl. gen_rewrite. rewrite IHp. reflexivity. simpl. gen_rewrite. rewrite IHp. reflexivity. simpl. gen_rewrite. simpl. intros. rewrite <- same_gen. induction p. simpl. generalize IHp. clear IHp. gen_rewrite. intro IHp. rewrite IHp. reflexivity. simpl. generalize IHp. clear IHp. gen_rewrite. intro IHp. rewrite IHp. reflexivity. simpl. gen_rewrite. Qed. Ltac ring_gen := match goal with |- ?g => let lterm := lterm_goal g in match eval red in (list_reifyl (lterm:=lterm)) with | (?fv, ?lexpr) => (*idtac "variables:";idtac fv; idtac "terms:"; idtac lterm; idtac "reifications:"; idtac lexpr; *) reify_goal fv lexpr lterm; match goal with |- ?g => apply (@ring_correct Z _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ (@gen_phiZ _ _ _ _ _ _ _ _ _) _ (@comm _ _ _ _ _ _ _ _ _ _) Zeq_bool Zeqb_ok N (fun n:N => n) (@pow_N _ _ _ _ _ _ _ _ _)); [apply mkpow_th; reflexivity |vm_compute; reflexivity] end end end. Ltac non_commutative_ring:= intros; ring_gen. (* simplification *) Ltac ring_simplify_aux lterm fv lexpr hyp := match lterm with | ?t0::?lterm => match lexpr with | ?e::?le => (* e:PExpr Z est la réification de t0:R *) let t := constr:(@Ncring_polynom.norm_subst Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (@eq Z) Zops Zeq_bool e) in (* t:Pol Z *) let te := constr:(@Ncring_polynom.Pphi Z _ 0 1 _+_ _*_ _-_ -_ _==_ _ Ncring_initial.gen_phiZ fv t) in let eq1 := fresh "ring" in let nft := eval vm_compute in t in let t':= fresh "t" in pose (t' := nft); assert (eq1 : t = t'); [vm_cast_no_check (eq_refl t')| let eq2 := fresh "ring" in assert (eq2:(@Ncring_polynom.PEeval Z _ 0 1 _+_ _*_ _-_ -_ _==_ _ Ncring_initial.gen_phiZ N (fun n:N => n) (@Ring_theory.pow_N _ 1 multiplication) fv e) == te); [apply (@Ncring_polynom.norm_subst_ok Z _ 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (@eq Z) _ _ 0 1 _+_ _*_ _-_ -_ _==_ _ _ Ncring_initial.gen_phiZ _ (@comm _ 0 1 _+_ _*_ _-_ -_ _==_ _ _) _ Zeqb_ok); apply mkpow_th; reflexivity | match hyp with | 1%nat => rewrite eq2 | ?H => try rewrite eq2 in H end]; let P:= fresh "P" in match hyp with | 1%nat => idtac "ok"; rewrite eq1; pattern (@Ncring_polynom.Pphi Z _ 0 1 _+_ _*_ _-_ -_ _==_ _ Ncring_initial.gen_phiZ fv t'); match goal with |- (?p ?t) => set (P:=p) end; unfold t' in *; clear t' eq1 eq2; simpl | ?H => rewrite eq1 in H; pattern (@Ncring_polynom.Pphi Z _ 0 1 _+_ _*_ _-_ -_ _==_ _ Ncring_initial.gen_phiZ fv t') in H; match type of H with | (?p ?t) => set (P:=p) in H end; unfold t' in *; clear t' eq1 eq2; simpl in H end; unfold P in *; clear P ]; ring_simplify_aux lterm fv le hyp | nil => idtac end | nil => idtac end. Ltac set_variables fv := match fv with | nil => idtac | ?t::?fv => let v := fresh "X" in set (v:=t) in *; set_variables fv end. Ltac deset n:= match n with | 0%nat => idtac | S ?n1 => match goal with | h:= ?v : ?t |- ?g => unfold h in *; clear h; deset n1 end end. (* a est soit un terme de l'anneau, soit une liste de termes. J'ai pas réussi à un décomposer les Vlists obtenues avec ne_constr_list dans Tactic Notation *) Ltac ring_simplify_gen a hyp := let lterm := match a with | _::_ => a | _ => constr:(a::nil) end in match eval red in (list_reifyl (lterm:=lterm)) with | (?fv, ?lexpr) => idtac lterm; idtac fv; idtac lexpr; let n := eval compute in (length fv) in idtac n; let lt:=fresh "lt" in set (lt:= lterm); let lv:=fresh "fv" in set (lv:= fv); (* les termes de fv sont remplacés par des variables pour pouvoir utiliser simpl ensuite sans risquer des simplifications indésirables *) set_variables fv; let lterm1 := eval unfold lt in lt in let lv1 := eval unfold lv in lv in idtac lterm1; idtac lv1; ring_simplify_aux lterm1 lv1 lexpr hyp; clear lt lv; (* on remet les termes de fv *) deset n end. Tactic Notation "non_commutative_ring_simplify" constr(lterm):= ring_simplify_gen lterm 1%nat. Tactic Notation "non_commutative_ring_simplify" constr(lterm) "in" ident(H):= ring_simplify_gen lterm H. coq-8.11.0/plugins/setoid_ring/Ring_base.v0000644000175000017500000000166113612664533020341 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* A -> A. Notation "_+_" := addition. Notation "x + y" := (addition x y). Class Multiplication {A B : Type} := multiplication : A -> B -> B. Notation "_*_" := multiplication. Notation "x * y" := (multiplication x y). Class Subtraction (A : Type) := subtraction : A -> A -> A. Notation "_-_" := subtraction. Notation "x - y" := (subtraction x y). Class Opposite (A : Type) := opposite : A -> A. Notation "-_" := opposite. Notation "- x" := (opposite(x)). Class Equality {A : Type}:= equality : A -> A -> Prop. Notation "_==_" := equality. Notation "x == y" := (equality x y) (at level 70, no associativity). Class Bracket (A B: Type):= bracket : A -> B. Notation "[ x ]" := (bracket(x)). Class Power {A B: Type} := power : A -> B -> A. Notation "x ^ y" := (power x y). coq-8.11.0/plugins/setoid_ring/newring_plugin.mlpack0000644000175000017500000000003613612664533022474 0ustar treinentreinenNewring_ast Newring G_newring coq-8.11.0/plugins/setoid_ring/Ncring_initial.v0000644000175000017500000001435213612664533021402 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 1 | xO p => (1 + 1) * (gen_phiPOS1 p) | xI p => 1 + ((1 + 1) * (gen_phiPOS1 p)) end. Fixpoint gen_phiPOS (p:positive) : R := match p with | xH => 1 | xO xH => (1 + 1) | xO p => (1 + 1) * (gen_phiPOS p) | xI xH => 1 + (1 +1) | xI p => 1 + ((1 + 1) * (gen_phiPOS p)) end. Definition gen_phiZ1 z := match z with | Zpos p => gen_phiPOS1 p | Z0 => 0 | Zneg p => -(gen_phiPOS1 p) end. Definition gen_phiZ z := match z with | Zpos p => gen_phiPOS p | Z0 => 0 | Zneg p => -(gen_phiPOS p) end. Declare Scope ZMORPHISM. Notation "[ x ]" := (gen_phiZ x) : ZMORPHISM. Open Scope ZMORPHISM. Definition get_signZ z := match z with | Zneg p => Some (Zpos p) | _ => None end. Ltac norm := gen_rewrite. Ltac add_push := Ncring.gen_add_push. Ltac rsimpl := simpl. Lemma same_gen : forall x, gen_phiPOS1 x == gen_phiPOS x. Proof. induction x;rsimpl. rewrite IHx. destruct x;simpl;norm. rewrite IHx;destruct x;simpl;norm. reflexivity. Qed. Lemma ARgen_phiPOS_Psucc : forall x, gen_phiPOS1 (Pos.succ x) == 1 + (gen_phiPOS1 x). Proof. induction x;rsimpl;norm. rewrite IHx. gen_rewrite. add_push 1. reflexivity. Qed. Lemma ARgen_phiPOS_add : forall x y, gen_phiPOS1 (x + y) == (gen_phiPOS1 x) + (gen_phiPOS1 y). Proof. induction x;destruct y;simpl;norm. rewrite Pos.add_carry_spec. rewrite ARgen_phiPOS_Psucc. rewrite IHx;norm. add_push (gen_phiPOS1 y);add_push 1;reflexivity. rewrite IHx;norm;add_push (gen_phiPOS1 y);reflexivity. rewrite ARgen_phiPOS_Psucc;norm;add_push 1;reflexivity. rewrite IHx;norm;add_push(gen_phiPOS1 y); add_push 1;reflexivity. rewrite IHx;norm;add_push(gen_phiPOS1 y);reflexivity. add_push 1;reflexivity. rewrite ARgen_phiPOS_Psucc;norm;add_push 1;reflexivity. Qed. Lemma ARgen_phiPOS_mult : forall x y, gen_phiPOS1 (x * y) == gen_phiPOS1 x * gen_phiPOS1 y. Proof. induction x;intros;simpl;norm. rewrite ARgen_phiPOS_add;simpl;rewrite IHx;norm. rewrite IHx;reflexivity. Qed. (*morphisms are extensionally equal*) Lemma same_genZ : forall x, [x] == gen_phiZ1 x. Proof. destruct x;rsimpl; try rewrite same_gen; reflexivity. Qed. Lemma gen_Zeqb_ok : forall x y, Zeq_bool x y = true -> [x] == [y]. Proof. intros x y H7. assert (H10 := Zeq_bool_eq x y H7);unfold IDphi in H10. rewrite H10;reflexivity. Qed. Lemma gen_phiZ1_add_pos_neg : forall x y, gen_phiZ1 (Z.pos_sub x y) == gen_phiPOS1 x + -gen_phiPOS1 y. Proof. intros x y. generalize (Z.pos_sub_discr x y). destruct (Z.pos_sub x y) as [|p|p]; intros; subst. - now rewrite ring_opp_def. - rewrite ARgen_phiPOS_add;simpl;norm. add_push (gen_phiPOS1 p). rewrite ring_opp_def;norm. - rewrite ARgen_phiPOS_add;simpl;norm. rewrite ring_opp_def;norm. Qed. Lemma match_compOpp : forall x (B:Type) (be bl bg:B), match CompOpp x with Eq => be | Lt => bl | Gt => bg end = match x with Eq => be | Lt => bg | Gt => bl end. Proof. destruct x;simpl;intros;trivial. Qed. Lemma gen_phiZ_add : forall x y, [x + y] == [x] + [y]. Proof. intros x y; repeat rewrite same_genZ; generalize x y;clear x y. induction x;destruct y;simpl;norm. apply ARgen_phiPOS_add. apply gen_phiZ1_add_pos_neg. rewrite gen_phiZ1_add_pos_neg. rewrite ring_add_comm. reflexivity. rewrite ARgen_phiPOS_add. rewrite ring_opp_add. reflexivity. Qed. Lemma gen_phiZ_opp : forall x, [- x] == - [x]. Proof. intros x. repeat rewrite same_genZ. generalize x ;clear x. induction x;simpl;norm. rewrite ring_opp_opp. reflexivity. Qed. Lemma gen_phiZ_mul : forall x y, [x * y] == [x] * [y]. Proof. intros x y;repeat rewrite same_genZ. destruct x;destruct y;simpl;norm; rewrite ARgen_phiPOS_mult;try (norm;fail). rewrite ring_opp_opp ;reflexivity. Qed. Lemma gen_phiZ_ext : forall x y : Z, x = y -> [x] == [y]. Proof. intros;subst;reflexivity. Qed. Declare Equivalent Keys bracket gen_phiZ. (*proof that [.] satisfies morphism specifications*) Global Instance gen_phiZ_morph : (@Ring_morphism (Z:Type) R _ _ _ _ _ _ _ Zops Zr _ _ _ _ _ _ _ _ _ gen_phiZ) . (* beurk!*) apply Build_Ring_morphism; simpl;try reflexivity. apply gen_phiZ_add. intros. rewrite ring_sub_def. replace (x-y)%Z with (x + (-y))%Z. now rewrite gen_phiZ_add, gen_phiZ_opp, ring_sub_def. reflexivity. apply gen_phiZ_mul. apply gen_phiZ_opp. apply gen_phiZ_ext. Defined. End ZMORPHISM. Instance multiplication_phi_ring{R:Type}`{Ring R} : Multiplication := {multiplication x y := (gen_phiZ x) * y}. coq-8.11.0/plugins/setoid_ring/ZArithRing.v0000644000175000017500000000323413612664533020467 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t | _ => constr:(NotConstant) end. Ltac isZpow_coef t := match t with | Zpos ?p => isPcst p | Z0 => constr:(true) | _ => constr:(false) end. Notation N_of_Z := Z.to_N (only parsing). Ltac Zpow_tac t := match isZpow_coef t with | true => constr:(N_of_Z t) | _ => constr:(NotConstant) end. Ltac Zpower_neg := repeat match goal with | [|- ?G] => match G with | context c [Z.pow _ (Zneg _)] => let t := context c [Z0] in change t end end. Add Ring Zr : Zth (decidable Zeq_bool_eq, constants [Zcst], preprocess [Zpower_neg;unfold Z.succ], power_tac Zpower_theory [Zpow_tac], (* The following two options are not needed; they are the default choice when the set of coefficient is the usual ring Z *) div (InitialRing.Ztriv_div_th (@Eqsth Z) (@IDphi Z)), sign get_signZ_th). coq-8.11.0/plugins/setoid_ring/Field_tac.v0000644000175000017500000004641213612664533020325 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* match t with | rO => fun _ => constr:(@FEO C) | rI => fun _ => constr:(@FEI C) | (radd ?t1 ?t2) => fun _ => let e1 := mkP t1 in let e2 := mkP t2 in constr:(@FEadd C e1 e2) | (rmul ?t1 ?t2) => fun _ => let e1 := mkP t1 in let e2 := mkP t2 in constr:(@FEmul C e1 e2) | (rsub ?t1 ?t2) => fun _ => let e1 := mkP t1 in let e2 := mkP t2 in constr:(@FEsub C e1 e2) | (ropp ?t1) => fun _ => let e1 := mkP t1 in constr:(@FEopp C e1) | (rdiv ?t1 ?t2) => fun _ => let e1 := mkP t1 in let e2 := mkP t2 in constr:(@FEdiv C e1 e2) | (rinv ?t1) => fun _ => let e1 := mkP t1 in constr:(@FEinv C e1) | (rpow ?t1 ?n) => match CstPow n with | InitialRing.NotConstant => fun _ => let p := Find_at t fv in constr:(@FEX C p) | ?c => fun _ => let e1 := mkP t1 in constr:(@FEpow C e1 c) end | _ => fun _ => let p := Find_at t fv in constr:(@FEX C p) end | ?c => fun _ => constr:(@FEc C c) end in f () in mkP t. (* We do not assume that Cst recognizes the rO and rI terms as constants, as *) (* the tactic could be used to discriminate occurrences of an opaque *) (* constant phi, with (phi 0) not convertible to 0 for instance *) Ltac FFV Cst CstPow rO rI add mul sub opp div inv pow t fv := let rec TFV t fv := match Cst t with | InitialRing.NotConstant => match t with | rO => fv | rI => fv | (add ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv) | (mul ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv) | (sub ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv) | (opp ?t1) => TFV t1 fv | (div ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv) | (inv ?t1) => TFV t1 fv | (pow ?t1 ?n) => match CstPow n with | InitialRing.NotConstant => AddFvTail t fv | _ => TFV t1 fv end | _ => AddFvTail t fv end | _ => fv end in TFV t fv. (* packaging the field structure *) (* TODO: inline PackField into field_lookup *) Ltac PackField F req Cst_tac Pow_tac L1 L2 L3 L4 cond_ok pre post := let FLD := match type of L1 with | context [req (@FEeval ?R ?rO ?rI ?radd ?rmul ?rsub ?ropp ?rdiv ?rinv ?C ?phi ?Cpow ?Cp_phi ?rpow _ _) _ ] => (fun proj => proj Cst_tac Pow_tac pre post req rO rI radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok) | _ => fail 1 "field anomaly: bad correctness lemma (parse)" end in F FLD. Ltac get_FldPre FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => pre). Ltac get_FldPost FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => post). Ltac get_L1 FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => L1). Ltac get_SimplifyEqLemma FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => L2). Ltac get_SimplifyLemma FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => L3). Ltac get_L4 FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => L4). Ltac get_CondLemma FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => cond_ok). Ltac get_FldEq FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => req). Ltac get_FldCarrier FLD := let req := get_FldEq FLD in relation_carrier req. Ltac get_RingFV FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => FV Cst_tac Pow_tac r0 r1 radd rmul rsub ropp rpow). Ltac get_FFV FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => FFV Cst_tac Pow_tac r0 r1 radd rmul rsub ropp rdiv rinv rpow). Ltac get_RingMeta FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => mkPolexpr C Cst_tac Pow_tac r0 r1 radd rmul rsub ropp rpow). Ltac get_Meta FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => mkFieldexpr C Cst_tac Pow_tac r0 r1 radd rmul rsub ropp rdiv rinv rpow). Ltac get_Hyp_tac FLD := FLD ltac: (fun Cst_tac Pow_tac pre post req r0 r1 radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok => let mkPol := mkPolexpr C Cst_tac Pow_tac r0 r1 radd rmul rsub ropp rpow in fun fv lH => mkHyp_tac C req ltac:(fun t => mkPol t fv) lH). Ltac get_FEeval FLD := let L1 := get_L1 FLD in match type of L1 with | context [(@FEeval ?R ?r0 ?r1 ?add ?mul ?sub ?opp ?div ?inv ?C ?phi ?Cpow ?powphi ?pow _ _)] => constr:(@FEeval R r0 r1 add mul sub opp div inv C phi Cpow powphi pow) | _ => fail 1 "field anomaly: bad correctness lemma (get_FEeval)" end. (* simplifying the non-zero condition... *) Ltac fold_field_cond req := let rec fold_concl t := match t with ?x /\ ?y => let fx := fold_concl x in let fy := fold_concl y in constr:(fx/\fy) | req ?x ?y -> False => constr:(~ req x y) | _ => t end in let ft := fold_concl Get_goal in change ft. Ltac simpl_PCond FLD := let req := get_FldEq FLD in let lemma := get_CondLemma FLD in try (apply lemma; intros lock lock_def; vm_compute; rewrite lock_def; clear lock_def lock); protect_fv "field_cond"; fold_field_cond req; try exact I. Ltac simpl_PCond_BEURK FLD := let req := get_FldEq FLD in let lemma := get_CondLemma FLD in (apply lemma; intros lock lock_def; vm_compute; rewrite lock_def; clear lock_def lock); protect_fv "field_cond"; fold_field_cond req. (* Rewriting (field_simplify) *) Ltac Field_norm_gen f n FLD lH rl := let mkFV := get_RingFV FLD in let mkFFV := get_FFV FLD in let mkFE := get_Meta FLD in let fv0 := FV_hypo_tac mkFV ltac:(get_FldEq FLD) lH in let lemma_tac fv kont := let lemma := get_SimplifyLemma FLD in (* reify equations of the context *) let lpe := get_Hyp_tac FLD fv lH in let vlpe := fresh "hyps" in pose (vlpe := lpe); let prh := proofHyp_tac lH in (* compute the normal form of the reified hyps *) let vlmp := fresh "hyps'" in let vlmp_eq := fresh "hyps_eq" in let mk_monpol := get_MonPol lemma in compute_assertion vlmp_eq vlmp (mk_monpol vlpe); (* partially instantiate the lemma *) let lem := fresh "f_rw_lemma" in (assert (lem := lemma n vlpe fv prh vlmp vlmp_eq) || fail "type error when building the rewriting lemma"); (* continuation will call main_tac for all reified terms *) kont lem; (* at the end, cleanup *) (clear lem vlmp_eq vlmp vlpe||idtac"Field_norm_gen:cleanup failed") in (* each instance of the lemma is simplified then passed to f *) let main_tac H := protect_fv "field" in H; f H in (* generate and use equations for each expression *) ReflexiveRewriteTactic mkFFV mkFE lemma_tac main_tac fv0 rl; try simpl_PCond FLD. Ltac Field_simplify_gen f FLD lH rl := get_FldPre FLD (); Field_norm_gen f ring_subst_niter FLD lH rl; get_FldPost FLD (). Ltac Field_simplify := Field_simplify_gen ltac:(fun H => rewrite H). Tactic Notation (at level 0) "field_simplify" constr_list(rl) := let G := Get_goal in field_lookup (PackField Field_simplify) [] rl G. Tactic Notation (at level 0) "field_simplify" "[" constr_list(lH) "]" constr_list(rl) := let G := Get_goal in field_lookup (PackField Field_simplify) [lH] rl G. Tactic Notation "field_simplify" constr_list(rl) "in" hyp(H):= let G := Get_goal in let t := type of H in let g := fresh "goal" in set (g:= G); revert H; field_lookup (PackField Field_simplify) [] rl t; intro H; unfold g;clear g. Tactic Notation "field_simplify" "["constr_list(lH) "]" constr_list(rl) "in" hyp(H):= let G := Get_goal in let t := type of H in let g := fresh "goal" in set (g:= G); revert H; field_lookup (PackField Field_simplify) [lH] rl t; intro H; unfold g;clear g. (* Ltac Field_simplify_in hyp:= Field_simplify_gen ltac:(fun H => rewrite H in hyp). Tactic Notation (at level 0) "field_simplify" constr_list(rl) "in" hyp(h) := let t := type of h in field_lookup (Field_simplify_in h) [] rl t. Tactic Notation (at level 0) "field_simplify" "[" constr_list(lH) "]" constr_list(rl) "in" hyp(h) := let t := type of h in field_lookup (Field_simplify_in h) [lH] rl t. *) (** Generic tactic for solving equations *) Ltac Field_Scheme Simpl_tac n lemma FLD lH := let req := get_FldEq FLD in let mkFV := get_RingFV FLD in let mkFFV := get_FFV FLD in let mkFE := get_Meta FLD in let Main_eq t1 t2 := let fv := FV_hypo_tac mkFV req lH in let fv := mkFFV t1 fv in let fv := mkFFV t2 fv in let lpe := get_Hyp_tac FLD fv lH in let prh := proofHyp_tac lH in let vlpe := fresh "list_hyp" in let fe1 := mkFE t1 fv in let fe2 := mkFE t2 fv in pose (vlpe := lpe); let nlemma := fresh "field_lemma" in (assert (nlemma := lemma n fv vlpe fe1 fe2 prh) || fail "field anomaly:failed to build lemma"); ProveLemmaHyps nlemma ltac:(fun ilemma => apply ilemma || fail "field anomaly: failed in applying lemma"; [ Simpl_tac | simpl_PCond FLD]); clear nlemma; subst vlpe in OnEquation req Main_eq. (* solve completely a field equation, leaving non-zero conditions to be proved (field) *) Ltac FIELD FLD lH rl := let Simpl := vm_compute; reflexivity || fail "not a valid field equation" in let lemma := get_L1 FLD in get_FldPre FLD (); Field_Scheme Simpl Ring_tac.ring_subst_niter lemma FLD lH; try exact I; get_FldPost FLD(). Tactic Notation (at level 0) "field" := let G := Get_goal in field_lookup (PackField FIELD) [] G. Tactic Notation (at level 0) "field" "[" constr_list(lH) "]" := let G := Get_goal in field_lookup (PackField FIELD) [lH] G. (* transforms a field equation to an equivalent (simplified) ring equation, and leaves non-zero conditions to be proved (field_simplify_eq) *) Ltac FIELD_SIMPL FLD lH rl := let Simpl := (protect_fv "field") in let lemma := get_SimplifyEqLemma FLD in get_FldPre FLD (); Field_Scheme Simpl Ring_tac.ring_subst_niter lemma FLD lH; get_FldPost FLD (). Tactic Notation (at level 0) "field_simplify_eq" := let G := Get_goal in field_lookup (PackField FIELD_SIMPL) [] G. Tactic Notation (at level 0) "field_simplify_eq" "[" constr_list(lH) "]" := let G := Get_goal in field_lookup (PackField FIELD_SIMPL) [lH] G. (* Same as FIELD_SIMPL but in hypothesis *) Ltac Field_simplify_eq n FLD lH := let req := get_FldEq FLD in let mkFV := get_RingFV FLD in let mkFFV := get_FFV FLD in let mkFE := get_Meta FLD in let lemma := get_L4 FLD in let hyp := fresh "hyp" in intro hyp; OnEquationHyp req hyp ltac:(fun t1 t2 => let fv := FV_hypo_tac mkFV req lH in let fv := mkFFV t1 fv in let fv := mkFFV t2 fv in let lpe := get_Hyp_tac FLD fv lH in let prh := proofHyp_tac lH in let fe1 := mkFE t1 fv in let fe2 := mkFE t2 fv in let vlpe := fresh "vlpe" in ProveLemmaHyps (lemma n fv lpe fe1 fe2 prh) ltac:(fun ilemma => match type of ilemma with | req _ _ -> _ -> ?EQ => let tmp := fresh "tmp" in assert (tmp : EQ); [ apply ilemma; [ exact hyp | simpl_PCond_BEURK FLD] | protect_fv "field" in tmp; revert tmp ]; clear hyp end)). Ltac FIELD_SIMPL_EQ FLD lH rl := get_FldPre FLD (); Field_simplify_eq Ring_tac.ring_subst_niter FLD lH; get_FldPost FLD (). Tactic Notation (at level 0) "field_simplify_eq" "in" hyp(H) := let t := type of H in generalize H; field_lookup (PackField FIELD_SIMPL_EQ) [] t; [ try exact I | clear H;intro H]. Tactic Notation (at level 0) "field_simplify_eq" "[" constr_list(lH) "]" "in" hyp(H) := let t := type of H in generalize H; field_lookup (PackField FIELD_SIMPL_EQ) [lH] t; [ try exact I |clear H;intro H]. (* More generic tactics to build variants of field *) (* This tactic reifies c and pass to F: - the FLD structure gathering all info in the field DB - the atom list - the expression (FExpr) *) Ltac gen_with_field F c := let MetaExpr FLD _ rl := let R := get_FldCarrier FLD in let mkFFV := get_FFV FLD in let mkFE := get_Meta FLD in let csr := match rl with | List.cons ?r _ => r | _ => fail 1 "anomaly: ill-formed list" end in let fv := mkFFV csr (@List.nil R) in let expr := mkFE csr fv in F FLD fv expr in field_lookup (PackField MetaExpr) [] (c=c). (* pushes the equation expr = ope(expr) in the goal, and discharge it with field *) Ltac prove_field_eqn ope FLD fv expr := let res := ope expr in let expr' := fresh "input_expr" in pose (expr' := expr); let res' := fresh "result" in pose (res' := res); let lemma := get_L1 FLD in let lemma := constr:(lemma O fv List.nil expr' res' I List.nil (eq_refl _)) in let ty := type of lemma in let lhs := match ty with forall _, ?lhs=_ -> _ => lhs end in let rhs := match ty with forall _, _=_ -> forall _, ?rhs=_ -> _ => rhs end in let lhs' := fresh "lhs" in let lhs_eq := fresh "lhs_eq" in let rhs' := fresh "rhs" in let rhs_eq := fresh "rhs_eq" in compute_assertion lhs_eq lhs' lhs; compute_assertion rhs_eq rhs' rhs; let H := fresh "fld_eqn" in refine (_ (lemma lhs' lhs_eq rhs' rhs_eq _ _)); (* main goal *) [intro H;protect_fv "field" in H; revert H (* ring-nf(lhs') = ring-nf(rhs') *) | vm_compute; reflexivity || fail "field cannot prove this equality" (* denominator condition *) | simpl_PCond FLD]; clear lhs_eq rhs_eq; subst lhs' rhs'. Ltac prove_with_field ope c := gen_with_field ltac:(prove_field_eqn ope) c. (* Prove an equation x=ope(x) and rewrite with it *) Ltac prove_rw ope x := prove_with_field ope x; [ let H := fresh "Heq_maple" in intro H; rewrite H; clear H |..]. (* Apply ope (FExpr->FExpr) on an expression *) Ltac reduce_field_expr ope kont FLD fv expr := let evfun := get_FEeval FLD in let res := ope expr in let c := (eval simpl_field_expr in (evfun fv res)) in kont c. (* Hack to let a Ltac return a term in the context of a primitive tactic *) Ltac return_term x := generalize (eq_refl x). Ltac get_term := match goal with | |- ?x = _ -> _ => x end. (* Turn an operation on field expressions (FExpr) into a reduction on terms (in the field carrier). Because of field_lookup, the tactic cannot return a term directly, so it is returned via the conclusion of the goal (return_term). *) Ltac reduce_field_ope ope c := gen_with_field ltac:(reduce_field_expr ope return_term) c. (* Adding a new field *) Ltac ring_of_field f := match type of f with | almost_field_theory _ _ _ _ _ _ _ _ _ => constr:(AF_AR f) | field_theory _ _ _ _ _ _ _ _ _ => constr:(F_R f) | semi_field_theory _ _ _ _ _ _ _ => constr:(SF_SR f) end. Ltac coerce_to_almost_field set ext f := match type of f with | almost_field_theory _ _ _ _ _ _ _ _ _ => f | field_theory _ _ _ _ _ _ _ _ _ => constr:(F2AF set ext f) | semi_field_theory _ _ _ _ _ _ _ => constr:(SF2AF set f) end. Ltac field_elements set ext fspec pspec sspec dspec rk := let afth := coerce_to_almost_field set ext fspec in let rspec := ring_of_field fspec in ring_elements set ext rspec pspec sspec dspec rk ltac:(fun arth ext_r morph p_spec s_spec d_spec f => f afth ext_r morph p_spec s_spec d_spec). Ltac field_lemmas set ext inv_m fspec pspec sspec dspec rk := let get_lemma := match pspec with None => fun x y => x | _ => fun x y => y end in let simpl_eq_lemma := get_lemma Field_simplify_eq_correct Field_simplify_eq_pow_correct in let simpl_eq_in_lemma := get_lemma Field_simplify_eq_in_correct Field_simplify_eq_pow_in_correct in let rw_lemma := get_lemma Field_rw_correct Field_rw_pow_correct in field_elements set ext fspec pspec sspec dspec rk ltac:(fun afth ext_r morph p_spec s_spec d_spec => match morph with | _ => let field_ok1 := constr:(Field_correct set ext_r inv_m afth morph) in match p_spec with | mkhypo ?pp_spec => let field_ok2 := constr:(field_ok1 _ _ _ pp_spec) in match s_spec with | mkhypo ?ss_spec => match d_spec with | mkhypo ?dd_spec => let field_ok := constr:(field_ok2 _ dd_spec) in let mk_lemma lemma := constr:(lemma _ _ _ _ _ _ _ _ _ _ set ext_r inv_m afth _ _ _ _ _ _ _ _ _ morph _ _ _ pp_spec _ ss_spec _ dd_spec) in let field_simpl_eq_ok := mk_lemma simpl_eq_lemma in let field_simpl_ok := mk_lemma rw_lemma in let field_simpl_eq_in := mk_lemma simpl_eq_in_lemma in let cond1_ok := constr:(Pcond_simpl_gen set ext_r afth morph pp_spec dd_spec) in let cond2_ok := constr:(Pcond_simpl_complete set ext_r afth morph pp_spec dd_spec) in (fun f => f afth ext_r morph field_ok field_simpl_ok field_simpl_eq_ok field_simpl_eq_in cond1_ok cond2_ok) | _ => fail 4 "field: bad coefficient division specification" end | _ => fail 3 "field: bad sign specification" end | _ => fail 2 "field: bad power specification" end | _ => fail 1 "field internal error : field_lemmas, please report" end). coq-8.11.0/plugins/setoid_ring/Rings_Q.v0000644000175000017500000000267013612664533020013 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* R->R) (ropp : R -> R). Variable req : R -> R -> Prop. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y). Notation "x - y " := (rsub x y). Notation "- x" := (ropp x). Notation "x == y" := (req x y). Variable Rsth : Setoid_Theory R req. Add Parametric Relation : R req reflexivity proved by (@Equivalence_Reflexive _ _ Rsth) symmetry proved by (@Equivalence_Symmetric _ _ Rsth) transitivity proved by (@Equivalence_Transitive _ _ Rsth) as R_setoid3. Ltac rrefl := gen_reflexivity Rsth. Variable Reqe : ring_eq_ext radd rmul ropp req. Add Morphism radd with signature (req ==> req ==> req) as radd_ext3. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext3. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext3. Proof. exact (Ropp_ext Reqe). Qed. Fixpoint gen_phiPOS1 (p:positive) : R := match p with | xH => 1 | xO p => (1 + 1) * (gen_phiPOS1 p) | xI p => 1 + ((1 + 1) * (gen_phiPOS1 p)) end. Fixpoint gen_phiPOS (p:positive) : R := match p with | xH => 1 | xO xH => (1 + 1) | xO p => (1 + 1) * (gen_phiPOS p) | xI xH => 1 + (1 +1) | xI p => 1 + ((1 + 1) * (gen_phiPOS p)) end. Definition gen_phiZ1 z := match z with | Zpos p => gen_phiPOS1 p | Z0 => 0 | Zneg p => -(gen_phiPOS1 p) end. Definition gen_phiZ z := match z with | Zpos p => gen_phiPOS p | Z0 => 0 | Zneg p => -(gen_phiPOS p) end. Notation "[ x ]" := (gen_phiZ x). Definition get_signZ z := match z with | Zneg p => Some (Zpos p) | _ => None end. Lemma get_signZ_th : sign_theory Z.opp Zeq_bool get_signZ. Proof. constructor. destruct c;intros;try discriminate. injection H as [= <-]. simpl. unfold Zeq_bool. rewrite Z.compare_refl. trivial. Qed. Section ALMOST_RING. Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req. Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext3. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed. Ltac norm := gen_srewrite Rsth Reqe ARth. Ltac add_push := gen_add_push radd Rsth Reqe ARth. Lemma same_gen : forall x, gen_phiPOS1 x == gen_phiPOS x. Proof. induction x;simpl. rewrite IHx;destruct x;simpl;norm. rewrite IHx;destruct x;simpl;norm. rrefl. Qed. Lemma ARgen_phiPOS_Psucc : forall x, gen_phiPOS1 (Pos.succ x) == 1 + (gen_phiPOS1 x). Proof. induction x;simpl;norm. rewrite IHx;norm. add_push 1;rrefl. Qed. Lemma ARgen_phiPOS_add : forall x y, gen_phiPOS1 (x + y) == (gen_phiPOS1 x) + (gen_phiPOS1 y). Proof. induction x;destruct y;simpl;norm. rewrite Pos.add_carry_spec. rewrite ARgen_phiPOS_Psucc. rewrite IHx;norm. add_push (gen_phiPOS1 y);add_push 1;rrefl. rewrite IHx;norm;add_push (gen_phiPOS1 y);rrefl. rewrite ARgen_phiPOS_Psucc;norm;add_push 1;rrefl. rewrite IHx;norm;add_push(gen_phiPOS1 y); add_push 1;rrefl. rewrite IHx;norm;add_push(gen_phiPOS1 y);rrefl. add_push 1;rrefl. rewrite ARgen_phiPOS_Psucc;norm;add_push 1;rrefl. Qed. Lemma ARgen_phiPOS_mult : forall x y, gen_phiPOS1 (x * y) == gen_phiPOS1 x * gen_phiPOS1 y. Proof. induction x;intros;simpl;norm. rewrite ARgen_phiPOS_add;simpl;rewrite IHx;norm. rewrite IHx;rrefl. Qed. End ALMOST_RING. Variable Rth : ring_theory 0 1 radd rmul rsub ropp req. Let ARth := Rth_ARth Rsth Reqe Rth. Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext4. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed. Ltac norm := gen_srewrite Rsth Reqe ARth. Ltac add_push := gen_add_push radd Rsth Reqe ARth. (*morphisms are extensionally equal*) Lemma same_genZ : forall x, [x] == gen_phiZ1 x. Proof. destruct x;simpl; try rewrite (same_gen ARth);rrefl. Qed. Lemma gen_Zeqb_ok : forall x y, Zeq_bool x y = true -> [x] == [y]. Proof. intros x y H. assert (H1 := Zeq_bool_eq x y H);unfold IDphi in H1. rewrite H1;rrefl. Qed. Lemma gen_phiZ1_pos_sub : forall x y, gen_phiZ1 (Z.pos_sub x y) == gen_phiPOS1 x + -gen_phiPOS1 y. Proof. intros x y. rewrite Z.pos_sub_spec. case Pos.compare_spec; intros H; simpl. rewrite H. rewrite (Ropp_def Rth);rrefl. rewrite <- (Pos.sub_add y x H) at 2. rewrite Pos.add_comm. rewrite (ARgen_phiPOS_add ARth);simpl;norm. rewrite (Ropp_def Rth);norm. rewrite <- (Pos.sub_add x y H) at 2. rewrite (ARgen_phiPOS_add ARth);simpl;norm. add_push (gen_phiPOS1 (x-y));rewrite (Ropp_def Rth); norm. Qed. Lemma gen_phiZ_add : forall x y, [x + y] == [x] + [y]. Proof. intros x y; repeat rewrite same_genZ; generalize x y;clear x y. destruct x, y; simpl; norm. apply (ARgen_phiPOS_add ARth). apply gen_phiZ1_pos_sub. rewrite gen_phiZ1_pos_sub. apply (Radd_comm Rth). rewrite (ARgen_phiPOS_add ARth); norm. Qed. Lemma gen_phiZ_mul : forall x y, [x * y] == [x] * [y]. Proof. intros x y;repeat rewrite same_genZ. destruct x;destruct y;simpl;norm; rewrite (ARgen_phiPOS_mult ARth);try (norm;fail). rewrite (Ropp_opp Rsth Reqe Rth);rrefl. Qed. Lemma gen_phiZ_ext : forall x y : Z, x = y -> [x] == [y]. Proof. intros;subst;rrefl. Qed. (*proof that [.] satisfies morphism specifications*) Lemma gen_phiZ_morph : ring_morph 0 1 radd rmul rsub ropp req Z0 (Zpos xH) Z.add Z.mul Z.sub Z.opp Zeq_bool gen_phiZ. Proof. assert ( SRmorph : semi_morph 0 1 radd rmul req Z0 (Zpos xH) Z.add Z.mul Zeq_bool gen_phiZ). apply mkRmorph;simpl;try rrefl. apply gen_phiZ_add. apply gen_phiZ_mul. apply gen_Zeqb_ok. apply (Smorph_morph Rsth Reqe Rth Zth SRmorph gen_phiZ_ext). Qed. End ZMORPHISM. (** N is a semi-ring and a setoid*) Lemma Nsth : Setoid_Theory N (@eq N). Proof (Eqsth N). Lemma Nseqe : sring_eq_ext N.add N.mul (@eq N). Proof (Eq_s_ext N.add N.mul). Lemma Nth : semi_ring_theory 0%N 1%N N.add N.mul (@eq N). Proof. constructor. exact N.add_0_l. exact N.add_comm. exact N.add_assoc. exact N.mul_1_l. exact N.mul_0_l. exact N.mul_comm. exact N.mul_assoc. exact N.mul_add_distr_r. Qed. Definition Nsub := SRsub N.add. Definition Nopp := (@SRopp N). Lemma Neqe : ring_eq_ext N.add N.mul Nopp (@eq N). Proof (SReqe_Reqe Nseqe). Lemma Nath : almost_ring_theory 0%N 1%N N.add N.mul Nsub Nopp (@eq N). Proof (SRth_ARth Nsth Nth). Lemma Neqb_ok : forall x y, N.eqb x y = true -> x = y. Proof. exact (fun x y => proj1 (N.eqb_eq x y)). Qed. (**Same as above : definition of two, extensionally equal, generic morphisms *) (**from N to any semi-ring*) Section NMORPHISM. Variable R : Type. Variable (rO rI : R) (radd rmul: R->R->R). Variable req : R -> R -> Prop. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y). Variable Rsth : Setoid_Theory R req. Add Parametric Relation : R req reflexivity proved by (@Equivalence_Reflexive _ _ Rsth) symmetry proved by (@Equivalence_Symmetric _ _ Rsth) transitivity proved by (@Equivalence_Transitive _ _ Rsth) as R_setoid4. Ltac rrefl := gen_reflexivity Rsth. Variable SReqe : sring_eq_ext radd rmul req. Variable SRth : semi_ring_theory 0 1 radd rmul req. Let ARth := SRth_ARth Rsth SRth. Let Reqe := SReqe_Reqe SReqe. Let ropp := (@SRopp R). Let rsub := (@SRsub R radd). Notation "x - y " := (rsub x y). Notation "- x" := (ropp x). Notation "x == y" := (req x y). Add Morphism radd with signature (req ==> req ==> req) as radd_ext4. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext4. Proof. exact (Rmul_ext Reqe). Qed. Ltac norm := gen_srewrite_sr Rsth Reqe ARth. Definition gen_phiN1 x := match x with | N0 => 0 | Npos x => gen_phiPOS1 1 radd rmul x end. Definition gen_phiN x := match x with | N0 => 0 | Npos x => gen_phiPOS 1 radd rmul x end. Notation "[ x ]" := (gen_phiN x). Lemma same_genN : forall x, [x] == gen_phiN1 x. Proof. destruct x;simpl. reflexivity. now rewrite (same_gen Rsth Reqe ARth). Qed. Lemma gen_phiN_add : forall x y, [x + y] == [x] + [y]. Proof. intros x y;repeat rewrite same_genN. destruct x;destruct y;simpl;norm. apply (ARgen_phiPOS_add Rsth Reqe ARth). Qed. Lemma gen_phiN_mult : forall x y, [x * y] == [x] * [y]. Proof. intros x y;repeat rewrite same_genN. destruct x;destruct y;simpl;norm. apply (ARgen_phiPOS_mult Rsth Reqe ARth). Qed. Lemma gen_phiN_sub : forall x y, [Nsub x y] == [x] - [y]. Proof. exact gen_phiN_add. Qed. (*gen_phiN satisfies morphism specifications*) Lemma gen_phiN_morph : ring_morph 0 1 radd rmul rsub ropp req 0%N 1%N N.add N.mul Nsub Nopp N.eqb gen_phiN. Proof. constructor; simpl; try reflexivity. apply gen_phiN_add. apply gen_phiN_sub. apply gen_phiN_mult. intros x y EQ. apply N.eqb_eq in EQ. now subst. Qed. End NMORPHISM. (* Words on N : initial structure for almost-rings. *) Definition Nword := list N. Definition NwO : Nword := nil. Definition NwI : Nword := 1%N :: nil. Definition Nwcons n (w : Nword) : Nword := match w, n with | nil, 0%N => nil | _, _ => n :: w end. Fixpoint Nwadd (w1 w2 : Nword) {struct w1} : Nword := match w1, w2 with | n1::w1', n2:: w2' => (n1+n2)%N :: Nwadd w1' w2' | nil, _ => w2 | _, nil => w1 end. Definition Nwopp (w:Nword) : Nword := Nwcons 0%N w. Definition Nwsub w1 w2 := Nwadd w1 (Nwopp w2). Fixpoint Nwscal (n : N) (w : Nword) {struct w} : Nword := match w with | m :: w' => (n*m)%N :: Nwscal n w' | nil => nil end. Fixpoint Nwmul (w1 w2 : Nword) {struct w1} : Nword := match w1 with | 0%N::w1' => Nwopp (Nwmul w1' w2) | n1::w1' => Nwsub (Nwscal n1 w2) (Nwmul w1' w2) | nil => nil end. Fixpoint Nw_is0 (w : Nword) : bool := match w with | nil => true | 0%N :: w' => Nw_is0 w' | _ => false end. Fixpoint Nweq_bool (w1 w2 : Nword) {struct w1} : bool := match w1, w2 with | n1::w1', n2::w2' => if N.eqb n1 n2 then Nweq_bool w1' w2' else false | nil, _ => Nw_is0 w2 | _, nil => Nw_is0 w1 end. Section NWORDMORPHISM. Variable R : Type. Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R). Variable req : R -> R -> Prop. Notation "0" := rO. Notation "1" := rI. Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y). Notation "x - y " := (rsub x y). Notation "- x" := (ropp x). Notation "x == y" := (req x y). Variable Rsth : Setoid_Theory R req. Add Parametric Relation : R req reflexivity proved by (@Equivalence_Reflexive _ _ Rsth) symmetry proved by (@Equivalence_Symmetric _ _ Rsth) transitivity proved by (@Equivalence_Transitive _ _ Rsth) as R_setoid5. Ltac rrefl := gen_reflexivity Rsth. Variable Reqe : ring_eq_ext radd rmul ropp req. Add Morphism radd with signature (req ==> req ==> req) as radd_ext5. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext5. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext5. Proof. exact (Ropp_ext Reqe). Qed. Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req. Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext7. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed. Ltac norm := gen_srewrite Rsth Reqe ARth. Ltac add_push := gen_add_push radd Rsth Reqe ARth. Fixpoint gen_phiNword (w : Nword) : R := match w with | nil => 0 | n :: nil => gen_phiN rO rI radd rmul n | N0 :: w' => - gen_phiNword w' | n::w' => gen_phiN rO rI radd rmul n - gen_phiNword w' end. Lemma gen_phiNword0_ok : forall w, Nw_is0 w = true -> gen_phiNword w == 0. Proof. induction w; simpl; intros; auto. reflexivity. destruct a. destruct w. reflexivity. rewrite IHw; trivial. apply (ARopp_zero Rsth Reqe ARth). discriminate. Qed. Lemma gen_phiNword_cons : forall w n, gen_phiNword (n::w) == gen_phiN rO rI radd rmul n - gen_phiNword w. induction w. destruct n; simpl; norm. intros. destruct n; norm. Qed. Lemma gen_phiNword_Nwcons : forall w n, gen_phiNword (Nwcons n w) == gen_phiN rO rI radd rmul n - gen_phiNword w. destruct w; intros. destruct n; norm. unfold Nwcons. rewrite gen_phiNword_cons. reflexivity. Qed. Lemma gen_phiNword_ok : forall w1 w2, Nweq_bool w1 w2 = true -> gen_phiNword w1 == gen_phiNword w2. induction w1; intros. simpl. rewrite (gen_phiNword0_ok _ H). reflexivity. rewrite gen_phiNword_cons. destruct w2. simpl in H. destruct a; try discriminate. rewrite (gen_phiNword0_ok _ H). norm. simpl in H. rewrite gen_phiNword_cons. case_eq (N.eqb a n); intros H0. rewrite H0 in H. apply N.eqb_eq in H0. rewrite <- H0. rewrite (IHw1 _ H). reflexivity. rewrite H0 in H; discriminate H. Qed. Lemma Nwadd_ok : forall x y, gen_phiNword (Nwadd x y) == gen_phiNword x + gen_phiNword y. induction x; intros. simpl. norm. destruct y. simpl Nwadd; norm. simpl Nwadd. repeat rewrite gen_phiNword_cons. rewrite (fun sreq => gen_phiN_add Rsth sreq (ARth_SRth ARth)) by (destruct Reqe; constructor; trivial). rewrite IHx. norm. add_push (- gen_phiNword x); reflexivity. Qed. Lemma Nwopp_ok : forall x, gen_phiNword (Nwopp x) == - gen_phiNword x. simpl. unfold Nwopp; simpl. intros. rewrite gen_phiNword_Nwcons; norm. Qed. Lemma Nwscal_ok : forall n x, gen_phiNword (Nwscal n x) == gen_phiN rO rI radd rmul n * gen_phiNword x. induction x; intros. norm. simpl Nwscal. repeat rewrite gen_phiNword_cons. rewrite (fun sreq => gen_phiN_mult Rsth sreq (ARth_SRth ARth)) by (destruct Reqe; constructor; trivial). rewrite IHx. norm. Qed. Lemma Nwmul_ok : forall x y, gen_phiNword (Nwmul x y) == gen_phiNword x * gen_phiNword y. induction x; intros. norm. destruct a. simpl Nwmul. rewrite Nwopp_ok. rewrite IHx. rewrite gen_phiNword_cons. norm. simpl Nwmul. unfold Nwsub. rewrite Nwadd_ok. rewrite Nwscal_ok. rewrite Nwopp_ok. rewrite IHx. rewrite gen_phiNword_cons. norm. Qed. (* Proof that [.] satisfies morphism specifications *) Lemma gen_phiNword_morph : ring_morph 0 1 radd rmul rsub ropp req NwO NwI Nwadd Nwmul Nwsub Nwopp Nweq_bool gen_phiNword. constructor. reflexivity. reflexivity. exact Nwadd_ok. intros. unfold Nwsub. rewrite Nwadd_ok. rewrite Nwopp_ok. norm. exact Nwmul_ok. exact Nwopp_ok. exact gen_phiNword_ok. Qed. End NWORDMORPHISM. Section GEN_DIV. Variables (R : Type) (rO : R) (rI : R) (radd : R -> R -> R) (rmul : R -> R -> R) (rsub : R -> R -> R) (ropp : R -> R) (req : R -> R -> Prop) (C : Type) (cO : C) (cI : C) (cadd : C -> C -> C) (cmul : C -> C -> C) (csub : C -> C -> C) (copp : C -> C) (ceqb : C -> C -> bool) (phi : C -> R). Variable Rsth : Setoid_Theory R req. Variable Reqe : ring_eq_ext radd rmul ropp req. Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req. Variable morph : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi. (* Useful tactics *) Add Parametric Relation : R req reflexivity proved by (@Equivalence_Reflexive _ _ Rsth) symmetry proved by (@Equivalence_Symmetric _ _ Rsth) transitivity proved by (@Equivalence_Transitive _ _ Rsth) as R_set1. Ltac rrefl := gen_reflexivity Rsth. Add Morphism radd with signature (req ==> req ==> req) as radd_ext. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext. Proof. exact (Ropp_ext Reqe). Qed. Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed. Ltac rsimpl := gen_srewrite Rsth Reqe ARth. Definition triv_div x y := if ceqb x y then (cI, cO) else (cO, x). Ltac Esimpl :=repeat (progress ( match goal with | |- context [phi cO] => rewrite (morph0 morph) | |- context [phi cI] => rewrite (morph1 morph) | |- context [phi (cadd ?x ?y)] => rewrite ((morph_add morph) x y) | |- context [phi (cmul ?x ?y)] => rewrite ((morph_mul morph) x y) | |- context [phi (csub ?x ?y)] => rewrite ((morph_sub morph) x y) | |- context [phi (copp ?x)] => rewrite ((morph_opp morph) x) end)). Lemma triv_div_th : Ring_theory.div_theory req cadd cmul phi triv_div. Proof. constructor. intros a b;unfold triv_div. assert (X:= morph_eq morph a b);destruct (ceqb a b). Esimpl. rewrite X; trivial. rsimpl. Esimpl; rsimpl. Qed. Variable zphi : Z -> R. Lemma Ztriv_div_th : div_theory req Z.add Z.mul zphi Z.quotrem. Proof. constructor. intros; generalize (Z.quotrem_eq a b); case Z.quotrem; intros; subst. rewrite Z.mul_comm; rsimpl. Qed. Variable nphi : N -> R. Lemma Ntriv_div_th : div_theory req N.add N.mul nphi N.div_eucl. constructor. intros; generalize (N.div_eucl_spec a b); case N.div_eucl; intros; subst. rewrite N.mul_comm; rsimpl. Qed. End GEN_DIV. (* syntaxification of constants in an abstract ring: the inverse of gen_phiPOS *) Ltac inv_gen_phi_pos rI add mul t := let rec inv_cst t := match t with rI => constr:(1%positive) | (add rI rI) => constr:(2%positive) | (add rI (add rI rI)) => constr:(3%positive) | (mul (add rI rI) ?p) => (* 2p *) match inv_cst p with NotConstant => constr:(NotConstant) | 1%positive => constr:(NotConstant) (* 2*1 is not convertible to 2 *) | ?p => constr:(xO p) end | (add rI (mul (add rI rI) ?p)) => (* 1+2p *) match inv_cst p with NotConstant => constr:(NotConstant) | 1%positive => constr:(NotConstant) | ?p => constr:(xI p) end | _ => constr:(NotConstant) end in inv_cst t. (* The (partial) inverse of gen_phiNword *) Ltac inv_gen_phiNword rO rI add mul opp t := match t with rO => constr:(NwO) | _ => match inv_gen_phi_pos rI add mul t with NotConstant => constr:(NotConstant) | ?p => constr:(Npos p::nil) end end. (* The inverse of gen_phiN *) Ltac inv_gen_phiN rO rI add mul t := match t with rO => constr:(0%N) | _ => match inv_gen_phi_pos rI add mul t with NotConstant => constr:(NotConstant) | ?p => constr:(Npos p) end end. (* The inverse of gen_phiZ *) Ltac inv_gen_phiZ rO rI add mul opp t := match t with rO => constr:(0%Z) | (opp ?p) => match inv_gen_phi_pos rI add mul p with NotConstant => constr:(NotConstant) | ?p => constr:(Zneg p) end | _ => match inv_gen_phi_pos rI add mul t with NotConstant => constr:(NotConstant) | ?p => constr:(Zpos p) end end. (* A simple tactic recognizing only 0 and 1. The inv_gen_phiX above are only optimisations that directly returns the reified constant instead of resorting to the constant propagation of the simplification algorithm. *) Ltac inv_gen_phi rO rI cO cI t := match t with | rO => cO | rI => cI end. (* A simple tactic recognizing no constant *) Ltac inv_morph_nothing t := constr:(NotConstant). Ltac coerce_to_almost_ring set ext rspec := match type of rspec with | ring_theory _ _ _ _ _ _ _ => constr:(Rth_ARth set ext rspec) | semi_ring_theory _ _ _ _ _ => constr:(SRth_ARth set rspec) | almost_ring_theory _ _ _ _ _ _ _ => rspec | _ => fail 1 "not a valid ring theory" end. Ltac coerce_to_ring_ext ext := match type of ext with | ring_eq_ext _ _ _ _ => ext | sring_eq_ext _ _ _ => constr:(SReqe_Reqe ext) | _ => fail 1 "not a valid ring_eq_ext theory" end. Ltac abstract_ring_morphism set ext rspec := match type of rspec with | ring_theory _ _ _ _ _ _ _ => constr:(gen_phiZ_morph set ext rspec) | semi_ring_theory _ _ _ _ _ => constr:(gen_phiN_morph set ext rspec) | almost_ring_theory _ _ _ _ _ _ _ => constr:(gen_phiNword_morph set ext rspec) | _ => fail 1 "bad ring structure" end. Record hypo : Type := mkhypo { hypo_type : Type; hypo_proof : hypo_type }. Ltac gen_ring_pow set arth pspec := match pspec with | None => match type of arth with | @almost_ring_theory ?R ?rO ?rI ?radd ?rmul ?rsub ?ropp ?req => constr:(mkhypo (@pow_N_th R rI rmul req set)) | _ => fail 1 "gen_ring_pow" end | Some ?t => constr:(t) end. Ltac gen_ring_sign morph sspec := match sspec with | None => match type of morph with | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req Z ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceqb ?phi => constr:(@mkhypo (sign_theory copp ceqb get_signZ) get_signZ_th) | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req ?C ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceqb ?phi => constr:(mkhypo (@get_sign_None_th C copp ceqb)) | _ => fail 2 "ring anomaly : default_sign_spec" end | Some ?t => constr:(t) end. Ltac default_div_spec set reqe arth morph := match type of morph with | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req Z ?c0 ?c1 Z.add Z.mul ?csub ?copp ?ceq_b ?phi => constr:(mkhypo (Ztriv_div_th set phi)) | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req N ?c0 ?c1 N.add N.mul ?csub ?copp ?ceq_b ?phi => constr:(mkhypo (Ntriv_div_th set phi)) | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req ?C ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceq_b ?phi => constr:(mkhypo (triv_div_th set reqe arth morph)) | _ => fail 1 "ring anomaly : default_sign_spec" end. Ltac gen_ring_div set reqe arth morph dspec := match dspec with | None => default_div_spec set reqe arth morph | Some ?t => constr:(t) end. Ltac ring_elements set ext rspec pspec sspec dspec rk := let arth := coerce_to_almost_ring set ext rspec in let ext_r := coerce_to_ring_ext ext in let morph := match rk with | Abstract => abstract_ring_morphism set ext rspec | @Computational ?reqb_ok => match type of arth with | almost_ring_theory ?rO ?rI ?add ?mul ?sub ?opp _ => constr:(IDmorph rO rI add mul sub opp set _ reqb_ok) | _ => fail 2 "ring anomaly" end | @Morphism ?m => match type of m with | ring_morph _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ => m | @semi_morph _ _ _ _ _ _ _ _ _ _ _ _ _ => constr:(SRmorph_Rmorph set m) | _ => fail 2 "ring anomaly" end | _ => fail 1 "ill-formed ring kind" end in let p_spec := gen_ring_pow set arth pspec in let s_spec := gen_ring_sign morph sspec in let d_spec := gen_ring_div set ext_r arth morph dspec in fun f => f arth ext_r morph p_spec s_spec d_spec. (* Given a ring structure and the kind of morphism, returns 2 lemmas (one for ring, and one for ring_simplify). *) Ltac ring_lemmas set ext rspec pspec sspec dspec rk := let gen_lemma2 := match pspec with | None => constr:(ring_rw_correct) | Some _ => constr:(ring_rw_pow_correct) end in ring_elements set ext rspec pspec sspec dspec rk ltac:(fun arth ext_r morph p_spec s_spec d_spec => lazymatch type of morph with | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req ?C ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceq_b ?phi => let gen_lemma2_0 := constr:(gen_lemma2 R r0 rI radd rmul rsub ropp req set ext_r arth C c0 c1 cadd cmul csub copp ceq_b phi morph) in lazymatch p_spec with | @mkhypo (power_theory _ _ _ ?Cp_phi ?rpow) ?pp_spec => let gen_lemma2_1 := constr:(gen_lemma2_0 _ Cp_phi rpow pp_spec) in lazymatch d_spec with | @mkhypo (div_theory _ _ _ _ ?cdiv) ?dd_spec => let gen_lemma2_2 := constr:(gen_lemma2_1 cdiv dd_spec) in lazymatch s_spec with | @mkhypo (sign_theory _ _ ?get_sign) ?ss_spec => let lemma2 := constr:(gen_lemma2_2 get_sign ss_spec) in let lemma1 := constr:(ring_correct set ext_r arth morph pp_spec dd_spec) in fun f => f arth ext_r morph lemma1 lemma2 | _ => fail "ring: bad sign specification" end | _ => fail "ring: bad coefficient division specification" end | _ => fail "ring: bad power specification" end | _ => fail "ring internal error: ring_lemmas, please report" end). (* Tactic for constant *) Ltac isnatcst t := match t with O => constr:(true) | S ?p => isnatcst p | _ => constr:(false) end. Ltac isPcst t := match t with | xI ?p => isPcst p | xO ?p => isPcst p | xH => constr:(true) (* nat -> positive *) | Pos.of_succ_nat ?n => isnatcst n | _ => constr:(false) end. Ltac isNcst t := match t with N0 => constr:(true) | Npos ?p => isPcst p | _ => constr:(false) end. Ltac isZcst t := match t with Z0 => constr:(true) | Zpos ?p => isPcst p | Zneg ?p => isPcst p (* injection nat -> Z *) | Z.of_nat ?n => isnatcst n (* injection N -> Z *) | Z.of_N ?n => isNcst n (* *) | _ => constr:(false) end. coq-8.11.0/plugins/setoid_ring/Rings_R.v0000644000175000017500000000310413612664533020005 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 0%R. discrR. Qed. Instance Rdi : (Integral_domain (Rcr:=Rcri)). constructor. exact Rmult_integral. exact R_one_zero. Defined. coq-8.11.0/plugins/setoid_ring/Ring_theory.v0000644000175000017500000004533313612664533020745 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* R -> R. Variable req : R -> R -> Prop. Variable Rsth : Equivalence req. Infix "*" := rmul. Infix "==" := req. Hypothesis mul_ext : Proper (req ==> req ==> req) rmul. Hypothesis mul_assoc : forall x y z, x * (y * z) == (x * y) * z. Fixpoint pow_pos (x:R) (i:positive) : R := match i with | xH => x | xO i => let p := pow_pos x i in p * p | xI i => let p := pow_pos x i in x * (p * p) end. Lemma pow_pos_swap x j : pow_pos x j * x == x * pow_pos x j. Proof. induction j; simpl; rewrite <- ?mul_assoc. - f_equiv. now do 2 (rewrite IHj, mul_assoc). - now do 2 (rewrite IHj, mul_assoc). - reflexivity. Qed. Lemma pow_pos_succ x j : pow_pos x (Pos.succ j) == x * pow_pos x j. Proof. induction j; simpl; try reflexivity. rewrite IHj, <- mul_assoc; f_equiv. now rewrite mul_assoc, pow_pos_swap, mul_assoc. Qed. Lemma pow_pos_add x i j : pow_pos x (i + j) == pow_pos x i * pow_pos x j. Proof. induction i using Pos.peano_ind. - now rewrite Pos.add_1_l, pow_pos_succ. - now rewrite Pos.add_succ_l, !pow_pos_succ, IHi, mul_assoc. Qed. Definition pow_N (x:R) (p:N) := match p with | N0 => rI | Npos p => pow_pos x p end. Definition id_phi_N (x:N) : N := x. Lemma pow_N_pow_N x n : pow_N x (id_phi_N n) == pow_N x n. Proof. reflexivity. Qed. End Power. Section DEFINITIONS. Variable R : Type. Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R). Variable req : R -> R -> Prop. Notation "0" := rO. Notation "1" := rI. Infix "==" := req. Infix "+" := radd. Infix "*" := rmul. Infix "-" := rsub. Notation "- x" := (ropp x). (** Semi Ring *) Record semi_ring_theory : Prop := mk_srt { SRadd_0_l : forall n, 0 + n == n; SRadd_comm : forall n m, n + m == m + n ; SRadd_assoc : forall n m p, n + (m + p) == (n + m) + p; SRmul_1_l : forall n, 1*n == n; SRmul_0_l : forall n, 0*n == 0; SRmul_comm : forall n m, n*m == m*n; SRmul_assoc : forall n m p, n*(m*p) == (n*m)*p; SRdistr_l : forall n m p, (n + m)*p == n*p + m*p }. (** Almost Ring *) (*Almost ring are no ring : Ropp_def is missing **) Record almost_ring_theory : Prop := mk_art { ARadd_0_l : forall x, 0 + x == x; ARadd_comm : forall x y, x + y == y + x; ARadd_assoc : forall x y z, x + (y + z) == (x + y) + z; ARmul_1_l : forall x, 1 * x == x; ARmul_0_l : forall x, 0 * x == 0; ARmul_comm : forall x y, x * y == y * x; ARmul_assoc : forall x y z, x * (y * z) == (x * y) * z; ARdistr_l : forall x y z, (x + y) * z == (x * z) + (y * z); ARopp_mul_l : forall x y, -(x * y) == -x * y; ARopp_add : forall x y, -(x + y) == -x + -y; ARsub_def : forall x y, x - y == x + -y }. (** Ring *) Record ring_theory : Prop := mk_rt { Radd_0_l : forall x, 0 + x == x; Radd_comm : forall x y, x + y == y + x; Radd_assoc : forall x y z, x + (y + z) == (x + y) + z; Rmul_1_l : forall x, 1 * x == x; Rmul_comm : forall x y, x * y == y * x; Rmul_assoc : forall x y z, x * (y * z) == (x * y) * z; Rdistr_l : forall x y z, (x + y) * z == (x * z) + (y * z); Rsub_def : forall x y, x - y == x + -y; Ropp_def : forall x, x + (- x) == 0 }. (** Equality is extensional *) Record sring_eq_ext : Prop := mk_seqe { (* SRing operators are compatible with equality *) SRadd_ext : Proper (req ==> req ==> req) radd; SRmul_ext : Proper (req ==> req ==> req) rmul }. Record ring_eq_ext : Prop := mk_reqe { (* Ring operators are compatible with equality *) Radd_ext : Proper (req ==> req ==> req) radd; Rmul_ext : Proper (req ==> req ==> req) rmul; Ropp_ext : Proper (req ==> req) ropp }. (** Interpretation morphisms definition*) Section MORPHISM. Variable C:Type. Variable (cO cI : C) (cadd cmul csub : C->C->C) (copp : C->C). Variable ceqb : C->C->bool. (* [phi] est un morphisme de [C] dans [R] *) Variable phi : C -> R. Infix "+!" := cadd. Infix "-!" := csub. Infix "*!" := cmul. Notation "-! x" := (copp x). Infix "?=!" := ceqb. Notation "[ x ]" := (phi x). (*for semi rings*) Record semi_morph : Prop := mkRmorph { Smorph0 : [cO] == 0; Smorph1 : [cI] == 1; Smorph_add : forall x y, [x +! y] == [x]+[y]; Smorph_mul : forall x y, [x *! y] == [x]*[y]; Smorph_eq : forall x y, x?=!y = true -> [x] == [y] }. (* for rings*) Record ring_morph : Prop := mkmorph { morph0 : [cO] == 0; morph1 : [cI] == 1; morph_add : forall x y, [x +! y] == [x]+[y]; morph_sub : forall x y, [x -! y] == [x]-[y]; morph_mul : forall x y, [x *! y] == [x]*[y]; morph_opp : forall x, [-!x] == -[x]; morph_eq : forall x y, x?=!y = true -> [x] == [y] }. Section SIGN. Variable get_sign : C -> option C. Record sign_theory : Prop := mksign_th { sign_spec : forall c c', get_sign c = Some c' -> c ?=! -! c' = true }. End SIGN. Definition get_sign_None (c:C) := @None C. Lemma get_sign_None_th : sign_theory get_sign_None. Proof. constructor;intros;discriminate. Qed. Section DIV. Variable cdiv: C -> C -> C*C. Record div_theory : Prop := mkdiv_th { div_eucl_th : forall a b, let (q,r) := cdiv a b in [a] == [b *! q +! r] }. End DIV. End MORPHISM. (** Identity is a morphism *) Variable Rsth : Equivalence req. Variable reqb : R->R->bool. Hypothesis morph_req : forall x y, (reqb x y) = true -> x == y. Definition IDphi (x:R) := x. Lemma IDmorph : ring_morph rO rI radd rmul rsub ropp reqb IDphi. Proof. now apply (mkmorph rO rI radd rmul rsub ropp reqb IDphi). Qed. (** Specification of the power function *) Section POWER. Variable Cpow : Type. Variable Cp_phi : N -> Cpow. Variable rpow : R -> Cpow -> R. Record power_theory : Prop := mkpow_th { rpow_pow_N : forall r n, req (rpow r (Cp_phi n)) (pow_N rI rmul r n) }. End POWER. Definition pow_N_th := mkpow_th id_phi_N (pow_N rI rmul) (pow_N_pow_N rI rmul Rsth). End DEFINITIONS. Section ALMOST_RING. Variable R : Type. Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R). Variable req : R -> R -> Prop. Notation "0" := rO. Notation "1" := rI. Infix "==" := req. Infix "+" := radd. Infix "* " := rmul. (** Leibniz equality leads to a setoid theory and is extensional*) Lemma Eqsth : Equivalence (@eq R). Proof. exact eq_equivalence. Qed. Lemma Eq_s_ext : sring_eq_ext radd rmul (@eq R). Proof. constructor;solve_proper. Qed. Lemma Eq_ext : ring_eq_ext radd rmul ropp (@eq R). Proof. constructor;solve_proper. Qed. Variable Rsth : Equivalence req. Section SEMI_RING. Variable SReqe : sring_eq_ext radd rmul req. Add Morphism radd with signature (req ==> req ==> req) as radd_ext1. Proof. exact (SRadd_ext SReqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext1. Proof. exact (SRmul_ext SReqe). Qed. Variable SRth : semi_ring_theory 0 1 radd rmul req. (** Every semi ring can be seen as an almost ring, by taking : [-x = x] and [x - y = x + y] *) Definition SRopp (x:R) := x. Notation "- x" := (SRopp x). Definition SRsub x y := x + -y. Infix "-" := SRsub. Lemma SRopp_ext : forall x y, x == y -> -x == -y. Proof. intros x y H; exact H. Qed. Lemma SReqe_Reqe : ring_eq_ext radd rmul SRopp req. Proof. constructor. - exact (SRadd_ext SReqe). - exact (SRmul_ext SReqe). - exact SRopp_ext. Qed. Lemma SRopp_mul_l : forall x y, -(x * y) == -x * y. Proof. reflexivity. Qed. Lemma SRopp_add : forall x y, -(x + y) == -x + -y. Proof. reflexivity. Qed. Lemma SRsub_def : forall x y, x - y == x + -y. Proof. reflexivity. Qed. Lemma SRth_ARth : almost_ring_theory 0 1 radd rmul SRsub SRopp req. Proof (mk_art 0 1 radd rmul SRsub SRopp req (SRadd_0_l SRth) (SRadd_comm SRth) (SRadd_assoc SRth) (SRmul_1_l SRth) (SRmul_0_l SRth) (SRmul_comm SRth) (SRmul_assoc SRth) (SRdistr_l SRth) SRopp_mul_l SRopp_add SRsub_def). (** Identity morphism for semi-ring equipped with their almost-ring structure*) Variable reqb : R->R->bool. Hypothesis morph_req : forall x y, (reqb x y) = true -> x == y. Definition SRIDmorph : ring_morph 0 1 radd rmul SRsub SRopp req 0 1 radd rmul SRsub SRopp reqb (@IDphi R). Proof. now apply mkmorph. Qed. (* a semi_morph can be extended to a ring_morph for the almost_ring derived from a semi_ring, provided the ring is a setoid (we only need reflexivity) *) Variable C : Type. Variable (cO cI : C) (cadd cmul: C->C->C). Variable (ceqb : C -> C -> bool). Variable phi : C -> R. Variable Smorph : semi_morph rO rI radd rmul req cO cI cadd cmul ceqb phi. Lemma SRmorph_Rmorph : ring_morph rO rI radd rmul SRsub SRopp req cO cI cadd cmul cadd (fun x => x) ceqb phi. Proof. case Smorph; now constructor. Qed. End SEMI_RING. Infix "-" := rsub. Notation "- x" := (ropp x). Variable Reqe : ring_eq_ext radd rmul ropp req. Add Morphism radd with signature (req ==> req ==> req) as radd_ext2. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext2. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext2. Proof. exact (Ropp_ext Reqe). Qed. Section RING. Variable Rth : ring_theory 0 1 radd rmul rsub ropp req. (** Rings are almost rings*) Lemma Rmul_0_l x : 0 * x == 0. Proof. setoid_replace (0*x) with ((0+1)*x + -x). now rewrite (Radd_0_l Rth), (Rmul_1_l Rth), (Ropp_def Rth). rewrite (Rdistr_l Rth), (Rmul_1_l Rth). rewrite <- (Radd_assoc Rth), (Ropp_def Rth). now rewrite (Radd_comm Rth), (Radd_0_l Rth). Qed. Lemma Ropp_mul_l x y : -(x * y) == -x * y. Proof. rewrite <-(Radd_0_l Rth (- x * y)). rewrite (Radd_comm Rth), <-(Ropp_def Rth (x*y)). rewrite (Radd_assoc Rth), <- (Rdistr_l Rth). rewrite (Radd_comm Rth (-x)), (Ropp_def Rth). now rewrite Rmul_0_l, (Radd_0_l Rth). Qed. Lemma Ropp_add x y : -(x + y) == -x + -y. Proof. rewrite <- ((Radd_0_l Rth) (-(x+y))). rewrite <- ((Ropp_def Rth) x). rewrite <- ((Radd_0_l Rth) (x + - x + - (x + y))). rewrite <- ((Ropp_def Rth) y). rewrite ((Radd_comm Rth) x). rewrite ((Radd_comm Rth) y). rewrite <- ((Radd_assoc Rth) (-y)). rewrite <- ((Radd_assoc Rth) (- x)). rewrite ((Radd_assoc Rth) y). rewrite ((Radd_comm Rth) y). rewrite <- ((Radd_assoc Rth) (- x)). rewrite ((Radd_assoc Rth) y). rewrite ((Radd_comm Rth) y), (Ropp_def Rth). rewrite ((Radd_comm Rth) (-x) 0), (Radd_0_l Rth). now apply (Radd_comm Rth). Qed. Lemma Ropp_opp x : - -x == x. Proof. rewrite <- (Radd_0_l Rth (- -x)). rewrite <- (Ropp_def Rth x). rewrite <- (Radd_assoc Rth), (Ropp_def Rth). rewrite ((Radd_comm Rth) x); now apply (Radd_0_l Rth). Qed. Lemma Rth_ARth : almost_ring_theory 0 1 radd rmul rsub ropp req. Proof (mk_art 0 1 radd rmul rsub ropp req (Radd_0_l Rth) (Radd_comm Rth) (Radd_assoc Rth) (Rmul_1_l Rth) Rmul_0_l (Rmul_comm Rth) (Rmul_assoc Rth) (Rdistr_l Rth) Ropp_mul_l Ropp_add (Rsub_def Rth)). (** Every semi morphism between two rings is a morphism*) Variable C : Type. Variable (cO cI : C) (cadd cmul csub: C->C->C) (copp : C -> C). Variable (ceq : C -> C -> Prop) (ceqb : C -> C -> bool). Variable phi : C -> R. Infix "+!" := cadd. Infix "*!" := cmul. Infix "-!" := csub. Notation "-! x" := (copp x). Notation "?=!" := ceqb. Notation "[ x ]" := (phi x). Variable Csth : Equivalence ceq. Variable Ceqe : ring_eq_ext cadd cmul copp ceq. Add Parametric Relation : C ceq reflexivity proved by (@Equivalence_Reflexive _ _ Csth) symmetry proved by (@Equivalence_Symmetric _ _ Csth) transitivity proved by (@Equivalence_Transitive _ _ Csth) as C_setoid. Add Morphism cadd with signature (ceq ==> ceq ==> ceq) as cadd_ext. Proof. exact (Radd_ext Ceqe). Qed. Add Morphism cmul with signature (ceq ==> ceq ==> ceq) as cmul_ext. Proof. exact (Rmul_ext Ceqe). Qed. Add Morphism copp with signature (ceq ==> ceq) as copp_ext. Proof. exact (Ropp_ext Ceqe). Qed. Variable Cth : ring_theory cO cI cadd cmul csub copp ceq. Variable Smorph : semi_morph 0 1 radd rmul req cO cI cadd cmul ceqb phi. Variable phi_ext : forall x y, ceq x y -> [x] == [y]. Add Morphism phi with signature (ceq ==> req) as phi_ext1. Proof. exact phi_ext. Qed. Lemma Smorph_opp x : [-!x] == -[x]. Proof. rewrite <- (Radd_0_l Rth [-!x]). rewrite <- ((Ropp_def Rth) [x]). rewrite ((Radd_comm Rth) [x]). rewrite <- (Radd_assoc Rth). rewrite <- (Smorph_add Smorph). rewrite (Ropp_def Cth). rewrite (Smorph0 Smorph). rewrite (Radd_comm Rth (-[x])). now apply (Radd_0_l Rth). Qed. Lemma Smorph_sub x y : [x -! y] == [x] - [y]. Proof. rewrite (Rsub_def Cth), (Rsub_def Rth). now rewrite (Smorph_add Smorph), Smorph_opp. Qed. Lemma Smorph_morph : ring_morph 0 1 radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi. Proof (mkmorph 0 1 radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi (Smorph0 Smorph) (Smorph1 Smorph) (Smorph_add Smorph) Smorph_sub (Smorph_mul Smorph) Smorph_opp (Smorph_eq Smorph)). End RING. (** Useful lemmas on almost ring *) Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req. Lemma ARth_SRth : semi_ring_theory 0 1 radd rmul req. Proof. elim ARth; intros. constructor; trivial. Qed. Instance ARsub_ext : Proper (req ==> req ==> req) rsub. Proof. intros x1 x2 Ex y1 y2 Ey. now rewrite !(ARsub_def ARth), Ex, Ey. Qed. Ltac mrewrite := repeat first [ rewrite (ARadd_0_l ARth) | rewrite <- ((ARadd_comm ARth) 0) | rewrite (ARmul_1_l ARth) | rewrite <- ((ARmul_comm ARth) 1) | rewrite (ARmul_0_l ARth) | rewrite <- ((ARmul_comm ARth) 0) | rewrite (ARdistr_l ARth) | reflexivity | match goal with | |- context [?z * (?x + ?y)] => rewrite ((ARmul_comm ARth) z (x+y)) end]. Lemma ARadd_0_r x : x + 0 == x. Proof. mrewrite. Qed. Lemma ARmul_1_r x : x * 1 == x. Proof. mrewrite. Qed. Lemma ARmul_0_r x : x * 0 == 0. Proof. mrewrite. Qed. Lemma ARdistr_r x y z : z * (x + y) == z*x + z*y. Proof. mrewrite. now rewrite !(ARmul_comm ARth z). Qed. Lemma ARadd_assoc1 x y z : (x + y) + z == (y + z) + x. Proof. now rewrite <-(ARadd_assoc ARth x), (ARadd_comm ARth x). Qed. Lemma ARadd_assoc2 x y z : (y + x) + z == (y + z) + x. Proof. now rewrite <- !(ARadd_assoc ARth), ((ARadd_comm ARth) x). Qed. Lemma ARmul_assoc1 x y z : (x * y) * z == (y * z) * x. Proof. now rewrite <- ((ARmul_assoc ARth) x), ((ARmul_comm ARth) x). Qed. Lemma ARmul_assoc2 x y z : (y * x) * z == (y * z) * x. Proof. now rewrite <- !(ARmul_assoc ARth), ((ARmul_comm ARth) x). Qed. Lemma ARopp_mul_r x y : - (x * y) == x * -y. Proof. rewrite ((ARmul_comm ARth) x y), (ARopp_mul_l ARth). now apply (ARmul_comm ARth). Qed. Lemma ARopp_zero : -0 == 0. Proof. now rewrite <- (ARmul_0_r 0), (ARopp_mul_l ARth), !ARmul_0_r. Qed. End ALMOST_RING. Section AddRing. (* Variable R : Type. Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R). Variable req : R -> R -> Prop. *) Inductive ring_kind : Type := | Abstract | Computational (R:Type) (req : R -> R -> Prop) (reqb : R -> R -> bool) (_ : forall x y, (reqb x y) = true -> req x y) | Morphism (R : Type) (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R) (req : R -> R -> Prop) (C : Type) (cO cI : C) (cadd cmul csub : C->C->C) (copp : C->C) (ceqb : C->C->bool) phi (_ : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi). End AddRing. (** Some simplification tactics*) Ltac gen_reflexivity Rsth := apply (Seq_refl _ _ Rsth). Ltac gen_srewrite Rsth Reqe ARth := repeat first [ gen_reflexivity Rsth | progress rewrite (ARopp_zero Rsth Reqe ARth) | rewrite (ARadd_0_l ARth) | rewrite (ARadd_0_r Rsth ARth) | rewrite (ARmul_1_l ARth) | rewrite (ARmul_1_r Rsth ARth) | rewrite (ARmul_0_l ARth) | rewrite (ARmul_0_r Rsth ARth) | rewrite (ARdistr_l ARth) | rewrite (ARdistr_r Rsth Reqe ARth) | rewrite (ARadd_assoc ARth) | rewrite (ARmul_assoc ARth) | progress rewrite (ARopp_add ARth) | progress rewrite (ARsub_def ARth) | progress rewrite <- (ARopp_mul_l ARth) | progress rewrite <- (ARopp_mul_r Rsth Reqe ARth) ]. Ltac gen_srewrite_sr Rsth Reqe ARth := repeat first [ gen_reflexivity Rsth | progress rewrite (ARopp_zero Rsth Reqe ARth) | rewrite (ARadd_0_l ARth) | rewrite (ARadd_0_r Rsth ARth) | rewrite (ARmul_1_l ARth) | rewrite (ARmul_1_r Rsth ARth) | rewrite (ARmul_0_l ARth) | rewrite (ARmul_0_r Rsth ARth) | rewrite (ARdistr_l ARth) | rewrite (ARdistr_r Rsth Reqe ARth) | rewrite (ARadd_assoc ARth) | rewrite (ARmul_assoc ARth) ]. Ltac gen_add_push add Rsth Reqe ARth x := repeat (match goal with | |- context [add (add ?y x) ?z] => progress rewrite (ARadd_assoc2 Rsth Reqe ARth x y z) | |- context [add (add x ?y) ?z] => progress rewrite (ARadd_assoc1 Rsth ARth x y z) | |- context [(add x ?y)] => progress rewrite (ARadd_comm ARth x y) end). Ltac gen_mul_push mul Rsth Reqe ARth x := repeat (match goal with | |- context [mul (mul ?y x) ?z] => progress rewrite (ARmul_assoc2 Rsth Reqe ARth x y z) | |- context [mul (mul x ?y) ?z] => progress rewrite (ARmul_assoc1 Rsth ARth x y z) | |- context [(mul x ?y)] => progress rewrite (ARmul_comm ARth x y) end). coq-8.11.0/plugins/setoid_ring/Ring_polynom.v0000644000175000017500000012533013612664533021124 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* R->R) (ropp : R->R). Variable req : R -> R -> Prop. (* Ring properties *) Variable Rsth : Equivalence req. Variable Reqe : ring_eq_ext radd rmul ropp req. Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req. (* Coefficients *) Variable C: Type. Variable (cO cI: C) (cadd cmul csub : C->C->C) (copp : C->C). Variable ceqb : C->C->bool. Variable phi : C -> R. Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi. (* Power coefficients *) Variable Cpow : Type. Variable Cp_phi : N -> Cpow. Variable rpow : R -> Cpow -> R. Variable pow_th : power_theory rI rmul req Cp_phi rpow. (* division is ok *) Variable cdiv: C -> C -> C * C. Variable div_th: div_theory req cadd cmul phi cdiv. (* R notations *) Notation "0" := rO. Notation "1" := rI. Infix "+" := radd. Infix "*" := rmul. Infix "-" := rsub. Notation "- x" := (ropp x). Infix "==" := req. Infix "^" := (pow_pos rmul). (* C notations *) Infix "+!" := cadd. Infix "*!" := cmul. Infix "-! " := csub. Notation "-! x" := (copp x). Infix "?=!" := ceqb. Notation "[ x ]" := (phi x). (* Useful tactics *) Add Morphism radd with signature (req ==> req ==> req) as radd_ext. Proof. exact (Radd_ext Reqe). Qed. Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext. Proof. exact (Rmul_ext Reqe). Qed. Add Morphism ropp with signature (req ==> req) as ropp_ext. Proof. exact (Ropp_ext Reqe). Qed. Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext. Proof. exact (ARsub_ext Rsth Reqe ARth). Qed. Ltac rsimpl := gen_srewrite Rsth Reqe ARth. Ltac add_push := gen_add_push radd Rsth Reqe ARth. Ltac mul_push := gen_mul_push rmul Rsth Reqe ARth. Ltac add_permut_rec t := match t with | ?x + ?y => add_permut_rec y || add_permut_rec x | _ => add_push t; apply (Radd_ext Reqe); [|reflexivity] end. Ltac add_permut := repeat (reflexivity || match goal with |- ?t == _ => add_permut_rec t end). Ltac mul_permut_rec t := match t with | ?x * ?y => mul_permut_rec y || mul_permut_rec x | _ => mul_push t; apply (Rmul_ext Reqe); [|reflexivity] end. Ltac mul_permut := repeat (reflexivity || match goal with |- ?t == _ => mul_permut_rec t end). (* Definition of multivariable polynomials with coefficients in C : Type [Pol] represents [X1 ... Xn]. The representation is Horner's where a [n] variable polynomial (C[X1..Xn]) is seen as a polynomial on [X1] which coefficients are polynomials with [n-1] variables (C[X2..Xn]). There are several optimisations to make the repr compacter: - [Pc c] is the constant polynomial of value c == c*X1^0*..*Xn^0 - [Pinj j Q] is a polynomial constant w.r.t the [j] first variables. variable indices are shifted of j in Q. == X1^0 *..* Xj^0 * Q{X1 <- Xj+1;..; Xn-j <- Xn} - [PX P i Q] is an optimised Horner form of P*X^i + Q with P not the null polynomial == P * X1^i + Q{X1 <- X2; ..; Xn-1 <- Xn} In addition: - polynomials of the form (PX (PX P i (Pc 0)) j Q) are forbidden since they can be represented by the simpler form (PX P (i+j) Q) - (Pinj i (Pinj j P)) is (Pinj (i+j) P) - (Pinj i (Pc c)) is (Pc c) *) Inductive Pol : Type := | Pc : C -> Pol | Pinj : positive -> Pol -> Pol | PX : Pol -> positive -> Pol -> Pol. Definition P0 := Pc cO. Definition P1 := Pc cI. Fixpoint Peq (P P' : Pol) {struct P'} : bool := match P, P' with | Pc c, Pc c' => c ?=! c' | Pinj j Q, Pinj j' Q' => match j ?= j' with | Eq => Peq Q Q' | _ => false end | PX P i Q, PX P' i' Q' => match i ?= i' with | Eq => if Peq P P' then Peq Q Q' else false | _ => false end | _, _ => false end. Infix "?==" := Peq. Definition mkPinj j P := match P with | Pc _ => P | Pinj j' Q => Pinj (j + j') Q | _ => Pinj j P end. Definition mkPinj_pred j P:= match j with | xH => P | xO j => Pinj (Pos.pred_double j) P | xI j => Pinj (xO j) P end. Definition mkPX P i Q := match P with | Pc c => if c ?=! cO then mkPinj xH Q else PX P i Q | Pinj _ _ => PX P i Q | PX P' i' Q' => if Q' ?== P0 then PX P' (i' + i) Q else PX P i Q end. Definition mkXi i := PX P1 i P0. Definition mkX := mkXi 1. (** Opposite of addition *) Fixpoint Popp (P:Pol) : Pol := match P with | Pc c => Pc (-! c) | Pinj j Q => Pinj j (Popp Q) | PX P i Q => PX (Popp P) i (Popp Q) end. Notation "-- P" := (Popp P). (** Addition et subtraction *) Fixpoint PaddC (P:Pol) (c:C) : Pol := match P with | Pc c1 => Pc (c1 +! c) | Pinj j Q => Pinj j (PaddC Q c) | PX P i Q => PX P i (PaddC Q c) end. Fixpoint PsubC (P:Pol) (c:C) : Pol := match P with | Pc c1 => Pc (c1 -! c) | Pinj j Q => Pinj j (PsubC Q c) | PX P i Q => PX P i (PsubC Q c) end. Section PopI. Variable Pop : Pol -> Pol -> Pol. Variable Q : Pol. Fixpoint PaddI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PaddC Q c) | Pinj j' Q' => match Z.pos_sub j' j with | Zpos k => mkPinj j (Pop (Pinj k Q') Q) | Z0 => mkPinj j (Pop Q' Q) | Zneg k => mkPinj j' (PaddI k Q') end | PX P i Q' => match j with | xH => PX P i (Pop Q' Q) | xO j => PX P i (PaddI (Pos.pred_double j) Q') | xI j => PX P i (PaddI (xO j) Q') end end. Fixpoint PsubI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PaddC (--Q) c) | Pinj j' Q' => match Z.pos_sub j' j with | Zpos k => mkPinj j (Pop (Pinj k Q') Q) | Z0 => mkPinj j (Pop Q' Q) | Zneg k => mkPinj j' (PsubI k Q') end | PX P i Q' => match j with | xH => PX P i (Pop Q' Q) | xO j => PX P i (PsubI (Pos.pred_double j) Q') | xI j => PX P i (PsubI (xO j) Q') end end. Variable P' : Pol. Fixpoint PaddX (i':positive) (P:Pol) : Pol := match P with | Pc c => PX P' i' P | Pinj j Q' => match j with | xH => PX P' i' Q' | xO j => PX P' i' (Pinj (Pos.pred_double j) Q') | xI j => PX P' i' (Pinj (xO j) Q') end | PX P i Q' => match Z.pos_sub i i' with | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' | Z0 => mkPX (Pop P P') i Q' | Zneg k => mkPX (PaddX k P) i Q' end end. Fixpoint PsubX (i':positive) (P:Pol) : Pol := match P with | Pc c => PX (--P') i' P | Pinj j Q' => match j with | xH => PX (--P') i' Q' | xO j => PX (--P') i' (Pinj (Pos.pred_double j) Q') | xI j => PX (--P') i' (Pinj (xO j) Q') end | PX P i Q' => match Z.pos_sub i i' with | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' | Z0 => mkPX (Pop P P') i Q' | Zneg k => mkPX (PsubX k P) i Q' end end. End PopI. Fixpoint Padd (P P': Pol) {struct P'} : Pol := match P' with | Pc c' => PaddC P c' | Pinj j' Q' => PaddI Padd Q' j' P | PX P' i' Q' => match P with | Pc c => PX P' i' (PaddC Q' c) | Pinj j Q => match j with | xH => PX P' i' (Padd Q Q') | xO j => PX P' i' (Padd (Pinj (Pos.pred_double j) Q) Q') | xI j => PX P' i' (Padd (Pinj (xO j) Q) Q') end | PX P i Q => match Z.pos_sub i i' with | Zpos k => mkPX (Padd (PX P k P0) P') i' (Padd Q Q') | Z0 => mkPX (Padd P P') i (Padd Q Q') | Zneg k => mkPX (PaddX Padd P' k P) i (Padd Q Q') end end end. Infix "++" := Padd. Fixpoint Psub (P P': Pol) {struct P'} : Pol := match P' with | Pc c' => PsubC P c' | Pinj j' Q' => PsubI Psub Q' j' P | PX P' i' Q' => match P with | Pc c => PX (--P') i' (*(--(PsubC Q' c))*) (PaddC (--Q') c) | Pinj j Q => match j with | xH => PX (--P') i' (Psub Q Q') | xO j => PX (--P') i' (Psub (Pinj (Pos.pred_double j) Q) Q') | xI j => PX (--P') i' (Psub (Pinj (xO j) Q) Q') end | PX P i Q => match Z.pos_sub i i' with | Zpos k => mkPX (Psub (PX P k P0) P') i' (Psub Q Q') | Z0 => mkPX (Psub P P') i (Psub Q Q') | Zneg k => mkPX (PsubX Psub P' k P) i (Psub Q Q') end end end. Infix "--" := Psub. (** Multiplication *) Fixpoint PmulC_aux (P:Pol) (c:C) : Pol := match P with | Pc c' => Pc (c' *! c) | Pinj j Q => mkPinj j (PmulC_aux Q c) | PX P i Q => mkPX (PmulC_aux P c) i (PmulC_aux Q c) end. Definition PmulC P c := if c ?=! cO then P0 else if c ?=! cI then P else PmulC_aux P c. Section PmulI. Variable Pmul : Pol -> Pol -> Pol. Variable Q : Pol. Fixpoint PmulI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PmulC Q c) | Pinj j' Q' => match Z.pos_sub j' j with | Zpos k => mkPinj j (Pmul (Pinj k Q') Q) | Z0 => mkPinj j (Pmul Q' Q) | Zneg k => mkPinj j' (PmulI k Q') end | PX P' i' Q' => match j with | xH => mkPX (PmulI xH P') i' (Pmul Q' Q) | xO j' => mkPX (PmulI j P') i' (PmulI (Pos.pred_double j') Q') | xI j' => mkPX (PmulI j P') i' (PmulI (xO j') Q') end end. End PmulI. Fixpoint Pmul (P P'' : Pol) {struct P''} : Pol := match P'' with | Pc c => PmulC P c | Pinj j' Q' => PmulI Pmul Q' j' P | PX P' i' Q' => match P with | Pc c => PmulC P'' c | Pinj j Q => let QQ' := match j with | xH => Pmul Q Q' | xO j => Pmul (Pinj (Pos.pred_double j) Q) Q' | xI j => Pmul (Pinj (xO j) Q) Q' end in mkPX (Pmul P P') i' QQ' | PX P i Q=> let QQ' := Pmul Q Q' in let PQ' := PmulI Pmul Q' xH P in let QP' := Pmul (mkPinj xH Q) P' in let PP' := Pmul P P' in (mkPX (mkPX PP' i P0 ++ QP') i' P0) ++ mkPX PQ' i QQ' end end. Infix "**" := Pmul. (** Monomial **) (** A monomial is X1^k1...Xi^ki. Its representation is a simplified version of the polynomial representation: - [mon0] correspond to the polynom [P1]. - [(zmon j M)] corresponds to [(Pinj j ...)], i.e. skip j variable indices. - [(vmon i M)] is X^i*M with X the current variable, its corresponds to (PX P1 i ...)] *) Inductive Mon: Set := | mon0: Mon | zmon: positive -> Mon -> Mon | vmon: positive -> Mon -> Mon. Definition mkZmon j M := match M with mon0 => mon0 | _ => zmon j M end. Definition zmon_pred j M := match j with xH => M | _ => mkZmon (Pos.pred j) M end. Definition mkVmon i M := match M with | mon0 => vmon i mon0 | zmon j m => vmon i (zmon_pred j m) | vmon i' m => vmon (i+i') m end. Fixpoint CFactor (P: Pol) (c: C) {struct P}: Pol * Pol := match P with | Pc c1 => let (q,r) := cdiv c1 c in (Pc r, Pc q) | Pinj j1 P1 => let (R,S) := CFactor P1 c in (mkPinj j1 R, mkPinj j1 S) | PX P1 i Q1 => let (R1, S1) := CFactor P1 c in let (R2, S2) := CFactor Q1 c in (mkPX R1 i R2, mkPX S1 i S2) end. Fixpoint MFactor (P: Pol) (c: C) (M: Mon) {struct P}: Pol * Pol := match P, M with _, mon0 => if (ceqb c cI) then (Pc cO, P) else CFactor P c | Pc _, _ => (P, Pc cO) | Pinj j1 P1, zmon j2 M1 => match j1 ?= j2 with Eq => let (R,S) := MFactor P1 c M1 in (mkPinj j1 R, mkPinj j1 S) | Lt => let (R,S) := MFactor P1 c (zmon (j2 - j1) M1) in (mkPinj j1 R, mkPinj j1 S) | Gt => (P, Pc cO) end | Pinj _ _, vmon _ _ => (P, Pc cO) | PX P1 i Q1, zmon j M1 => let M2 := zmon_pred j M1 in let (R1, S1) := MFactor P1 c M in let (R2, S2) := MFactor Q1 c M2 in (mkPX R1 i R2, mkPX S1 i S2) | PX P1 i Q1, vmon j M1 => match i ?= j with Eq => let (R1,S1) := MFactor P1 c (mkZmon xH M1) in (mkPX R1 i Q1, S1) | Lt => let (R1,S1) := MFactor P1 c (vmon (j - i) M1) in (mkPX R1 i Q1, S1) | Gt => let (R1,S1) := MFactor P1 c (mkZmon xH M1) in (mkPX R1 i Q1, mkPX S1 (i-j) (Pc cO)) end end. Definition POneSubst (P1: Pol) (cM1: C * Mon) (P2: Pol): option Pol := let (c,M1) := cM1 in let (Q1,R1) := MFactor P1 c M1 in match R1 with (Pc c) => if c ?=! cO then None else Some (Padd Q1 (Pmul P2 R1)) | _ => Some (Padd Q1 (Pmul P2 R1)) end. Fixpoint PNSubst1 (P1: Pol) (cM1: C * Mon) (P2: Pol) (n: nat) : Pol := match POneSubst P1 cM1 P2 with Some P3 => match n with S n1 => PNSubst1 P3 cM1 P2 n1 | _ => P3 end | _ => P1 end. Definition PNSubst (P1: Pol) (cM1: C * Mon) (P2: Pol) (n: nat): option Pol := match POneSubst P1 cM1 P2 with Some P3 => match n with S n1 => Some (PNSubst1 P3 cM1 P2 n1) | _ => None end | _ => None end. Fixpoint PSubstL1 (P1: Pol) (LM1: list ((C * Mon) * Pol)) (n: nat) : Pol := match LM1 with cons (M1,P2) LM2 => PSubstL1 (PNSubst1 P1 M1 P2 n) LM2 n | _ => P1 end. Fixpoint PSubstL (P1: Pol) (LM1: list ((C * Mon) * Pol)) (n: nat) : option Pol := match LM1 with cons (M1,P2) LM2 => match PNSubst P1 M1 P2 n with Some P3 => Some (PSubstL1 P3 LM2 n) | None => PSubstL P1 LM2 n end | _ => None end. Fixpoint PNSubstL (P1: Pol) (LM1: list ((C * Mon) * Pol)) (m n: nat) : Pol := match PSubstL P1 LM1 n with Some P3 => match m with S m1 => PNSubstL P3 LM1 m1 n | _ => P3 end | _ => P1 end. (** Evaluation of a polynomial towards R *) Local Notation hd := (List.hd 0). Fixpoint Pphi(l:list R) (P:Pol) : R := match P with | Pc c => [c] | Pinj j Q => Pphi (jump j l) Q | PX P i Q => Pphi l P * (hd l) ^ i + Pphi (tail l) Q end. Reserved Notation "P @ l " (at level 10, no associativity). Notation "P @ l " := (Pphi l P). Definition Pequiv (P Q : Pol) := forall l, P@l == Q@l. Infix "===" := Pequiv (at level 70, no associativity). Instance Pequiv_eq : Equivalence Pequiv. Proof. unfold Pequiv; split; red; intros; [reflexivity|now symmetry|now etransitivity]. Qed. Instance Pphi_ext : Proper (eq ==> Pequiv ==> req) Pphi. Proof. now intros l l' <- P Q H. Qed. Instance Pinj_ext : Proper (eq ==> Pequiv ==> Pequiv) Pinj. Proof. intros i j <- P P' HP l. simpl. now rewrite HP. Qed. Instance PX_ext : Proper (Pequiv ==> eq ==> Pequiv ==> Pequiv) PX. Proof. intros P P' HP p p' <- Q Q' HQ l. simpl. now rewrite HP, HQ. Qed. (** Evaluation of a monomial towards R *) Fixpoint Mphi(l:list R) (M: Mon) : R := match M with | mon0 => rI | zmon j M1 => Mphi (jump j l) M1 | vmon i M1 => Mphi (tail l) M1 * (hd l) ^ i end. Notation "M @@ l" := (Mphi l M) (at level 10, no associativity). (** Proofs *) Ltac destr_pos_sub := match goal with |- context [Z.pos_sub ?x ?y] => generalize (Z.pos_sub_discr x y); destruct (Z.pos_sub x y) end. Lemma jump_add' i j (l:list R) : jump (i + j) l = jump j (jump i l). Proof. rewrite Pos.add_comm. apply jump_add. Qed. Lemma Peq_ok P P' : (P ?== P') = true -> P === P'. Proof. unfold Pequiv. revert P';induction P;destruct P';simpl; intros H l; try easy. - now apply (morph_eq CRmorph). - destruct (Pos.compare_spec p p0); [ subst | easy | easy ]. now rewrite IHP. - specialize (IHP1 P'1); specialize (IHP2 P'2). destruct (Pos.compare_spec p p0); [ subst | easy | easy ]. destruct (P2 ?== P'1); [|easy]. rewrite H in *. now rewrite IHP1, IHP2. Qed. Lemma Peq_spec P P' : BoolSpec (P === P') True (P ?== P'). Proof. generalize (Peq_ok P P'). destruct (P ?== P'); auto. Qed. Lemma Pphi0 l : P0@l == 0. Proof. simpl;apply (morph0 CRmorph). Qed. Lemma Pphi1 l : P1@l == 1. Proof. simpl;apply (morph1 CRmorph). Qed. Lemma mkPinj_ok j l P : (mkPinj j P)@l == P@(jump j l). Proof. destruct P;simpl;rsimpl. now rewrite jump_add'. Qed. Instance mkPinj_ext : Proper (eq ==> Pequiv ==> Pequiv) mkPinj. Proof. intros i j <- P Q H l. now rewrite !mkPinj_ok. Qed. Lemma pow_pos_add x i j : x^(j + i) == x^i * x^j. Proof. rewrite Pos.add_comm. apply (pow_pos_add Rsth (Rmul_ext Reqe) (ARmul_assoc ARth)). Qed. Lemma ceqb_spec c c' : BoolSpec ([c] == [c']) True (c ?=! c'). Proof. generalize (morph_eq CRmorph c c'). destruct (c ?=! c'); auto. Qed. Lemma mkPX_ok l P i Q : (mkPX P i Q)@l == P@l * (hd l)^i + Q@(tail l). Proof. unfold mkPX. destruct P. - case ceqb_spec; intros H; simpl; try reflexivity. rewrite H, (morph0 CRmorph), mkPinj_ok; rsimpl. - reflexivity. - case Peq_spec; intros H; simpl; try reflexivity. rewrite H, Pphi0, Pos.add_comm, pow_pos_add; rsimpl. Qed. Instance mkPX_ext : Proper (Pequiv ==> eq ==> Pequiv ==> Pequiv) mkPX. Proof. intros P P' HP i i' <- Q Q' HQ l. now rewrite !mkPX_ok, HP, HQ. Qed. Hint Rewrite Pphi0 Pphi1 mkPinj_ok mkPX_ok (morph0 CRmorph) (morph1 CRmorph) (morph0 CRmorph) (morph_add CRmorph) (morph_mul CRmorph) (morph_sub CRmorph) (morph_opp CRmorph) : Esimpl. (* Quicker than autorewrite with Esimpl :-) *) Ltac Esimpl := try rewrite_db Esimpl; rsimpl; simpl. Lemma PaddC_ok c P l : (PaddC P c)@l == P@l + [c]. Proof. revert l;induction P;simpl;intros;Esimpl;trivial. rewrite IHP2;rsimpl. Qed. Lemma PsubC_ok c P l : (PsubC P c)@l == P@l - [c]. Proof. revert l;induction P;simpl;intros. - Esimpl. - rewrite IHP;rsimpl. - rewrite IHP2;rsimpl. Qed. Lemma PmulC_aux_ok c P l : (PmulC_aux P c)@l == P@l * [c]. Proof. revert l;induction P;simpl;intros;Esimpl;trivial. rewrite IHP1, IHP2;rsimpl. add_permut. mul_permut. Qed. Lemma PmulC_ok c P l : (PmulC P c)@l == P@l * [c]. Proof. unfold PmulC. case ceqb_spec; intros H. - rewrite H; Esimpl. - case ceqb_spec; intros H'. + rewrite H'; Esimpl. + apply PmulC_aux_ok. Qed. Lemma Popp_ok P l : (--P)@l == - P@l. Proof. revert l;induction P;simpl;intros. - Esimpl. - apply IHP. - rewrite IHP1, IHP2;rsimpl. Qed. Hint Rewrite PaddC_ok PsubC_ok PmulC_ok Popp_ok : Esimpl. Lemma PaddX_ok P' P k l : (forall P l, (P++P')@l == P@l + P'@l) -> (PaddX Padd P' k P) @ l == P@l + P'@l * (hd l)^k. Proof. intros IHP'. revert k l. induction P;simpl;intros. - add_permut. - destruct p; simpl; rewrite ?jump_pred_double; add_permut. - destr_pos_sub; intros ->; Esimpl. + rewrite IHP';rsimpl. add_permut. + rewrite IHP', pow_pos_add;simpl;Esimpl. add_permut. + rewrite IHP1, pow_pos_add;rsimpl. add_permut. Qed. Lemma Padd_ok P' P l : (P ++ P')@l == P@l + P'@l. Proof. revert P l; induction P';simpl;intros;Esimpl. - revert p l; induction P;simpl;intros. + Esimpl; add_permut. + destr_pos_sub; intros ->;Esimpl. * now rewrite IHP'. * rewrite IHP';Esimpl. now rewrite jump_add'. * rewrite IHP. now rewrite jump_add'. + destruct p0;simpl. * rewrite IHP2;simpl. rsimpl. * rewrite IHP2;simpl. rewrite jump_pred_double. rsimpl. * rewrite IHP'. rsimpl. - destruct P;simpl. + Esimpl. add_permut. + destruct p0;simpl;Esimpl; rewrite IHP'2; simpl. * rsimpl. add_permut. * rewrite jump_pred_double. rsimpl. add_permut. * rsimpl. add_permut. + destr_pos_sub; intros ->; Esimpl. * rewrite IHP'1, IHP'2;rsimpl. add_permut. * rewrite IHP'1, IHP'2;simpl;Esimpl. rewrite pow_pos_add;rsimpl. add_permut. * rewrite PaddX_ok by trivial; rsimpl. rewrite IHP'2, pow_pos_add; rsimpl. add_permut. Qed. Lemma Psub_opp P' P : P -- P' === P ++ (--P'). Proof. revert P; induction P'; simpl; intros. - intro l; Esimpl. - revert p; induction P; simpl; intros; try reflexivity. + destr_pos_sub; intros ->; now apply mkPinj_ext. + destruct p0; now apply PX_ext. - destruct P; simpl; try reflexivity. + destruct p0; now apply PX_ext. + destr_pos_sub; intros ->; apply mkPX_ext; auto. revert p1. induction P2; simpl; intros; try reflexivity. destr_pos_sub; intros ->; now apply mkPX_ext. Qed. Lemma Psub_ok P' P l : (P -- P')@l == P@l - P'@l. Proof. rewrite Psub_opp, Padd_ok, Popp_ok. rsimpl. Qed. Lemma PmulI_ok P' : (forall P l, (Pmul P P') @ l == P @ l * P' @ l) -> forall P p l, (PmulI Pmul P' p P) @ l == P @ l * P' @ (jump p l). Proof. intros IHP'. induction P;simpl;intros. - Esimpl; mul_permut. - destr_pos_sub; intros ->;Esimpl. + now rewrite IHP'. + now rewrite IHP', jump_add'. + now rewrite IHP, jump_add'. - destruct p0;Esimpl; rewrite ?IHP1, ?IHP2; rsimpl. + f_equiv. mul_permut. + rewrite jump_pred_double. f_equiv. mul_permut. + rewrite IHP'. f_equiv. mul_permut. Qed. Lemma Pmul_ok P P' l : (P**P')@l == P@l * P'@l. Proof. revert P l;induction P';simpl;intros. - apply PmulC_ok. - apply PmulI_ok;trivial. - destruct P. + rewrite (ARmul_comm ARth). Esimpl. + Esimpl. f_equiv. rewrite IHP'1; Esimpl. destruct p0;rewrite IHP'2;Esimpl. rewrite jump_pred_double; Esimpl. + rewrite Padd_ok, !mkPX_ok, Padd_ok, !mkPX_ok, !IHP'1, !IHP'2, PmulI_ok; trivial. simpl. Esimpl. add_permut; f_equiv; mul_permut. Qed. Lemma mkZmon_ok M j l : (mkZmon j M) @@ l == (zmon j M) @@ l. Proof. destruct M; simpl; rsimpl. Qed. Lemma zmon_pred_ok M j l : (zmon_pred j M) @@ (tail l) == (zmon j M) @@ l. Proof. destruct j; simpl; rewrite ?mkZmon_ok; simpl; rsimpl. rewrite jump_pred_double; rsimpl. Qed. Lemma mkVmon_ok M i l : (mkVmon i M)@@l == M@@l * (hd l)^i. Proof. destruct M;simpl;intros;rsimpl. - rewrite zmon_pred_ok;simpl;rsimpl. - rewrite pow_pos_add;rsimpl. Qed. Ltac destr_factor := match goal with | H : context [CFactor ?P _] |- context [CFactor ?P ?c] => destruct (CFactor P c); destr_factor; rewrite H; clear H | H : context [MFactor ?P _ _] |- context [MFactor ?P ?c ?M] => specialize (H M); destruct (MFactor P c M); destr_factor; rewrite H; clear H | _ => idtac end. Lemma Mcphi_ok P c l : let (Q,R) := CFactor P c in P@l == Q@l + [c] * R@l. Proof. revert l. induction P as [c0 | j P IH | P1 IH1 i P2 IH2]; intros l; Esimpl. - assert (H := (div_eucl_th div_th) c0 c). destruct cdiv as (q,r). rewrite H; Esimpl. add_permut. - destr_factor. Esimpl. - destr_factor. Esimpl. add_permut. Qed. Lemma Mphi_ok P (cM: C * Mon) l : let (c,M) := cM in let (Q,R) := MFactor P c M in P@l == Q@l + [c] * M@@l * R@l. Proof. destruct cM as (c,M). revert M l. induction P; destruct M; intros l; simpl; auto; try (case ceqb_spec; intro He); try (case Pos.compare_spec; intros He); rewrite ?He; destr_factor; simpl; Esimpl. - assert (H := div_eucl_th div_th c0 c). destruct cdiv as (q,r). rewrite H; Esimpl. add_permut. - assert (H := Mcphi_ok P c). destr_factor. Esimpl. - now rewrite <- jump_add, Pos.sub_add. - assert (H2 := Mcphi_ok P2 c). assert (H3 := Mcphi_ok P3 c). destr_factor. Esimpl. add_permut. - rewrite zmon_pred_ok. simpl. add_permut. - rewrite mkZmon_ok. simpl. add_permut. mul_permut. - add_permut. mul_permut. rewrite <- pow_pos_add, Pos.add_comm, Pos.sub_add by trivial; rsimpl. - rewrite mkZmon_ok. simpl. Esimpl. add_permut. mul_permut. rewrite <- pow_pos_add, Pos.sub_add by trivial; rsimpl. Qed. Lemma POneSubst_ok P1 cM1 P2 P3 l : POneSubst P1 cM1 P2 = Some P3 -> [fst cM1] * (snd cM1)@@l == P2@l -> P1@l == P3@l. Proof. destruct cM1 as (cc,M1). unfold POneSubst. assert (H := Mphi_ok P1 (cc, M1) l). simpl in H. destruct MFactor as (R1,S1); simpl. rewrite H. clear H. intros EQ EQ'. replace P3 with (R1 ++ P2 ** S1). - rewrite EQ', Padd_ok, Pmul_ok; rsimpl. - revert EQ. destruct S1; try now injection 1. case ceqb_spec; now inversion 2. Qed. Lemma PNSubst1_ok n P1 cM1 P2 l : [fst cM1] * (snd cM1)@@l == P2@l -> P1@l == (PNSubst1 P1 cM1 P2 n)@l. Proof. revert P1. induction n; simpl; intros P1; generalize (POneSubst_ok P1 cM1 P2); destruct POneSubst; intros; rewrite <- ?IHn; auto; reflexivity. Qed. Lemma PNSubst_ok n P1 cM1 P2 l P3 : PNSubst P1 cM1 P2 n = Some P3 -> [fst cM1] * (snd cM1)@@l == P2@l -> P1@l == P3@l. Proof. unfold PNSubst. assert (H := POneSubst_ok P1 cM1 P2); destruct POneSubst; try discriminate. destruct n; inversion_clear 1. intros. rewrite <- PNSubst1_ok; auto. Qed. Fixpoint MPcond (LM1: list (C * Mon * Pol)) (l: list R) : Prop := match LM1 with | (M1,P2) :: LM2 => ([fst M1] * (snd M1)@@l == P2@l) /\ MPcond LM2 l | _ => True end. Lemma PSubstL1_ok n LM1 P1 l : MPcond LM1 l -> P1@l == (PSubstL1 P1 LM1 n)@l. Proof. revert P1; induction LM1 as [|(M2,P2) LM2 IH]; simpl; intros. - reflexivity. - rewrite <- IH by intuition; now apply PNSubst1_ok. Qed. Lemma PSubstL_ok n LM1 P1 P2 l : PSubstL P1 LM1 n = Some P2 -> MPcond LM1 l -> P1@l == P2@l. Proof. revert P1. induction LM1 as [|(M2,P2') LM2 IH]; simpl; intros. - discriminate. - assert (H':=PNSubst_ok n P3 M2 P2'). destruct PNSubst. * injection H as [= <-]. rewrite <- PSubstL1_ok; intuition. * now apply IH. Qed. Lemma PNSubstL_ok m n LM1 P1 l : MPcond LM1 l -> P1@l == (PNSubstL P1 LM1 m n)@l. Proof. revert LM1 P1. induction m; simpl; intros; assert (H' := PSubstL_ok n LM1 P2); destruct PSubstL; auto; try reflexivity. rewrite <- IHm; auto. Qed. (** Definition of polynomial expressions *) Inductive PExpr : Type := | PEO : PExpr | PEI : PExpr | PEc : C -> PExpr | PEX : positive -> PExpr | PEadd : PExpr -> PExpr -> PExpr | PEsub : PExpr -> PExpr -> PExpr | PEmul : PExpr -> PExpr -> PExpr | PEopp : PExpr -> PExpr | PEpow : PExpr -> N -> PExpr. Register PExpr as plugins.setoid_ring.pexpr. Register PEc as plugins.setoid_ring.const. Register PEX as plugins.setoid_ring.var. Register PEadd as plugins.setoid_ring.add. Register PEsub as plugins.setoid_ring.sub. Register PEmul as plugins.setoid_ring.mul. Register PEopp as plugins.setoid_ring.opp. Register PEpow as plugins.setoid_ring.pow. (** evaluation of polynomial expressions towards R *) Definition mk_X j := mkPinj_pred j mkX. (** evaluation of polynomial expressions towards R *) Fixpoint PEeval (l:list R) (pe:PExpr) {struct pe} : R := match pe with | PEO => rO | PEI => rI | PEc c => phi c | PEX j => nth 0 j l | PEadd pe1 pe2 => (PEeval l pe1) + (PEeval l pe2) | PEsub pe1 pe2 => (PEeval l pe1) - (PEeval l pe2) | PEmul pe1 pe2 => (PEeval l pe1) * (PEeval l pe2) | PEopp pe1 => - (PEeval l pe1) | PEpow pe1 n => rpow (PEeval l pe1) (Cp_phi n) end. Strategy expand [PEeval]. (** Correctness proofs *) Lemma mkX_ok p l : nth 0 p l == (mk_X p) @ l. Proof. destruct p;simpl;intros;Esimpl;trivial. - now rewrite <-jump_tl, nth_jump. - now rewrite <- nth_jump, nth_pred_double. Qed. Hint Rewrite Padd_ok Psub_ok : Esimpl. Section POWER. Variable subst_l : Pol -> Pol. Fixpoint Ppow_pos (res P:Pol) (p:positive) : Pol := match p with | xH => subst_l (res ** P) | xO p => Ppow_pos (Ppow_pos res P p) P p | xI p => subst_l ((Ppow_pos (Ppow_pos res P p) P p) ** P) end. Definition Ppow_N P n := match n with | N0 => P1 | Npos p => Ppow_pos P1 P p end. Lemma Ppow_pos_ok l : (forall P, subst_l P@l == P@l) -> forall res P p, (Ppow_pos res P p)@l == res@l * (pow_pos Pmul P p)@l. Proof. intros subst_l_ok res P p. revert res. induction p;simpl;intros; rewrite ?subst_l_ok, ?Pmul_ok, ?IHp; mul_permut. Qed. Lemma Ppow_N_ok l : (forall P, subst_l P@l == P@l) -> forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l. Proof. destruct n;simpl. - reflexivity. - rewrite Ppow_pos_ok by trivial. Esimpl. Qed. End POWER. (** Normalization and rewriting *) Section NORM_SUBST_REC. Variable n : nat. Variable lmp:list (C*Mon*Pol). Let subst_l P := PNSubstL P lmp n n. Let Pmul_subst P1 P2 := subst_l (P1 ** P2). Let Ppow_subst := Ppow_N subst_l. Fixpoint norm_aux (pe:PExpr) : Pol := match pe with | PEO => Pc cO | PEI => Pc cI | PEc c => Pc c | PEX j => mk_X j | PEadd (PEopp pe1) pe2 => (norm_aux pe2) -- (norm_aux pe1) | PEadd pe1 (PEopp pe2) => (norm_aux pe1) -- (norm_aux pe2) | PEadd pe1 pe2 => (norm_aux pe1) ++ (norm_aux pe2) | PEsub pe1 pe2 => (norm_aux pe1) -- (norm_aux pe2) | PEmul pe1 pe2 => (norm_aux pe1) ** (norm_aux pe2) | PEopp pe1 => -- (norm_aux pe1) | PEpow pe1 n => Ppow_N (fun p => p) (norm_aux pe1) n end. Definition norm_subst pe := subst_l (norm_aux pe). (** Internally, [norm_aux] is expanded in a large number of cases. To speed-up proofs, we use an alternative definition. *) Definition get_PEopp pe := match pe with | PEopp pe' => Some pe' | _ => None end. Lemma norm_aux_PEadd pe1 pe2 : norm_aux (PEadd pe1 pe2) = match get_PEopp pe1, get_PEopp pe2 with | Some pe1', _ => (norm_aux pe2) -- (norm_aux pe1') | None, Some pe2' => (norm_aux pe1) -- (norm_aux pe2') | None, None => (norm_aux pe1) ++ (norm_aux pe2) end. Proof. simpl (norm_aux (PEadd _ _)). destruct pe1; [ | | | | | | | reflexivity | ]; destruct pe2; simpl get_PEopp; reflexivity. Qed. Lemma norm_aux_PEopp pe : match get_PEopp pe with | Some pe' => norm_aux pe = -- (norm_aux pe') | None => True end. Proof. now destruct pe. Qed. Arguments norm_aux !pe : simpl nomatch. Lemma norm_aux_spec l pe : PEeval l pe == (norm_aux pe)@l. Proof. intros. induction pe; cbn. - now rewrite (morph0 CRmorph). - now rewrite (morph1 CRmorph). - reflexivity. - apply mkX_ok. - rewrite IHpe1, IHpe2. assert (H1 := norm_aux_PEopp pe1). assert (H2 := norm_aux_PEopp pe2). rewrite norm_aux_PEadd. do 2 destruct get_PEopp; rewrite ?H1, ?H2; Esimpl; add_permut. - rewrite IHpe1, IHpe2. Esimpl. - rewrite IHpe1, IHpe2. now rewrite Pmul_ok. - rewrite IHpe. Esimpl. - rewrite Ppow_N_ok by reflexivity. rewrite (rpow_pow_N pow_th). destruct n0; simpl; Esimpl. induction p;simpl; now rewrite ?IHp, ?IHpe, ?Pms_ok, ?Pmul_ok. Qed. Lemma norm_subst_spec : forall l pe, MPcond lmp l -> PEeval l pe == (norm_subst pe)@l. Proof. intros;unfold norm_subst. unfold subst_l;rewrite <- PNSubstL_ok;trivial. apply norm_aux_spec. Qed. End NORM_SUBST_REC. Fixpoint interp_PElist (l:list R) (lpe:list (PExpr*PExpr)) {struct lpe} : Prop := match lpe with | nil => True | (me,pe)::lpe => match lpe with | nil => PEeval l me == PEeval l pe | _ => PEeval l me == PEeval l pe /\ interp_PElist l lpe end end. Fixpoint mon_of_pol (P:Pol) : option (C * Mon) := match P with | Pc c => if (c ?=! cO) then None else Some (c, mon0) | Pinj j P => match mon_of_pol P with | None => None | Some (c,m) => Some (c, mkZmon j m) end | PX P i Q => if Peq Q P0 then match mon_of_pol P with | None => None | Some (c,m) => Some (c, mkVmon i m) end else None end. Fixpoint mk_monpol_list (lpe:list (PExpr * PExpr)) : list (C*Mon*Pol) := match lpe with | nil => nil | (me,pe)::lpe => match mon_of_pol (norm_subst 0 nil me) with | None => mk_monpol_list lpe | Some m => (m,norm_subst 0 nil pe):: mk_monpol_list lpe end end. Lemma mon_of_pol_ok : forall P m, mon_of_pol P = Some m -> forall l, [fst m] * Mphi l (snd m) == P@l. Proof. induction P;simpl;intros;Esimpl. assert (H1 := (morph_eq CRmorph) c cO). destruct (c ?=! cO). discriminate. inversion H;trivial;Esimpl. generalize H;clear H;case_eq (mon_of_pol P). intros (c1,P2) H0 H1; inversion H1; Esimpl. generalize (IHP (c1, P2) H0 (jump p l)). rewrite mkZmon_ok;simpl;auto. intros; discriminate. generalize H;clear H;change match P3 with | Pc c => c ?=! cO | Pinj _ _ => false | PX _ _ _ => false end with (P3 ?== P0). assert (H := Peq_ok P3 P0). destruct (P3 ?== P0). case_eq (mon_of_pol P2);try intros (cc, pp); intros. inversion H1. simpl. rewrite mkVmon_ok;simpl. rewrite H;trivial;Esimpl. generalize (IHP1 _ H0); simpl; intros HH; rewrite HH; rsimpl. discriminate. intros;discriminate. Qed. Lemma interp_PElist_ok : forall l lpe, interp_PElist l lpe -> MPcond (mk_monpol_list lpe) l. Proof. induction lpe;simpl. trivial. destruct a;simpl;intros. assert (HH:=mon_of_pol_ok (norm_subst 0 nil p)); destruct (mon_of_pol (norm_subst 0 nil p)). split. rewrite <- norm_subst_spec by exact I. destruct lpe;try destruct H;rewrite <- H; rewrite (norm_subst_spec 0 nil); try exact I;apply HH;trivial. apply IHlpe. destruct lpe;simpl;trivial. destruct H. exact H0. apply IHlpe. destruct lpe;simpl;trivial. destruct H. exact H0. Qed. Lemma norm_subst_ok : forall n l lpe pe, interp_PElist l lpe -> PEeval l pe == (norm_subst n (mk_monpol_list lpe) pe)@l. Proof. intros;apply norm_subst_spec. apply interp_PElist_ok;trivial. Qed. Lemma ring_correct : forall n l lpe pe1 pe2, interp_PElist l lpe -> (let lmp := mk_monpol_list lpe in norm_subst n lmp pe1 ?== norm_subst n lmp pe2) = true -> PEeval l pe1 == PEeval l pe2. Proof. simpl;intros. do 2 (rewrite (norm_subst_ok n l lpe);trivial). apply Peq_ok;trivial. Qed. (** Generic evaluation of polynomial towards R avoiding parenthesis *) Variable get_sign : C -> option C. Variable get_sign_spec : sign_theory copp ceqb get_sign. Section EVALUATION. (* [mkpow x p] = x^p *) Variable mkpow : R -> positive -> R. (* [mkpow x p] = -(x^p) *) Variable mkopp_pow : R -> positive -> R. (* [mkmult_pow r x p] = r * x^p *) Variable mkmult_pow : R -> R -> positive -> R. Fixpoint mkmult_rec (r:R) (lm:list (R*positive)) {struct lm}: R := match lm with | nil => r | cons (x,p) t => mkmult_rec (mkmult_pow r x p) t end. Definition mkmult1 lm := match lm with | nil => 1 | cons (x,p) t => mkmult_rec (mkpow x p) t end. Definition mkmultm1 lm := match lm with | nil => ropp rI | cons (x,p) t => mkmult_rec (mkopp_pow x p) t end. Definition mkmult_c_pos c lm := if c ?=! cI then mkmult1 (rev' lm) else mkmult_rec [c] (rev' lm). Definition mkmult_c c lm := match get_sign c with | None => mkmult_c_pos c lm | Some c' => if c' ?=! cI then mkmultm1 (rev' lm) else mkmult_rec [c] (rev' lm) end. Definition mkadd_mult rP c lm := match get_sign c with | None => rP + mkmult_c_pos c lm | Some c' => rP - mkmult_c_pos c' lm end. Definition add_pow_list (r:R) n l := match n with | N0 => l | Npos p => (r,p)::l end. Fixpoint add_mult_dev (rP:R) (P:Pol) (fv:list R) (n:N) (lm:list (R*positive)) {struct P} : R := match P with | Pc c => let lm := add_pow_list (hd fv) n lm in mkadd_mult rP c lm | Pinj j Q => add_mult_dev rP Q (jump j fv) N0 (add_pow_list (hd fv) n lm) | PX P i Q => let rP := add_mult_dev rP P fv (N.add (Npos i) n) lm in if Q ?== P0 then rP else add_mult_dev rP Q (tail fv) N0 (add_pow_list (hd fv) n lm) end. Fixpoint mult_dev (P:Pol) (fv : list R) (n:N) (lm:list (R*positive)) {struct P} : R := (* P@l * (hd 0 l)^n * lm *) match P with | Pc c => mkmult_c c (add_pow_list (hd fv) n lm) | Pinj j Q => mult_dev Q (jump j fv) N0 (add_pow_list (hd fv) n lm) | PX P i Q => let rP := mult_dev P fv (N.add (Npos i) n) lm in if Q ?== P0 then rP else let lmq := add_pow_list (hd fv) n lm in add_mult_dev rP Q (tail fv) N0 lmq end. Definition Pphi_avoid fv P := mult_dev P fv N0 nil. Fixpoint r_list_pow (l:list (R*positive)) : R := match l with | nil => rI | cons (r,p) l => pow_pos rmul r p * r_list_pow l end. Hypothesis mkpow_spec : forall r p, mkpow r p == pow_pos rmul r p. Hypothesis mkopp_pow_spec : forall r p, mkopp_pow r p == - (pow_pos rmul r p). Hypothesis mkmult_pow_spec : forall r x p, mkmult_pow r x p == r * pow_pos rmul x p. Lemma mkmult_rec_ok : forall lm r, mkmult_rec r lm == r * r_list_pow lm. Proof. induction lm;intros;simpl;Esimpl. destruct a as (x,p);Esimpl. rewrite IHlm. rewrite mkmult_pow_spec. Esimpl. Qed. Lemma mkmult1_ok : forall lm, mkmult1 lm == r_list_pow lm. Proof. destruct lm;simpl;Esimpl. destruct p. rewrite mkmult_rec_ok;rewrite mkpow_spec;Esimpl. Qed. Lemma mkmultm1_ok : forall lm, mkmultm1 lm == - r_list_pow lm. Proof. destruct lm;simpl;Esimpl. destruct p;rewrite mkmult_rec_ok. rewrite mkopp_pow_spec;Esimpl. Qed. Lemma r_list_pow_rev : forall l, r_list_pow (rev' l) == r_list_pow l. Proof. assert (forall l lr : list (R * positive), r_list_pow (rev_append l lr) == r_list_pow lr * r_list_pow l). induction l;intros;simpl;Esimpl. destruct a;rewrite IHl;Esimpl. rewrite (ARmul_comm ARth (pow_pos rmul r p)). reflexivity. intros;unfold rev'. rewrite H;simpl;Esimpl. Qed. Lemma mkmult_c_pos_ok : forall c lm, mkmult_c_pos c lm == [c]* r_list_pow lm. Proof. intros;unfold mkmult_c_pos;simpl. assert (H := (morph_eq CRmorph) c cI). rewrite <- r_list_pow_rev; destruct (c ?=! cI). rewrite H;trivial;Esimpl. apply mkmult1_ok. apply mkmult_rec_ok. Qed. Lemma mkmult_c_ok : forall c lm, mkmult_c c lm == [c] * r_list_pow lm. Proof. intros;unfold mkmult_c;simpl. case_eq (get_sign c);intros. assert (H1 := (morph_eq CRmorph) c0 cI). destruct (c0 ?=! cI). rewrite (morph_eq CRmorph _ _ (sign_spec get_sign_spec _ H)). Esimpl. rewrite H1;trivial. rewrite <- r_list_pow_rev;trivial;Esimpl. apply mkmultm1_ok. rewrite <- r_list_pow_rev; apply mkmult_rec_ok. apply mkmult_c_pos_ok. Qed. Lemma mkadd_mult_ok : forall rP c lm, mkadd_mult rP c lm == rP + [c]*r_list_pow lm. Proof. intros;unfold mkadd_mult. case_eq (get_sign c);intros. rewrite (morph_eq CRmorph _ _ (sign_spec get_sign_spec _ H));Esimpl. rewrite mkmult_c_pos_ok;Esimpl. rewrite mkmult_c_pos_ok;Esimpl. Qed. Lemma add_pow_list_ok : forall r n l, r_list_pow (add_pow_list r n l) == pow_N rI rmul r n * r_list_pow l. Proof. destruct n;simpl;intros;Esimpl. Qed. Lemma add_mult_dev_ok : forall P rP fv n lm, add_mult_dev rP P fv n lm == rP + P@fv*pow_N rI rmul (hd fv) n * r_list_pow lm. Proof. induction P;simpl;intros. rewrite mkadd_mult_ok. rewrite add_pow_list_ok; Esimpl. rewrite IHP. simpl. rewrite add_pow_list_ok; Esimpl. change (match P3 with | Pc c => c ?=! cO | Pinj _ _ => false | PX _ _ _ => false end) with (Peq P3 P0). change match n with | N0 => Npos p | Npos q => Npos (p + q) end with (N.add (Npos p) n);trivial. assert (H := Peq_ok P3 P0). destruct (P3 ?== P0). rewrite (H eq_refl). rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. add_permut. mul_permut. rewrite IHP2. rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. add_permut. mul_permut. Qed. Lemma mult_dev_ok : forall P fv n lm, mult_dev P fv n lm == P@fv * pow_N rI rmul (hd fv) n * r_list_pow lm. Proof. induction P;simpl;intros;Esimpl. rewrite mkmult_c_ok;rewrite add_pow_list_ok;Esimpl. rewrite IHP. simpl;rewrite add_pow_list_ok;Esimpl. change (match P3 with | Pc c => c ?=! cO | Pinj _ _ => false | PX _ _ _ => false end) with (Peq P3 P0). change match n with | N0 => Npos p | Npos q => Npos (p + q) end with (N.add (Npos p) n);trivial. assert (H := Peq_ok P3 P0). destruct (P3 ?== P0). rewrite (H eq_refl). rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. mul_permut. rewrite add_mult_dev_ok. rewrite IHP1; rewrite add_pow_list_ok. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. add_permut; mul_permut. Qed. Lemma Pphi_avoid_ok : forall P fv, Pphi_avoid fv P == P@fv. Proof. unfold Pphi_avoid;intros;rewrite mult_dev_ok;simpl;Esimpl. Qed. End EVALUATION. Definition Pphi_pow := let mkpow x p := match p with xH => x | _ => rpow x (Cp_phi (Npos p)) end in let mkopp_pow x p := ropp (mkpow x p) in let mkmult_pow r x p := rmul r (mkpow x p) in Pphi_avoid mkpow mkopp_pow mkmult_pow. Lemma local_mkpow_ok r p : match p with | xI _ => rpow r (Cp_phi (Npos p)) | xO _ => rpow r (Cp_phi (Npos p)) | 1 => r end == pow_pos rmul r p. Proof. destruct p; now rewrite ?(rpow_pow_N pow_th). Qed. Lemma Pphi_pow_ok : forall P fv, Pphi_pow fv P == P@fv. Proof. unfold Pphi_pow;intros;apply Pphi_avoid_ok;intros; now rewrite ?local_mkpow_ok. Qed. Lemma ring_rw_pow_correct : forall n lH l, interp_PElist l lH -> forall lmp, mk_monpol_list lH = lmp -> forall pe npe, norm_subst n lmp pe = npe -> PEeval l pe == Pphi_pow l npe. Proof. intros n lH l H1 lmp Heq1 pe npe Heq2. rewrite Pphi_pow_ok, <- Heq2, <- Heq1. apply norm_subst_ok. trivial. Qed. Fixpoint mkmult_pow (r x:R) (p: positive) {struct p} : R := match p with | xH => r*x | xO p => mkmult_pow (mkmult_pow r x p) x p | xI p => mkmult_pow (mkmult_pow (r*x) x p) x p end. Definition mkpow x p := match p with | xH => x | xO p => mkmult_pow x x (Pos.pred_double p) | xI p => mkmult_pow x x (xO p) end. Definition mkopp_pow x p := match p with | xH => -x | xO p => mkmult_pow (-x) x (Pos.pred_double p) | xI p => mkmult_pow (-x) x (xO p) end. Definition Pphi_dev := Pphi_avoid mkpow mkopp_pow mkmult_pow. Lemma mkmult_pow_ok p r x : mkmult_pow r x p == r * x^p. Proof. revert r; induction p;intros;simpl;Esimpl;rewrite !IHp;Esimpl. Qed. Lemma mkpow_ok p x : mkpow x p == x^p. Proof. destruct p;simpl;intros;Esimpl. - rewrite !mkmult_pow_ok;Esimpl. - rewrite mkmult_pow_ok;Esimpl. change x with (x^1) at 1. now rewrite <- pow_pos_add, Pos.add_1_r, Pos.succ_pred_double. Qed. Lemma mkopp_pow_ok p x : mkopp_pow x p == - x^p. Proof. destruct p;simpl;intros;Esimpl. - rewrite !mkmult_pow_ok;Esimpl. - rewrite mkmult_pow_ok;Esimpl. change x with (x^1) at 1. now rewrite <- pow_pos_add, Pos.add_1_r, Pos.succ_pred_double. Qed. Lemma Pphi_dev_ok : forall P fv, Pphi_dev fv P == P@fv. Proof. unfold Pphi_dev;intros;apply Pphi_avoid_ok. - intros;apply mkpow_ok. - intros;apply mkopp_pow_ok. - intros;apply mkmult_pow_ok. Qed. Lemma ring_rw_correct : forall n lH l, interp_PElist l lH -> forall lmp, mk_monpol_list lH = lmp -> forall pe npe, norm_subst n lmp pe = npe -> PEeval l pe == Pphi_dev l npe. Proof. intros n lH l H1 lmp Heq1 pe npe Heq2. rewrite Pphi_dev_ok. rewrite <- Heq2;rewrite <- Heq1. apply norm_subst_ok. trivial. Qed. End MakeRingPol. Arguments PEO {C}. Arguments PEI {C}. coq-8.11.0/plugins/setoid_ring/NArithRing.v0000644000175000017500000000162713612664533020457 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t | _ => constr:(NotConstant) end. Add Ring Nr : Nth (decidable Neqb_ok, constants [Ncst]). coq-8.11.0/plugins/setoid_ring/plugin_base.dune0000644000175000017500000000021513612664533021420 0ustar treinentreinen(library (name newring_plugin) (public_name coq.plugins.setoid_ring) (synopsis "Coq's setoid ring plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/setoid_ring/newring.mli0000644000175000017500000000241513612664533020433 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Id.t -> unit Proofview.tactic val protect_tac : string -> unit Proofview.tactic val add_theory : Id.t -> constr_expr -> constr_expr ring_mod list -> unit val print_rings : unit -> unit val ring_lookup : Geninterp.Val.t -> constr list -> constr list -> constr -> unit Proofview.tactic val add_field_theory : Id.t -> constr_expr -> constr_expr field_mod list -> unit val print_fields : unit -> unit val field_lookup : Geninterp.Val.t -> constr list -> constr list -> constr -> unit Proofview.tactic coq-8.11.0/plugins/fourier/0000755000175000017500000000000013612664533015422 5ustar treinentreinencoq-8.11.0/plugins/fourier/plugin_base.dune0000644000175000017500000000020513612664533020564 0ustar treinentreinen(library (name fourier_plugin) (public_name coq.plugins.fourier) (synopsis "Coq's fourier plugin") (libraries coq.plugins.ltac)) coq-8.11.0/plugins/derive/0000755000175000017500000000000013612664533015225 5ustar treinentreinencoq-8.11.0/plugins/derive/Derive.v0000644000175000017500000000004313612664533016627 0ustar treinentreinenDeclare ML Module "derive_plugin". coq-8.11.0/plugins/derive/derive_plugin.mlpack0000644000175000017500000000002013612664533021242 0ustar treinentreinenDerive G_derive coq-8.11.0/plugins/derive/derive.ml0000644000175000017500000000452513612664533017043 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* TCons ( env , sigma , f_type , (fun sigma ef -> let f_type = EConstr.Unsafe.to_constr f_type in let ef = EConstr.Unsafe.to_constr ef in let env' = Environ.push_named (LocalDef (annotR f, ef, f_type)) env in let sigma, suchthat = Constrintern.interp_type_evars ~program_mode:false env' sigma suchthat in TCons ( env' , sigma , suchthat , (fun sigma _ -> TNil sigma)))))) in let info = Lemmas.Info.make ~proof_ending:(Lemmas.Proof_ending.(End_derive {f; name})) ~kind () in let lemma = Lemmas.start_dependent_lemma ~name ~poly ~info goals in Lemmas.pf_map (Proof_global.map_proof begin fun p -> Util.pi1 @@ Proof.run_tactic env Proofview.(tclFOCUS 1 2 shelve) p end) lemma coq-8.11.0/plugins/derive/derive.mli0000644000175000017500000000176013612664533017212 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constrexpr.constr_expr -> Names.Id.t -> Lemmas.t coq-8.11.0/plugins/derive/g_derive.mlg0000644000175000017500000000201113612664533017504 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* { Derive.start_deriving f suchthat lemma } END coq-8.11.0/plugins/derive/plugin_base.dune0000644000175000017500000000020213612664533020364 0ustar treinentreinen(library (name derive_plugin) (public_name coq.plugins.derive) (synopsis "Coq's derive plugin") (libraries coq.plugins.ltac)) coq-8.11.0/kernel/0000755000175000017500000000000013612664533013546 5ustar treinentreinencoq-8.11.0/kernel/indtypes.mli0000644000175000017500000000152013612664533016106 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* MutInd.t -> mutual_inductive_entry -> mutual_inductive_body coq-8.11.0/kernel/environ.mli0000644000175000017500000003347613612664533015746 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (Vmvalues.values * Id.Set.t) -> unit val force_lazy_val : lazy_val -> (Vmvalues.values * Id.Set.t) option val dummy_lazy_val : unit -> lazy_val (** Linking information for the native compiler *) type link_info = | Linked of string | LinkedInteractive of string | NotLinked type key = int CEphemeron.key option ref type constant_key = Opaqueproof.opaque constant_body * (link_info ref * key) type mind_key = mutual_inductive_body * link_info ref module Globals : sig type t type view = { constants : constant_key Cmap_env.t ; inductives : mind_key Mindmap_env.t ; modules : module_body MPmap.t ; modtypes : module_type_body MPmap.t } val view : t -> view end type stratification = { env_universes : UGraph.t; env_sprop_allowed : bool; env_universes_lbound : Univ.Level.t; env_engagement : engagement } type named_context_val = private { env_named_ctx : Constr.named_context; env_named_map : (Constr.named_declaration * lazy_val) Id.Map.t; } type rel_context_val = private { env_rel_ctx : Constr.rel_context; env_rel_map : (Constr.rel_declaration * lazy_val) Range.t; } type env = private { env_globals : Globals.t; env_named_context : named_context_val; (* section variables *) env_rel_context : rel_context_val; env_nb_rel : int; env_stratification : stratification; env_typing_flags : typing_flags; retroknowledge : Retroknowledge.retroknowledge; indirect_pterms : Opaqueproof.opaquetab; native_symbols : Nativevalues.symbols DPmap.t; } val oracle : env -> Conv_oracle.oracle val set_oracle : env -> Conv_oracle.oracle -> env val eq_named_context_val : named_context_val -> named_context_val -> bool val empty_env : env val universes : env -> UGraph.t val universes_lbound : env -> Univ.Level.t val set_universes_lbound : env -> Univ.Level.t -> env val rel_context : env -> Constr.rel_context val named_context : env -> Constr.named_context val named_context_val : env -> named_context_val val opaque_tables : env -> Opaqueproof.opaquetab val set_opaque_tables : env -> Opaqueproof.opaquetab -> env val engagement : env -> engagement val typing_flags : env -> typing_flags val is_impredicative_set : env -> bool val type_in_type : env -> bool val deactivated_guard : env -> bool val indices_matter : env -> bool val check_template : env -> bool val is_impredicative_sort : env -> Sorts.t -> bool val is_impredicative_univ : env -> Univ.Universe.t -> bool (** is the local context empty *) val empty_context : env -> bool (** {5 Context of de Bruijn variables ([rel_context]) } *) val nb_rel : env -> int val push_rel : Constr.rel_declaration -> env -> env val push_rel_context : Constr.rel_context -> env -> env val push_rec_types : rec_declaration -> env -> env (** Looks up in the context of local vars referred by indice ([rel_context]) raises [Not_found] if the index points out of the context *) val lookup_rel : int -> env -> Constr.rel_declaration val lookup_rel_val : int -> env -> lazy_val val evaluable_rel : int -> env -> bool val env_of_rel : int -> env -> env (** {6 Recurrence on [rel_context] } *) val fold_rel_context : (env -> Constr.rel_declaration -> 'a -> 'a) -> env -> init:'a -> 'a (** {5 Context of variables (section variables and goal assumptions) } *) val named_context_of_val : named_context_val -> Constr.named_context val val_of_named_context : Constr.named_context -> named_context_val val empty_named_context_val : named_context_val val ids_of_named_context_val : named_context_val -> Id.Set.t (** [map_named_val f ctxt] apply [f] to the body and the type of each declarations. *** /!\ *** [f t] should be convertible with t, and preserve the name *) val map_named_val : (named_declaration -> named_declaration) -> named_context_val -> named_context_val val push_named : Constr.named_declaration -> env -> env val push_named_context : Constr.named_context -> env -> env val push_named_context_val : Constr.named_declaration -> named_context_val -> named_context_val (** Looks up in the context of local vars referred by names ([named_context]) raises [Not_found] if the Id.t is not found *) val lookup_named : variable -> env -> Constr.named_declaration val lookup_named_val : variable -> env -> lazy_val val lookup_named_ctxt : variable -> named_context_val -> Constr.named_declaration val evaluable_named : variable -> env -> bool val named_type : variable -> env -> types val named_body : variable -> env -> constr option (** {6 Recurrence on [named_context]: older declarations processed first } *) val fold_named_context : (env -> Constr.named_declaration -> 'a -> 'a) -> env -> init:'a -> 'a (** Recurrence on [named_context] starting from younger decl *) val fold_named_context_reverse : ('a -> Constr.named_declaration -> 'a) -> init:'a -> env -> 'a (** This forgets named and rel contexts *) val reset_context : env -> env (** This forgets rel context and sets a new named context *) val reset_with_named_context : named_context_val -> env -> env (** This removes the [n] last declarations from the rel context *) val pop_rel_context : int -> env -> env (** Useful for printing *) val fold_constants : (Constant.t -> Opaqueproof.opaque constant_body -> 'a -> 'a) -> env -> 'a -> 'a val fold_inductives : (MutInd.t -> Declarations.mutual_inductive_body -> 'a -> 'a) -> env -> 'a -> 'a (** {5 Global constants } {6 Add entries to global environment } *) val add_constant : Constant.t -> Opaqueproof.opaque constant_body -> env -> env val add_constant_key : Constant.t -> Opaqueproof.opaque constant_body -> link_info -> env -> env val lookup_constant_key : Constant.t -> env -> constant_key (** Looks up in the context of global constant names raises an anomaly if the required path is not found *) val lookup_constant : Constant.t -> env -> Opaqueproof.opaque constant_body val evaluable_constant : Constant.t -> env -> bool val mem_constant : Constant.t -> env -> bool (** New-style polymorphism *) val polymorphic_constant : Constant.t -> env -> bool val polymorphic_pconstant : pconstant -> env -> bool val type_in_type_constant : Constant.t -> env -> bool (** {6 ... } *) (** [constant_value env c] raises [NotEvaluableConst Opaque] if [c] is opaque, [NotEvaluableConst NoBody] if it has no body, [NotEvaluableConst IsProj] if [c] is a projection, [NotEvaluableConst (IsPrimitive p)] if [c] is primitive [p] and an anomaly if it does not exist in [env] *) type const_evaluation_result = | NoBody | Opaque | IsPrimitive of CPrimitives.t exception NotEvaluableConst of const_evaluation_result val constant_type : env -> Constant.t puniverses -> types constrained val constant_value_and_type : env -> Constant.t puniverses -> constr option * types * Univ.Constraint.t (** The universe context associated to the constant, empty if not polymorphic *) val constant_context : env -> Constant.t -> Univ.AUContext.t (* These functions should be called under the invariant that [env] already contains the constraints corresponding to the constant application. *) val constant_value_in : env -> Constant.t puniverses -> constr val constant_type_in : env -> Constant.t puniverses -> types val constant_opt_value_in : env -> Constant.t puniverses -> constr option val is_primitive : env -> Constant.t -> bool (** {6 Primitive projections} *) (** Checks that the number of parameters is correct. *) val lookup_projection : Names.Projection.t -> env -> types val get_projection : env -> inductive -> proj_arg:int -> Names.Projection.Repr.t option val get_projections : env -> inductive -> Names.Projection.Repr.t array option (** {5 Inductive types } *) val lookup_mind_key : MutInd.t -> env -> mind_key val add_mind_key : MutInd.t -> mind_key -> env -> env val add_mind : MutInd.t -> mutual_inductive_body -> env -> env (** Looks up in the context of global inductive names raises an anomaly if the required path is not found *) val lookup_mind : MutInd.t -> env -> mutual_inductive_body val mem_mind : MutInd.t -> env -> bool (** The universe context associated to the inductive, empty if not polymorphic *) val mind_context : env -> MutInd.t -> Univ.AUContext.t (** New-style polymorphism *) val polymorphic_ind : inductive -> env -> bool val polymorphic_pind : pinductive -> env -> bool val type_in_type_ind : inductive -> env -> bool (** Old-style polymorphism *) val template_polymorphic_ind : inductive -> env -> bool val template_polymorphic_variables : inductive -> env -> Univ.Level.t list val template_polymorphic_pind : pinductive -> env -> bool val template_checked_ind : inductive -> env -> bool (** {5 Modules } *) val add_modtype : module_type_body -> env -> env (** [shallow_add_module] does not add module components *) val shallow_add_module : module_body -> env -> env val lookup_module : ModPath.t -> env -> module_body val lookup_modtype : ModPath.t -> env -> module_type_body (** {5 Universe constraints } *) (** Add universe constraints to the environment. @raise UniverseInconsistency . *) val add_constraints : Univ.Constraint.t -> env -> env (** Check constraints are satifiable in the environment. *) val check_constraints : Univ.Constraint.t -> env -> bool val push_context : ?strict:bool -> Univ.UContext.t -> env -> env val push_context_set : ?strict:bool -> Univ.ContextSet.t -> env -> env val push_constraints_to_env : 'a Univ.constrained -> env -> env val push_subgraph : Univ.ContextSet.t -> env -> env (** [push_subgraph univs env] adds the universes and constraints in [univs] to [env] as [push_context_set ~strict:false univs env], and also checks that they do not imply new transitive constraints between pre-existing universes in [env]. *) val set_engagement : engagement -> env -> env val set_typing_flags : typing_flags -> env -> env val make_sprop_cumulative : env -> env val set_allow_sprop : bool -> env -> env val sprop_allowed : env -> bool val universes_of_global : env -> GlobRef.t -> AUContext.t (** {6 Sets of referred section variables } [global_vars_set env c] returns the list of [id]'s occurring either directly as [Var id] in [c] or indirectly as a section variable dependent in a global reference occurring in [c] *) val global_vars_set : env -> constr -> Id.Set.t val vars_of_global : env -> GlobRef.t -> Id.Set.t (** closure of the input id set w.r.t. dependency *) val really_needed : env -> Id.Set.t -> Id.Set.t (** like [really_needed] but computes a well ordered named context *) val keep_hyps : env -> Id.Set.t -> Constr.named_context (** {5 Unsafe judgments. } We introduce here the pre-type of judgments, which is actually only a datatype to store a term with its type and the type of its type. *) type ('constr, 'types) punsafe_judgment = { uj_val : 'constr; uj_type : 'types } val on_judgment : ('a -> 'b) -> ('a, 'a) punsafe_judgment -> ('b, 'b) punsafe_judgment val on_judgment_value : ('c -> 'c) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment val on_judgment_type : ('t -> 't) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment type unsafe_judgment = (constr, types) punsafe_judgment val make_judge : 'constr -> 'types -> ('constr, 'types) punsafe_judgment val j_val : ('constr, 'types) punsafe_judgment -> 'constr val j_type : ('constr, 'types) punsafe_judgment -> 'types type 'types punsafe_type_judgment = { utj_val : 'types; utj_type : Sorts.t } type unsafe_type_judgment = types punsafe_type_judgment exception Hyp_not_found (** [apply_to_hyp sign id f] split [sign] into [tail::(id,_,_)::head] and return [tail::(f head (id,_,_) (rev tail))::head]. the value associated to id should not change *) val apply_to_hyp : named_context_val -> variable -> (Constr.named_context -> Constr.named_declaration -> Constr.named_context -> Constr.named_declaration) -> named_context_val val remove_hyps : Id.Set.t -> (Constr.named_declaration -> Constr.named_declaration) -> (lazy_val -> lazy_val) -> named_context_val -> named_context_val val is_polymorphic : env -> Names.GlobRef.t -> bool val is_template_polymorphic : env -> GlobRef.t -> bool val get_template_polymorphic_variables : env -> GlobRef.t -> Univ.Level.t list val is_template_checked : env -> GlobRef.t -> bool val is_type_in_type : env -> GlobRef.t -> bool (** Native compiler *) val no_link_info : link_info (** Primitives *) val set_retroknowledge : env -> Retroknowledge.retroknowledge -> env val set_native_symbols : env -> Nativevalues.symbols DPmap.t -> env val add_native_symbols : DirPath.t -> Nativevalues.symbols -> env -> env coq-8.11.0/kernel/esubst.ml0000644000175000017500000001531313612664533015410 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* n *) (* i.e under n binders *) let el_id = ELID (* compose a relocation of magnitude n *) let el_shft_rec n = function | ELSHFT(el,k) -> ELSHFT(el,k+n) | el -> ELSHFT(el,n) let el_shft n el = if Int.equal n 0 then el else el_shft_rec n el (* cross n binders *) let el_liftn_rec n = function | ELID -> ELID | ELLFT(k,el) -> ELLFT(n+k, el) | el -> ELLFT(n, el) let el_liftn n el = if Int.equal n 0 then el else el_liftn_rec n el let el_lift el = el_liftn_rec 1 el (* relocation of de Bruijn n in an explicit lift *) let rec reloc_rel n = function | ELID -> n | ELLFT(k,el) -> if n <= k then n else (reloc_rel (n-k) el) + k | ELSHFT(el,k) -> (reloc_rel (n+k) el) let rec is_lift_id = function | ELID -> true | ELSHFT(e,n) -> Int.equal n 0 && is_lift_id e | ELLFT (_,e) -> is_lift_id e (*********************) (* Substitutions *) (*********************) (* (bounded) explicit substitutions of type 'a *) type 'a subs = | ESID of int (* ESID(n) = %n END bounded identity *) | CONS of 'a array * 'a subs (* CONS([|t1..tn|],S) = (S.t1...tn) parallel substitution beware of the order *) | SHIFT of int * 'a subs (* SHIFT(n,S) = (^n o S) terms in S are relocated *) (* with n vars *) | LIFT of int * 'a subs (* LIFT(n,S) = (%n S) stands for ((^n o S).n...1) *) (* operations of subs: collapses constructors when possible. * Needn't be recursive if we always use these functions *) let subs_id i = ESID i let subs_cons(x,s) = if Int.equal (Array.length x) 0 then s else CONS(x,s) let subs_liftn n = function | ESID p -> ESID (p+n) (* bounded identity lifted extends by p *) | LIFT (p,lenv) -> LIFT (p+n, lenv) | lenv -> LIFT (n,lenv) let subs_lift a = subs_liftn 1 a let subs_liftn n a = if Int.equal n 0 then a else subs_liftn n a let subs_shft = function | (0, s) -> s | (n, SHIFT (k,s1)) -> SHIFT (k+n, s1) | (n, s) -> SHIFT (n,s) let subs_shft s = if Int.equal (fst s) 0 then snd s else subs_shft s let subs_shift_cons = function (0, s, t) -> CONS(t,s) | (k, SHIFT(n,s1), t) -> CONS(t,SHIFT(k+n, s1)) | (k, s, t) -> CONS(t,SHIFT(k, s));; (* Tests whether a substitution is equal to the identity *) let rec is_subs_id = function ESID _ -> true | LIFT(_,s) -> is_subs_id s | SHIFT(0,s) -> is_subs_id s | CONS(x,s) -> Int.equal (Array.length x) 0 && is_subs_id s | _ -> false (* Expands de Bruijn k in the explicit substitution subs * lams accumulates de shifts to perform when retrieving the i-th value * the rules used are the following: * * [id]k --> k * [S.t]1 --> t * [S.t]k --> [S](k-1) if k > 1 * [^n o S] k --> [^n]([S]k) * [(%n S)] k --> k if k <= n * [(%n S)] k --> [^n]([S](k-n)) * * the result is (Inr (k+lams,p)) when the variable is just relocated * where p is None if the variable points inside subs and Some(k) if the * variable points k bindings beyond subs. *) let rec exp_rel lams k subs = match subs with | CONS (def,_) when k <= Array.length def -> Inl(lams,def.(Array.length def - k)) | CONS (v,l) -> exp_rel lams (k - Array.length v) l | LIFT (n,_) when k<=n -> Inr(lams+k,None) | LIFT (n,l) -> exp_rel (n+lams) (k-n) l | SHIFT (n,s) -> exp_rel (n+lams) k s | ESID n when k<=n -> Inr(lams+k,None) | ESID n -> Inr(lams+k,Some (k-n)) let expand_rel k subs = exp_rel 0 k subs let rec subs_map f = function | ESID _ as s -> s | CONS (x, s) -> CONS (Array.map f x, subs_map f s) | SHIFT (n, s) -> SHIFT (n, subs_map f s) | LIFT (n, s) -> LIFT (n, subs_map f s) let rec lift_subst mk_cl s1 s2 = match s1 with | ELID -> subs_map (fun c -> mk_cl ELID c) s2 | ELSHFT(s, k) -> subs_shft(k, lift_subst mk_cl s s2) | ELLFT (k, s) -> match s2 with | CONS(x,s') -> CONS(CArray.Fun1.map mk_cl s1 x, lift_subst mk_cl s1 s') | ESID n -> lift_subst mk_cl s (ESID (n + k)) | SHIFT(k',s') -> if k if k s1 | ESID _, _ -> s2 | SHIFT(k,s), _ -> subs_shft(k, comp mk_cl s s2) | _, CONS(x,s') -> CONS(Array.Fun1.map (fun s t -> mk_cl(s,t)) s1 x, comp mk_cl s1 s') | CONS(x,s), SHIFT(k,s') -> let lg = Array.length x in if k == lg then comp mk_cl s s' else if k > lg then comp mk_cl s (SHIFT(k-lg, s')) else comp mk_cl (CONS(Array.sub x 0 (lg-k), s)) s' | CONS(x,s), LIFT(k,s') -> let lg = Array.length x in if k == lg then CONS(x, comp mk_cl s s') else if k > lg then CONS(x, comp mk_cl s (LIFT(k-lg, s'))) else CONS(Array.sub x (lg-k) k, comp mk_cl (CONS(Array.sub x 0 (lg-k),s)) s') | LIFT(k,s), SHIFT(k',s') -> if k if k mp | MEapply (expr,_) -> mp_from_mexpr expr | MEwith (expr,_) -> mp_from_mexpr expr let is_modular = function | SFBmodule _ | SFBmodtype _ -> true | SFBconst _ | SFBmind _ -> false (** Split a [structure_body] at some label corresponding to a modular definition or not. *) let split_struc k m struc = let rec split rev_before = function | [] -> raise Not_found | (k',b)::after when Label.equal k k' && (is_modular b) == (m : bool) -> List.rev rev_before,b,after | h::tail -> split (h::rev_before) tail in split [] struc let discr_resolver mtb = match mtb.mod_type with | NoFunctor _ -> mtb.mod_delta | MoreFunctor _ -> empty_delta_resolver let rec rebuild_mp mp l = match l with | []-> mp | i::r -> rebuild_mp (MPdot(mp,Label.of_id i)) r let (+++) = Univ.ContextSet.union let rec check_with_def env struc (idl,(c,ctx)) mp equiv = let lab,idl = match idl with | [] -> assert false | id::idl -> Label.of_id id, idl in try let modular = not (List.is_empty idl) in let before,spec,after = split_struc lab modular struc in let env' = Modops.add_structure mp before equiv env in if List.is_empty idl then (* Toplevel definition *) let cb = match spec with | SFBconst cb -> cb | _ -> error_not_a_constant lab in (* In the spirit of subtyping.check_constant, we accept any implementations of parameters and opaques terms, as long as they have the right type *) let c', univs, ctx' = match cb.const_universes, ctx with | Monomorphic _, None -> let c',cst = match cb.const_body with | Undef _ | OpaqueDef _ -> let j = Typeops.infer env' c in assert (j.uj_val == c); (* relevances should already be correct here *) let typ = cb.const_type in let cst' = Reduction.infer_conv_leq env' (Environ.universes env') j.uj_type typ in j.uj_val, cst' | Def cs -> let c' = Mod_subst.force_constr cs in c, Reduction.infer_conv env' (Environ.universes env') c c' | Primitive _ -> error_incorrect_with_constraint lab in c', Monomorphic Univ.ContextSet.empty, cst | Polymorphic uctx, Some ctx -> let () = if not (UGraph.check_subtype ~lbound:(Environ.universes_lbound env) (Environ.universes env) uctx ctx) then error_incorrect_with_constraint lab in (** Terms are compared in a context with De Bruijn universe indices *) let env' = Environ.push_context ~strict:false (Univ.AUContext.repr uctx) env in let cst = match cb.const_body with | Undef _ | OpaqueDef _ -> let j = Typeops.infer env' c in assert (j.uj_val == c); (* relevances should already be correct here *) let typ = cb.const_type in let cst' = Reduction.infer_conv_leq env' (Environ.universes env') j.uj_type typ in cst' | Def cs -> let c' = Mod_subst.force_constr cs in let cst' = Reduction.infer_conv env' (Environ.universes env') c c' in cst' | Primitive _ -> error_incorrect_with_constraint lab in if not (Univ.Constraint.is_empty cst) then error_incorrect_with_constraint lab; c, Polymorphic ctx, Univ.Constraint.empty | _ -> error_incorrect_with_constraint lab in let def = Def (Mod_subst.from_val c') in (* let ctx' = Univ.UContext.make (newus, cst) in *) let cb' = { cb with const_body = def; const_universes = univs ; const_body_code = Option.map Cemitcodes.from_val (Cbytegen.compile_constant_body ~fail_on_error:false env' cb.const_universes def) } in before@(lab,SFBconst(cb'))::after, c', ctx' else (* Definition inside a sub-module *) let mb = match spec with | SFBmodule mb -> mb | _ -> error_not_a_module (Label.to_string lab) in begin match mb.mod_expr with | Abstract -> let struc = Modops.destr_nofunctor mb.mod_type in let struc',c',cst = check_with_def env' struc (idl,(c,ctx)) (MPdot(mp,lab)) mb.mod_delta in let mb' = { mb with mod_type = NoFunctor struc'; mod_type_alg = None } in before@(lab,SFBmodule mb')::after, c', cst | _ -> error_generative_module_expected lab end with | Not_found -> error_no_such_label lab | Reduction.NotConvertible -> error_incorrect_with_constraint lab let rec check_with_mod env struc (idl,mp1) mp equiv = let lab,idl = match idl with | [] -> assert false | id::idl -> Label.of_id id, idl in try let before,spec,after = split_struc lab true struc in let env' = Modops.add_structure mp before equiv env in let old = match spec with | SFBmodule mb -> mb | _ -> error_not_a_module (Label.to_string lab) in if List.is_empty idl then (* Toplevel module definition *) let mb_mp1 = lookup_module mp1 env in let mtb_mp1 = module_type_of_module mb_mp1 in let cst = match old.mod_expr with | Abstract -> let mtb_old = module_type_of_module old in let chk_cst = Subtyping.check_subtypes env' mtb_mp1 mtb_old in Univ.ContextSet.add_constraints chk_cst old.mod_constraints | Algebraic (NoFunctor (MEident(mp'))) -> check_modpath_equiv env' mp1 mp'; old.mod_constraints | _ -> error_generative_module_expected lab in let mp' = MPdot (mp,lab) in let new_mb = strengthen_and_subst_mb mb_mp1 mp' false in let new_mb' = { new_mb with mod_mp = mp'; mod_expr = Algebraic (NoFunctor (MEident mp1)); mod_constraints = cst } in let new_equiv = add_delta_resolver equiv new_mb.mod_delta in (* we propagate the new equality in the rest of the signature with the identity substitution accompanied by the new resolver*) let id_subst = map_mp mp' mp' new_mb.mod_delta in let new_after = subst_structure id_subst after in before@(lab,SFBmodule new_mb')::new_after, new_equiv, cst else (* Module definition of a sub-module *) let mp' = MPdot (mp,lab) in let old = match spec with | SFBmodule msb -> msb | _ -> error_not_a_module (Label.to_string lab) in begin match old.mod_expr with | Abstract -> let struc = destr_nofunctor old.mod_type in let struc',equiv',cst = check_with_mod env' struc (idl,mp1) mp' old.mod_delta in let new_mb = { old with mod_type = NoFunctor struc'; mod_type_alg = None; mod_delta = equiv' } in let new_equiv = add_delta_resolver equiv equiv' in let id_subst = map_mp mp' mp' equiv' in let new_after = subst_structure id_subst after in before@(lab,SFBmodule new_mb)::new_after, new_equiv, cst | Algebraic (NoFunctor (MEident mp0)) -> let mpnew = rebuild_mp mp0 idl in check_modpath_equiv env' mpnew mp; before@(lab,spec)::after, equiv, Univ.ContextSet.empty | _ -> error_generative_module_expected lab end with | Not_found -> error_no_such_label lab | Reduction.NotConvertible -> error_incorrect_with_constraint lab let check_with env mp (sign,alg,reso,cst) = function |WithDef(idl, (c, ctx)) -> let struc = destr_nofunctor sign in let struc', c', cst' = check_with_def env struc (idl, (c, ctx)) mp reso in let wd' = WithDef (idl, (c', ctx)) in NoFunctor struc', MEwith (alg,wd'), reso, Univ.ContextSet.add_constraints cst' cst |WithMod(idl,mp1) as wd -> let struc = destr_nofunctor sign in let struc',reso',cst' = check_with_mod env struc (idl,mp1) mp reso in NoFunctor struc', MEwith (alg,wd), reso', cst+++cst' let translate_apply env inl (sign,alg,reso,cst1) mp1 mkalg = let farg_id, farg_b, fbody_b = destr_functor sign in let mtb = module_type_of_module (lookup_module mp1 env) in let cst2 = Subtyping.check_subtypes env mtb farg_b in let mp_delta = discr_resolver mtb in let mp_delta = inline_delta_resolver env inl mp1 farg_id farg_b mp_delta in let subst = map_mbid farg_id mp1 mp_delta in let body = subst_signature subst fbody_b in let alg' = mkalg alg mp1 in let reso' = subst_codom_delta_resolver subst reso in body,alg',reso', Univ.ContextSet.add_constraints cst2 cst1 (** Translation of a module struct entry : - We translate to a module when a [module_path] is given, otherwise to a module type. - The first output is the expanded signature - The second output is the algebraic expression, kept for the extraction. *) let mk_alg_app alg arg = MEapply (alg,arg) let rec translate_mse env mpo inl = function |MEident mp1 as me -> let mb = match mpo with |Some mp -> strengthen_and_subst_mb (lookup_module mp1 env) mp false |None -> let mt = lookup_modtype mp1 env in module_body_of_type mt.mod_mp mt in mb.mod_type, me, mb.mod_delta, Univ.ContextSet.empty |MEapply (fe,mp1) -> translate_apply env inl (translate_mse env mpo inl fe) mp1 mk_alg_app |MEwith(me, with_decl) -> assert (mpo == None); (* No 'with' syntax for modules *) let mp = mp_from_mexpr me in check_with env mp (translate_mse env None inl me) with_decl let mk_mod mp e ty cst reso = { mod_mp = mp; mod_expr = e; mod_type = ty; mod_type_alg = None; mod_constraints = cst; mod_delta = reso; mod_retroknowledge = ModBodyRK []; } let mk_modtype mp ty cst reso = let mb = mk_mod mp Abstract ty cst reso in { mb with mod_expr = (); mod_retroknowledge = ModTypeRK } let rec translate_mse_funct env mpo inl mse = function |[] -> let sign,alg,reso,cst = translate_mse env mpo inl mse in sign, NoFunctor alg, reso, cst |(mbid, ty) :: params -> let mp_id = MPbound mbid in let mtb = translate_modtype env mp_id inl ([],ty) in let env' = add_module_type mp_id mtb env in let sign,alg,reso,cst = translate_mse_funct env' mpo inl mse params in let alg' = MoreFunctor (mbid,mtb,alg) in MoreFunctor (mbid, mtb, sign), alg',reso, cst +++ mtb.mod_constraints and translate_modtype env mp inl (params,mte) = let sign,alg,reso,cst = translate_mse_funct env None inl mte params in let mtb = mk_modtype (mp_from_mexpr mte) sign cst reso in let mtb' = subst_modtype_and_resolver mtb mp in { mtb' with mod_type_alg = Some alg } (** [finalize_module] : from an already-translated (or interactive) implementation and an (optional) signature entry, produces a final [module_body] *) let finalize_module env mp (sign,alg,reso,cst) restype = match restype with |None -> let impl = match alg with Some e -> Algebraic e | None -> FullStruct in mk_mod mp impl sign cst reso |Some (params_mte,inl) -> let res_mtb = translate_modtype env mp inl params_mte in let auto_mtb = mk_modtype mp sign Univ.ContextSet.empty reso in let cst' = Subtyping.check_subtypes env auto_mtb res_mtb in let impl = match alg with Some e -> Algebraic e | None -> Struct sign in { res_mtb with mod_mp = mp; mod_expr = impl; mod_retroknowledge = ModBodyRK []; (** cst from module body typing, cst' from subtyping, constraints from module type. *) mod_constraints = Univ.ContextSet.add_constraints cst' (cst +++ res_mtb.mod_constraints) } let translate_module env mp inl = function |MType (params,ty) -> let mtb = translate_modtype env mp inl (params,ty) in module_body_of_type mp mtb |MExpr (params,mse,oty) -> let (sg,alg,reso,cst) = translate_mse_funct env (Some mp) inl mse params in let restype = Option.map (fun ty -> ((params,ty),inl)) oty in finalize_module env mp (sg,Some alg,reso,cst) restype (** We now forbid any Include of functors with restricted signatures. Otherwise, we could end with the creation of undesired axioms (see #3746). Note that restricted non-functorized modules are ok, thanks to strengthening. *) let rec unfunct = function |NoFunctor me -> me |MoreFunctor(_,_,me) -> unfunct me let rec forbid_incl_signed_functor env = function |MEapply(fe,_) -> forbid_incl_signed_functor env fe |MEwith _ -> assert false (* No 'with' syntax for modules *) |MEident mp1 -> let mb = lookup_module mp1 env in match mb.mod_type, mb.mod_type_alg, mb.mod_expr with |MoreFunctor _, Some _, _ -> (* functor + restricted signature = error *) error_include_restricted_functor mp1 |MoreFunctor _, None, Algebraic me -> (* functor, no signature yet, a definition which may be restricted *) forbid_incl_signed_functor env (unfunct me) |_ -> () let rec translate_mse_inclmod env mp inl = function |MEident mp1 -> let mb = strengthen_and_subst_mb (lookup_module mp1 env) mp true in let sign = clean_bounded_mod_expr mb.mod_type in sign,(),mb.mod_delta,Univ.ContextSet.empty |MEapply (fe,arg) -> let ftrans = translate_mse_inclmod env mp inl fe in translate_apply env inl ftrans arg (fun _ _ -> ()) |MEwith _ -> assert false (* No 'with' syntax for modules *) let translate_mse_incl is_mod env mp inl me = if is_mod then let () = forbid_incl_signed_functor env me in translate_mse_inclmod env mp inl me else let mtb = translate_modtype env mp inl ([],me) in let sign = clean_bounded_mod_expr mtb.mod_type in sign,(),mtb.mod_delta,mtb.mod_constraints coq-8.11.0/kernel/cbytegen.mli0000644000175000017500000000244513612664533016056 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ?universes:int -> env -> constr -> (bytecodes * bytecodes * fv) option (** init, fun, fv *) val compile_constant_body : fail_on_error:bool -> env -> universes -> (Constr.t Mod_subst.substituted, 'opaque) constant_def -> body_code option (** Shortcut of the previous function used during module strengthening *) val compile_alias : Names.Constant.t -> body_code (** Dump the bytecode after compilation (for debugging purposes) *) val dump_bytecode : bool ref coq-8.11.0/kernel/uint63_63.ml0000644000175000017500000001505313612664533015544 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (float [@unboxed]) = "coq_uint63_to_float_byte" "coq_uint63_to_float" [@@noalloc] let hash i = i [@@ocaml.inline always] (* conversion of an uint63 to a string *) let to_string i = Int64.to_string (to_uint64 i) let of_string s = let i64 = Int64.of_string s in if Int64.compare Int64.zero i64 <= 0 && Int64.compare i64 maxuint63 <= 0 then Int64.to_int i64 else raise (Failure "Int64.of_string") (* Compiles an unsigned int to OCaml code *) let compile i = Printf.sprintf "Uint63.of_int (%i)" i let zero = 0 let one = 1 (* logical shift *) let l_sl x y = if 0 <= y && y < 63 then x lsl y else 0 let l_sr x y = if 0 <= y && y < 63 then x lsr y else 0 let l_and x y = x land y [@@ocaml.inline always] let l_or x y = x lor y [@@ocaml.inline always] let l_xor x y = x lxor y [@@ocaml.inline always] (* addition of int63 *) let add x y = x + y [@@ocaml.inline always] (* subtraction *) let sub x y = x - y [@@ocaml.inline always] (* multiplication *) let mul x y = x * y [@@ocaml.inline always] (* division *) let div (x : int) (y : int) = if y = 0 then 0 else Int64.to_int (Int64.div (to_uint64 x) (to_uint64 y)) (* modulo *) let rem (x : int) (y : int) = if y = 0 then 0 else Int64.to_int (Int64.rem (to_uint64 x) (to_uint64 y)) let diveucl x y = (div x y, rem x y) let addmuldiv p x y = l_or (l_sl x p) (l_sr y (uint_size - p)) (* comparison *) let lt (x : int) (y : int) = (x lxor 0x4000000000000000) < (y lxor 0x4000000000000000) [@@ocaml.inline always] let le (x : int) (y : int) = (x lxor 0x4000000000000000) <= (y lxor 0x4000000000000000) [@@ocaml.inline always] let to_int_min n m = if lt n m then n else m [@@ocaml.inline always] (* division of two numbers by one *) (* precondition: xh < y *) (* outputs: q, r s.t. x = q * y + r, r < y *) let div21 xh xl y = (* nh might temporarily grow as large as 2*y - 1 in the loop body, so we store it as a 64-bit unsigned integer *) let nh = ref xh in let nl = ref xl in let q = ref 0 in for _i = 0 to 62 do (* invariants: 0 <= nh < y, nl = (xl*2^i) % 2^63, (q*y + nh) * 2^(63-i) + (xl % 2^(63-i)) = (xh%y) * 2^63 + xl *) nh := Int64.logor (Int64.shift_left !nh 1) (Int64.of_int (!nl lsr 62)); nl := !nl lsl 1; q := !q lsl 1; (* TODO: use "Int64.unsigned_compare !nh y >= 0", once OCaml 4.08 becomes the minimal required version *) if Int64.compare !nh 0L < 0 || Int64.compare !nh y >= 0 then begin q := !q lor 1; nh := Int64.sub !nh y; end done; !q, Int64.to_int !nh let div21 xh xl y = let xh = to_uint64 xh in let y = to_uint64 y in if Int64.compare y xh <= 0 then 0, 0 else div21 xh xl y (* exact multiplication *) (* TODO: check that none of these additions could be a logical or *) (* size = 32 + 31 (hx << 31 + lx) * (hy << 31 + ly) = hxhy << 62 + (hxly + lxhy) << 31 + lxly *) (* precondition : (x lsr 62 = 0 || y lsr 62 = 0) *) let mulc_aux x y = let lx = x land maxuint31 in let ly = y land maxuint31 in let hx = x lsr 31 in let hy = y lsr 31 in (* hx and hy are 32 bits value but at most one is 32 *) let hxy = hx * hy in (* 63 bits *) let hxly = hx * ly in (* 63 bits *) let lxhy = lx * hy in (* 63 bits *) let lxy = lx * ly in (* 62 bits *) let l = lxy lor (hxy lsl 62) (* 63 bits *) in let h = hxy lsr 1 in (* 62 bits *) let hl = hxly + lxhy in (* We can have a carry *) let h = if lt hl hxly then h + (1 lsl 31) else h in let hl'= hl lsl 31 in let l = l + hl' in let h = if lt l hl' then h + 1 else h in let h = h + (hl lsr 32) in (h,l) let mulc x y = if (x lsr 62 = 0 || y lsr 62 = 0) then mulc_aux x y else let yl = y lxor (1 lsl 62) in let (h,l) = mulc_aux x yl in (* h << 63 + l = x * yl x * y = x * (1 << 62 + yl) = x << 62 + x*yl = x << 62 + h << 63 + l *) let l' = l + (x lsl 62) in let h = if lt l' l then h + 1 else h in (h + (x lsr 1), l') let equal (x : int) (y : int) = x = y [@@ocaml.inline always] let compare (x:int) (y:int) = let x = x lxor 0x4000000000000000 in let y = y lxor 0x4000000000000000 in if x > y then 1 else if y > x then -1 else 0 (* head tail *) let head0 x = let r = ref 0 in let x = ref x in if !x land 0x7FFFFFFF00000000 = 0 then r := !r + 31 else x := !x lsr 31; if !x land 0xFFFF0000 = 0 then (x := !x lsl 16; r := !r + 16); if !x land 0xFF000000 = 0 then (x := !x lsl 8; r := !r + 8); if !x land 0xF0000000 = 0 then (x := !x lsl 4; r := !r + 4); if !x land 0xC0000000 = 0 then (x := !x lsl 2; r := !r + 2); if !x land 0x80000000 = 0 then (x := !x lsl 1; r := !r + 1); if !x land 0x80000000 = 0 then ( r := !r + 1); !r;; let tail0 x = let r = ref 0 in let x = ref x in if !x land 0xFFFFFFFF = 0 then (x := !x lsr 32; r := !r + 32); if !x land 0xFFFF = 0 then (x := !x lsr 16; r := !r + 16); if !x land 0xFF = 0 then (x := !x lsr 8; r := !r + 8); if !x land 0xF = 0 then (x := !x lsr 4; r := !r + 4); if !x land 0x3 = 0 then (x := !x lsr 2; r := !r + 2); if !x land 0x1 = 0 then ( r := !r + 1); !r let is_uint63 t = Obj.is_int t [@@ocaml.inline always] (* Arithmetic with explicit carries *) (* Analog of Numbers.Abstract.Cyclic.carry *) type 'a carry = C0 of 'a | C1 of 'a let addc x y = let r = x + y in if lt r x then C1 r else C0 r let addcarryc x y = let r = x + y + 1 in if le r x then C1 r else C0 r let subc x y = let r = x - y in if le y x then C0 r else C1 r let subcarryc x y = let r = x - y - 1 in if lt y x then C0 r else C1 r coq-8.11.0/kernel/nativeconv.mli0000644000175000017500000000203513612664533016425 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* evars -> types kernel_conversion_function (** A conversion function parametrized by a universe comparator. Used outside of the kernel. *) val native_conv_gen : conv_pb -> evars -> (types, 'a) generic_conversion_function coq-8.11.0/kernel/csymtable.mli0000644000175000017500000000163313612664533016237 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr -> Vmvalues.values val set_opaque_const : Constant.t -> unit val set_transparent_const : Constant.t -> unit val get_global_data : unit -> Vmvalues.vm_global coq-8.11.0/kernel/nativelibrary.ml0000644000175000017500000000554013612664533016757 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let env' = add_structure mp struc empty_delta_resolver env in List.fold_left (translate_field prefix mp env') acc struc | MoreFunctor _ -> acc and translate_field prefix mp env acc (l,x) = match x with | SFBconst cb -> let con = Constant.make2 mp l in (if !Flags.debug then let msg = Printf.sprintf "Compiling constant %s..." (Constant.to_string con) in Feedback.msg_debug (Pp.str msg)); compile_constant_field env prefix con acc cb | SFBmind mb -> (if !Flags.debug then let id = mb.mind_packets.(0).mind_typename in let msg = Printf.sprintf "Compiling inductive %s..." (Id.to_string id) in Feedback.msg_debug (Pp.str msg)); compile_mind_field mp l acc mb | SFBmodule md -> let mp = md.mod_mp in (if !Flags.debug then let msg = Printf.sprintf "Compiling module %s..." (ModPath.to_string mp) in Feedback.msg_debug (Pp.str msg)); translate_mod prefix mp env md.mod_type acc | SFBmodtype mdtyp -> let mp = mdtyp.mod_mp in (if !Flags.debug then let msg = Printf.sprintf "Compiling module type %s..." (ModPath.to_string mp) in Feedback.msg_debug (Pp.str msg)); translate_mod prefix mp env mdtyp.mod_type acc let dump_library mp dp env mod_expr = if !Flags.debug then Feedback.msg_debug (Pp.str "Compiling library..."); match mod_expr with | NoFunctor struc -> let env = add_structure mp struc empty_delta_resolver env in let prefix = mod_uid_of_dirpath dp ^ "." in let t0 = Sys.time () in clear_global_tbl (); clear_symbols (); let mlcode = List.fold_left (translate_field prefix mp env) [] struc in let t1 = Sys.time () in let time_info = Format.sprintf "Time spent generating this code: %.5fs" (t1-.t0) in let mlcode = add_header_comment (List.rev mlcode) time_info in mlcode, get_symbols () | _ -> assert false coq-8.11.0/kernel/inductive.ml0000644000175000017500000014363513612664533016106 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* = Array.length mib.mind_packets then user_err Pp.(str "Inductive.lookup_mind_specif: invalid inductive index"); (mib, mib.mind_packets.(tyi)) let find_rectype env c = let (t, l) = decompose_app (whd_all env c) in match kind t with | Ind ind -> (ind, l) | _ -> raise Not_found let find_inductive env c = let (t, l) = decompose_app (whd_all env c) in match kind t with | Ind ind when (fst (lookup_mind_specif env (out_punivs ind))).mind_finite <> CoFinite -> (ind, l) | _ -> raise Not_found let find_coinductive env c = let (t, l) = decompose_app (whd_all env c) in match kind t with | Ind ind when (fst (lookup_mind_specif env (out_punivs ind))).mind_finite == CoFinite -> (ind, l) | _ -> raise Not_found let inductive_params (mib,_) = mib.mind_nparams let inductive_paramdecls (mib,u) = Vars.subst_instance_context u mib.mind_params_ctxt let instantiate_inductive_constraints mib u = Univ.AUContext.instantiate u (Declareops.inductive_polymorphic_context mib) (************************************************************************) (* Build the substitution that replaces Rels by the appropriate *) (* inductives *) let ind_subst mind mib u = let ntypes = mib.mind_ntypes in let make_Ik k = mkIndU ((mind,ntypes-k-1),u) in List.init ntypes make_Ik (* Instantiate inductives in constructor type *) let constructor_instantiate mind u mib c = let s = ind_subst mind mib u in substl s (subst_instance_constr u c) let instantiate_params full t u args sign = let fail () = anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch.") in let (rem_args, subs, ty) = Context.Rel.fold_outside (fun decl (largs,subs,ty) -> match (decl, largs, kind ty) with | (LocalAssum _, a::args, Prod(_,_,t)) -> (args, a::subs, t) | (LocalDef (_,b,_), _, LetIn(_,_,_,t)) -> (largs, (substl subs (subst_instance_constr u b))::subs, t) | (_,[],_) -> if full then fail() else ([], subs, ty) | _ -> fail ()) sign ~init:(args,[],t) in let () = if not (List.is_empty rem_args) then fail () in substl subs ty let full_inductive_instantiate mib u params sign = let dummy = Sorts.prop in let t = Term.mkArity (Vars.subst_instance_context u sign,dummy) in fst (Term.destArity (instantiate_params true t u params mib.mind_params_ctxt)) let full_constructor_instantiate ((mind,_),u,(mib,_),params) t = let inst_ind = constructor_instantiate mind u mib t in instantiate_params true inst_ind u params mib.mind_params_ctxt (************************************************************************) (************************************************************************) (* Functions to build standard types related to inductive *) (* Computing the actual sort of an applied or partially applied inductive type: I_i: forall uniformparams:utyps, forall otherparams:otyps, Type(a) uniformargs : utyps otherargs : otyps I_1:forall ...,s_1;...I_n:forall ...,s_n |- sort(C_kj(uniformargs)) = s_kj s'_k = max(..s_kj..) merge(..s'_k..) = ..s''_k.. -------------------------------------------------------------------- Gamma |- I_i uniformargs otherargs : phi(s''_i) where - if p=0, phi() = Prop - if p=1, phi(s) = s - if p<>1, phi(s) = sup(Set,s) Remark: Set (predicative) is encoded as Type(0) *) (* Template polymorphism *) (* cons_subst add the mapping [u |-> su] in subst if [u] is not *) (* in the domain or add [u |-> sup x su] if [u] is already mapped *) (* to [x]. *) let cons_subst u su subst = try Univ.LMap.add u (Univ.sup (Univ.LMap.find u subst) su) subst with Not_found -> Univ.LMap.add u su subst (* remember_subst updates the mapping [u |-> x] by [u |-> sup x u] *) (* if it is presents and returns the substitution unchanged if not.*) let remember_subst u subst = try let su = Universe.make u in Univ.LMap.add u (Univ.sup (Univ.LMap.find u subst) su) subst with Not_found -> subst (* Bind expected levels of parameters to actual levels *) (* Propagate the new levels in the signature *) let make_subst env = let rec make subst = function | LocalDef _ :: sign, exp, args -> make subst (sign, exp, args) | _d::sign, None::exp, args -> let args = match args with _::args -> args | [] -> [] in make subst (sign, exp, args) | _d::sign, Some u::exp, a::args -> (* We recover the level of the argument, but we don't change the *) (* level in the corresponding type in the arity; this level in the *) (* arity is a global level which, at typing time, will be enforce *) (* to be greater than the level of the argument; this is probably *) (* a useless extra constraint *) let s = Sorts.univ_of_sort (snd (dest_arity env (Lazy.force a))) in make (cons_subst u s subst) (sign, exp, args) | LocalAssum (_na,_t) :: sign, Some u::exp, [] -> (* No more argument here: we add the remaining universes to the *) (* substitution (when [u] is distinct from all other universes in the *) (* template, it is identity substitution otherwise (ie. when u is *) (* already in the domain of the substitution) [remember_subst] will *) (* update its image [x] by [sup x u] in order not to forget the *) (* dependency in [u] that remains to be fulfilled. *) make (remember_subst u subst) (sign, exp, []) | _sign, [], _ -> (* Uniform parameters are exhausted *) subst | [], _, _ -> assert false in make Univ.LMap.empty exception SingletonInductiveBecomesProp of Id.t let instantiate_universes env ctx ar argsorts = let args = Array.to_list argsorts in let subst = make_subst env (ctx,ar.template_param_levels,args) in let level = Univ.subst_univs_universe (Univ.make_subst subst) ar.template_level in let ty = (* Singleton type not containing types are interpretable in Prop *) if is_type0m_univ level then Sorts.prop (* Non singleton type not containing types are interpretable in Set *) else if is_type0_univ level then Sorts.set (* This is a Type with constraints *) else Sorts.sort_of_univ level in (ctx, ty) (* Type of an inductive type *) let relevance_of_inductive env ind = let _, mip = lookup_mind_specif env ind in mip.mind_relevance let type_of_inductive_gen ?(polyprop=true) env ((_,mip),u) paramtyps = match mip.mind_arity with | RegularArity a -> subst_instance_constr u a.mind_user_arity | TemplateArity ar -> let ctx = List.rev mip.mind_arity_ctxt in let ctx,s = instantiate_universes env ctx ar paramtyps in (* The Ocaml extraction cannot handle (yet?) "Prop-polymorphism", i.e. the situation where a non-Prop singleton inductive becomes Prop when applied to Prop params *) if not polyprop && not (is_type0m_univ ar.template_level) && Sorts.is_prop s then raise (SingletonInductiveBecomesProp mip.mind_typename); Term.mkArity (List.rev ctx,s) let type_of_inductive env pind = type_of_inductive_gen env pind [||] let constrained_type_of_inductive env ((mib,_mip),u as pind) = let ty = type_of_inductive env pind in let cst = instantiate_inductive_constraints mib u in (ty, cst) let constrained_type_of_inductive_knowing_parameters env ((mib,_mip),u as pind) args = let ty = type_of_inductive_gen env pind args in let cst = instantiate_inductive_constraints mib u in (ty, cst) let type_of_inductive_knowing_parameters env ?(polyprop=true) mip args = type_of_inductive_gen ~polyprop env mip args (* The max of an array of universes *) let cumulate_constructor_univ u = let open Sorts in function | SProp | Prop -> (* SProp is non cumulative but allowed in constructors of any inductive (except non-sprop primitive records) *) u | Set -> Universe.sup Universe.type0 u | Type u' -> Universe.sup u u' let max_inductive_sort = Array.fold_left cumulate_constructor_univ Universe.type0m (************************************************************************) (* Type of a constructor *) let type_of_constructor (cstr, u) (mib,mip) = let ind = inductive_of_constructor cstr in let specif = mip.mind_user_lc in let i = index_of_constructor cstr in let nconstr = Array.length mip.mind_consnames in if i > nconstr then user_err Pp.(str "Not enough constructors in the type."); constructor_instantiate (fst ind) u mib specif.(i-1) let constrained_type_of_constructor (_cstr,u as cstru) (mib,_mip as ind) = let ty = type_of_constructor cstru ind in let cst = instantiate_inductive_constraints mib u in (ty, cst) let arities_of_specif (kn,u) (mib,mip) = let specif = mip.mind_nf_lc in let map (ctx, c) = let cty = Term.it_mkProd_or_LetIn c ctx in constructor_instantiate kn u mib cty in Array.map map specif let arities_of_constructors ind specif = arities_of_specif (fst (fst ind), snd ind) specif let type_of_constructors (ind,u) (mib,mip) = let specif = mip.mind_user_lc in Array.map (constructor_instantiate (fst ind) u mib) specif (************************************************************************) (* Type of case predicates *) (* Get type of inductive, with parameters instantiated *) let inductive_sort_family mip = match mip.mind_arity with | RegularArity s -> Sorts.family s.mind_sort | TemplateArity _ -> Sorts.InType let mind_arity mip = mip.mind_arity_ctxt, inductive_sort_family mip let get_instantiated_arity (_ind,u) (mib,mip) params = let sign, s = mind_arity mip in full_inductive_instantiate mib u params sign, s let elim_sort (_,mip) = mip.mind_kelim let is_private (mib,_) = mib.mind_private = Some true let is_primitive_record (mib,_) = match mib.mind_record with | PrimRecord _ -> true | NotRecord | FakeRecord -> false let build_dependent_inductive ind (_,mip) params = let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in Term.applist (mkIndU ind, List.map (lift mip.mind_nrealdecls) params @ Context.Rel.to_extended_list mkRel 0 realargs) (* This exception is local *) exception LocalArity of (Sorts.family * Sorts.family * Sorts.family * arity_error) option let check_allowed_sort ksort specif = if not (Sorts.family_leq ksort (elim_sort specif)) then let s = inductive_sort_family (snd specif) in raise (LocalArity (Some(elim_sort specif, ksort,s,error_elim_explain ksort s))) let is_correct_arity env c pj ind specif params = let arsign,_ = get_instantiated_arity ind specif params in let rec srec env pt ar = let pt' = whd_all env pt in match kind pt', ar with | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> let () = try conv env a1 a1' with NotConvertible -> raise (LocalArity None) in srec (push_rel (LocalAssum (na1,a1)) env) t ar' (* The last Prod domain is the type of the scrutinee *) | Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *) let env' = push_rel (LocalAssum (na1,a1)) env in let ksort = match kind (whd_all env' a2) with | Sort s -> Sorts.family s | _ -> raise (LocalArity None) in let dep_ind = build_dependent_inductive ind specif params in let _ = try conv env a1 dep_ind with NotConvertible -> raise (LocalArity None) in check_allowed_sort ksort specif | _, (LocalDef _ as d)::ar' -> srec (push_rel d env) (lift 1 pt') ar' | _ -> raise (LocalArity None) in try srec env pj.uj_type (List.rev arsign) with LocalArity kinds -> error_elim_arity env ind c pj kinds (************************************************************************) (* Type of case branches *) (* [p] is the predicate, [i] is the constructor number (starting from 0), and [cty] is the type of the constructor (params not instantiated) *) let build_branches_type (ind,u) (_,mip as specif) params p = let build_one_branch i (ctx, c) = let cty = Term.it_mkProd_or_LetIn c ctx in let typi = full_constructor_instantiate (ind,u,specif,params) cty in let (cstrsign,ccl) = Term.decompose_prod_assum typi in let nargs = Context.Rel.length cstrsign in let (_,allargs) = decompose_app ccl in let (lparams,vargs) = List.chop (inductive_params specif) allargs in let cargs = let cstr = ith_constructor_of_inductive ind (i+1) in let dep_cstr = Term.applist (mkConstructU (cstr,u),lparams@(Context.Rel.to_extended_list mkRel 0 cstrsign)) in vargs @ [dep_cstr] in let base = Term.lambda_appvect_assum (mip.mind_nrealdecls+1) (lift nargs p) (Array.of_list cargs) in Term.it_mkProd_or_LetIn base cstrsign in Array.mapi build_one_branch mip.mind_nf_lc (* [p] is the predicate, [c] is the match object, [realargs] is the list of real args of the inductive type *) let build_case_type env n p c realargs = whd_betaiota env (Term.lambda_appvect_assum (n+1) p (Array.of_list (realargs@[c]))) let type_case_branches env (pind,largs) pj c = let specif = lookup_mind_specif env (fst pind) in let nparams = inductive_params specif in let (params,realargs) = List.chop nparams largs in let p = pj.uj_val in let () = is_correct_arity env c pj pind specif params in let lc = build_branches_type pind specif params p in let ty = build_case_type env (snd specif).mind_nrealdecls p c realargs in (lc, ty) (************************************************************************) (* Checking the case annotation is relevant *) let check_case_info env (indsp,u) r ci = let (mib,mip as spec) = lookup_mind_specif env indsp in if not (eq_ind indsp ci.ci_ind) || not (Int.equal mib.mind_nparams ci.ci_npar) || not (Array.equal Int.equal mip.mind_consnrealdecls ci.ci_cstr_ndecls) || not (Array.equal Int.equal mip.mind_consnrealargs ci.ci_cstr_nargs) || not (ci.ci_relevance == r) || is_primitive_record spec then raise (TypeError(env,WrongCaseInfo((indsp,u),ci))) (************************************************************************) (************************************************************************) (* Guard conditions for fix and cofix-points *) (* Check if t is a subterm of Rel n, and gives its specification, assuming lst already gives index of subterms with corresponding specifications of recursive arguments *) (* A powerful notion of subterm *) (* To each inductive definition corresponds an array describing the structure of recursive arguments for each constructor, we call it the recursive spec of the type (it has type recargs vect). For checking the guard, we start from the decreasing argument (Rel n) with its recursive spec. During checking the guardness condition, we collect patterns variables corresponding to subterms of n, each of them with its recursive spec. They are organised in a list lst of type (int * recargs) list which is sorted with respect to the first argument. *) (*************************************************************) (* Environment annotated with marks on recursive arguments *) (* tells whether it is a strict or loose subterm *) type size = Large | Strict (* merging information *) let size_glb s1 s2 = match s1,s2 with Strict, Strict -> Strict | _ -> Large (* possible specifications for a term: - Not_subterm: when the size of a term is not related to the recursive argument of the fixpoint - Subterm: when the term is a subterm of the recursive argument the wf_paths argument specifies which subterms are recursive - Dead_code: when the term has been built by elimination over an empty type *) type subterm_spec = Subterm of (size * wf_paths) | Dead_code | Not_subterm let eq_wf_paths = Rtree.equal Declareops.eq_recarg let inter_recarg r1 r2 = match r1, r2 with | Norec, Norec -> Some r1 | Mrec i1, Mrec i2 | Imbr i1, Imbr i2 | Mrec i1, Imbr i2 -> if Names.eq_ind i1 i2 then Some r1 else None | Imbr i1, Mrec i2 -> if Names.eq_ind i1 i2 then Some r2 else None | _ -> None let inter_wf_paths = Rtree.inter Declareops.eq_recarg inter_recarg Norec let incl_wf_paths = Rtree.incl Declareops.eq_recarg inter_recarg Norec let spec_of_tree t = if eq_wf_paths t mk_norec then Not_subterm else Subterm (Strict, t) let inter_spec s1 s2 = match s1, s2 with | _, Dead_code -> s1 | Dead_code, _ -> s2 | Not_subterm, _ -> s1 | _, Not_subterm -> s2 | Subterm (a1,t1), Subterm (a2,t2) -> Subterm (size_glb a1 a2, inter_wf_paths t1 t2) let subterm_spec_glb = Array.fold_left inter_spec Dead_code type guard_env = { env : env; (* dB of last fixpoint *) rel_min : int; (* dB of variables denoting subterms *) genv : subterm_spec Lazy.t list; } let make_renv env recarg tree = { env = env; rel_min = recarg+2; (* recarg = 0 ==> Rel 1 -> recarg; Rel 2 -> fix *) genv = [Lazy.from_val(Subterm(Large,tree))] } let push_var renv (x,ty,spec) = { env = push_rel (LocalAssum (x,ty)) renv.env; rel_min = renv.rel_min+1; genv = spec:: renv.genv } let assign_var_spec renv (i,spec) = { renv with genv = List.assign renv.genv (i-1) spec } let push_var_renv renv (x,ty) = push_var renv (x,ty,lazy Not_subterm) (* Fetch recursive information about a variable p *) let subterm_var p renv = try Lazy.force (List.nth renv.genv (p-1)) with Failure _ | Invalid_argument _ -> Not_subterm let push_ctxt_renv renv ctxt = let n = Context.Rel.length ctxt in { env = push_rel_context ctxt renv.env; rel_min = renv.rel_min+n; genv = iterate (fun ge -> lazy Not_subterm::ge) n renv.genv } let push_fix_renv renv (_,v,_ as recdef) = let n = Array.length v in { env = push_rec_types recdef renv.env; rel_min = renv.rel_min+n; genv = iterate (fun ge -> lazy Not_subterm::ge) n renv.genv } (* Definition and manipulation of the stack *) type stack_element = |SClosure of guard_env*constr |SArg of subterm_spec Lazy.t let push_stack_closures renv l stack = List.fold_right (fun h b -> (SClosure (renv,h))::b) l stack let push_stack_args l stack = List.fold_right (fun h b -> (SArg h)::b) l stack (******************************) (* {6 Computing the recursive subterms of a term (propagation of size information through Cases).} *) let lookup_subterms env ind = let (_,mip) = lookup_mind_specif env ind in mip.mind_recargs let match_inductive ind ra = match ra with | (Mrec i | Imbr i) -> eq_ind ind i | Norec -> false (* In {match c as z in ci y_s return P with |C_i x_s => t end} [branches_specif renv c_spec ci] returns an array of x_s specs knowing c_spec. *) let branches_specif renv c_spec ci = let car = (* We fetch the regular tree associated to the inductive of the match. This is just to get the number of constructors (and constructor arities) that fit the match branches without forcing c_spec. Note that c_spec might be more precise than [v] below, because of nested inductive types. *) let (_,mip) = lookup_mind_specif renv.env ci.ci_ind in let v = dest_subterms mip.mind_recargs in Array.map List.length v in Array.mapi (fun i nca -> (* i+1-th cstructor has arity nca *) let lvra = lazy (match Lazy.force c_spec with Subterm (_,t) when match_inductive ci.ci_ind (dest_recarg t) -> let vra = Array.of_list (dest_subterms t).(i) in assert (Int.equal nca (Array.length vra)); Array.map spec_of_tree vra | Dead_code -> Array.make nca Dead_code | _ -> Array.make nca Not_subterm) in List.init nca (fun j -> lazy (Lazy.force lvra).(j))) car let check_inductive_codomain env p = let absctx, ar = dest_lam_assum env p in let env = push_rel_context absctx env in let arctx, s = dest_prod_assum env ar in let env = push_rel_context arctx env in let i,_l' = decompose_app (whd_all env s) in isInd i (* The following functions are almost duplicated from indtypes.ml, except that they carry here a poorer environment (containing less information). *) let ienv_push_var (env, lra) (x,a,ra) = (push_rel (LocalAssum (x,a)) env, (Norec,ra)::lra) let ienv_push_inductive (env, ra_env) ((mind,u),lpar) = let mib = Environ.lookup_mind mind env in let ntypes = mib.mind_ntypes in let push_ind specif env = let r = specif.mind_relevance in let anon = Context.make_annot Anonymous r in let decl = LocalAssum (anon, hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) in push_rel decl env in let env = Array.fold_right push_ind mib.mind_packets env in let rc = Array.mapi (fun j t -> (Imbr (mind,j),t)) (Rtree.mk_rec_calls ntypes) in let lra_ind = Array.rev_to_list rc in let ra_env = List.map (fun (r,t) -> (r,Rtree.lift ntypes t)) ra_env in (env, lra_ind @ ra_env) let rec ienv_decompose_prod (env,_ as ienv) n c = if Int.equal n 0 then (ienv,c) else let c' = whd_all env c in match kind c' with Prod(na,a,b) -> let ienv' = ienv_push_var ienv (na,a,mk_norec) in ienv_decompose_prod ienv' (n-1) b | _ -> assert false let dummy_univ = Level.(make (UGlobal.make (DirPath.make [Id.of_string "implicit"]) 0)) let dummy_implicit_sort = mkType (Universe.make dummy_univ) let lambda_implicit_lift n a = let anon = Context.make_annot Anonymous Sorts.Relevant in let lambda_implicit a = mkLambda (anon, dummy_implicit_sort, a) in iterate lambda_implicit n (lift n a) (* This removes global parameters of the inductive types in lc (for nested inductive types only ) *) let abstract_mind_lc ntyps npars lc = let lc = Array.map (fun (ctx, c) -> Term.it_mkProd_or_LetIn c ctx) lc in if Int.equal npars 0 then lc else let make_abs = List.init ntyps (function i -> lambda_implicit_lift npars (mkRel (i+1))) in Array.map (substl make_abs) lc (* [get_recargs_approx env tree ind args] builds an approximation of the recargs tree for ind, knowing args. The argument tree is used to know when candidate nested types should be traversed, pruning the tree otherwise. This code is very close to check_positive in indtypes.ml, but does no positivity check and does not compute the number of recursive arguments. *) let get_recargs_approx env tree ind args = let rec build_recargs (env, ra_env as ienv) tree c = let x,largs = decompose_app (whd_all env c) in match kind x with | Prod (na,b,d) -> assert (List.is_empty largs); build_recargs (ienv_push_var ienv (na, b, mk_norec)) tree d | Rel k -> (* Free variables are allowed and assigned Norec *) (try snd (List.nth ra_env (k-1)) with Failure _ | Invalid_argument _ -> mk_norec) | Ind ind_kn -> (* When the inferred tree allows it, we consider that we have a potential nested inductive type *) begin match dest_recarg tree with | Imbr kn' | Mrec kn' when eq_ind (fst ind_kn) kn' -> build_recargs_nested ienv tree (ind_kn, largs) | _ -> mk_norec end | _err -> mk_norec and build_recargs_nested (env,_ra_env as ienv) tree (((mind,i),u), largs) = (* If the inferred tree already disallows recursion, no need to go further *) if eq_wf_paths tree mk_norec then tree else let mib = Environ.lookup_mind mind env in let auxnpar = mib.mind_nparams_rec in let nonrecpar = mib.mind_nparams - auxnpar in let (lpar,_) = List.chop auxnpar largs in let auxntyp = mib.mind_ntypes in (* Extends the environment with a variable corresponding to the inductive def *) let (env',_ as ienv') = ienv_push_inductive ienv ((mind,u),lpar) in (* Parameters expressed in env' *) let lpar' = List.map (lift auxntyp) lpar in (* In case of mutual inductive types, we use the recargs tree which was computed statically. This is fine because nested inductive types with mutually recursive containers are not supported. *) let trees = if Int.equal auxntyp 1 then [|dest_subterms tree|] else Array.map (fun mip -> dest_subterms mip.mind_recargs) mib.mind_packets in let mk_irecargs j specif = (* The nested inductive type with parameters removed *) let auxlcvect = abstract_mind_lc auxntyp auxnpar specif.mind_nf_lc in let paths = Array.mapi (fun k c -> let c' = hnf_prod_applist env' c lpar' in (* skip non-recursive parameters *) let (ienv',c') = ienv_decompose_prod ienv' nonrecpar c' in build_recargs_constructors ienv' trees.(j).(k) c') auxlcvect in mk_paths (Imbr (mind,j)) paths in let irecargs = Array.mapi mk_irecargs mib.mind_packets in (Rtree.mk_rec irecargs).(i) and build_recargs_constructors ienv trees c = let rec recargs_constr_rec (env,_ra_env as ienv) trees lrec c = let x,largs = decompose_app (whd_all env c) in match kind x with | Prod (na,b,d) -> let () = assert (List.is_empty largs) in let recarg = build_recargs ienv (List.hd trees) b in let ienv' = ienv_push_var ienv (na,b,mk_norec) in recargs_constr_rec ienv' (List.tl trees) (recarg::lrec) d | _hd -> List.rev lrec in recargs_constr_rec ienv trees [] c in (* starting with ra_env = [] seems safe because any unbounded Rel will be assigned Norec *) build_recargs_nested (env,[]) tree (ind, args) (* [restrict_spec env spec p] restricts the size information in spec to what is allowed to flow through a match with predicate p in environment env. *) let restrict_spec env spec p = if spec = Not_subterm then spec else let absctx, ar = dest_lam_assum env p in (* Optimization: if the predicate is not dependent, no restriction is needed and we avoid building the recargs tree. *) if noccur_with_meta 1 (Context.Rel.length absctx) ar then spec else let env = push_rel_context absctx env in let arctx, s = dest_prod_assum env ar in let env = push_rel_context arctx env in let i,args = decompose_app (whd_all env s) in match kind i with | Ind i -> begin match spec with | Dead_code -> spec | Subterm(st,tree) -> let recargs = get_recargs_approx env tree i args in let recargs = inter_wf_paths tree recargs in Subterm(st,recargs) | _ -> assert false end | _ -> Not_subterm (* [subterm_specif renv t] computes the recursive structure of [t] and compare its size with the size of the initial recursive argument of the fixpoint we are checking. [renv] collects such information about variables. *) let rec subterm_specif renv stack t = (* maybe reduction is not always necessary! *) let f,l = decompose_app (whd_all renv.env t) in match kind f with | Rel k -> subterm_var k renv | Case (ci,p,c,lbr) -> let stack' = push_stack_closures renv l stack in let cases_spec = branches_specif renv (lazy_subterm_specif renv [] c) ci in let stl = Array.mapi (fun i br' -> let stack_br = push_stack_args (cases_spec.(i)) stack' in subterm_specif renv stack_br br') lbr in let spec = subterm_spec_glb stl in restrict_spec renv.env spec p | Fix ((recindxs,i),(_,typarray,bodies as recdef)) -> (* when proving that the fixpoint f(x)=e is less than n, it is enough to prove that e is less than n assuming f is less than n furthermore when f is applied to a term which is strictly less than n, one may assume that x itself is strictly less than n *) if not (check_inductive_codomain renv.env typarray.(i)) then Not_subterm else let (ctxt,clfix) = dest_prod renv.env typarray.(i) in let oind = let env' = push_rel_context ctxt renv.env in try Some(fst(find_inductive env' clfix)) with Not_found -> None in (match oind with None -> Not_subterm (* happens if fix is polymorphic *) | Some (ind, _) -> let nbfix = Array.length typarray in let recargs = lookup_subterms renv.env ind in (* pushing the fixpoints *) let renv' = push_fix_renv renv recdef in let renv' = (* Why Strict here ? To be general, it could also be Large... *) assign_var_spec renv' (nbfix-i, lazy (Subterm(Strict,recargs))) in let decrArg = recindxs.(i) in let theBody = bodies.(i) in let nbOfAbst = decrArg+1 in let sign,strippedBody = Term.decompose_lam_n_assum nbOfAbst theBody in (* pushing the fix parameters *) let stack' = push_stack_closures renv l stack in let renv'' = push_ctxt_renv renv' sign in let renv'' = if List.length stack' < nbOfAbst then renv'' else let decrArg = List.nth stack' decrArg in let arg_spec = stack_element_specif decrArg in assign_var_spec renv'' (1, arg_spec) in subterm_specif renv'' [] strippedBody) | Lambda (x,a,b) -> let () = assert (List.is_empty l) in let spec,stack' = extract_stack stack in subterm_specif (push_var renv (x,a,spec)) stack' b (* Metas and evars are considered OK *) | (Meta _|Evar _) -> Dead_code | Proj (p, c) -> let subt = subterm_specif renv stack c in (match subt with | Subterm (_s, wf) -> (* We take the subterm specs of the constructor of the record *) let wf_args = (dest_subterms wf).(0) in (* We extract the tree of the projected argument *) let n = Projection.arg p in spec_of_tree (List.nth wf_args n) | Dead_code -> Dead_code | Not_subterm -> Not_subterm) | Var _ | Sort _ | Cast _ | Prod _ | LetIn _ | App _ | Const _ | Ind _ | Construct _ | CoFix _ | Int _ | Float _ -> Not_subterm (* Other terms are not subterms *) and lazy_subterm_specif renv stack t = lazy (subterm_specif renv stack t) and stack_element_specif = function |SClosure (h_renv,h) -> lazy_subterm_specif h_renv [] h |SArg x -> x and extract_stack = function | [] -> Lazy.from_val Not_subterm , [] | h::t -> stack_element_specif h, t (* Check term c can be applied to one of the mutual fixpoints. *) let check_is_subterm x tree = match Lazy.force x with | Subterm (Strict,tree') -> incl_wf_paths tree tree' | Dead_code -> true | _ -> false (************************************************************************) exception FixGuardError of env * guard_error let error_illegal_rec_call renv fx (arg_renv,arg) = let (_,le_vars,lt_vars) = List.fold_left (fun (i,le,lt) sbt -> match Lazy.force sbt with (Subterm(Strict,_) | Dead_code) -> (i+1, le, i::lt) | (Subterm(Large,_)) -> (i+1, i::le, lt) | _ -> (i+1, le ,lt)) (1,[],[]) renv.genv in raise (FixGuardError (renv.env, RecursionOnIllegalTerm(fx,(arg_renv.env, arg), le_vars,lt_vars))) let error_partial_apply renv fx = raise (FixGuardError (renv.env,NotEnoughArgumentsForFixCall fx)) let filter_stack_domain env p stack = let absctx, ar = dest_lam_assum env p in (* Optimization: if the predicate is not dependent, no restriction is needed and we avoid building the recargs tree. *) if noccur_with_meta 1 (Context.Rel.length absctx) ar then stack else let env = push_rel_context absctx env in let rec filter_stack env ar stack = let t = whd_all env ar in match stack, kind t with | elt :: stack', Prod (n,a,c0) -> let d = LocalAssum (n,a) in let ctx, a = dest_prod_assum env a in let env = push_rel_context ctx env in let ty, args = decompose_app (whd_all env a) in let elt = match kind ty with | Ind ind -> let spec' = stack_element_specif elt in (match (Lazy.force spec') with | Not_subterm | Dead_code -> elt | Subterm(s,path) -> let recargs = get_recargs_approx env path ind args in let path = inter_wf_paths path recargs in SArg (lazy (Subterm(s,path)))) | _ -> (SArg (lazy Not_subterm)) in elt :: filter_stack (push_rel d env) c0 stack' | _,_ -> List.fold_right (fun _ l -> SArg (lazy Not_subterm) :: l) stack [] in filter_stack env ar stack let judgment_of_fixpoint (_, types, bodies) = Array.map2 (fun typ body -> { uj_val = body ; uj_type = typ }) types bodies (* Check if [def] is a guarded fixpoint body with decreasing arg. given [recpos], the decreasing arguments of each mutually defined fixpoint. *) let check_one_fix renv recpos trees def = let nfi = Array.length recpos in (* Checks if [t] only make valid recursive calls [stack] is the list of constructor's argument specification and arguments that will be applied after reduction. example u in t where we have (match .. with |.. => t end) u *) let rec check_rec_call renv stack t = (* if [t] does not make recursive calls, it is guarded: *) if noccur_with_meta renv.rel_min nfi t then () else let (f,l) = decompose_app (whd_betaiotazeta renv.env t) in match kind f with | Rel p -> (* Test if [p] is a fixpoint (recursive call) *) if renv.rel_min <= p && p < renv.rel_min+nfi then begin List.iter (check_rec_call renv []) l; (* the position of the invoked fixpoint: *) let glob = renv.rel_min+nfi-1-p in (* the decreasing arg of the rec call: *) let np = recpos.(glob) in let stack' = push_stack_closures renv l stack in if List.length stack' <= np then error_partial_apply renv glob else (* Retrieve the expected tree for the argument *) (* Check the decreasing arg is smaller *) let z = List.nth stack' np in if not (check_is_subterm (stack_element_specif z) trees.(glob)) then begin match z with |SClosure (z,z') -> error_illegal_rec_call renv glob (z,z') |SArg _ -> error_partial_apply renv glob end end else begin match lookup_rel p renv.env with | LocalAssum _ -> List.iter (check_rec_call renv []) l | LocalDef (_,c,_) -> try List.iter (check_rec_call renv []) l with FixGuardError _ -> check_rec_call renv stack (Term.applist(lift p c,l)) end | Case (ci,p,c_0,lrest) -> begin try List.iter (check_rec_call renv []) (c_0::p::l); (* compute the recarg info for the arguments of each branch *) let case_spec = branches_specif renv (lazy_subterm_specif renv [] c_0) ci in let stack' = push_stack_closures renv l stack in let stack' = filter_stack_domain renv.env p stack' in lrest |> Array.iteri (fun k br' -> let stack_br = push_stack_args case_spec.(k) stack' in check_rec_call renv stack_br br') with (FixGuardError _ as exn) -> let exn = CErrors.push exn in (* we try hard to reduce the match away by looking for a constructor in c_0 (we unfold definitions too) *) let c_0 = whd_all renv.env c_0 in let hd, _ = decompose_app c_0 in match kind hd with | Construct _ -> (* the call to whd_betaiotazeta will reduce the apparent iota redex away *) check_rec_call renv [] (Term.applist (mkCase (ci,p,c_0,lrest), l)) | _ -> Exninfo.iraise exn end (* Enables to traverse Fixpoint definitions in a more intelligent way, ie, the rule : if - g = fix g (y1:T1)...(yp:Tp) {struct yp} := e & - f is guarded with respect to the set of pattern variables S in a1 ... am & - f is guarded with respect to the set of pattern variables S in T1 ... Tp & - ap is a sub-term of the formal argument of f & - f is guarded with respect to the set of pattern variables S+{yp} in e then f is guarded with respect to S in (g a1 ... am). Eduardo 7/9/98 *) | Fix ((recindxs,i),(_,typarray,bodies as recdef)) -> let decrArg = recindxs.(i) in begin try List.iter (check_rec_call renv []) l; Array.iter (check_rec_call renv []) typarray; let renv' = push_fix_renv renv recdef in let stack' = push_stack_closures renv l stack in bodies |> Array.iteri (fun j body -> if Int.equal i j && (List.length stack' > decrArg) then let recArg = List.nth stack' decrArg in let arg_sp = stack_element_specif recArg in let illformed () = error_ill_formed_rec_body renv.env NotEnoughAbstractionInFixBody (pi1 recdef) i (push_rec_types recdef renv.env) (judgment_of_fixpoint recdef) in check_nested_fix_body illformed renv' (decrArg+1) arg_sp body else check_rec_call renv' [] body) with (FixGuardError _ as exn) -> let exn = CErrors.push exn in (* we try hard to reduce the fix away by looking for a constructor in l[decrArg] (we unfold definitions too) *) if List.length l <= decrArg then Exninfo.iraise exn; let recArg = List.nth l decrArg in let recArg = whd_all renv.env recArg in let hd, _ = decompose_app recArg in match kind hd with | Construct _ -> let before, after = CList.(firstn decrArg l, skipn (decrArg+1) l) in check_rec_call renv [] (Term.applist (mkFix ((recindxs,i),recdef), (before @ recArg :: after))) | _ -> Exninfo.iraise exn end | Const (kn,_u as cu) -> if evaluable_constant kn renv.env then try List.iter (check_rec_call renv []) l with (FixGuardError _ ) -> let value = (Term.applist(constant_value_in renv.env cu, l)) in check_rec_call renv stack value else List.iter (check_rec_call renv []) l | Lambda (x,a,b) -> let () = assert (List.is_empty l) in check_rec_call renv [] a ; let spec, stack' = extract_stack stack in check_rec_call (push_var renv (x,a,spec)) stack' b | Prod (x,a,b) -> let () = assert (List.is_empty l && List.is_empty stack) in check_rec_call renv [] a; check_rec_call (push_var_renv renv (x,a)) [] b | CoFix (_i,(_,typarray,bodies as recdef)) -> List.iter (check_rec_call renv []) l; Array.iter (check_rec_call renv []) typarray; let renv' = push_fix_renv renv recdef in Array.iter (check_rec_call renv' []) bodies | (Ind _ | Construct _) -> List.iter (check_rec_call renv []) l | Proj (p, c) -> begin try List.iter (check_rec_call renv []) l; check_rec_call renv [] c with (FixGuardError _ as exn) -> let exn = CErrors.push exn in (* we try hard to reduce the proj away by looking for a constructor in c (we unfold definitions too) *) let c = whd_all renv.env c in let hd, _ = decompose_app c in match kind hd with | Construct _ -> check_rec_call renv [] (Term.applist (mkProj(Projection.unfold p,c), l)) | _ -> Exninfo.iraise exn end | Var id -> begin let open! Context.Named.Declaration in match lookup_named id renv.env with | LocalAssum _ -> List.iter (check_rec_call renv []) l | LocalDef (_,c,_) -> try List.iter (check_rec_call renv []) l with (FixGuardError _) -> check_rec_call renv stack (Term.applist(c,l)) end | Sort _ | Int _ | Float _ -> assert (List.is_empty l) (* l is not checked because it is considered as the meta's context *) | (Evar _ | Meta _) -> () | (App _ | LetIn _ | Cast _) -> assert false (* beta zeta reduction *) and check_nested_fix_body illformed renv decr recArgsDecrArg body = if Int.equal decr 0 then check_rec_call (assign_var_spec renv (1,recArgsDecrArg)) [] body else match kind (whd_all renv.env body) with | Lambda (x,a,b) -> check_rec_call renv [] a; let renv' = push_var_renv renv (x,a) in check_nested_fix_body illformed renv' (decr-1) recArgsDecrArg b | _ -> illformed () in check_rec_call renv [] def let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) = let nbfix = Array.length bodies in if Int.equal nbfix 0 || not (Int.equal (Array.length nvect) nbfix) || not (Int.equal (Array.length types) nbfix) || not (Int.equal (Array.length names) nbfix) || bodynum < 0 || bodynum >= nbfix then anomaly (Pp.str "Ill-formed fix term."); let fixenv = push_rec_types recdef env in let vdefj = judgment_of_fixpoint recdef in let raise_err env i err = error_ill_formed_rec_body env err names i fixenv vdefj in (* Check the i-th definition with recarg k *) let find_ind i k def = (* check fi does not appear in the k+1 first abstractions, gives the type of the k+1-eme abstraction (must be an inductive) *) let rec check_occur env n def = match kind (whd_all env def) with | Lambda (x,a,b) -> if noccur_with_meta n nbfix a then let env' = push_rel (LocalAssum (x,a)) env in if Int.equal n (k + 1) then (* get the inductive type of the fixpoint *) let (mind, _) = try find_inductive env a with Not_found -> raise_err env i (RecursionNotOnInductiveType a) in let mib,_ = lookup_mind_specif env (out_punivs mind) in if mib.mind_finite != Finite then raise_err env i (RecursionNotOnInductiveType a); (mind, (env', b)) else check_occur env' (n+1) b else anomaly ~label:"check_one_fix" (Pp.str "Bad occurrence of recursive call.") | _ -> raise_err env i NotEnoughAbstractionInFixBody in let ((ind, _), _) as res = check_occur fixenv 1 def in let _, mip = lookup_mind_specif env ind in (* recursive sprop means non record with projections -> squashed *) if mip.mind_relevance == Sorts.Irrelevant && not (Environ.is_type_in_type env (GlobRef.IndRef ind)) then begin if names.(i).Context.binder_relevance == Sorts.Relevant then raise_err env i FixpointOnIrrelevantInductive end; res in (* Do it on every fixpoint *) let rv = Array.map2_i find_ind nvect bodies in (Array.map fst rv, Array.map snd rv) let check_fix env ((nvect,_),(names,_,bodies as recdef) as fix) = let flags = Environ.typing_flags env in if flags.check_guarded then let (minds, rdef) = inductive_of_mutfix env fix in let get_tree (kn,i) = let mib = Environ.lookup_mind kn env in mib.mind_packets.(i).mind_recargs in let trees = Array.map (fun (mind,_) -> get_tree mind) minds in for i = 0 to Array.length bodies - 1 do let (fenv,body) = rdef.(i) in let renv = make_renv fenv nvect.(i) trees.(i) in try check_one_fix renv nvect trees body with FixGuardError (fixenv,err) -> error_ill_formed_rec_body fixenv err names i (push_rec_types recdef env) (judgment_of_fixpoint recdef) done else () (* let cfkey = CProfile.declare_profile "check_fix";; let check_fix env fix = CProfile.profile3 cfkey check_fix env fix;; *) (************************************************************************) (* Co-fixpoints. *) exception CoFixGuardError of env * guard_error let anomaly_ill_typed () = anomaly ~label:"check_one_cofix" (Pp.str "too many arguments applied to constructor.") let rec codomain_is_coind env c = let b = whd_all env c in match kind b with | Prod (x,a,b) -> codomain_is_coind (push_rel (LocalAssum (x,a)) env) b | _ -> (try find_coinductive env b with Not_found -> raise (CoFixGuardError (env, CodomainNotInductiveType b))) let check_one_cofix env nbfix def deftype = let rec check_rec_call env alreadygrd n tree vlra t = if not (noccur_with_meta n nbfix t) then let c,args = decompose_app (whd_all env t) in match kind c with | Rel p when n <= p && p < n+nbfix -> (* recursive call: must be guarded and no nested recursive call allowed *) if not alreadygrd then raise (CoFixGuardError (env,UnguardedRecursiveCall t)) else if not(List.for_all (noccur_with_meta n nbfix) args) then raise (CoFixGuardError (env,NestedRecursiveOccurrences)) | Construct ((_,i as cstr_kn),_u) -> let lra = vlra.(i-1) in let mI = inductive_of_constructor cstr_kn in let (mib,_mip) = lookup_mind_specif env mI in let realargs = List.skipn mib.mind_nparams args in let rec process_args_of_constr = function | (t::lr), (rar::lrar) -> if eq_wf_paths rar mk_norec then if noccur_with_meta n nbfix t then process_args_of_constr (lr, lrar) else raise (CoFixGuardError (env,RecCallInNonRecArgOfConstructor t)) else begin check_rec_call env true n rar (dest_subterms rar) t; process_args_of_constr (lr, lrar) end | [],_ -> () | _ -> anomaly_ill_typed () in process_args_of_constr (realargs, lra) | Lambda (x,a,b) -> let () = assert (List.is_empty args) in if noccur_with_meta n nbfix a then let env' = push_rel (LocalAssum (x,a)) env in check_rec_call env' alreadygrd (n+1) tree vlra b else raise (CoFixGuardError (env,RecCallInTypeOfAbstraction a)) | CoFix (_j,(_,varit,vdefs as recdef)) -> if List.for_all (noccur_with_meta n nbfix) args then if Array.for_all (noccur_with_meta n nbfix) varit then let nbfix = Array.length vdefs in let env' = push_rec_types recdef env in (Array.iter (check_rec_call env' alreadygrd (n+nbfix) tree vlra) vdefs; List.iter (check_rec_call env alreadygrd n tree vlra) args) else raise (CoFixGuardError (env,RecCallInTypeOfDef c)) else raise (CoFixGuardError (env,UnguardedRecursiveCall c)) | Case (_,p,tm,vrest) -> begin let tree = match restrict_spec env (Subterm (Strict, tree)) p with | Dead_code -> assert false | Subterm (_, tree') -> tree' | _ -> raise (CoFixGuardError (env, ReturnPredicateNotCoInductive c)) in if (noccur_with_meta n nbfix p) then if (noccur_with_meta n nbfix tm) then if (List.for_all (noccur_with_meta n nbfix) args) then let vlra = dest_subterms tree in Array.iter (check_rec_call env alreadygrd n tree vlra) vrest else raise (CoFixGuardError (env,RecCallInCaseFun c)) else raise (CoFixGuardError (env,RecCallInCaseArg c)) else raise (CoFixGuardError (env,RecCallInCasePred c)) end | Meta _ -> () | Evar _ -> List.iter (check_rec_call env alreadygrd n tree vlra) args | Rel _ | Var _ | Sort _ | Cast _ | Prod _ | LetIn _ | App _ | Const _ | Ind _ | Fix _ | Proj _ | Int _ | Float _ -> raise (CoFixGuardError (env,NotGuardedForm t)) in let ((mind, _),_) = codomain_is_coind env deftype in let vlra = lookup_subterms env mind in check_rec_call env false 1 vlra (dest_subterms vlra) def (* The function which checks that the whole block of definitions satisfies the guarded condition *) let check_cofix env (_bodynum,(names,types,bodies as recdef)) = let flags = Environ.typing_flags env in if flags.check_guarded then let nbfix = Array.length bodies in for i = 0 to nbfix-1 do let fixenv = push_rec_types recdef env in try check_one_cofix fixenv nbfix bodies.(i) types.(i) with CoFixGuardError (errenv,err) -> error_ill_formed_rec_body errenv err names i fixenv (judgment_of_fixpoint recdef) done else () coq-8.11.0/kernel/.merlin.in0000644000175000017500000000015513612664533015443 0ustar treinentreinenFLG -rectypes -thread -safe-string -w +a-4-44 S ../clib B ../clib S ../config B ../config S ../lib B ../lib coq-8.11.0/kernel/indTyping.mli0000644000175000017500000000331313612664533016216 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* mutual_inductive_entry -> env * universes * Univ.Variance.t array option * Names.Id.t array option option * Constr.rel_context * ((inductive_arity * Constr.types array) * (Constr.rel_context * (Constr.rel_context * Constr.types) array) * Sorts.family) array (* Utility function to compute the actual universe parameters of a template polymorphic inductive *) val template_polymorphic_univs : template_check:bool -> ctor_levels:Univ.LSet.t -> Univ.ContextSet.t -> Constr.rel_context -> Univ.Universe.t -> Univ.Level.t option list * Univ.LSet.t coq-8.11.0/kernel/safe_typing.mli0000644000175000017500000001736013612664533016570 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* required:vodigest -> bool (** {6 Safe environments } *) (** Since we are now able to type terms, we can define an abstract type of safe environments, where objects are typed before being added. We also provide functionality for modules : [start_module], [end_module], etc. *) type safe_environment type section_data val empty_environment : safe_environment val is_initial : safe_environment -> bool val env_of_safe_env : safe_environment -> Environ.env val sections_of_safe_env : safe_environment -> section_data Section.t (** The safe_environment state monad *) type safe_transformer0 = safe_environment -> safe_environment type 'a safe_transformer = safe_environment -> 'a * safe_environment (** {6 Stm machinery } *) type private_constants val empty_private_constants : private_constants val concat_private : private_constants -> private_constants -> private_constants (** [concat_private e1 e2] adds the constants of [e1] to [e2], i.e. constants in [e1] must be more recent than those of [e2]. *) val mk_pure_proof : Constr.constr -> private_constants Entries.proof_output val inline_private_constants : Environ.env -> private_constants Entries.proof_output -> Constr.constr Univ.in_universe_context_set val push_private_constants : Environ.env -> private_constants -> Environ.env (** Push the constants in the environment if not already there. *) val universes_of_private : private_constants -> Univ.ContextSet.t val is_curmod_library : safe_environment -> bool (* safe_environment has functional data affected by lazy computations, * thus this function returns a new safe_environment *) val join_safe_environment : ?except:Future.UUIDSet.t -> safe_environment -> safe_environment val is_joined_environment : safe_environment -> bool (** {6 Enriching a safe environment } *) (** Insertion of global axioms or definitions *) type global_declaration = | ConstantEntry : Entries.constant_entry -> global_declaration | OpaqueEntry : private_constants Entries.const_entry_body Entries.opaque_entry -> global_declaration type side_effect_declaration = | DefinitionEff : Entries.definition_entry -> side_effect_declaration | OpaqueEff : Constr.constr Entries.opaque_entry -> side_effect_declaration type exported_private_constant = Constant.t val export_private_constants : private_constants Entries.proof_output -> (Constr.constr Univ.in_universe_context_set * exported_private_constant list) safe_transformer (** returns the main constant *) val add_constant : Label.t -> global_declaration -> Constant.t safe_transformer (** Similar to add_constant but also returns a certificate *) val add_private_constant : Label.t -> side_effect_declaration -> (Constant.t * private_constants) safe_transformer (** Adding an inductive type *) val add_mind : Label.t -> Entries.mutual_inductive_entry -> MutInd.t safe_transformer (** Adding a module or a module type *) val add_module : Label.t -> Entries.module_entry -> Declarations.inline -> (ModPath.t * Mod_subst.delta_resolver) safe_transformer val add_modtype : Label.t -> Entries.module_type_entry -> Declarations.inline -> ModPath.t safe_transformer (** Adding universe constraints *) val push_context_set : bool -> Univ.ContextSet.t -> safe_transformer0 val add_constraints : Univ.Constraint.t -> safe_transformer0 (* (\** Generator of universes *\) *) (* val next_universe : int safe_transformer *) (** Setting the type theory flavor *) val set_engagement : Declarations.engagement -> safe_transformer0 val set_indices_matter : bool -> safe_transformer0 val set_typing_flags : Declarations.typing_flags -> safe_transformer0 val set_share_reduction : bool -> safe_transformer0 val set_check_guarded : bool -> safe_transformer0 val set_check_positive : bool -> safe_transformer0 val set_check_universes : bool -> safe_transformer0 val set_VM : bool -> safe_transformer0 val set_native_compiler : bool -> safe_transformer0 val make_sprop_cumulative : safe_transformer0 val set_allow_sprop : bool -> safe_transformer0 val check_engagement : Environ.env -> Declarations.set_predicativity -> unit (** {6 Interactive section functions } *) val open_section : safe_transformer0 val close_section : safe_transformer0 val sections_are_opened : safe_environment -> bool (** Insertion of local declarations (Local or Variables) *) val push_named_assum : (Id.t * Constr.types) -> safe_transformer0 val push_named_def : Id.t * Entries.section_def_entry -> safe_transformer0 (** Add local universes to a polymorphic section *) val push_section_context : (Name.t array * Univ.UContext.t) -> safe_transformer0 (** {6 Interactive module functions } *) val start_module : Label.t -> ModPath.t safe_transformer val start_modtype : Label.t -> ModPath.t safe_transformer val add_module_parameter : MBId.t -> Entries.module_struct_entry -> Declarations.inline -> Mod_subst.delta_resolver safe_transformer (** Traditional mode: check at end of module that no future was created. *) val allow_delayed_constants : bool ref (** The optional result type is given without its functorial part *) val end_module : Label.t -> (Entries.module_struct_entry * Declarations.inline) option -> (ModPath.t * MBId.t list * Mod_subst.delta_resolver) safe_transformer val end_modtype : Label.t -> (ModPath.t * MBId.t list) safe_transformer val add_include : Entries.module_struct_entry -> bool -> Declarations.inline -> Mod_subst.delta_resolver safe_transformer val current_modpath : safe_environment -> ModPath.t val current_dirpath : safe_environment -> DirPath.t (** {6 Libraries : loading and saving compilation units } *) type compiled_library type native_library = Nativecode.global list val module_of_library : compiled_library -> Declarations.module_body val start_library : DirPath.t -> ModPath.t safe_transformer val export : ?except:Future.UUIDSet.t -> output_native_objects:bool -> safe_environment -> DirPath.t -> ModPath.t * compiled_library * native_library (* Constraints are non empty iff the file is a vi2vo *) val import : compiled_library -> Univ.ContextSet.t -> vodigest -> ModPath.t safe_transformer (** {6 Safe typing judgments } *) type judgment val j_val : judgment -> Constr.constr val j_type : judgment -> Constr.constr (** The safe typing of a term returns a typing judgment. *) val typing : safe_environment -> Constr.constr -> judgment (** {6 Queries } *) val exists_objlabel : Label.t -> safe_environment -> bool val delta_of_senv : safe_environment -> Mod_subst.delta_resolver * Mod_subst.delta_resolver val constant_of_delta_kn_senv : safe_environment -> KerName.t -> Constant.t val mind_of_delta_kn_senv : safe_environment -> KerName.t -> MutInd.t (** {6 Retroknowledge / Native compiler } *) val register_inline : Constant.t -> safe_transformer0 val register_inductive : inductive -> 'a CPrimitives.prim_ind -> safe_transformer0 val set_strategy : Names.Constant.t Names.tableKey -> Conv_oracle.level -> safe_transformer0 coq-8.11.0/kernel/clambda.mli0000644000175000017500000000300713612664533015634 0ustar treinentreinenopen Names open Constr open Vmvalues open Environ type lambda = | Lrel of Name.t * int | Lvar of Id.t | Levar of Evar.t * lambda array | Lprod of lambda * lambda | Llam of Name.t Context.binder_annot array * lambda | Llet of Name.t Context.binder_annot * lambda * lambda | Lapp of lambda * lambda array | Lconst of pconstant | Lprim of pconstant option * CPrimitives.t * lambda array (* No check if None *) | Lcase of case_info * reloc_table * lambda * lambda * lam_branches | Lif of lambda * lambda * lambda | Lfix of (int array * int) * fix_decl | Lcofix of int * fix_decl | Lint of int | Lmakeblock of int * lambda array | Luint of Uint63.t | Lfloat of Float64.t | Lval of structured_values | Lsort of Sorts.t | Lind of pinductive | Lproj of Projection.Repr.t * lambda and lam_branches = { constant_branches : lambda array; nonconstant_branches : (Name.t Context.binder_annot array * lambda) array } and fix_decl = Name.t Context.binder_annot array * lambda array * lambda array exception TooLargeInductive of Pp.t val lambda_of_constr : optimize:bool -> env -> Constr.t -> lambda val decompose_Llam : lambda -> Name.t Context.binder_annot array * lambda val get_alias : env -> Constant.t -> Constant.t (** Dump the VM lambda code after compilation (for debugging purposes) *) val dump_lambda : bool ref coq-8.11.0/kernel/cemitcodes.ml0000644000175000017500000004154713612664533016232 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Vmvalues.tcode = "coq_tcode_of_code" (* Relocation information *) type reloc_info = | Reloc_annot of annot_switch | Reloc_const of structured_constant | Reloc_getglobal of Names.Constant.t | Reloc_proj_name of Projection.Repr.t let eq_reloc_info r1 r2 = match r1, r2 with | Reloc_annot sw1, Reloc_annot sw2 -> eq_annot_switch sw1 sw2 | Reloc_annot _, _ -> false | Reloc_const c1, Reloc_const c2 -> eq_structured_constant c1 c2 | Reloc_const _, _ -> false | Reloc_getglobal c1, Reloc_getglobal c2 -> Constant.equal c1 c2 | Reloc_getglobal _, _ -> false | Reloc_proj_name p1, Reloc_proj_name p2 -> Projection.Repr.equal p1 p2 | Reloc_proj_name _, _ -> false let hash_reloc_info r = let open Hashset.Combine in match r with | Reloc_annot sw -> combinesmall 1 (hash_annot_switch sw) | Reloc_const c -> combinesmall 2 (hash_structured_constant c) | Reloc_getglobal c -> combinesmall 3 (Constant.hash c) | Reloc_proj_name p -> combinesmall 4 (Projection.Repr.hash p) module RelocTable = Hashtbl.Make(struct type t = reloc_info let equal = eq_reloc_info let hash = hash_reloc_info end) (** We use arrays for on-disk representation. On 32-bit machines, this means we can only have a maximum amount of about 4.10^6 relocations, which seems quite a lot, but potentially reachable if e.g. compiling big proofs. This would prevent VM computing with these terms on 32-bit architectures. Maybe we should use a more robust data structure? *) type patches = { reloc_infos : (reloc_info * int array) array; } let patch_char4 buff pos c1 c2 c3 c4 = Bytes.unsafe_set buff pos c1; Bytes.unsafe_set buff (pos + 1) c2; Bytes.unsafe_set buff (pos + 2) c3; Bytes.unsafe_set buff (pos + 3) c4 let patch1 buff pos n = patch_char4 buff pos (Char.unsafe_chr n) (Char.unsafe_chr (n asr 8)) (Char.unsafe_chr (n asr 16)) (Char.unsafe_chr (n asr 24)) let patch_int buff reloc = (* copy code *before* patching because of nested evaluations: the code we are patching might be called (and thus "concurrently" patched) and results in wrong results. Side-effects... *) let buff = Bytes.of_string buff in let iter (reloc, npos) = Array.iter (fun pos -> patch1 buff pos reloc) npos in let () = CArray.iter iter reloc in buff let patch buff pl f = (** Order seems important here? *) let reloc = CArray.map (fun (r, pos) -> (f r, pos)) pl.reloc_infos in let buff = patch_int buff reloc in tcode_of_code buff (* Buffering of bytecode *) type label_definition = Label_defined of int | Label_undefined of (int * int) list type env = { mutable out_buffer : Bytes.t; mutable out_position : int; mutable label_table : label_definition array; (* le ieme element de la table = Label_defined n signifie que l'on a deja rencontrer le label i et qu'il est a l'offset n. = Label_undefined l signifie que l'on a pas encore rencontrer ce label, le premier entier indique ou est l'entier a patcher dans la string, le deuxieme son origine *) reloc_info : int list RelocTable.t; } let out_word env b1 b2 b3 b4 = let p = env.out_position in if p >= Bytes.length env.out_buffer then begin let len = Bytes.length env.out_buffer in let new_len = if len <= Sys.max_string_length / 2 then 2 * len else if len = Sys.max_string_length then invalid_arg "String.create" (* Pas la bonne exception .... *) else Sys.max_string_length in let new_buffer = Bytes.create new_len in Bytes.blit env.out_buffer 0 new_buffer 0 len; env.out_buffer <- new_buffer end; patch_char4 env.out_buffer p (Char.unsafe_chr b1) (Char.unsafe_chr b2) (Char.unsafe_chr b3) (Char.unsafe_chr b4); env.out_position <- p + 4 let out env opcode = out_word env opcode 0 0 0 let is_immed i = Uint63.le (Uint63.of_int i) Uint63.maxuint31 let out_int env n = out_word env n (n asr 8) (n asr 16) (n asr 24) (* Handling of local labels and backpatching *) let extend_label_table env needed = let new_size = ref(Array.length env.label_table) in while needed >= !new_size do new_size := 2 * !new_size done; let new_table = Array.make !new_size (Label_undefined []) in Array.blit env.label_table 0 new_table 0 (Array.length env.label_table); env.label_table <- new_table let backpatch env (pos, orig) = let displ = (env.out_position - orig) asr 2 in Bytes.set env.out_buffer pos @@ Char.unsafe_chr displ; Bytes.set env.out_buffer (pos+1) @@ Char.unsafe_chr (displ asr 8); Bytes.set env.out_buffer (pos+2) @@ Char.unsafe_chr (displ asr 16); Bytes.set env.out_buffer (pos+3) @@ Char.unsafe_chr (displ asr 24) let define_label env lbl = if lbl >= Array.length env.label_table then extend_label_table env lbl; match (env.label_table).(lbl) with Label_defined _ -> raise(Failure "CEmitcode.define_label") | Label_undefined patchlist -> List.iter (fun p -> backpatch env p) patchlist; (env.label_table).(lbl) <- Label_defined env.out_position let out_label_with_orig env orig lbl = if lbl >= Array.length env.label_table then extend_label_table env lbl; match (env.label_table).(lbl) with Label_defined def -> out_int env ((def - orig) asr 2) | Label_undefined patchlist -> (* spiwack: patchlist is supposed to be non-empty all the time thus I commented that out. If there is no problem I suggest removing it for next release (cur: 8.1) *) (*if patchlist = [] then *) (env.label_table).(lbl) <- Label_undefined((env.out_position, orig) :: patchlist); out_int env 0 let out_label env l = out_label_with_orig env env.out_position l (* Relocation information *) let enter env info = let pos = env.out_position in let old = try RelocTable.find env.reloc_info info with Not_found -> [] in RelocTable.replace env.reloc_info info (pos :: old) let slot_for_const env c = enter env (Reloc_const c); out_int env 0 let slot_for_annot env a = enter env (Reloc_annot a); out_int env 0 let slot_for_getglobal env p = enter env (Reloc_getglobal p); out_int env 0 let slot_for_proj_name env p = enter env (Reloc_proj_name p); out_int env 0 (* Emission of one instruction *) let nocheck_prim_op = function | Int63add -> opADDINT63 | Int63sub -> opSUBINT63 | Int63lt -> opLTINT63 | Int63le -> opLEINT63 | _ -> assert false let check_prim_op = function | Int63head0 -> opCHECKHEAD0INT63 | Int63tail0 -> opCHECKTAIL0INT63 | Int63add -> opCHECKADDINT63 | Int63sub -> opCHECKSUBINT63 | Int63mul -> opCHECKMULINT63 | Int63div -> opCHECKDIVINT63 | Int63mod -> opCHECKMODINT63 | Int63lsr -> opCHECKLSRINT63 | Int63lsl -> opCHECKLSLINT63 | Int63land -> opCHECKLANDINT63 | Int63lor -> opCHECKLORINT63 | Int63lxor -> opCHECKLXORINT63 | Int63addc -> opCHECKADDCINT63 | Int63subc -> opCHECKSUBCINT63 | Int63addCarryC -> opCHECKADDCARRYCINT63 | Int63subCarryC -> opCHECKSUBCARRYCINT63 | Int63mulc -> opCHECKMULCINT63 | Int63diveucl -> opCHECKDIVEUCLINT63 | Int63div21 -> opCHECKDIV21INT63 | Int63addMulDiv -> opCHECKADDMULDIVINT63 | Int63eq -> opCHECKEQINT63 | Int63lt -> opCHECKLTINT63 | Int63le -> opCHECKLEINT63 | Int63compare -> opCHECKCOMPAREINT63 | Float64opp -> opCHECKOPPFLOAT | Float64abs -> opCHECKABSFLOAT | Float64eq -> opCHECKEQFLOAT | Float64lt -> opCHECKLTFLOAT | Float64le -> opCHECKLEFLOAT | Float64compare -> opCHECKCOMPAREFLOAT | Float64classify -> opCHECKCLASSIFYFLOAT | Float64add -> opCHECKADDFLOAT | Float64sub -> opCHECKSUBFLOAT | Float64mul -> opCHECKMULFLOAT | Float64div -> opCHECKDIVFLOAT | Float64sqrt -> opCHECKSQRTFLOAT | Float64ofInt63 -> opCHECKFLOATOFINT63 | Float64normfr_mantissa -> opCHECKFLOATNORMFRMANTISSA | Float64frshiftexp -> opCHECKFRSHIFTEXP | Float64ldshiftexp -> opCHECKLDSHIFTEXP | Float64next_up -> opCHECKNEXTUPFLOAT | Float64next_down -> opCHECKNEXTDOWNFLOAT let emit_instr env = function | Klabel lbl -> define_label env lbl | Kacc n -> if n < 8 then out env(opACC0 + n) else (out env opACC; out_int env n) | Kenvacc n -> if n >= 1 && n <= 4 then out env(opENVACC1 + n - 1) else (out env opENVACC; out_int env n) | Koffsetclosure ofs -> if Int.equal ofs (-2) || Int.equal ofs 0 || Int.equal ofs 2 then out env (opOFFSETCLOSURE0 + ofs / 2) else (out env opOFFSETCLOSURE; out_int env ofs) | Kpush -> out env opPUSH | Kpop n -> out env opPOP; out_int env n | Kpush_retaddr lbl -> out env opPUSH_RETADDR; out_label env lbl | Kapply n -> if n <= 4 then out env(opAPPLY1 + n - 1) else (out env opAPPLY; out_int env n) | Kappterm(n, sz) -> if n < 4 then (out env(opAPPTERM1 + n - 1); out_int env sz) else (out env opAPPTERM; out_int env n; out_int env sz) | Kreturn n -> out env opRETURN; out_int env n | Kjump -> out env opRETURN; out_int env 0 | Krestart -> out env opRESTART | Kgrab n -> out env opGRAB; out_int env n | Kgrabrec(rec_arg) -> out env opGRABREC; out_int env rec_arg | Kclosure(lbl, n) -> out env opCLOSURE; out_int env n; out_label env lbl | Kclosurerec(nfv,init,lbl_types,lbl_bodies) -> out env opCLOSUREREC;out_int env (Array.length lbl_bodies); out_int env nfv; out_int env init; let org = env.out_position in Array.iter (out_label_with_orig env org) lbl_types; let org = env.out_position in Array.iter (out_label_with_orig env org) lbl_bodies | Kclosurecofix(nfv,init,lbl_types,lbl_bodies) -> out env opCLOSURECOFIX;out_int env (Array.length lbl_bodies); out_int env nfv; out_int env init; let org = env.out_position in Array.iter (out_label_with_orig env org) lbl_types; let org = env.out_position in Array.iter (out_label_with_orig env org) lbl_bodies | Kgetglobal q -> out env opGETGLOBAL; slot_for_getglobal env q | Kconst (Const_b0 i) when is_immed i -> if i >= 0 && i <= 3 then out env (opCONST0 + i) else (out env opCONSTINT; out_int env i) | Kconst c -> out env opGETGLOBAL; slot_for_const env c | Kmakeblock(n, t) -> if Int.equal n 0 then invalid_arg "emit_instr : block size = 0" else if n < 4 then (out env(opMAKEBLOCK1 + n - 1); out_int env t) else (out env opMAKEBLOCK; out_int env n; out_int env t) | Kmakeprod -> out env opMAKEPROD | Kmakeswitchblock(typlbl,swlbl,annot,sz) -> out env opMAKESWITCHBLOCK; out_label env typlbl; out_label env swlbl; slot_for_annot env annot;out_int env sz | Kswitch (tbl_const, tbl_block) -> let lenb = Array.length tbl_block in let lenc = Array.length tbl_const in assert (lenb < 0x100 && lenc < 0x1000000); out env opSWITCH; out_word env lenc (lenc asr 8) (lenc asr 16) (lenb); (* out_int env (Array.length tbl_const + (Array.length tbl_block lsl 23)); *) let org = env.out_position in Array.iter (out_label_with_orig env org) tbl_const; Array.iter (out_label_with_orig env org) tbl_block | Kpushfields n -> out env opPUSHFIELDS;out_int env n | Kfield n -> if n <= 1 then out env (opGETFIELD0+n) else (out env opGETFIELD;out_int env n) | Ksetfield n -> if n <= 1 then out env (opSETFIELD0+n) else (out env opSETFIELD;out_int env n) | Ksequence _ -> invalid_arg "Cemitcodes.emit_instr" | Kproj p -> out env opPROJ; out_int env (Projection.Repr.arg p); slot_for_proj_name env p | Kensurestackcapacity size -> out env opENSURESTACKCAPACITY; out_int env size | Kbranch lbl -> out env opBRANCH; out_label env lbl | Kprim (op,None) -> out env (nocheck_prim_op op) | Kprim(op,Some (q,_u)) -> out env (check_prim_op op); slot_for_getglobal env q | Kareint 1 -> out env opISINT | Kareint 2 -> out env opAREINT2; | Kstop -> out env opSTOP | Kareint _ -> assert false (* Emission of a current list and remaining lists of instructions. Include some peephole optimization. *) let rec emit env insns remaining = match insns with | [] -> (match remaining with [] -> () | (first::rest) -> emit env first rest) (* Peephole optimizations *) | Kpush :: Kacc n :: c -> if n < 8 then out env(opPUSHACC0 + n) else (out env opPUSHACC; out_int env n); emit env c remaining | Kpush :: Kenvacc n :: c -> if n >= 1 && n <= 4 then out env(opPUSHENVACC1 + n - 1) else (out env opPUSHENVACC; out_int env n); emit env c remaining | Kpush :: Koffsetclosure ofs :: c -> if Int.equal ofs (-2) || Int.equal ofs 0 || Int.equal ofs 2 then out env(opPUSHOFFSETCLOSURE0 + ofs / 2) else (out env opPUSHOFFSETCLOSURE; out_int env ofs); emit env c remaining | Kpush :: Kgetglobal id :: c -> out env opPUSHGETGLOBAL; slot_for_getglobal env id; emit env c remaining | Kpush :: Kconst (Const_b0 i) :: c when is_immed i -> if i >= 0 && i <= 3 then out env (opPUSHCONST0 + i) else (out env opPUSHCONSTINT; out_int env i); emit env c remaining | Kpush :: Kconst const :: c -> out env opPUSHGETGLOBAL; slot_for_const env const; emit env c remaining | Kpop n :: Kjump :: c -> out env opRETURN; out_int env n; emit env c remaining | Ksequence(c1,c2)::c -> emit env c1 (c2::c::remaining) (* Default case *) | instr :: c -> emit_instr env instr; emit env c remaining (* Initialization *) type to_patch = emitcodes * patches * fv (* Substitution *) let subst_strcst s sc = match sc with | Const_sort _ | Const_b0 _ | Const_univ_level _ | Const_val _ | Const_uint _ | Const_float _ -> sc | Const_ind ind -> let kn,i = ind in Const_ind (subst_mind s kn, i) let subst_reloc s ri = match ri with | Reloc_annot a -> let (kn,i) = a.ci.ci_ind in let ci = {a.ci with ci_ind = (subst_mind s kn,i)} in Reloc_annot {a with ci = ci} | Reloc_const sc -> Reloc_const (subst_strcst s sc) | Reloc_getglobal kn -> Reloc_getglobal (subst_constant s kn) | Reloc_proj_name p -> Reloc_proj_name (subst_proj_repr s p) let subst_patches subst p = let infos = CArray.map (fun (r, pos) -> (subst_reloc subst r, pos)) p.reloc_infos in { reloc_infos = infos; } let subst_to_patch s (code,pl,fv) = code, subst_patches s pl, fv type body_code = | BCdefined of to_patch | BCalias of Names.Constant.t | BCconstant type to_patch_substituted = | PBCdefined of to_patch substituted | PBCalias of Names.Constant.t substituted | PBCconstant let from_val = function | BCdefined tp -> PBCdefined (from_val tp) | BCalias cu -> PBCalias (from_val cu) | BCconstant -> PBCconstant let force = function | PBCdefined tp -> BCdefined (force subst_to_patch tp) | PBCalias cu -> BCalias (force subst_constant cu) | PBCconstant -> BCconstant let subst_to_patch_subst s = function | PBCdefined tp -> PBCdefined (subst_substituted s tp) | PBCalias cu -> PBCalias (subst_substituted s cu) | PBCconstant -> PBCconstant let repr_body_code = function | PBCdefined tp -> let (s, tp) = repr_substituted tp in (s, BCdefined tp) | PBCalias cu -> let (s, cu) = repr_substituted cu in (s, BCalias cu) | PBCconstant -> (None, BCconstant) let to_memory (init_code, fun_code, fv) = let env = { out_buffer = Bytes.create 1024; out_position = 0; label_table = Array.make 16 (Label_undefined []); reloc_info = RelocTable.create 91; } in emit env init_code []; emit env fun_code []; (** Later uses of this string are all purely functional *) let code = Bytes.sub_string env.out_buffer 0 env.out_position in let code = CString.hcons code in let fold reloc npos accu = (reloc, Array.of_list npos) :: accu in let reloc = RelocTable.fold fold env.reloc_info [] in let reloc = { reloc_infos = CArray.of_list reloc } in Array.iter (fun lbl -> (match lbl with Label_defined _ -> assert true | Label_undefined patchlist -> assert (patchlist = []))) env.label_table; (code, reloc, fv) coq-8.11.0/kernel/float64.ml0000644000175000017500000001166513612664533015370 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* f let is_infinity f = f = infinity let is_neg_infinity f = f = neg_infinity (* Printing a binary64 float in 17 decimal places and parsing it again will yield the same float. We assume [to_string_raw] is not given a [nan] as input. *) let to_string_raw f = Printf.sprintf "%.17g" f (* OCaml gives a sign to nan values which should not be displayed as all NaNs are considered equal here *) let to_string f = if is_nan f then "nan" else to_string_raw f let of_string = float_of_string (* Compiles a float to OCaml code *) let compile f = let s = if is_nan f then "nan" else if is_neg_infinity f then "neg_infinity" else Printf.sprintf "%h" f in Printf.sprintf "Float64.of_float (%s)" s let of_float f = f let sign f = copysign 1. f < 0. let opp = ( ~-. ) let abs = abs_float type float_comparison = FEq | FLt | FGt | FNotComparable let eq x y = x = y [@@ocaml.inline always] let lt x y = x < y [@@ocaml.inline always] let le x y = x <= y [@@ocaml.inline always] (* inspired by lib/util.ml; see also #10471 *) let pervasives_compare = compare let compare x y = if x < y then FLt else ( if x > y then FGt else ( if x = y then FEq else FNotComparable (* NaN case *) ) ) [@@ocaml.inline always] type float_class = | PNormal | NNormal | PSubn | NSubn | PZero | NZero | PInf | NInf | NaN let classify x = match classify_float x with | FP_normal -> if 0. < x then PNormal else NNormal | FP_subnormal -> if 0. < x then PSubn else NSubn | FP_zero -> if 0. < 1. /. x then PZero else NZero | FP_infinite -> if 0. < x then PInf else NInf | FP_nan -> NaN [@@ocaml.inline always] external mul : float -> float -> float = "coq_fmul_byte" "coq_fmul" [@@unboxed] [@@noalloc] external add : float -> float -> float = "coq_fadd_byte" "coq_fadd" [@@unboxed] [@@noalloc] external sub : float -> float -> float = "coq_fsub_byte" "coq_fsub" [@@unboxed] [@@noalloc] external div : float -> float -> float = "coq_fdiv_byte" "coq_fdiv" [@@unboxed] [@@noalloc] external sqrt : float -> float = "coq_fsqrt_byte" "coq_fsqrt" [@@unboxed] [@@noalloc] let of_int63 x = Uint63.to_float x [@@ocaml.inline always] let prec = 53 let normfr_mantissa f = let f = abs f in if f >= 0.5 && f < 1. then Uint63.of_float (ldexp f prec) else Uint63.zero [@@ocaml.inline always] let eshift = 2101 (* 2*emax + prec *) (* When calling frexp on a nan or an infinity, the returned value inside the exponent is undefined. Therefore we must always set it to a fixed value (here 0). *) let frshiftexp f = match classify_float f with | FP_zero | FP_infinite | FP_nan -> (f, Uint63.zero) | FP_normal | FP_subnormal -> let (m, e) = frexp f in m, Uint63.of_int (e + eshift) [@@ocaml.inline always] let ldshiftexp f e = ldexp f (Uint63.to_int_min e (2 * eshift) - eshift) [@@ocaml.inline always] external next_up : float -> float = "coq_next_up_byte" "coq_next_up" [@@unboxed] [@@noalloc] external next_down : float -> float = "coq_next_down_byte" "coq_next_down" [@@unboxed] [@@noalloc] let equal f1 f2 = match classify_float f1 with | FP_normal | FP_subnormal | FP_infinite -> (f1 = f2) | FP_nan -> is_nan f2 | FP_zero -> f1 = f2 && 1. /. f1 = 1. /. f2 (* OCaml consider 0. = -0. *) [@@ocaml.inline always] let hash = (* Hashtbl.hash already considers all NaNs as equal, cf. https://github.com/ocaml/ocaml/commit/aea227fdebe0b5361fd3e1d0aaa42cf929052269 and http://caml.inria.fr/pub/docs/manual-ocaml/libref/Hashtbl.html *) Hashtbl.hash let total_compare f1 f2 = (* pervasives_compare considers all NaNs as equal, which is fine here, but also considers -0. and +0. as equal *) if f1 = 0. && f2 = 0. then pervasives_compare (1. /. f1) (1. /. f2) else pervasives_compare f1 f2 let is_float64 t = Obj.tag t = Obj.double_tag [@@ocaml.inline always] (*** Test at runtime that no harmful double rounding seems to be performed with an intermediate 80 bits representation (x87). *) let () = let b = ldexp 1. 53 in let s = add 1. (ldexp 1. (-52)) in if add b s <= b || add b 1. <> b || ldexp 1. (-1074) <= 0. then failwith "Detected non IEEE-754 compliant architecture (or wrong \ rounding mode). Use of Float is thus unsafe." coq-8.11.0/kernel/term_typing.ml0000644000175000017500000003324713612664533016452 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constr.t -> 'a -> (Constr.t * Univ.ContextSet.t * int) let skip_trusted_seff sl b e = let rec aux sl b e acc = let open Context.Rel.Declaration in if Int.equal sl 0 then b, e, acc else match kind b with | LetIn (n,c,ty,bo) -> aux (sl - 1) bo (Environ.push_rel (LocalDef (n,c,ty)) e) (`Let(n,c,ty)::acc) | App(hd,arg) -> begin match kind hd with | Lambda (n,ty,bo) -> aux (sl - 1) bo (Environ.push_rel (LocalAssum (n,ty)) e) (`Cut(n,ty,arg)::acc) | _ -> assert false end | _ -> assert false in aux sl b e [] let rec unzip ctx j = match ctx with | [] -> j | `Let (n,c,ty) :: ctx -> unzip ctx { j with uj_val = mkLetIn (n,c,ty,j.uj_val) } | `Cut (n,ty,arg) :: ctx -> unzip ctx { j with uj_val = mkApp (mkLambda (n,ty,j.uj_val),arg) } let feedback_completion_typecheck = Option.iter (fun state_id -> Feedback.feedback ~id:state_id Feedback.Complete) type typing_context = | MonoTyCtx of Environ.env * unsafe_type_judgment * Univ.ContextSet.t * Id.Set.t * Stateid.t option | PolyTyCtx of Environ.env * unsafe_type_judgment * Univ.universe_level_subst * Univ.AUContext.t * Id.Set.t * Stateid.t option let infer_declaration env (dcl : constant_entry) = match dcl with | ParameterEntry (ctx,(t,uctx),nl) -> let env = match uctx with | Monomorphic_entry uctx -> push_context_set ~strict:true uctx env | Polymorphic_entry (_, uctx) -> push_context ~strict:false uctx env in let j = Typeops.infer env t in let usubst, univs = Declareops.abstract_universes uctx in let r = Typeops.assumption_of_judgment env j in let t = Vars.subst_univs_level_constr usubst j.uj_val in { Cooking.cook_body = Undef nl; cook_type = t; cook_universes = univs; cook_relevance = r; cook_inline = false; cook_context = ctx; } (** Primitives cannot be universe polymorphic *) | PrimitiveEntry ({ prim_entry_type = otyp; prim_entry_univs = uctxt; prim_entry_content = op_t; }) -> let env = push_context_set ~strict:true uctxt env in let ty = match otyp with | Some typ -> let typ = Typeops.infer_type env typ in Typeops.check_primitive_type env op_t typ.utj_val; typ.utj_val | None -> match op_t with | CPrimitives.OT_op op -> Typeops.type_of_prim env op | CPrimitives.OT_type _ -> mkSet in let cd = match op_t with | CPrimitives.OT_op op -> Declarations.Primitive op | CPrimitives.OT_type _ -> Undef None in { Cooking.cook_body = cd; cook_type = ty; cook_universes = Monomorphic uctxt; cook_inline = false; cook_context = None; cook_relevance = Sorts.Relevant; } | DefinitionEntry c -> let { const_entry_type = typ; _ } = c in let { const_entry_body = body; const_entry_feedback = feedback_id; _ } = c in let env, usubst, univs = match c.const_entry_universes with | Monomorphic_entry ctx -> let env = push_context_set ~strict:true ctx env in env, Univ.empty_level_subst, Monomorphic ctx | Polymorphic_entry (nas, uctx) -> (** [ctx] must contain local universes, such that it has no impact on the rest of the graph (up to transitivity). *) let env = push_context ~strict:false uctx env in let sbst, auctx = Univ.abstract_universes nas uctx in let sbst = Univ.make_instance_subst sbst in env, sbst, Polymorphic auctx in let j = Typeops.infer env body in let typ = match typ with | None -> Vars.subst_univs_level_constr usubst j.uj_type | Some t -> let tj = Typeops.infer_type env t in let _ = Typeops.judge_of_cast env j DEFAULTcast tj in Vars.subst_univs_level_constr usubst tj.utj_val in let def = Vars.subst_univs_level_constr usubst j.uj_val in let def = Def (Mod_subst.from_val def) in feedback_completion_typecheck feedback_id; { Cooking.cook_body = def; cook_type = typ; cook_universes = univs; cook_relevance = Retypeops.relevance_of_term env j.uj_val; cook_inline = c.const_entry_inline_code; cook_context = c.const_entry_secctx; } (** Definition is opaque (Qed), so we delay the typing of its body. *) let infer_opaque env = function | ({ opaque_entry_type = typ; opaque_entry_universes = Monomorphic_entry univs; _ } as c) -> let env = push_context_set ~strict:true univs env in let { opaque_entry_feedback = feedback_id; _ } = c in let tyj = Typeops.infer_type env typ in let context = MonoTyCtx (env, tyj, univs, c.opaque_entry_secctx, feedback_id) in let def = OpaqueDef () in { Cooking.cook_body = def; cook_type = tyj.utj_val; cook_universes = Monomorphic univs; cook_relevance = Sorts.relevance_of_sort tyj.utj_type; cook_inline = false; cook_context = Some c.opaque_entry_secctx; }, context | ({ opaque_entry_type = typ; opaque_entry_universes = Polymorphic_entry (nas, uctx); _ } as c) -> let { opaque_entry_feedback = feedback_id; _ } = c in let env = push_context ~strict:false uctx env in let tj = Typeops.infer_type env typ in let sbst, auctx = Univ.abstract_universes nas uctx in let usubst = Univ.make_instance_subst sbst in let context = PolyTyCtx (env, tj, usubst, auctx, c.opaque_entry_secctx, feedback_id) in let def = OpaqueDef () in let typ = Vars.subst_univs_level_constr usubst tj.utj_val in { Cooking.cook_body = def; cook_type = typ; cook_universes = Polymorphic auctx; cook_relevance = Sorts.relevance_of_sort tj.utj_type; cook_inline = false; cook_context = Some c.opaque_entry_secctx; }, context let check_section_variables env declared_set typ body = let ids_typ = global_vars_set env typ in let ids_def = global_vars_set env body in let inferred_set = Environ.really_needed env (Id.Set.union ids_typ ids_def) in if not (Id.Set.subset inferred_set declared_set) then let l = Id.Set.elements (Id.Set.diff inferred_set declared_set) in let n = List.length l in let declared_vars = Pp.pr_sequence Id.print (Id.Set.elements declared_set) in let inferred_vars = Pp.pr_sequence Id.print (Id.Set.elements inferred_set) in let missing_vars = Pp.pr_sequence Id.print (List.rev l) in user_err Pp.(prlist str ["The following section "; (String.plural n "variable"); " "; (String.conjugate_verb_to_be n); " used but not declared:"] ++ fnl () ++ missing_vars ++ str "." ++ fnl () ++ fnl () ++ str "You can either update your proof to not depend on " ++ missing_vars ++ str ", or you can update your Proof line from" ++ fnl () ++ str "Proof using " ++ declared_vars ++ fnl () ++ str "to" ++ fnl () ++ str "Proof using " ++ inferred_vars) let build_constant_declaration env result = let open Cooking in let typ = result.cook_type in (* We try to postpone the computation of used section variables *) let hyps, def = let context_ids = List.map NamedDecl.get_id (named_context env) in let def = result.cook_body in match result.cook_context with | None -> if List.is_empty context_ids then (* Empty section context: no need to check *) Id.Set.empty, def else (* No declared section vars, and non-empty section context: we must look at the body NOW, if any *) let ids_typ = global_vars_set env typ in let ids_def = match def with | Undef _ | Primitive _ -> Id.Set.empty | Def cs -> global_vars_set env (Mod_subst.force_constr cs) | OpaqueDef _ -> (* Opaque definitions always come with their section variables *) assert false in Environ.really_needed env (Id.Set.union ids_typ ids_def), def | Some declared -> let needed = Environ.really_needed env declared in (* Transitive closure ensured by the upper layers *) let () = assert (Id.Set.equal needed declared) in (* We use the declared set and chain a check of correctness *) declared, match def with | Undef _ | Primitive _ | OpaqueDef _ as x -> x (* nothing to check *) | Def cs as x -> let () = check_section_variables env declared typ (Mod_subst.force_constr cs) in x in let univs = result.cook_universes in let hyps = List.filter (fun d -> Id.Set.mem (NamedDecl.get_id d) hyps) (Environ.named_context env) in let tps = let res = Cbytegen.compile_constant_body ~fail_on_error:false env univs def in Option.map Cemitcodes.from_val res in { const_hyps = hyps; const_body = def; const_type = typ; const_body_code = tps; const_universes = univs; const_relevance = result.cook_relevance; const_inline_code = result.cook_inline; const_typing_flags = Environ.typing_flags env } let check_delayed (type a) (handle : a effect_handler) tyenv (body : a proof_output) = match tyenv with | MonoTyCtx (env, tyj, univs, declared, feedback_id) -> let ((body, uctx), side_eff) = body in (* don't redeclare universes which are declared for the type *) let uctx = Univ.ContextSet.diff uctx univs in let (body, uctx', valid_signatures) = handle env body side_eff in let uctx = Univ.ContextSet.union uctx uctx' in let env = push_context_set uctx env in let body,env,ectx = skip_trusted_seff valid_signatures body env in let j = Typeops.infer env body in let j = unzip ectx j in let _ = Typeops.judge_of_cast env j DEFAULTcast tyj in let c = j.uj_val in let () = check_section_variables env declared tyj.utj_val body in feedback_completion_typecheck feedback_id; c, Opaqueproof.PrivateMonomorphic uctx | PolyTyCtx (env, tj, usubst, auctx, declared, feedback_id) -> let ((body, ctx), side_eff) = body in let body, ctx', _ = handle env body side_eff in let ctx = Univ.ContextSet.union ctx ctx' in (** [ctx] must contain local universes, such that it has no impact on the rest of the graph (up to transitivity). *) let env = push_subgraph ctx env in let private_univs = on_snd (Univ.subst_univs_level_constraints usubst) ctx in let j = Typeops.infer env body in let _ = Typeops.judge_of_cast env j DEFAULTcast tj in let () = check_section_variables env declared tj.utj_val body in let def = Vars.subst_univs_level_constr usubst j.uj_val in let () = feedback_completion_typecheck feedback_id in def, Opaqueproof.PrivatePolymorphic (Univ.AUContext.size auctx, private_univs) (*s Global and local constant declaration. *) let translate_constant env _kn ce = build_constant_declaration env (infer_declaration env ce) let translate_opaque env _kn ce = let def, ctx = infer_opaque env ce in build_constant_declaration env def, ctx let translate_local_assum env t = let j = Typeops.infer env t in let t = Typeops.assumption_of_judgment env j in j.uj_val, t let translate_recipe env _kn r = let open Cooking in let result = Cooking.cook_constant r in let univs = result.cook_universes in let res = Cbytegen.compile_constant_body ~fail_on_error:false env univs result.cook_body in let tps = Option.map Cemitcodes.from_val res in let hyps = Option.get result.cook_context in (* Trust the set of section hypotheses generated by Cooking *) let hyps = List.filter (fun d -> Id.Set.mem (NamedDecl.get_id d) hyps) (Environ.named_context env) in { const_hyps = hyps; const_body = result.cook_body; const_type = result.cook_type; const_body_code = tps; const_universes = univs; const_relevance = result.cook_relevance; const_inline_code = result.cook_inline; const_typing_flags = Environ.typing_flags env } let translate_local_def env _id centry = let open Cooking in let centry = { const_entry_body = centry.secdef_body; const_entry_secctx = centry.secdef_secctx; const_entry_feedback = centry.secdef_feedback; const_entry_type = centry.secdef_type; const_entry_universes = Monomorphic_entry Univ.ContextSet.empty; const_entry_inline_code = false; } in let decl = infer_declaration env (DefinitionEntry centry) in let typ = decl.cook_type in let () = match decl.cook_universes with | Monomorphic ctx -> assert (Univ.ContextSet.is_empty ctx) | Polymorphic _ -> assert false in let c = match decl.cook_body with | Def c -> Mod_subst.force_constr c | Undef _ | Primitive _ | OpaqueDef _ -> assert false in c, decl.cook_relevance, typ coq-8.11.0/kernel/vars.mli0000644000175000017500000001535213612664533015232 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr -> bool (** [closed0 M] is true iff [M] is a (de Bruijn) closed term *) val closed0 : constr -> bool (** [noccurn n M] returns true iff [Rel n] does NOT occur in term [M] *) val noccurn : int -> constr -> bool (** [noccur_between n m M] returns true iff [Rel p] does NOT occur in term [M] for n <= p < n+m *) val noccur_between : int -> int -> constr -> bool (** Checking function for terms containing existential- or meta-variables. The function [noccur_with_meta] does not consider meta-variables applied to some terms (intended to be its local context) (for existential variables, it is necessarily the case) *) val noccur_with_meta : int -> int -> constr -> bool (** {6 Relocation and substitution } *) (** [exliftn el c] lifts [c] with lifting [el] *) val exliftn : Esubst.lift -> constr -> constr (** [liftn n k c] lifts by [n] indexes above or equal to [k] in [c] *) val liftn : int -> int -> constr -> constr (** [lift n c] lifts by [n] the positive indexes in [c] *) val lift : int -> constr -> constr (** The type [substl] is the type of substitutions [u₁..un] of type some context Δ and defined in some environment Γ. Typing of substitutions is defined by: - Γ ⊢ ∅ : ∅, - Γ ⊢ u₁..u{_n-1} : Δ and Γ ⊢ u{_n} : An\[u₁..u{_n-1}\] implies Γ ⊢ u₁..u{_n} : Δ,x{_n}:A{_n} - Γ ⊢ u₁..u{_n-1} : Δ and Γ ⊢ un : A{_n}\[u₁..u{_n-1}\] implies Γ ⊢ u₁..u{_n} : Δ,x{_n}:=c{_n}:A{_n} when Γ ⊢ u{_n} ≡ c{_n}\[u₁..u{_n-1}\] Note that [u₁..un] is represented as a list with [un] at the head of the list, i.e. as [[un;...;u₁]]. *) type substl = constr list (** Let [Γ] be a context interleaving declarations [x₁:T₁..xn:Tn] and definitions [y₁:=c₁..yp:=cp] in some context [Γ₀]. Let [u₁..un] be an {e instance} of [Γ], i.e. an instance in [Γ₀] of the [xi]. Then, [subst_of_rel_context_instance Γ u₁..un] returns the corresponding {e substitution} of [Γ], i.e. the appropriate interleaving [σ] of the [u₁..un] with the [c₁..cp], all of them in [Γ₀], so that a derivation [Γ₀, Γ, Γ₁|- t:T] can be instantiated into a derivation [Γ₀, Γ₁ |- t[σ]:T[σ]] using [substnl σ |Γ₁| t]. Note that the instance [u₁..un] is represented starting with [u₁], as if usable in [applist] while the substitution is represented the other way round, i.e. ending with either [u₁] or [c₁], as if usable for [substl]. *) val subst_of_rel_context_instance : Constr.rel_context -> constr list -> substl (** For compatibility: returns the substitution reversed *) val adjust_subst_to_rel_context : Constr.rel_context -> constr list -> constr list (** Take an index in an instance of a context and returns its index wrt to the full context (e.g. 2 in [x:A;y:=b;z:C] is 3, i.e. a reference to z) *) val adjust_rel_to_rel_context : ('a, 'b) Context.Rel.pt -> int -> int (** [substnl [a₁;...;an] k c] substitutes in parallel [a₁],...,[an] for respectively [Rel(k+1)],...,[Rel(k+n)] in [c]; it relocates accordingly indexes in [an],...,[a1] and [c]. In terms of typing, if Γ ⊢ a{_n}..a₁ : Δ and Γ, Δ, Γ' ⊢ c : T with |Γ'|=k, then Γ, Γ' ⊢ [substnl [a₁;...;an] k c] : [substnl [a₁;...;an] k T]. *) val substnl : substl -> int -> constr -> constr (** [substl σ c] is a short-hand for [substnl σ 0 c] *) val substl : substl -> constr -> constr (** [substl a c] is a short-hand for [substnl [a] 0 c] *) val subst1 : constr -> constr -> constr (** [substnl_decl [a₁;...;an] k Ω] substitutes in parallel [a₁], ..., [an] for respectively [Rel(k+1)], ..., [Rel(k+n)] in [Ω]; it relocates accordingly indexes in [a₁],...,[an] and [c]. In terms of typing, if Γ ⊢ a{_n}..a₁ : Δ and Γ, Δ, Γ', Ω ⊢ with |Γ'|=[k], then Γ, Γ', [substnl_decl [a₁;...;an]] k Ω ⊢. *) val substnl_decl : substl -> int -> Constr.rel_declaration -> Constr.rel_declaration (** [substl_decl σ Ω] is a short-hand for [substnl_decl σ 0 Ω] *) val substl_decl : substl -> Constr.rel_declaration -> Constr.rel_declaration (** [subst1_decl a Ω] is a short-hand for [substnl_decl [a] 0 Ω] *) val subst1_decl : constr -> Constr.rel_declaration -> Constr.rel_declaration (** [replace_vars k [(id₁,c₁);...;(idn,cn)] t] substitutes [Var idj] by [cj] in [t]. *) val replace_vars : (Id.t * constr) list -> constr -> constr (** [substn_vars k [id₁;...;idn] t] substitutes [Var idj] by [Rel j+k-1] in [t]. If two names are identical, the one of least index is kept. In terms of typing, if Γ,x{_n}:U{_n},...,x₁:U₁,Γ' ⊢ t:T, together with id{_j}:T{_j} and Γ,x{_n}:U{_n},...,x₁:U₁,Γ' ⊢ T{_j}\[id{_j+1}..id{_n}:=x{_j+1}..x{_n}\] ≡ Uj, then Γ\\{id₁,...,id{_n}\},x{_n}:U{_n},...,x₁:U₁,Γ' ⊢ [substn_vars (|Γ'|+1) [id₁;...;idn] t] : [substn_vars (|Γ'|+1) [id₁;...;idn] T]. *) val substn_vars : int -> Id.t list -> constr -> constr (** [subst_vars [id1;...;idn] t] is a short-hand for [substn_vars [id1;...;idn] 1 t]: it substitutes [Var idj] by [Rel j] in [t]. If two names are identical, the one of least index is kept. *) val subst_vars : Id.t list -> constr -> constr (** [subst_var id t] is a short-hand for [substn_vars [id] 1 t]: it substitutes [Var id] by [Rel 1] in [t]. *) val subst_var : Id.t -> constr -> constr (** {3 Substitution of universes} *) open Univ (** Level substitutions for polymorphism. *) val subst_univs_level_constr : universe_level_subst -> constr -> constr val subst_univs_level_context : Univ.universe_level_subst -> Constr.rel_context -> Constr.rel_context (** Instance substitution for polymorphism. *) val subst_instance_constr : Instance.t -> constr -> constr val subst_instance_context : Instance.t -> Constr.rel_context -> Constr.rel_context val univ_instantiate_constr : Instance.t -> constr univ_abstracted -> constr (** Ignores the constraints carried by [univ_abstracted]. *) val universes_of_constr : constr -> Univ.LSet.t coq-8.11.0/kernel/evar.mli0000644000175000017500000000276613612664533015221 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int (** Recover the underlying integer. *) val unsafe_of_int : int -> t (** This is not for dummies. Do not use this function if you don't know what you are doing. *) val equal : t -> t -> bool (** Equality over existential variables. *) val compare : t -> t -> int (** Comparison over existential variables. *) val hash : t -> int (** Hash over existential variables. *) val print : t -> Pp.t (** Printing representation *) module Set : Set.S with type elt = t module Map : CMap.ExtS with type key = t and module Set := Set coq-8.11.0/kernel/vconv.ml0000644000175000017500000001724213612664533015241 0ustar treinentreinenopen Util open Names open Environ open Reduction open Vm open Vmvalues open Csymtable (* Test la structure des piles *) let compare_zipper z1 z2 = match z1, z2 with | Zapp args1, Zapp args2 -> Int.equal (nargs args1) (nargs args2) | Zfix(_f1,args1), Zfix(_f2,args2) -> Int.equal (nargs args1) (nargs args2) | Zswitch _, Zswitch _ | Zproj _, Zproj _ -> true | Zapp _ , _ | Zfix _, _ | Zswitch _, _ | Zproj _, _ -> false let rec compare_stack stk1 stk2 = match stk1, stk2 with | [], [] -> true | z1::stk1, z2::stk2 -> if compare_zipper z1 z2 then compare_stack stk1 stk2 else false | _, _ -> false (* Conversion *) let conv_vect fconv vect1 vect2 cu = let n = Array.length vect1 in if Int.equal n (Array.length vect2) then let rcu = ref cu in for i = 0 to n - 1 do rcu := fconv vect1.(i) vect2.(i) !rcu done; !rcu else raise NotConvertible let rec conv_val env pb k v1 v2 cu = if v1 == v2 then cu else conv_whd env pb k (whd_val v1) (whd_val v2) cu and conv_whd env pb k whd1 whd2 cu = (* Pp.(msg_debug (str "conv_whd(" ++ pr_whd whd1 ++ str ", " ++ pr_whd whd2 ++ str ")")) ; *) match whd1, whd2 with | Vuniv_level _ , _ | _ , Vuniv_level _ -> (** Both of these are invalid since universes are handled via ** special cases in the code. **) assert false | Vprod p1, Vprod p2 -> let cu = conv_val env CONV k (dom p1) (dom p2) cu in conv_fun env pb k (codom p1) (codom p2) cu | Vfun f1, Vfun f2 -> conv_fun env CONV k f1 f2 cu | Vfix (f1,None), Vfix (f2,None) -> conv_fix env k f1 f2 cu | Vfix (f1,Some args1), Vfix(f2,Some args2) -> if nargs args1 <> nargs args2 then raise NotConvertible else conv_arguments env k args1 args2 (conv_fix env k f1 f2 cu) | Vcofix (cf1,_,None), Vcofix (cf2,_,None) -> conv_cofix env k cf1 cf2 cu | Vcofix (cf1,_,Some args1), Vcofix (cf2,_,Some args2) -> if nargs args1 <> nargs args2 then raise NotConvertible else conv_arguments env k args1 args2 (conv_cofix env k cf1 cf2 cu) | Vconstr_const i1, Vconstr_const i2 -> if Int.equal i1 i2 then cu else raise NotConvertible | Vconstr_block b1, Vconstr_block b2 -> let tag1 = btag b1 and tag2 = btag b2 in let sz = bsize b1 in if Int.equal tag1 tag2 && Int.equal sz (bsize b2) then let rcu = ref cu in for i = 0 to sz - 1 do rcu := conv_val env CONV k (bfield b1 i) (bfield b2 i) !rcu done; !rcu else raise NotConvertible | Vint64 i1, Vint64 i2 -> if Int64.equal i1 i2 then cu else raise NotConvertible | Vfloat64 f1, Vfloat64 f2 -> if Float64.(equal (of_float f1) (of_float f2)) then cu else raise NotConvertible | Vatom_stk(a1,stk1), Vatom_stk(a2,stk2) -> conv_atom env pb k a1 stk1 a2 stk2 cu | Vfun _, _ | _, Vfun _ -> (* on the fly eta expansion *) conv_val env CONV (k+1) (apply_whd k whd1) (apply_whd k whd2) cu | Vprod _, _ | Vfix _, _ | Vcofix _, _ | Vconstr_const _, _ | Vint64 _, _ | Vfloat64 _, _ | Vconstr_block _, _ | Vatom_stk _, _ -> raise NotConvertible and conv_atom env pb k a1 stk1 a2 stk2 cu = (* Pp.(msg_debug (str "conv_atom(" ++ pr_atom a1 ++ str ", " ++ pr_atom a2 ++ str ")")) ; *) match a1, a2 with | Aind ((mi,_i) as ind1) , Aind ind2 -> if eq_ind ind1 ind2 && compare_stack stk1 stk2 then let ulen = Univ.AUContext.size (Environ.mind_context env mi) in if ulen = 0 then conv_stack env k stk1 stk2 cu else match stk1 , stk2 with | Zapp args1 :: stk1' , Zapp args2 :: stk2' -> assert (ulen <= nargs args1); assert (ulen <= nargs args2); let u1 = Array.init ulen (fun i -> uni_lvl_val (arg args1 i)) in let u2 = Array.init ulen (fun i -> uni_lvl_val (arg args2 i)) in let u1 = Univ.Instance.of_array u1 in let u2 = Univ.Instance.of_array u2 in let cu = convert_instances ~flex:false u1 u2 cu in conv_arguments env ~from:ulen k args1 args2 (conv_stack env k stk1' stk2' cu) | _, _ -> assert false (* Should not happen if problem is well typed *) else raise NotConvertible | Aid ik1, Aid ik2 -> if Vmvalues.eq_id_key ik1 ik2 && compare_stack stk1 stk2 then conv_stack env k stk1 stk2 cu else raise NotConvertible | Asort s1, Asort s2 -> sort_cmp_universes env pb s1 s2 cu | Asort _ , _ | Aind _, _ | Aid _, _ -> raise NotConvertible and conv_stack env k stk1 stk2 cu = match stk1, stk2 with | [], [] -> cu | Zapp args1 :: stk1, Zapp args2 :: stk2 -> conv_stack env k stk1 stk2 (conv_arguments env k args1 args2 cu) | Zfix(f1,args1) :: stk1, Zfix(f2,args2) :: stk2 -> conv_stack env k stk1 stk2 (conv_arguments env k args1 args2 (conv_fix env k f1 f2 cu)) | Zswitch sw1 :: stk1, Zswitch sw2 :: stk2 -> if check_switch sw1 sw2 then let vt1,vt2 = type_of_switch sw1, type_of_switch sw2 in let rcu = ref (conv_val env CONV k vt1 vt2 cu) in let b1, b2 = branch_of_switch k sw1, branch_of_switch k sw2 in for i = 0 to Array.length b1 - 1 do rcu := conv_val env CONV (k + fst b1.(i)) (snd b1.(i)) (snd b2.(i)) !rcu done; conv_stack env k stk1 stk2 !rcu else raise NotConvertible | Zproj p1 :: stk1, Zproj p2 :: stk2 -> if Projection.Repr.equal p1 p2 then conv_stack env k stk1 stk2 cu else raise NotConvertible | [], _ | Zapp _ :: _, _ | Zfix _ :: _, _ | Zswitch _ :: _, _ | Zproj _ :: _, _ -> raise NotConvertible and conv_fun env pb k f1 f2 cu = if f1 == f2 then cu else let arity,b1,b2 = decompose_vfun2 k f1 f2 in conv_val env pb (k+arity) b1 b2 cu and conv_fix env k f1 f2 cu = if f1 == f2 then cu else if check_fix f1 f2 then let bf1, tf1 = reduce_fix k f1 in let bf2, tf2 = reduce_fix k f2 in let cu = conv_vect (conv_val env CONV k) tf1 tf2 cu in conv_vect (conv_fun env CONV (k + Array.length tf1)) bf1 bf2 cu else raise NotConvertible and conv_cofix env k cf1 cf2 cu = if cf1 == cf2 then cu else if check_cofix cf1 cf2 then let bcf1, tcf1 = reduce_cofix k cf1 in let bcf2, tcf2 = reduce_cofix k cf2 in let cu = conv_vect (conv_val env CONV k) tcf1 tcf2 cu in conv_vect (conv_val env CONV (k + Array.length tcf1)) bcf1 bcf2 cu else raise NotConvertible and conv_arguments env ?from:(from=0) k args1 args2 cu = if args1 == args2 then cu else let n = nargs args1 in if Int.equal n (nargs args2) then let rcu = ref cu in for i = from to n - 1 do rcu := conv_val env CONV k (arg args1 i) (arg args2 i) !rcu done; !rcu else raise NotConvertible let warn_bytecode_compiler_failed = let open Pp in CWarnings.create ~name:"bytecode-compiler-failed" ~category:"bytecode-compiler" (fun () -> strbrk "Bytecode compiler failed, " ++ strbrk "falling back to standard conversion") let vm_conv_gen cv_pb env univs t1 t2 = if not (typing_flags env).Declarations.enable_VM then Reduction.generic_conv cv_pb ~l2r:false (fun _ -> None) TransparentState.full env univs t1 t2 else try let v1 = val_of_constr env t1 in let v2 = val_of_constr env t2 in fst (conv_val env cv_pb (nb_rel env) v1 v2 univs) with Not_found | Invalid_argument _ -> warn_bytecode_compiler_failed (); Reduction.generic_conv cv_pb ~l2r:false (fun _ -> None) TransparentState.full env univs t1 t2 let vm_conv cv_pb env t1 t2 = let univs = Environ.universes env in let b = if cv_pb = CUMUL then Constr.leq_constr_univs univs t1 t2 else Constr.eq_constr_univs univs t1 t2 in if not b then let univs = (univs, checked_universes) in let _ = vm_conv_gen cv_pb env univs t1 t2 in () coq-8.11.0/kernel/conv_oracle.mli0000644000175000017500000000347713612664533016556 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constant.t) -> oracle -> bool -> 'a tableKey -> 'a tableKey -> bool (** Priority for the expansion of constant in the conversion test. * Higher levels means that the expansion is less prioritary. * (And Expand stands for -oo, and Opaque +oo.) * The default value (transparent constants) is [Level 0]. *) type level = Expand | Level of int | Opaque val transparent : level (** Check whether a level is transparent *) val is_transparent : level -> bool val get_strategy : oracle -> Constant.t tableKey -> level (** Sets the level of a constant. * Level of RelKey constant cannot be set. *) val set_strategy : oracle -> Constant.t tableKey -> level -> oracle (** Fold over the non-transparent levels of the oracle. Order unspecified. *) val fold_strategy : (Constant.t tableKey -> level -> 'a -> 'a) -> oracle -> 'a -> 'a val get_transp_state : oracle -> TransparentState.t coq-8.11.0/kernel/mod_typing.mli0000644000175000017500000000450513612664533016426 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ModPath.t -> inline -> module_entry -> module_body (** [translate_modtype] produces a [module_type_body] whose [mod_type_alg] cannot be [None] (and of course [mod_expr] is [Abstract]). *) val translate_modtype : env -> ModPath.t -> inline -> module_type_entry -> module_type_body (** Low-level function for translating a module struct entry : - We translate to a module when a [ModPath.t] is given, otherwise to a module type. - The first output is the expanded signature - The second output is the algebraic expression, kept mostly for the extraction. *) type 'alg translation = module_signature * 'alg * delta_resolver * Univ.ContextSet.t val translate_mse : env -> ModPath.t option -> inline -> module_struct_entry -> module_alg_expr translation (** From an already-translated (or interactive) implementation and an (optional) signature entry, produces a final [module_body] *) val finalize_module : env -> ModPath.t -> (module_expression option) translation -> (module_type_entry * inline) option -> module_body (** [translate_mse_incl] translate the mse of a module or module type given to an Include *) val translate_mse_incl : bool -> env -> ModPath.t -> inline -> module_struct_entry -> unit translation coq-8.11.0/kernel/indTyping.ml0000644000175000017500000004121413612664533016047 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* to remove?. *) (* [check_constructors_names id s cl] checks that all the constructors names appearing in [l] are not present in the set [s], and returns the new set of names. The name [id] is the name of the current inductive type, used when reporting the error. *) let check_constructors_names = let rec check idset = function | [] -> idset | c::cl -> if Id.Set.mem c idset then raise (InductiveError (SameNamesConstructors c)) else check (Id.Set.add c idset) cl in check (* [mind_check_names mie] checks the names of an inductive types declaration, and raises the corresponding exceptions when two types or two constructors have the same name. *) let mind_check_names mie = let rec check indset cstset = function | [] -> () | ind::inds -> let id = ind.mind_entry_typename in let cl = ind.mind_entry_consnames in if Id.Set.mem id indset then raise (InductiveError (SameNamesTypes id)) else let cstset' = check_constructors_names cstset cl in check (Id.Set.add id indset) cstset' inds in check Id.Set.empty Id.Set.empty mie.mind_entry_inds (* The above verification is not necessary from the kernel point of vue since inductive and constructors are not referred to by their name, but only by the name of the inductive packet and an index. *) (************************************************************************) (************************** Cumulativity checking************************) (************************************************************************) (* Check arities and constructors *) let check_subtyping_arity_constructor env subst arcn numparams is_arity = let numchecked = ref 0 in let basic_check ev tp = if !numchecked < numparams then () else Reduction.conv_leq ev tp (subst tp); numchecked := !numchecked + 1 in let check_typ typ typ_env = match typ with | LocalAssum (_, typ') -> begin try basic_check typ_env typ'; Environ.push_rel typ typ_env with Reduction.NotConvertible -> CErrors.anomaly ~label:"bad inductive subtyping relation" Pp.(str "Invalid subtyping relation") end | _ -> CErrors.anomaly Pp.(str "") in let typs, codom = Reduction.dest_prod env arcn in let last_env = Context.Rel.fold_outside check_typ typs ~init:env in if not is_arity then basic_check last_env codom else () let check_cumulativity univs variances env_ar params data = let uctx = match univs with | Monomorphic_entry _ -> raise (InductiveError BadUnivs) | Polymorphic_entry (_,uctx) -> uctx in let instance = UContext.instance uctx in if Instance.length instance != Array.length variances then raise (InductiveError BadUnivs); let numparams = Context.Rel.nhyps params in let new_levels = Array.init (Instance.length instance) (fun i -> Level.(make (UGlobal.make DirPath.empty i))) in let lmap = Array.fold_left2 (fun lmap u u' -> LMap.add u u' lmap) LMap.empty (Instance.to_array instance) new_levels in let dosubst = Vars.subst_univs_level_constr lmap in let instance_other = Instance.of_array new_levels in let constraints_other = Univ.subst_univs_level_constraints lmap (UContext.constraints uctx) in let uctx_other = Univ.UContext.make (instance_other, constraints_other) in let env = Environ.push_context uctx_other env_ar in let subtyp_constraints = Univ.enforce_leq_variance_instances variances instance instance_other Constraint.empty in let env = Environ.add_constraints subtyp_constraints env in (* process individual inductive types: *) List.iter (fun (arity,lc) -> check_subtyping_arity_constructor env dosubst arity numparams true; Array.iter (fun cnt -> check_subtyping_arity_constructor env dosubst cnt numparams false) lc) data (************************************************************************) (************************** Type checking *******************************) (************************************************************************) type univ_info = { ind_squashed : bool; ind_has_relevant_arg : bool; ind_min_univ : Universe.t option; (* Some for template *) ind_univ : Universe.t } let check_univ_leq ?(is_real_arg=false) env u info = let ind_univ = info.ind_univ in let info = if not info.ind_has_relevant_arg && is_real_arg && not (Univ.Universe.is_sprop u) then {info with ind_has_relevant_arg=true} else info in (* Inductive types provide explicit lifting from SProp to other universes, so allow SProp <= any. *) if type_in_type env || Univ.Universe.is_sprop u || UGraph.check_leq (universes env) u ind_univ then { info with ind_min_univ = Option.map (Universe.sup u) info.ind_min_univ } else if is_impredicative_univ env ind_univ then if Option.is_empty info.ind_min_univ then { info with ind_squashed = true } else raise (InductiveError BadUnivs) else raise (InductiveError BadUnivs) let check_context_univs ~ctor env info ctx = let check_one d (info,env) = let info = match d with | LocalAssum (_,t) -> (* could be retyping if it becomes available in the kernel *) let tj = Typeops.infer_type env t in check_univ_leq ~is_real_arg:ctor env (Sorts.univ_of_sort tj.utj_type) info | LocalDef _ -> info in info, push_rel d env in fst (Context.Rel.fold_outside ~init:(info,env) check_one ctx) let check_indices_matter env_params info indices = if not (indices_matter env_params) then info else check_context_univs ~ctor:false env_params info indices (* env_ar contains the inductives before the current ones in the block, and no parameters *) let check_arity env_params env_ar ind = let {utj_val=arity;utj_type=_} = Typeops.infer_type env_params ind.mind_entry_arity in let indices, ind_sort = Reduction.dest_arity env_params arity in let ind_min_univ = if ind.mind_entry_template then Some Universe.type0m else None in let univ_info = { ind_squashed=false; ind_has_relevant_arg=false; ind_min_univ; ind_univ=Sorts.univ_of_sort ind_sort; } in let univ_info = check_indices_matter env_params univ_info indices in (* We do not need to generate the universe of the arity with params; if later, after the validation of the inductive definition, full_arity is used as argument or subject to cast, an upper universe will be generated *) let arity = it_mkProd_or_LetIn arity (Environ.rel_context env_params) in let x = Context.make_annot (Name ind.mind_entry_typename) (Sorts.relevance_of_sort ind_sort) in push_rel (LocalAssum (x, arity)) env_ar, (arity, indices, univ_info) let check_constructor_univs env_ar_par info (args,_) = (* We ignore the output, positivity will check that it's the expected inductive type *) check_context_univs ~ctor:true env_ar_par info args let check_constructors env_ar_par isrecord params lc (arity,indices,univ_info) = let lc = Array.map_of_list (fun c -> (Typeops.infer_type env_ar_par c).utj_val) lc in let splayed_lc = Array.map (Reduction.dest_prod_assum env_ar_par) lc in let univ_info = match Array.length lc with (* Empty type: all OK *) | 0 -> univ_info (* SProp primitive records are OK, if we squash and become fakerecord also OK *) | 1 when isrecord -> univ_info (* Unit and identity types must squash if SProp *) | 1 -> check_univ_leq env_ar_par Univ.Universe.type0m univ_info (* More than 1 constructor: must squash if Prop/SProp *) | _ -> check_univ_leq env_ar_par Univ.Universe.type0 univ_info in let univ_info = Array.fold_left (check_constructor_univs env_ar_par) univ_info splayed_lc in (* generalize the constructors over the parameters *) let lc = Array.map (fun c -> Term.it_mkProd_or_LetIn c params) lc in (arity, lc), (indices, splayed_lc), univ_info let check_record data = List.for_all (fun (_,(_,splayed_lc),info) -> (* records must have all projections definable -> equivalent to not being squashed *) not info.ind_squashed (* relevant records must have at least 1 relevant argument *) && (Univ.Universe.is_sprop info.ind_univ || info.ind_has_relevant_arg) && (match splayed_lc with (* records must have 1 constructor with at least 1 argument, and no anonymous fields *) | [|ctx,_|] -> let module D = Context.Rel.Declaration in List.exists D.is_local_assum ctx && List.for_all (fun d -> not (D.is_local_assum d) || not (Name.is_anonymous (D.get_name d))) ctx | _ -> false)) data (* Allowed eliminations *) (* Previous comment: *) (* Unitary/empty Prop: elimination to all sorts are realizable *) (* unless the type is large. If it is large, forbids large elimination *) (* which otherwise allows simulating the inconsistent system Type:Type. *) (* -> this is now handled by is_smashed: *) (* - all_sorts in case of small, unitary Prop (not smashed) *) (* - logical_sorts in case of large, unitary Prop (smashed) *) let allowed_sorts {ind_squashed;ind_univ;ind_min_univ=_;ind_has_relevant_arg=_} = if not ind_squashed then InType else Sorts.family (Sorts.sort_of_univ ind_univ) (* For a level to be template polymorphic, it must be introduced by the definition (so have no constraint except lbound <= l) and not to be constrained from below, so any universe l' <= l can be used as an instance of l. All bounds from above, i.e. l <=/< r will be valid for any l' <= l. *) let unbounded_from_below u cstrs = Univ.Constraint.for_all (fun (l, d, r) -> match d with | Eq -> not (Univ.Level.equal l u) && not (Univ.Level.equal r u) | Lt | Le -> not (Univ.Level.equal r u)) cstrs (* Returns the list [x_1, ..., x_n] of levels contributing to template polymorphism. The elements x_k is None if the k-th parameter (starting from the most recent and ignoring let-definitions) is not contributing to the inductive type's sort or is Some u_k if its level is u_k and is contributing. *) let template_polymorphic_univs ~template_check ~ctor_levels uctx paramsctxt concl = let check_level l = if Univ.LSet.mem l (Univ.ContextSet.levels uctx) && unbounded_from_below l (Univ.ContextSet.constraints uctx) && not (Univ.LSet.mem l ctor_levels) then Some l else None in let univs = Univ.Universe.levels concl in let univs = if template_check then Univ.LSet.filter (fun l -> Option.has_some (check_level l) || Univ.Level.is_prop l) univs else univs (* Doesn't check the universes can be generalized *) in let fold acc = function | (LocalAssum (_, p)) -> (let c = Term.strip_prod_assum p in match kind c with | Sort (Type u) -> if template_check then (match Univ.Universe.level u with | Some l -> if Univ.LSet.mem l univs && not (Univ.Level.is_prop l) then Some l else None | None -> None) else Univ.Universe.level u | _ -> None) :: acc | LocalDef _ -> acc in let params = List.fold_left fold [] paramsctxt in params, univs let abstract_packets ~template_check univs usubst params ((arity,lc),(indices,splayed_lc),univ_info) = let arity = Vars.subst_univs_level_constr usubst arity in let lc = Array.map (Vars.subst_univs_level_constr usubst) lc in let indices = Vars.subst_univs_level_context usubst indices in let splayed_lc = Array.map (fun (args,out) -> let args = Vars.subst_univs_level_context usubst args in let out = Vars.subst_univs_level_constr usubst out in args,out) splayed_lc in let ind_univ = Univ.subst_univs_level_universe usubst univ_info.ind_univ in let arity = match univ_info.ind_min_univ with | None -> RegularArity {mind_user_arity = arity; mind_sort = Sorts.sort_of_univ ind_univ} | Some min_univ -> let ctx = match univs with | Monomorphic ctx -> ctx | Polymorphic _ -> CErrors.anomaly ~label:"polymorphic_template_ind" Pp.(strbrk "Template polymorphism and full polymorphism are incompatible.") in let ctor_levels = let add_levels c levels = Univ.LSet.union levels (Vars.universes_of_constr c) in let param_levels = List.fold_left (fun levels d -> match d with | LocalAssum _ -> levels | LocalDef (_,b,t) -> add_levels b (add_levels t levels)) Univ.LSet.empty params in Array.fold_left (fun levels (d,c) -> let levels = List.fold_left (fun levels d -> Context.Rel.Declaration.fold_constr add_levels d levels) levels d in add_levels c levels) param_levels splayed_lc in let param_levels, concl_levels = template_polymorphic_univs ~template_check ~ctor_levels ctx params min_univ in if template_check && List.for_all (fun x -> Option.is_empty x) param_levels && Univ.LSet.is_empty concl_levels then CErrors.user_err Pp.(strbrk "Ill-formed template inductive declaration: not polymorphic on any universe.") else TemplateArity {template_param_levels = param_levels; template_level = min_univ} in let kelim = allowed_sorts univ_info in (arity,lc), (indices,splayed_lc), kelim let typecheck_inductive env (mie:mutual_inductive_entry) = let () = match mie.mind_entry_inds with | [] -> CErrors.anomaly Pp.(str "empty inductive types declaration.") | _ -> () in (* Check unicity of names (redundant with safe_typing's add_field checks) *) mind_check_names mie; assert (List.is_empty (Environ.rel_context env)); let has_template_poly = List.exists (fun oie -> oie.mind_entry_template) mie.mind_entry_inds in (* universes *) let env_univs = match mie.mind_entry_universes with | Monomorphic_entry ctx -> let env = if has_template_poly then set_universes_lbound env Univ.Level.prop else env in push_context_set ctx env | Polymorphic_entry (_, ctx) -> push_context ctx env in (* Params *) let env_params, params = Typeops.check_context env_univs mie.mind_entry_params in (* Arities *) let env_ar, data = List.fold_left_map (check_arity env_params) env_univs mie.mind_entry_inds in let env_ar_par = push_rel_context params env_ar in (* Constructors *) let isrecord = match mie.mind_entry_record with | Some (Some _) -> true | Some None | None -> false in let data = List.map2 (fun ind data -> check_constructors env_ar_par isrecord params ind.mind_entry_lc data) mie.mind_entry_inds data in let record = mie.mind_entry_record in let data, record = match record with | None | Some None -> data, record | Some (Some _) -> if check_record data then data, record else (* if someone tried to declare a record as SProp but it can't be primitive we must squash. *) let data = List.map (fun (a,b,univs) -> a,b,check_univ_leq env_ar_par Univ.Universe.type0m univs) data in data, Some None in let () = match mie.mind_entry_variance with | None -> () | Some variances -> check_cumulativity mie.mind_entry_universes variances env_ar params (List.map pi1 data) in (* Abstract universes *) let usubst, univs = Declareops.abstract_universes mie.mind_entry_universes in let params = Vars.subst_univs_level_context usubst params in let template_check = Environ.check_template env in let data = List.map (abstract_packets ~template_check univs usubst params) data in let env_ar_par = let ctx = Environ.rel_context env_ar_par in let ctx = Vars.subst_univs_level_context usubst ctx in let env = Environ.pop_rel_context (Environ.nb_rel env_ar_par) env_ar_par in Environ.push_rel_context ctx env in env_ar_par, univs, mie.mind_entry_variance, record, params, Array.of_list data coq-8.11.0/kernel/cbytecodes.ml0000644000175000017500000001230013612664533016220 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* str "FVnamed(" ++ Id.print id ++ str ")" | FVrel i -> str "Rel(" ++ int i ++ str ")" | FVuniv_var v -> str "FVuniv(" ++ int v ++ str ")" | FVevar e -> str "FVevar(" ++ int (Evar.repr e) ++ str ")" let rec pp_instr i = match i with | Klabel _ | Ksequence _ -> assert false | Kacc n -> str "acc " ++ int n | Kenvacc n -> str "envacc " ++ int n | Koffsetclosure n -> str "offsetclosure " ++ int n | Kpush -> str "push" | Kpop n -> str "pop " ++ int n | Kpush_retaddr lbl -> str "push_retaddr " ++ pp_lbl lbl | Kapply n -> str "apply " ++ int n | Kappterm(n, m) -> str "appterm " ++ int n ++ str ", " ++ int m | Kreturn n -> str "return " ++ int n | Kjump -> str "jump" | Krestart -> str "restart" | Kgrab n -> str "grab " ++ int n | Kgrabrec n -> str "grabrec " ++ int n | Kclosure(lbl, n) -> str "closure " ++ pp_lbl lbl ++ str ", " ++ int n | Kclosurerec(fv,init,lblt,lblb) -> h 1 (str "closurerec " ++ int fv ++ str ", " ++ int init ++ str " types = " ++ prlist_with_sep spc pp_lbl (Array.to_list lblt) ++ str " bodies = " ++ prlist_with_sep spc pp_lbl (Array.to_list lblb)) | Kclosurecofix (fv,init,lblt,lblb) -> h 1 (str "closurecofix " ++ int fv ++ str ", " ++ int init ++ str " types = " ++ prlist_with_sep spc pp_lbl (Array.to_list lblt) ++ str " bodies = " ++ prlist_with_sep spc pp_lbl (Array.to_list lblb)) | Kgetglobal idu -> str "getglobal " ++ Constant.print idu | Kconst sc -> str "const " ++ pp_struct_const sc | Kmakeblock(n, m) -> str "makeblock " ++ int n ++ str ", " ++ int m | Kmakeprod -> str "makeprod" | Kmakeswitchblock(lblt,lbls,_,sz) -> str "makeswitchblock " ++ pp_lbl lblt ++ str ", " ++ pp_lbl lbls ++ str ", " ++ int sz | Kswitch(lblc,lblb) -> h 1 (str "switch " ++ prlist_with_sep spc pp_lbl (Array.to_list lblc) ++ str " | " ++ prlist_with_sep spc pp_lbl (Array.to_list lblb)) | Kpushfields n -> str "pushfields " ++ int n | Kfield n -> str "field " ++ int n | Ksetfield n -> str "setfield " ++ int n | Kstop -> str "stop" | Kbranch lbl -> str "branch " ++ pp_lbl lbl | Kproj p -> str "proj " ++ Projection.Repr.print p | Kensurestackcapacity size -> str "growstack " ++ int size | Kprim (op, id) -> str (CPrimitives.to_string op) ++ str " " ++ (match id with Some (id,_u) -> Constant.print id | None -> str "") | Kareint n -> str "areint " ++ int n and pp_bytecodes c = match c with | [] -> str "" | Klabel lbl :: c -> str "L" ++ int lbl ++ str ":" ++ fnl () ++ pp_bytecodes c | Ksequence (l1, l2) :: c -> pp_bytecodes l1 ++ pp_bytecodes l2 ++ pp_bytecodes c | i :: c -> pp_instr i ++ fnl () ++ pp_bytecodes c coq-8.11.0/kernel/reduction.ml0000644000175000017500000011457113612664533016105 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* true | Zupdate _::s -> is_empty_stack s | Zshift _::s -> is_empty_stack s | _ -> false (* Compute the lift to be performed on a term placed in a given stack *) let el_stack el stk = let n = List.fold_left (fun i z -> match z with Zshift n -> i+n | _ -> i) 0 stk in el_shft n el let compare_stack_shape stk1 stk2 = let rec compare_rec bal stk1 stk2 = match (stk1,stk2) with ([],[]) -> Int.equal bal 0 | ((Zupdate _|Zshift _)::s1, _) -> compare_rec bal s1 stk2 | (_, (Zupdate _|Zshift _)::s2) -> compare_rec bal stk1 s2 | (Zapp l1::s1, _) -> compare_rec (bal+Array.length l1) s1 stk2 | (_, Zapp l2::s2) -> compare_rec (bal-Array.length l2) stk1 s2 | (Zproj _p1::s1, Zproj _p2::s2) -> Int.equal bal 0 && compare_rec 0 s1 s2 | (ZcaseT(_c1,_,_,_)::s1, ZcaseT(_c2,_,_,_)::s2) -> Int.equal bal 0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2 | (Zfix(_,a1)::s1, Zfix(_,a2)::s2) -> Int.equal bal 0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2 | Zprimitive(op1,_,rargs1, _kargs1)::s1, Zprimitive(op2,_,rargs2, _kargs2)::s2 -> bal=0 && op1=op2 && List.length rargs1=List.length rargs2 && compare_rec 0 s1 s2 | [], _ :: _ | (Zproj _ | ZcaseT _ | Zfix _ | Zprimitive _) :: _, _ -> false in compare_rec 0 stk1 stk2 type lft_fconstr = lift * fconstr type lft_constr_stack_elt = Zlapp of (lift * fconstr) array | Zlproj of Projection.Repr.t * lift | Zlfix of (lift * fconstr) * lft_constr_stack | Zlcase of case_info * lift * constr * constr array * fconstr subs | Zlprimitive of CPrimitives.t * pconstant * lft_fconstr list * lft_fconstr next_native_args and lft_constr_stack = lft_constr_stack_elt list let rec zlapp v = function Zlapp v2 :: s -> zlapp (Array.append v v2) s | s -> Zlapp v :: s (** Hand-unrolling of the map function to bypass the call to the generic array allocation. Type annotation is required to tell OCaml that the array does not contain floats. *) let map_lift (l : lift) (v : fconstr array) = match v with | [||] -> assert false | [|c0|] -> [|(l, c0)|] | [|c0; c1|] -> [|(l, c0); (l, c1)|] | [|c0; c1; c2|] -> [|(l, c0); (l, c1); (l, c2)|] | [|c0; c1; c2; c3|] -> [|(l, c0); (l, c1); (l, c2); (l, c3)|] | v -> Array.Fun1.map (fun l t -> (l, t)) l v let pure_stack lfts stk = let rec pure_rec lfts stk = match stk with [] -> (lfts,[]) | zi::s -> (match (zi,pure_rec lfts s) with (Zupdate _,lpstk) -> lpstk | (Zshift n,(l,pstk)) -> (el_shft n l, pstk) | (Zapp a, (l,pstk)) -> (l,zlapp (map_lift l a) pstk) | (Zproj p, (l,pstk)) -> (l, Zlproj (p,l)::pstk) | (Zfix(fx,a),(l,pstk)) -> let (lfx,pa) = pure_rec l a in (l, Zlfix((lfx,fx),pa)::pstk) | (ZcaseT(ci,p,br,e),(l,pstk)) -> (l,Zlcase(ci,l,p,br,e)::pstk) | (Zprimitive(op,c,rargs,kargs),(l,pstk)) -> (l,Zlprimitive(op,c,List.map (fun t -> (l,t)) rargs, List.map (fun (k,t) -> (k,(l,t))) kargs)::pstk)) in snd (pure_rec lfts stk) (****************************************************************************) (* Reduction Functions *) (****************************************************************************) let whd_betaiota env t = match kind t with | (Sort _|Var _|Meta _|Evar _|Const _|Ind _|Construct _| Prod _|Lambda _|Fix _|CoFix _) -> t | App (c, _) -> begin match kind c with | Ind _ | Construct _ | Evar _ | Meta _ | Const _ | LetIn _ -> t | _ -> whd_val (create_clos_infos betaiota env) (create_tab ()) (inject t) end | _ -> whd_val (create_clos_infos betaiota env) (create_tab ()) (inject t) let nf_betaiota env t = norm_val (create_clos_infos betaiota env) (create_tab ()) (inject t) let whd_betaiotazeta env x = match kind x with | (Sort _|Var _|Meta _|Evar _|Const _|Ind _|Construct _| Prod _|Lambda _|Fix _|CoFix _|Int _|Float _) -> x | App (c, _) -> begin match kind c with | Ind _ | Construct _ | Evar _ | Meta _ | Const _ | Int _ | Float _ -> x | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos betaiotazeta env) (create_tab ()) (inject x) end | Rel _ | Cast _ | LetIn _ | Case _ | Proj _ -> whd_val (create_clos_infos betaiotazeta env) (create_tab ()) (inject x) let whd_all env t = match kind t with | (Sort _|Meta _|Evar _|Ind _|Construct _| Prod _|Lambda _|Fix _|CoFix _|Int _|Float _) -> t | App (c, _) -> begin match kind c with | Ind _ | Construct _ | Evar _ | Meta _ | Int _ | Float _ -> t | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Const _ |Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos all env) (create_tab ()) (inject t) end | Rel _ | Cast _ | LetIn _ | Case _ | Proj _ | Const _ | Var _ -> whd_val (create_clos_infos all env) (create_tab ()) (inject t) let whd_allnolet env t = match kind t with | (Sort _|Meta _|Evar _|Ind _|Construct _| Prod _|Lambda _|Fix _|CoFix _|LetIn _|Int _|Float _) -> t | App (c, _) -> begin match kind c with | Ind _ | Construct _ | Evar _ | Meta _ | LetIn _ | Int _ | Float _ -> t | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | App _ | Const _ | Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos allnolet env) (create_tab ()) (inject t) end | Rel _ | Cast _ | Case _ | Proj _ | Const _ | Var _ -> whd_val (create_clos_infos allnolet env) (create_tab ()) (inject t) (********************************************************************) (* Conversion *) (********************************************************************) (* Conversion utility functions *) (* functions of this type are called from the kernel *) type 'a kernel_conversion_function = env -> 'a -> 'a -> unit (* functions of this type can be called from outside the kernel *) type 'a extended_conversion_function = ?l2r:bool -> ?reds:TransparentState.t -> env -> ?evars:((existential->constr option) * UGraph.t) -> 'a -> 'a -> unit exception NotConvertible (* Convertibility of sorts *) (* The sort cumulativity is Prop <= Set <= Type 1 <= ... <= Type i <= ... and this holds whatever Set is predicative or impredicative *) type conv_pb = | CONV | CUMUL let is_cumul = function CUMUL -> true | CONV -> false type 'a universe_compare = { (* Might raise NotConvertible *) compare_sorts : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a; compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a; compare_cumul_instances : conv_pb -> Univ.Variance.t array -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a } type 'a universe_state = 'a * 'a universe_compare type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a -> 'b type 'a infer_conversion_function = env -> UGraph.t -> 'a -> 'a -> Univ.Constraint.t let sort_cmp_universes env pb s0 s1 (u, check) = (check.compare_sorts env pb s0 s1 u, check) (* [flex] should be true for constants, false for inductive types and constructors. *) let convert_instances ~flex u u' (s, check) = (check.compare_instances ~flex u u' s, check) let get_cumulativity_constraints cv_pb variance u u' = match cv_pb with | CONV -> Univ.enforce_eq_variance_instances variance u u' Univ.Constraint.empty | CUMUL -> Univ.enforce_leq_variance_instances variance u u' Univ.Constraint.empty let inductive_cumulativity_arguments (mind,ind) = mind.Declarations.mind_nparams + mind.Declarations.mind_packets.(ind).Declarations.mind_nrealargs let convert_inductives_gen cmp_instances cmp_cumul cv_pb (mind,ind) nargs u1 u2 s = match mind.Declarations.mind_variance with | None -> cmp_instances u1 u2 s | Some variances -> let num_param_arity = inductive_cumulativity_arguments (mind,ind) in if not (Int.equal num_param_arity nargs) then cmp_instances u1 u2 s else cmp_cumul cv_pb variances u1 u2 s let convert_inductives cv_pb ind nargs u1 u2 (s, check) = convert_inductives_gen (check.compare_instances ~flex:false) check.compare_cumul_instances cv_pb ind nargs u1 u2 s, check let constructor_cumulativity_arguments (mind, ind, ctor) = mind.Declarations.mind_nparams + mind.Declarations.mind_packets.(ind).Declarations.mind_consnrealargs.(ctor - 1) let convert_constructors_gen cmp_instances cmp_cumul (mind, ind, cns) nargs u1 u2 s = match mind.Declarations.mind_variance with | None -> cmp_instances u1 u2 s | Some _ -> let num_cnstr_args = constructor_cumulativity_arguments (mind,ind,cns) in if not (Int.equal num_cnstr_args nargs) then cmp_instances u1 u2 s else (** By invariant, both constructors have a common supertype, so they are convertible _at that type_. *) let variance = Array.make (Univ.Instance.length u1) Univ.Variance.Irrelevant in cmp_cumul CONV variance u1 u2 s let convert_constructors ctor nargs u1 u2 (s, check) = convert_constructors_gen (check.compare_instances ~flex:false) check.compare_cumul_instances ctor nargs u1 u2 s, check let conv_table_key infos k1 k2 cuniv = if k1 == k2 then cuniv else match k1, k2 with | ConstKey (cst, u), ConstKey (cst', u') when Constant.equal cst cst' -> if Univ.Instance.equal u u' then cuniv else let flex = evaluable_constant cst (info_env infos) && RedFlags.red_set (info_flags infos) (RedFlags.fCONST cst) in convert_instances ~flex u u' cuniv | VarKey id, VarKey id' when Id.equal id id' -> cuniv | RelKey n, RelKey n' when Int.equal n n' -> cuniv | _ -> raise NotConvertible exception IrregularPatternShape let unfold_ref_with_args infos tab fl v = match unfold_reference infos tab fl with | Def def -> Some (def, v) | Primitive op when check_native_args op v -> let c = match fl with ConstKey c -> c | _ -> assert false in let rargs, a, nargs, v = get_native_args1 op c v in Some (whd_stack infos tab a (Zupdate a::(Zprimitive(op,c,rargs,nargs)::v))) | Undef _ | OpaqueDef _ | Primitive _ -> None type conv_tab = { cnv_inf : clos_infos; relevances : Sorts.relevance Range.t; lft_tab : clos_tab; rgt_tab : clos_tab; } (** Invariant: for any tl ∈ lft_tab and tr ∈ rgt_tab, there is no mutable memory location contained both in tl and in tr. *) (** The same heap separation invariant must hold for the fconstr arguments passed to each respective side of the conversion function below. *) let push_relevance infos r = { infos with relevances = Range.cons r.Context.binder_relevance infos.relevances } let push_relevances infos nas = { infos with relevances = Array.fold_left (fun l x -> Range.cons x.Context.binder_relevance l) infos.relevances nas } let rec skip_pattern infos relevances n c1 c2 = if Int.equal n 0 then {infos with relevances}, c1, c2 else match kind c1, kind c2 with | Lambda (x, _, c1), Lambda (_, _, c2) -> skip_pattern infos (Range.cons x.Context.binder_relevance relevances) (pred n) c1 c2 | _ -> raise IrregularPatternShape let skip_pattern infos n c1 c2 = if Int.equal n 0 then infos, c1, c2 else skip_pattern infos infos.relevances n c1 c2 let is_irrelevant infos lft c = let env = info_env infos.cnv_inf in try Retypeops.relevance_of_fterm env infos.relevances lft c == Sorts.Irrelevant with _ -> false (* Conversion between [lft1]term1 and [lft2]term2 *) let rec ccnv cv_pb l2r infos lft1 lft2 term1 term2 cuniv = try eqappr cv_pb l2r infos (lft1, (term1,[])) (lft2, (term2,[])) cuniv with NotConvertible when is_irrelevant infos lft1 term1 && is_irrelevant infos lft2 term2 -> cuniv (* Conversion between [lft1](hd1 v1) and [lft2](hd2 v2) *) and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = Control.check_for_interrupt (); (* First head reduce both terms *) let ninfos = infos_with_reds infos.cnv_inf betaiotazeta in let (hd1, v1 as appr1) = whd_stack ninfos infos.lft_tab (fst st1) (snd st1) in let (hd2, v2 as appr2) = whd_stack ninfos infos.rgt_tab (fst st2) (snd st2) in let appr1 = (lft1, appr1) and appr2 = (lft2, appr2) in (** We delay the computation of the lifts that apply to the head of the term with [el_stack] inside the branches where they are actually used. *) match (fterm_of hd1, fterm_of hd2) with (* case of leaves *) | (FAtom a1, FAtom a2) -> (match kind a1, kind a2 with | (Sort s1, Sort s2) -> if not (is_empty_stack v1 && is_empty_stack v2) then anomaly (Pp.str "conversion was given ill-typed terms (Sort)."); sort_cmp_universes (info_env infos.cnv_inf) cv_pb s1 s2 cuniv | (Meta n, Meta m) -> if Int.equal n m then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible | _ -> raise NotConvertible) | (FEvar ((ev1,args1),env1), FEvar ((ev2,args2),env2)) -> if Evar.equal ev1 ev2 then let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let cuniv = convert_stacks l2r infos lft1 lft2 v1 v2 cuniv in convert_vect l2r infos el1 el2 (Array.map (mk_clos env1) args1) (Array.map (mk_clos env2) args2) cuniv else raise NotConvertible (* 2 index known to be bound to no constant *) | (FRel n, FRel m) -> let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in if Int.equal (reloc_rel n el1) (reloc_rel m el2) then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible (* 2 constants, 2 local defined vars or 2 defined rels *) | (FFlex fl1, FFlex fl2) -> (try let cuniv = conv_table_key infos.cnv_inf fl1 fl2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with NotConvertible | Univ.UniverseInconsistency _ -> (* else the oracle tells which constant is to be expanded *) let oracle = CClosure.oracle_of_infos infos.cnv_inf in let (app1,app2) = let aux appr1 lft1 fl1 tab1 v1 appr2 lft2 fl2 tab2 v2 = match unfold_ref_with_args infos.cnv_inf tab1 fl1 v1 with | Some t1 -> ((lft1, t1), appr2) | None -> match unfold_ref_with_args infos.cnv_inf tab2 fl2 v2 with | Some t2 -> (appr1, (lft2, t2)) | None -> raise NotConvertible in if Conv_oracle.oracle_order Univ.out_punivs oracle l2r fl1 fl2 then aux appr1 lft1 fl1 infos.lft_tab v1 appr2 lft2 fl2 infos.rgt_tab v2 else let (app2,app1) = aux appr2 lft2 fl2 infos.rgt_tab v2 appr1 lft1 fl1 infos.lft_tab v1 in (app1,app2) in eqappr cv_pb l2r infos app1 app2 cuniv) | (FProj (p1,c1), FProj (p2, c2)) -> (* Projections: prefer unfolding to first-order unification, which will happen naturally if the terms c1, c2 are not in constructor form *) (match unfold_projection infos.cnv_inf p1 with | Some s1 -> eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv | None -> match unfold_projection infos.cnv_inf p2 with | Some s2 -> eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv | None -> if Projection.Repr.equal (Projection.repr p1) (Projection.repr p2) && compare_stack_shape v1 v2 then let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let u1 = ccnv CONV l2r infos el1 el2 c1 c2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 u1 else (* Two projections in WHNF: unfold *) raise NotConvertible) | (FProj (p1,c1), t2) -> begin match unfold_projection infos.cnv_inf p1 with | Some s1 -> eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv | None -> begin match t2 with | FFlex fl2 -> begin match unfold_ref_with_args infos.cnv_inf infos.rgt_tab fl2 v2 with | Some t2 -> eqappr cv_pb l2r infos appr1 (lft2, t2) cuniv | None -> raise NotConvertible end | _ -> raise NotConvertible end end | (t1, FProj (p2,c2)) -> begin match unfold_projection infos.cnv_inf p2 with | Some s2 -> eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv | None -> begin match t1 with | FFlex fl1 -> begin match unfold_ref_with_args infos.cnv_inf infos.lft_tab fl1 v1 with | Some t1 -> eqappr cv_pb l2r infos (lft1, t1) appr2 cuniv | None -> raise NotConvertible end | _ -> raise NotConvertible end end (* other constructors *) | (FLambda _, FLambda _) -> (* Inconsistency: we tolerate that v1, v2 contain shift and update but we throw them away *) if not (is_empty_stack v1 && is_empty_stack v2) then anomaly (Pp.str "conversion was given ill-typed terms (FLambda)."); let (x1,ty1,bd1) = destFLambda mk_clos hd1 in let (_,ty2,bd2) = destFLambda mk_clos hd2 in let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let cuniv = ccnv CONV l2r infos el1 el2 ty1 ty2 cuniv in ccnv CONV l2r (push_relevance infos x1) (el_lift el1) (el_lift el2) bd1 bd2 cuniv | (FProd (x1, c1, c2, e), FProd (_, c'1, c'2, e')) -> if not (is_empty_stack v1 && is_empty_stack v2) then anomaly (Pp.str "conversion was given ill-typed terms (FProd)."); (* Luo's system *) let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let cuniv = ccnv CONV l2r infos el1 el2 c1 c'1 cuniv in ccnv cv_pb l2r (push_relevance infos x1) (el_lift el1) (el_lift el2) (mk_clos (subs_lift e) c2) (mk_clos (subs_lift e') c'2) cuniv (* Eta-expansion on the fly *) | (FLambda _, _) -> let () = match v1 with | [] -> () | _ -> anomaly (Pp.str "conversion was given unreduced term (FLambda).") in let (x1,_ty1,bd1) = destFLambda mk_clos hd1 in let infos = push_relevance infos x1 in eqappr CONV l2r infos (el_lift lft1, (bd1, [])) (el_lift lft2, (hd2, eta_expand_stack v2)) cuniv | (_, FLambda _) -> let () = match v2 with | [] -> () | _ -> anomaly (Pp.str "conversion was given unreduced term (FLambda).") in let (x2,_ty2,bd2) = destFLambda mk_clos hd2 in let infos = push_relevance infos x2 in eqappr CONV l2r infos (el_lift lft1, (hd1, eta_expand_stack v1)) (el_lift lft2, (bd2, [])) cuniv (* only one constant, defined var or defined rel *) | (FFlex fl1, c2) -> begin match unfold_ref_with_args infos.cnv_inf infos.lft_tab fl1 v1 with | Some (def1,v1) -> (** By virtue of the previous case analyses, we know [c2] is rigid. Conversion check to rigid terms eventually implies full weak-head reduction, so instead of repeatedly performing small-step unfoldings, we perform reduction with all flags on. *) let all = RedFlags.red_add_transparent all (RedFlags.red_transparent (info_flags infos.cnv_inf)) in let r1 = whd_stack (infos_with_reds infos.cnv_inf all) infos.lft_tab def1 v1 in eqappr cv_pb l2r infos (lft1, r1) appr2 cuniv | None -> (match c2 with | FConstruct ((ind2,_j2),_u2) -> (try let v2, v1 = eta_expand_ind_stack (info_env infos.cnv_inf) ind2 hd2 v2 (snd appr1) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) | _ -> raise NotConvertible) end | (c1, FFlex fl2) -> begin match unfold_ref_with_args infos.cnv_inf infos.rgt_tab fl2 v2 with | Some (def2, v2) -> (** Symmetrical case of above. *) let all = RedFlags.red_add_transparent all (RedFlags.red_transparent (info_flags infos.cnv_inf)) in let r2 = whd_stack (infos_with_reds infos.cnv_inf all) infos.rgt_tab def2 v2 in eqappr cv_pb l2r infos appr1 (lft2, r2) cuniv | None -> match c1 with | FConstruct ((ind1,_j1),_u1) -> (try let v1, v2 = eta_expand_ind_stack (info_env infos.cnv_inf) ind1 hd1 v1 (snd appr2) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) | _ -> raise NotConvertible end (* Inductive types: MutInd MutConstruct Fix Cofix *) | (FInd (ind1,u1), FInd (ind2,u2)) -> if eq_ind ind1 ind2 then if Univ.Instance.length u1 = 0 || Univ.Instance.length u2 = 0 then let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else let mind = Environ.lookup_mind (fst ind1) (info_env infos.cnv_inf) in let nargs = CClosure.stack_args_size v1 in if not (Int.equal nargs (CClosure.stack_args_size v2)) then raise NotConvertible else let cuniv = convert_inductives cv_pb (mind, snd ind1) nargs u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible | (FConstruct ((ind1,j1),u1), FConstruct ((ind2,j2),u2)) -> if Int.equal j1 j2 && eq_ind ind1 ind2 then if Univ.Instance.length u1 = 0 || Univ.Instance.length u2 = 0 then let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else let mind = Environ.lookup_mind (fst ind1) (info_env infos.cnv_inf) in let nargs = CClosure.stack_args_size v1 in if not (Int.equal nargs (CClosure.stack_args_size v2)) then raise NotConvertible else let cuniv = convert_constructors (mind, snd ind1, j1) nargs u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible (* Eta expansion of records *) | (FConstruct ((ind1,_j1),_u1), _) -> (try let v1, v2 = eta_expand_ind_stack (info_env infos.cnv_inf) ind1 hd1 v1 (snd appr2) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) | (_, FConstruct ((ind2,_j2),_u2)) -> (try let v2, v1 = eta_expand_ind_stack (info_env infos.cnv_inf) ind2 hd2 v2 (snd appr1) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) | (FFix (((op1, i1),(na1,tys1,cl1)),e1), FFix(((op2, i2),(_,tys2,cl2)),e2)) -> if Int.equal i1 i2 && Array.equal Int.equal op1 op2 then let n = Array.length cl1 in let fty1 = Array.map (mk_clos e1) tys1 in let fty2 = Array.map (mk_clos e2) tys2 in let fcl1 = Array.map (mk_clos (subs_liftn n e1)) cl1 in let fcl2 = Array.map (mk_clos (subs_liftn n e2)) cl2 in let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in let cuniv = let infos = push_relevances infos na1 in convert_vect l2r infos (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible | (FCoFix ((op1,(na1,tys1,cl1)),e1), FCoFix((op2,(_,tys2,cl2)),e2)) -> if Int.equal op1 op2 then let n = Array.length cl1 in let fty1 = Array.map (mk_clos e1) tys1 in let fty2 = Array.map (mk_clos e2) tys2 in let fcl1 = Array.map (mk_clos (subs_liftn n e1)) cl1 in let fcl2 = Array.map (mk_clos (subs_liftn n e2)) cl2 in let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in let cuniv = let infos = push_relevances infos na1 in convert_vect l2r infos (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible | FInt i1, FInt i2 -> if Uint63.equal i1 i2 then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible | FFloat f1, FFloat f2 -> if Float64.equal f1 f2 then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible (* Should not happen because both (hd1,v1) and (hd2,v2) are in whnf *) | ( (FLetIn _, _) | (FCaseT _,_) | (FApp _,_) | (FCLOS _,_) | (FLIFT _,_) | (_, FLetIn _) | (_,FCaseT _) | (_,FApp _) | (_,FCLOS _) | (_,FLIFT _) | (FLOCKED,_) | (_,FLOCKED) ) -> assert false | (FRel _ | FAtom _ | FInd _ | FFix _ | FCoFix _ | FProd _ | FEvar _ | FInt _ | FFloat _), _ -> raise NotConvertible and convert_stacks l2r infos lft1 lft2 stk1 stk2 cuniv = let f (l1, t1) (l2, t2) cuniv = ccnv CONV l2r infos l1 l2 t1 t2 cuniv in let rec cmp_rec pstk1 pstk2 cuniv = match (pstk1,pstk2) with | (z1::s1, z2::s2) -> let cu1 = cmp_rec s1 s2 cuniv in (match (z1,z2) with | (Zlapp a1,Zlapp a2) -> Array.fold_right2 f a1 a2 cu1 | (Zlproj (c1,_l1),Zlproj (c2,_l2)) -> if not (Projection.Repr.equal c1 c2) then raise NotConvertible else cu1 | (Zlfix(fx1,a1),Zlfix(fx2,a2)) -> let cu2 = f fx1 fx2 cu1 in cmp_rec a1 a2 cu2 | (Zlcase(ci1,l1,p1,br1,e1),Zlcase(ci2,l2,p2,br2,e2)) -> if not (eq_ind ci1.ci_ind ci2.ci_ind) then raise NotConvertible; let cu2 = f (l1, mk_clos e1 p1) (l2, mk_clos e2 p2) cu1 in convert_branches l2r infos ci1 e1 e2 l1 l2 br1 br2 cu2 | (Zlprimitive(op1,_,rargs1,kargs1),Zlprimitive(op2,_,rargs2,kargs2)) -> if not (CPrimitives.equal op1 op2) then raise NotConvertible else let cu2 = List.fold_right2 f rargs1 rargs2 cu1 in let fk (_,a1) (_,a2) cu = f a1 a2 cu in List.fold_right2 fk kargs1 kargs2 cu2 | ((Zlapp _ | Zlproj _ | Zlfix _| Zlcase _| Zlprimitive _), _) -> assert false) | _ -> cuniv in if compare_stack_shape stk1 stk2 then cmp_rec (pure_stack lft1 stk1) (pure_stack lft2 stk2) cuniv else raise NotConvertible and convert_vect l2r infos lft1 lft2 v1 v2 cuniv = let lv1 = Array.length v1 in let lv2 = Array.length v2 in if Int.equal lv1 lv2 then let rec fold n cuniv = if n >= lv1 then cuniv else let cuniv = ccnv CONV l2r infos lft1 lft2 v1.(n) v2.(n) cuniv in fold (n+1) cuniv in fold 0 cuniv else raise NotConvertible and convert_branches l2r infos ci e1 e2 lft1 lft2 br1 br2 cuniv = (** Skip comparison of the pattern types. We know that the two terms are living in a common type, thus this check is useless. *) let fold n c1 c2 cuniv = match skip_pattern infos n c1 c2 with | (infos, c1, c2) -> let lft1 = el_liftn n lft1 in let lft2 = el_liftn n lft2 in let e1 = subs_liftn n e1 in let e2 = subs_liftn n e2 in ccnv CONV l2r infos lft1 lft2 (mk_clos e1 c1) (mk_clos e2 c2) cuniv | exception IrregularPatternShape -> (** Might happen due to a shape invariant that is not enforced *) ccnv CONV l2r infos lft1 lft2 (mk_clos e1 c1) (mk_clos e2 c2) cuniv in Array.fold_right3 fold ci.ci_cstr_nargs br1 br2 cuniv let clos_gen_conv trans cv_pb l2r evars env univs t1 t2 = let reds = CClosure.RedFlags.red_add_transparent betaiotazeta trans in let infos = create_clos_infos ~evars reds env in let infos = { cnv_inf = infos; relevances = Range.empty; lft_tab = create_tab (); rgt_tab = create_tab (); } in ccnv cv_pb l2r infos el_id el_id (inject t1) (inject t2) univs let check_eq univs u u' = if not (UGraph.check_eq univs u u') then raise NotConvertible let check_leq univs u u' = if not (UGraph.check_leq univs u u') then raise NotConvertible let check_sort_cmp_universes env pb s0 s1 univs = let open Sorts in if not (type_in_type env) then let check_pb u0 u1 = match pb with | CUMUL -> check_leq univs u0 u1 | CONV -> check_eq univs u0 u1 in match (s0,s1) with | SProp, SProp | Prop, Prop | Set, Set -> () | SProp, _ | _, SProp -> raise NotConvertible | Prop, (Set | Type _) -> if not (is_cumul pb) then raise NotConvertible | Set, Prop -> raise NotConvertible | Set, Type u -> check_pb Univ.type0_univ u | Type _u, Prop -> raise NotConvertible | Type u, Set -> check_pb u Univ.type0_univ | Type u0, Type u1 -> check_pb u0 u1 let checked_sort_cmp_universes env pb s0 s1 univs = check_sort_cmp_universes env pb s0 s1 univs; univs let check_convert_instances ~flex:_ u u' univs = if UGraph.check_eq_instances univs u u' then univs else raise NotConvertible (* general conversion and inference functions *) let check_inductive_instances cv_pb variance u1 u2 univs = let csts = get_cumulativity_constraints cv_pb variance u1 u2 in if (UGraph.check_constraints csts univs) then univs else raise NotConvertible let checked_universes = { compare_sorts = checked_sort_cmp_universes; compare_instances = check_convert_instances; compare_cumul_instances = check_inductive_instances; } let infer_eq (univs, cstrs as cuniv) u u' = if UGraph.check_eq univs u u' then cuniv else univs, (Univ.enforce_eq u u' cstrs) let infer_leq (univs, cstrs as cuniv) u u' = if UGraph.check_leq univs u u' then cuniv else let cstrs', _ = UGraph.enforce_leq_alg u u' univs in univs, Univ.Constraint.union cstrs cstrs' let infer_cmp_universes env pb s0 s1 univs = if type_in_type env then univs else let open Sorts in let infer_pb u0 u1 = match pb with | CUMUL -> infer_leq univs u0 u1 | CONV -> infer_eq univs u0 u1 in match (s0,s1) with | SProp, SProp | Prop, Prop | Set, Set -> univs | SProp, _ | _, SProp -> raise NotConvertible | Prop, (Set | Type _) -> if not (is_cumul pb) then raise NotConvertible else univs | Set, Prop -> raise NotConvertible | Set, Type u -> infer_pb Univ.type0_univ u | Type u, Prop -> infer_pb u Univ.type0m_univ | Type u, Set -> infer_pb u Univ.type0_univ | Type u0, Type u1 -> infer_pb u0 u1 let infer_convert_instances ~flex u u' (univs,cstrs) = let cstrs' = if flex then if UGraph.check_eq_instances univs u u' then cstrs else raise NotConvertible else Univ.enforce_eq_instances u u' cstrs in (univs, cstrs') let infer_inductive_instances cv_pb variance u1 u2 (univs,csts') = let csts = get_cumulativity_constraints cv_pb variance u1 u2 in (univs, Univ.Constraint.union csts csts') let inferred_universes : (UGraph.t * Univ.Constraint.t) universe_compare = { compare_sorts = infer_cmp_universes; compare_instances = infer_convert_instances; compare_cumul_instances = infer_inductive_instances; } let gen_conv cv_pb l2r reds env evars univs t1 t2 = let b = if cv_pb = CUMUL then leq_constr_univs univs t1 t2 else eq_constr_univs univs t1 t2 in if b then () else let _ = clos_gen_conv reds cv_pb l2r evars env (univs, checked_universes) t1 t2 in () (* Profiling *) let gen_conv cv_pb ?(l2r=false) ?(reds=TransparentState.full) env ?(evars=(fun _->None), universes env) = let evars, univs = evars in if Flags.profile then let fconv_universes_key = CProfile.declare_profile "trans_fconv_universes" in CProfile.profile8 fconv_universes_key gen_conv cv_pb l2r reds env evars univs else gen_conv cv_pb l2r reds env evars univs let conv = gen_conv CONV let conv_leq = gen_conv CUMUL let generic_conv cv_pb ~l2r evars reds env univs t1 t2 = let (s, _) = clos_gen_conv reds cv_pb l2r evars env univs t1 t2 in s let infer_conv_universes cv_pb l2r evars reds env univs t1 t2 = let b, cstrs = if cv_pb == CUMUL then Constr.leq_constr_univs_infer univs t1 t2 else Constr.eq_constr_univs_infer univs t1 t2 in if b then cstrs else let univs = ((univs, Univ.Constraint.empty), inferred_universes) in let ((_,cstrs), _) = clos_gen_conv reds cv_pb l2r evars env univs t1 t2 in cstrs (* Profiling *) let infer_conv_universes = if Flags.profile then let infer_conv_universes_key = CProfile.declare_profile "infer_conv_universes" in CProfile.profile8 infer_conv_universes_key infer_conv_universes else infer_conv_universes let infer_conv ?(l2r=false) ?(evars=fun _ -> None) ?(ts=TransparentState.full) env univs t1 t2 = infer_conv_universes CONV l2r evars ts env univs t1 t2 let infer_conv_leq ?(l2r=false) ?(evars=fun _ -> None) ?(ts=TransparentState.full) env univs t1 t2 = infer_conv_universes CUMUL l2r evars ts env univs t1 t2 let default_conv cv_pb ?l2r:_ env t1 t2 = gen_conv cv_pb env t1 t2 let default_conv_leq = default_conv CUMUL (* let convleqkey = CProfile.declare_profile "Kernel_reduction.conv_leq";; let conv_leq env t1 t2 = CProfile.profile4 convleqkey conv_leq env t1 t2;; let convkey = CProfile.declare_profile "Kernel_reduction.conv";; let conv env t1 t2 = CProfile.profile4 convleqkey conv env t1 t2;; *) (* Application with on-the-fly reduction *) let beta_applist c l = let rec app subst c l = match kind c, l with | Lambda(_,_,c), arg::l -> app (arg::subst) c l | _ -> Term.applist (substl subst c, l) in app [] c l let beta_appvect c v = beta_applist c (Array.to_list v) let beta_app c a = beta_applist c [a] (* Compatibility *) let betazeta_appvect = Term.lambda_appvect_assum (********************************************************************) (* Special-Purpose Reduction *) (********************************************************************) (* pseudo-reduction rule: * [hnf_prod_app env (Prod(_,B)) N --> B[N] * with an HNF on the first argument to produce a product. * if this does not work, then we use the string S as part of our * error message. *) let hnf_prod_app env t n = match kind (whd_all env t) with | Prod (_,_,b) -> subst1 n b | _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product.") let hnf_prod_applist env t nl = List.fold_left (hnf_prod_app env) t nl let hnf_prod_applist_assum env n c l = let rec app n subst t l = if Int.equal n 0 then if l == [] then substl subst t else anomaly (Pp.str "Too many arguments.") else match kind (whd_allnolet env t), l with | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l | _, [] -> anomaly (Pp.str "Not enough arguments.") | _ -> anomaly (Pp.str "Not enough prod/let's.") in app n [] c l (* Dealing with arities *) let dest_prod env = let rec decrec env m c = let t = whd_all env c in match kind t with | Prod (n,a,c0) -> let d = LocalAssum (n,a) in decrec (push_rel d env) (Context.Rel.add d m) c0 | _ -> m,t in decrec env Context.Rel.empty let dest_lam env = let rec decrec env m c = let t = whd_all env c in match kind t with | Lambda (n,a,c0) -> let d = LocalAssum (n,a) in decrec (push_rel d env) (Context.Rel.add d m) c0 | _ -> m,t in decrec env Context.Rel.empty (* The same but preserving lets in the context, not internal ones. *) let dest_prod_assum env = let rec prodec_rec env l ty = let rty = whd_allnolet env ty in match kind rty with | Prod (x,t,c) -> let d = LocalAssum (x,t) in prodec_rec (push_rel d env) (Context.Rel.add d l) c | LetIn (x,b,t,c) -> let d = LocalDef (x,b,t) in prodec_rec (push_rel d env) (Context.Rel.add d l) c | _ -> let rty' = whd_all env rty in if Constr.equal rty' rty then l, rty else prodec_rec env l rty' in prodec_rec env Context.Rel.empty let dest_lam_assum env = let rec lamec_rec env l ty = let rty = whd_allnolet env ty in match kind rty with | Lambda (x,t,c) -> let d = LocalAssum (x,t) in lamec_rec (push_rel d env) (Context.Rel.add d l) c | LetIn (x,b,t,c) -> let d = LocalDef (x,b,t) in lamec_rec (push_rel d env) (Context.Rel.add d l) c | _ -> l,rty in lamec_rec env Context.Rel.empty exception NotArity let dest_arity env c = let l, c = dest_prod_assum env c in match kind c with | Sort s -> l,s | _ -> raise NotArity let is_arity env c = try let _ = dest_arity env c in true with NotArity -> false let eta_expand env t ty = let ctxt, _codom = dest_prod env ty in let ctxt',t = dest_lam env t in let d = Context.Rel.nhyps ctxt - Context.Rel.nhyps ctxt' in let eta_args = List.rev_map mkRel (List.interval 1 d) in let t = Term.applistc (Vars.lift d t) eta_args in let t = Term.it_mkLambda_or_LetIn t (List.firstn d ctxt) in Term.it_mkLambda_or_LetIn t ctxt' coq-8.11.0/kernel/byterun/0000755000175000017500000000000013612664533015236 5ustar treinentreinencoq-8.11.0/kernel/byterun/coq_memory.h0000644000175000017500000000325513612664533017566 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ #ifndef _COQ_MEMORY_ #define _COQ_MEMORY_ #include #include #include #include #include #define Coq_stack_size (4096 * sizeof(value)) #define Coq_stack_threshold (256 * sizeof(value)) /* see kernel/cbytegen.ml */ #define Coq_max_stack_size (256 * 1024) #define TRANSP 0 #define BOXED 1 /* stack */ extern value * coq_stack_low; extern value * coq_stack_high; extern value * coq_stack_threshold; /* global_data */ extern int coq_all_transp; extern int drawinstr; /* interp state */ extern value * coq_sp; /* Some predefined pointer code */ extern code_t accumulate; /* functions over global environment */ value coq_static_alloc(value size); /* ML */ value init_coq_vm(value unit); /* ML */ value re_init_coq_vm(value unit); /* ML */ void realloc_coq_stack(asize_t required_space); value coq_set_transp_value(value transp); /* ML */ value get_coq_transp_value(value unit); /* ML */ #endif /* _COQ_MEMORY_ */ value coq_set_drawinstr(value unit); coq-8.11.0/kernel/byterun/coq_uint63_native.h0000644000175000017500000001261513612664533020754 0ustar treinentreinen#define Is_uint63(v) (Is_long(v)) #define uint_of_value(val) (((uint64_t)(val)) >> 1) #define uint63_of_value(val) ((uint64_t)(val) >> 1) /* 2^63 * y + x as a value */ //#define Val_intint(x,y) ((value)(((uint64_t)(x)) << 1 + ((uint64_t)(y) << 64))) #define uint63_zero ((value) 1) /* 2*0 + 1 */ #define uint63_one() ((value) 3) /* 2*1 + 1 */ #define uint63_eq(x,y) ((x) == (y)) #define Uint63_eq(r,x,y) ((r) = uint63_eq(x,y)) #define Uint63_eq0(r,x) ((r) = ((x) == (uint64_t)1)) #define uint63_lt(x,y) ((uint64_t) (x) < (uint64_t) (y)) #define Uint63_lt(r,x,y) ((r) = uint63_lt(x,y)) #define uint63_leq(x,y) ((uint64_t) (x) <= (uint64_t) (y)) #define Uint63_leq(r,x,y) ((r) = uint63_leq(x,y)) #define Uint63_add(x,y) (accu = (value)((uint64_t) (x) + (uint64_t) (y) - 1)) #define Uint63_addcarry(x,y) (accu = (value)((uint64_t) (x) + (uint64_t) (y) + 1)) #define Uint63_sub(x,y) (accu = (value)((uint64_t) (x) - (uint64_t) (y) + 1)) #define Uint63_subcarry(x,y) (accu = (value)((uint64_t) (x) - (uint64_t) (y) - 1)) #define Uint63_mul(x,y) (accu = Val_long(uint63_of_value(x) * uint63_of_value(y))) #define Uint63_div(x,y) (accu = Val_long(uint63_of_value(x) / uint63_of_value(y))) #define Uint63_mod(x,y) (accu = Val_long(uint63_of_value(x) % uint63_of_value(y))) #define Uint63_lxor(x,y) (accu = (value)(((uint64_t)(x) ^ (uint64_t)(y)) | 1)) #define Uint63_lor(x,y) (accu = (value)((uint64_t)(x) | (uint64_t)(y))) #define Uint63_land(x,y) (accu = (value)((uint64_t)(x) & (uint64_t)(y))) /* TODO: is + or | better? OCAML uses + */ /* TODO: is - or ^ better? */ #define Uint63_lsl(x,y) do{ \ value uint63_lsl_y__ = (y); \ if (uint63_lsl_y__ < (uint64_t) 127) \ accu = (value)((((uint64_t)(x)-1) << uint63_of_value(uint63_lsl_y__)) | 1); \ else \ accu = uint63_zero; \ }while(0) #define Uint63_lsr(x,y) do{ \ value uint63_lsl_y__ = (y); \ if (uint63_lsl_y__ < (uint64_t) 127) \ accu = (value)(((uint64_t)(x) >> uint63_of_value(uint63_lsl_y__)) | 1); \ else \ accu = uint63_zero; \ }while(0) #define Uint63_lsl1(x) (accu = (value)((((uint64_t)(x)-1) << 1) +1)) #define Uint63_lsr1(x) (accu = (value)(((uint64_t)(x) >> 1) |1)) /* addmuldiv(p,x,y) = x * 2^p + y / 2 ^ (63 - p) */ /* (modulo 2^63) for p <= 63 */ value uint63_addmuldiv(uint64_t p, uint64_t x, uint64_t y) { uint64_t shiftby = uint63_of_value(p); value r = (value)((uint64_t)(x^1) << shiftby); r |= ((uint64_t)y >> (63-shiftby)) | 1; return r; } #define Uint63_addmuldiv(p, x, y) (accu = uint63_addmuldiv(p, x, y)) value uint63_head0(uint64_t x) { int r = 0; if (!(x & 0xFFFFFFFF00000000)) { x <<= 32; r += 32; } if (!(x & 0xFFFF000000000000)) { x <<= 16; r += 16; } if (!(x & 0xFF00000000000000)) { x <<= 8; r += 8; } if (!(x & 0xF000000000000000)) { x <<= 4; r += 4; } if (!(x & 0xC000000000000000)) { x <<= 2; r += 2; } if (!(x & 0x8000000000000000)) { x <<=1; r += 1; } return Val_int(r); } #define Uint63_head0(x) (accu = uint63_head0(x)) value uint63_tail0(value x) { int r = 0; x = (uint64_t)x >> 1; if (!(x & 0xFFFFFFFF)) { x >>= 32; r += 32; } if (!(x & 0x0000FFFF)) { x >>= 16; r += 16; } if (!(x & 0x000000FF)) { x >>= 8; r += 8; } if (!(x & 0x0000000F)) { x >>= 4; r += 4; } if (!(x & 0x00000003)) { x >>= 2; r += 2; } if (!(x & 0x00000001)) { x >>=1; r += 1; } return Val_int(r); } #define Uint63_tail0(x) (accu = uint63_tail0(x)) value uint63_mulc(value x, value y, value* h) { x = (uint64_t)x >> 1; y = (uint64_t)y >> 1; uint64_t lx = x & 0xFFFFFFFF; uint64_t ly = y & 0xFFFFFFFF; uint64_t hx = x >> 32; uint64_t hy = y >> 32; uint64_t hr = hx * hy; uint64_t lr = lx * ly; lx *= hy; ly *= hx; hr += (lx >> 32) + (ly >> 32); lx <<= 32; lr += lx; if (lr < lx) { hr++; } ly <<= 32; lr += ly; if (lr < ly) { hr++; } hr = (hr << 1) | (lr >> 63); *h = Val_int(hr); return Val_int(lr); } #define Uint63_mulc(x, y, h) (accu = uint63_mulc(x, y, h)) #define lt128(xh,xl,yh,yl) (uint63_lt(xh,yh) || (uint63_eq(xh,yh) && uint63_lt(xl,yl))) #define le128(xh,xl,yh,yl) (uint63_lt(xh,yh) || (uint63_eq(xh,yh) && uint63_leq(xl,yl))) #define maxuint63 ((uint64_t)0x7FFFFFFFFFFFFFFF) /* precondition: xh < y */ /* outputs r and sets ql to q s.t. x = q * y + r, r < y */ static value uint63_div21_aux(value xh, value xl, value y, value* ql) { uint64_t nh = uint63_of_value(xh); uint64_t nl = uint63_of_value(xl); y = uint63_of_value(y); uint64_t q = 0; for (int i = 0; i < 63; ++i) { // invariants: 0 <= nh < y, nl = (xl*2^i) % 2^64, // (q*y + nh) * 2^(63-i) + (xl % 2^(63-i)) = (xh%y) * 2^63 + xl nl <<= 1; nh = (nh << 1) | (nl >> 63); q <<= 1; if (nh >= y) { q |= 1; nh -= y; } } *ql = Val_int(q); return Val_int(nh); } value uint63_div21(value xh, value xl, value y, value* ql) { if (uint63_leq(y, xh)) { *ql = Val_int(0); return Val_int(0); } else { return uint63_div21_aux(xh, xl, y, ql); } } #define Uint63_div21(xh, xl, y, q) (accu = uint63_div21(xh, xl, y, q)) #define Uint63_to_double(x) Coq_copy_double((double) uint63_of_value(x)) double coq_uint63_to_float(value x) { return (double) uint63_of_value(x); } value coq_uint63_to_float_byte(value x) { return caml_copy_double(coq_uint63_to_float(x)); } #define Uint63_of_double(f) do{ \ accu = Val_long((uint64_t)(f)); \ }while(0) #define Uint63_of_int(x) (accu = (x)) #define Uint63_to_int_min(n, m) do { \ if (uint63_lt((n),(m))) \ accu = (n); \ else \ accu = (m); \ }while(0) coq-8.11.0/kernel/byterun/libcoqrun.clib0000644000175000017500000000006613612664533020071 0ustar treinentreinencoq_fix_code.o coq_memory.o coq_values.o coq_interp.o coq-8.11.0/kernel/byterun/coq_fix_code.h0000644000175000017500000000240013612664533020025 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ #ifndef _COQ_FIX_CODE_ #define _COQ_FIX_CODE_ #include void * coq_stat_alloc (asize_t sz); #ifdef THREADED_CODE extern char ** coq_instr_table; extern char * coq_instr_base; void init_arity(); #define VALINSTR(instr) ((opcode_t)(coq_instr_table[instr] - coq_instr_base)) #else #define VALINSTR(instr) instr #endif /* THREADED_CODE */ #define Is_instruction(pc,instr) (*pc == VALINSTR(instr)) value coq_tcode_of_code(value code); value coq_makeaccu (value i); value coq_pushpop (value i); value coq_accucond (value i); value coq_is_accumulate_code(value code); #endif /* _COQ_FIX_CODE_ */ coq-8.11.0/kernel/byterun/coq_float64.h0000644000175000017500000000377413612664533017543 0ustar treinentreinen#ifndef _COQ_FLOAT64_ #define _COQ_FLOAT64_ #include #define DECLARE_FREL(name, e) \ int coq_##name(double x, double y) { \ return e; \ } \ \ value coq_##name##_byte(value x, value y) { \ return coq_##name(Double_val(x), Double_val(y)); \ } #define DECLARE_FBINOP(name, e) \ double coq_##name(double x, double y) { \ return e; \ } \ \ value coq_##name##_byte(value x, value y) { \ return caml_copy_double(coq_##name(Double_val(x), Double_val(y))); \ } #define DECLARE_FUNOP(name, e) \ double coq_##name(double x) { \ return e; \ } \ \ value coq_##name##_byte(value x) { \ return caml_copy_double(coq_##name(Double_val(x))); \ } DECLARE_FREL(feq, x == y) DECLARE_FREL(flt, x < y) DECLARE_FREL(fle, x <= y) DECLARE_FBINOP(fmul, x * y) DECLARE_FBINOP(fadd, x + y) DECLARE_FBINOP(fsub, x - y) DECLARE_FBINOP(fdiv, x / y) DECLARE_FUNOP(fsqrt, sqrt(x)) DECLARE_FUNOP(next_up, nextafter(x, INFINITY)) DECLARE_FUNOP(next_down, nextafter(x, -INFINITY)) value coq_is_double(value x) { return Val_long(Is_double(x)); } #endif /* _COQ_FLOAT64_ */ coq-8.11.0/kernel/byterun/coq_uint63_emul.h0000644000175000017500000001171613612664533020431 0ustar treinentreinen# pragma once # include # include #define Is_uint63(v) (Tag_val(v) == Custom_tag) #define uint_of_value(val) (((uint32_t)(val)) >> 1) # define DECLARE_NULLOP(name) \ value uint63_##name() { \ static value* cb = 0; \ CAMLparam0(); \ if (!cb) cb = caml_named_value("uint63 " #name); \ CAMLreturn(*cb); \ } # define DECLARE_UNOP(name) \ value uint63_##name##_ml(value x) { \ static value* cb = 0; \ CAMLparam1(x); \ if (!cb) cb = caml_named_value("uint63 " #name); \ CAMLreturn(caml_callback(*cb, x)); \ } # define CALL_UNOP_NOASSIGN(name, x) \ value uint63_return_value__; \ value uint63_arg_x__ = (x); \ Setup_for_gc; \ uint63_return_value__ = uint63_##name##_ml(uint63_arg_x__); \ Restore_after_gc # define CALL_UNOP(name, x) do{ \ CALL_UNOP_NOASSIGN(name, x); \ accu = uint63_return_value__; \ }while(0) # define CALL_PREDICATE(r, name, x) do{ \ CALL_UNOP_NOASSIGN(name, x); \ (r) = Int_val(uint63_return_value__); \ }while(0) # define DECLARE_BINOP(name) \ value uint63_##name##_ml(value x, value y) { \ static value* cb = 0; \ CAMLparam2(x, y); \ if (!cb) cb = caml_named_value("uint63 " #name); \ CAMLreturn(caml_callback2(*cb, x, y)); \ } # define CALL_BINOP_NOASSIGN(name, x, y) \ value uint63_return_value__; \ value uint63_arg_x__ = (x); \ value uint63_arg_y__ = (y); \ Setup_for_gc; \ uint63_return_value__ = uint63_##name##_ml(uint63_arg_x__, uint63_arg_y__); \ Restore_after_gc # define CALL_BINOP(name, x, y) do{ \ CALL_BINOP_NOASSIGN(name, x, y); \ accu = uint63_return_value__; \ }while(0) # define CALL_RELATION(r, name, x, y) do{ \ CALL_BINOP_NOASSIGN(name, x, y); \ (r) = Int_val(uint63_return_value__); \ }while(0) # define DECLARE_TEROP(name) \ value uint63_##name##_ml(value x, value y, value z) { \ static value* cb = 0; \ CAMLparam3(x, y, z); \ if (!cb) cb = caml_named_value("uint63 " #name); \ CAMLreturn(caml_callback3(*cb, x, y, z)); \ } # define CALL_TEROP(name, x, y, z) do{ \ value uint63_return_value__; \ value uint63_arg_x__ = (x); \ value uint63_arg_y__ = (y); \ value uint63_arg_z__ = (z); \ Setup_for_gc; \ uint63_return_value__ = uint63_##name##_ml(uint63_arg_x__, uint63_arg_y__, uint63_arg_z__); \ Restore_after_gc; \ accu = uint63_return_value__; \ }while(0) DECLARE_NULLOP(one) DECLARE_BINOP(add) #define Uint63_add(x, y) CALL_BINOP(add, x, y) DECLARE_BINOP(addcarry) #define Uint63_addcarry(x, y) CALL_BINOP(addcarry, x, y) DECLARE_TEROP(addmuldiv) #define Uint63_addmuldiv(x, y, z) CALL_TEROP(addmuldiv, x, y, z) DECLARE_BINOP(div) #define Uint63_div(x, y) CALL_BINOP(div, x, y) DECLARE_BINOP(eq) #define Uint63_eq(r, x, y) CALL_RELATION(r, eq, x, y) DECLARE_UNOP(eq0) #define Uint63_eq0(r, x) CALL_PREDICATE(r, eq0, x) DECLARE_UNOP(head0) #define Uint63_head0(x) CALL_UNOP(head0, x) DECLARE_BINOP(land) #define Uint63_land(x, y) CALL_BINOP(land, x, y) DECLARE_BINOP(leq) #define Uint63_leq(r, x, y) CALL_RELATION(r, leq, x, y) DECLARE_BINOP(lor) #define Uint63_lor(x, y) CALL_BINOP(lor, x, y) DECLARE_BINOP(lsl) #define Uint63_lsl(x, y) CALL_BINOP(lsl, x, y) DECLARE_UNOP(lsl1) #define Uint63_lsl1(x) CALL_UNOP(lsl1, x) DECLARE_BINOP(lsr) #define Uint63_lsr(x, y) CALL_BINOP(lsr, x, y) DECLARE_UNOP(lsr1) #define Uint63_lsr1(x) CALL_UNOP(lsr1, x) DECLARE_BINOP(lt) #define Uint63_lt(r, x, y) CALL_RELATION(r, lt, x, y) DECLARE_BINOP(lxor) #define Uint63_lxor(x, y) CALL_BINOP(lxor, x, y) DECLARE_BINOP(mod) #define Uint63_mod(x, y) CALL_BINOP(mod, x, y) DECLARE_BINOP(mul) #define Uint63_mul(x, y) CALL_BINOP(mul, x, y) DECLARE_BINOP(sub) #define Uint63_sub(x, y) CALL_BINOP(sub, x, y) DECLARE_BINOP(subcarry) #define Uint63_subcarry(x, y) CALL_BINOP(subcarry, x, y) DECLARE_UNOP(tail0) #define Uint63_tail0(x) CALL_UNOP(tail0, x) DECLARE_TEROP(div21_ml) # define Uint63_div21(x, y, z, q) do{ \ value uint63_return_value__; \ value uint63_arg_x__ = (x); \ value uint63_arg_y__ = (y); \ value uint63_arg_z__ = (z); \ Setup_for_gc; \ uint63_return_value__ = \ uint63_div21_ml_ml(uint63_arg_x__, uint63_arg_y__, uint63_arg_z__); \ Restore_after_gc; \ *q = Field(uint63_return_value__, 0); \ accu = Field(uint63_return_value__, 1); \ }while(0) DECLARE_BINOP(mulc_ml) # define Uint63_mulc(x, y, h) do{ \ value uint63_return_value__; \ value uint63_arg_x__ = (x); \ value uint63_arg_y__ = (y); \ Setup_for_gc; \ uint63_return_value__ = \ uint63_mulc_ml_ml(uint63_arg_x__, uint63_arg_y__); \ Restore_after_gc; \ *(h) = Field(uint63_return_value__, 0); \ accu = Field(uint63_return_value__, 1); \ }while(0) DECLARE_UNOP(to_float) #define Uint63_to_double(x) CALL_UNOP(to_float, x) DECLARE_UNOP(of_float) #define Uint63_of_double(f) do{ \ Coq_copy_double(f); \ CALL_UNOP(of_float, accu); \ }while(0) DECLARE_UNOP(of_int) #define Uint63_of_int(x) CALL_UNOP(of_int, x) DECLARE_BINOP(to_int_min) #define Uint63_to_int_min(n, m) CALL_BINOP(to_int_min, n, m) coq-8.11.0/kernel/byterun/coq_interp.h0000644000175000017500000000216513612664533017556 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ value coq_push_ra(value tcode); value coq_push_val(value v); value coq_push_arguments(value args); value coq_push_vstack(value stk); value coq_interprete_ml(value tcode, value a, value t, value g, value e, value ea); value coq_interprete_byte(value* argv, int argn); value coq_interprete (code_t coq_pc, value coq_accu, value coq_atom_tbl, value coq_global_data, value coq_env, long coq_extra_args); value coq_eval_tcode (value tcode, value t, value g, value e); coq-8.11.0/kernel/byterun/coq_values.h0000644000175000017500000000370513612664533017555 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ #ifndef _COQ_VALUES_ #define _COQ_VALUES_ #include #include #include #define Default_tag 0 #define Accu_tag 0 #define ATOM_ID_TAG 0 #define ATOM_INDUCTIVE_TAG 1 #define ATOM_TYPE_TAG 2 #define ATOM_PROJ_TAG 3 #define ATOM_FIX_TAG 4 #define ATOM_SWITCH_TAG 5 #define ATOM_COFIX_TAG 6 #define ATOM_COFIXEVALUATED_TAG 7 /* Les blocs accumulate */ #define Is_accu(v) (Is_block(v) && (Tag_val(v) == Accu_tag)) #define IS_EVALUATED_COFIX(v) (Is_accu(v) && Is_block(Field(v,1)) && (Tag_val(Field(v,1)) == ATOM_COFIXEVALUATED_TAG)) #define Is_double(v) (Tag_val(v) == Double_tag) /* coq values for primitive operations */ #define coq_tag_C1 2 #define coq_tag_C0 1 #define coq_tag_pair 1 #define coq_true Val_int(0) #define coq_false Val_int(1) #define coq_Eq Val_int(0) #define coq_Lt Val_int(1) #define coq_Gt Val_int(2) #define coq_FEq Val_int(0) #define coq_FLt Val_int(1) #define coq_FGt Val_int(2) #define coq_FNotComparable Val_int(3) #define coq_PNormal Val_int(0) #define coq_NNormal Val_int(1) #define coq_PSubn Val_int(2) #define coq_NSubn Val_int(3) #define coq_PZero Val_int(4) #define coq_NZero Val_int(5) #define coq_PInf Val_int(6) #define coq_NInf Val_int(7) #define coq_NaN Val_int(8) #define FLOAT_EXP_SHIFT (2101) /* 2*emax + prec */ #endif /* _COQ_VALUES_ */ coq-8.11.0/kernel/byterun/coq_values.c0000644000175000017500000000561513612664533017552 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ #include #include #include "coq_fix_code.h" #include "coq_instruct.h" #include "coq_memory.h" #include "coq_values.h" #include /* KIND OF VALUES */ #define Setup_for_gc #define Restore_after_gc value coq_kind_of_closure(value v) { opcode_t * c; int is_app = 0; c = Code_val(v); if (Is_instruction(c, GRAB)) return Val_int(0); if (Is_instruction(c, RESTART)) {is_app = 1; c++;} if (Is_instruction(c, GRABREC)) return Val_int(1+is_app); if (Is_instruction(c, MAKEACCU)) return Val_int(3); return Val_int(0); } /* DESTRUCT ACCU */ value coq_closure_arity(value clos) { opcode_t * c = Code_val(clos); if (Is_instruction(c,RESTART)) { c++; if (Is_instruction(c,GRAB)) return Val_int(3 + c[1] - Wosize_val(clos)); else { if (Wosize_val(clos) != 2) caml_failwith("Coq Values : coq_closure_arity"); return Val_int(1); } } if (Is_instruction(c,GRAB)) return Val_int(1 + c[1]); return Val_int(1); } /* Fonction sur les fix */ value coq_offset(value v) { if (Tag_val(v) == Closure_tag) return Val_int(0); else return Val_long(-Wsize_bsize(Infix_offset_val(v))); } value coq_offset_closure(value v, value offset){ return (value)&Field(v, Int_val(offset)); } value coq_set_bytecode_field(value v, value i, value code) { // No write barrier because the bytecode does not live on the OCaml heap Field(v, Long_val(i)) = (value) Code_val(code); return Val_unit; } value coq_offset_tcode(value code,value offset){ CAMLparam1(code); CAMLlocal1(res); res = caml_alloc_small(1, Abstract_tag); Code_val(res) = Code_val(code) + Int_val(offset); CAMLreturn(res); } value coq_int_tcode(value pc, value offset) { code_t code = Code_val(pc); return Val_int(*((code_t) code + Int_val(offset))); } value coq_tcode_array(value tcodes) { CAMLparam1(tcodes); CAMLlocal2(res, tmp); int i; /* Assumes that the vector of types is small. This was implicit in the previous code which was building the type array using Alloc_small. */ res = caml_alloc_small(Wosize_val(tcodes), Default_tag); for (i = 0; i < Wosize_val(tcodes); i++) { tmp = caml_alloc_small(1, Abstract_tag); Code_val(tmp) = (code_t) Field(tcodes, i); Store_field(res, i, tmp); } CAMLreturn(res); } coq-8.11.0/kernel/byterun/dune0000644000175000017500000000067513612664533016124 0ustar treinentreinen(library (name byterun) (synopsis "Coq's Kernel Abstract Reduction Machine [C implementation]") (public_name coq.vm) (c_flags (:include %{project_root}/config/dune.c_flags)) (c_names coq_fix_code coq_memory coq_values coq_interp)) (rule (targets coq_instruct.h) (action (with-stdout-to %{targets} (run ../genOpcodeFiles.exe enum)))) (rule (targets coq_jumptbl.h) (action (with-stdout-to %{targets} (run ../genOpcodeFiles.exe jump)))) coq-8.11.0/kernel/byterun/coq_gc.h0000644000175000017500000000453713612664533016653 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ #ifndef _COQ_CAML_GC_ #define _COQ_CAML_GC_ #include #include #include typedef void (*scanning_action) (value, value *); CAMLextern char *young_ptr; CAMLextern char *young_limit; CAMLextern void (*scan_roots_hook) (scanning_action); CAMLextern void minor_collection (void); #define Caml_white (0 << 8) #define Caml_black (3 << 8) #ifdef HAS_OCP_MEMPROF /* This code is necessary to make the OCamlPro memory profiling branch of OCaml compile. */ #define Make_header(wosize, tag, color) \ caml_make_header(wosize, tag, color) #else #define Make_header(wosize, tag, color) \ (((header_t) (((header_t) (wosize) << 10) \ + (color) \ + (tag_t) (tag))) \ ) #endif #define Alloc_small(result, wosize, tag) do{ \ young_ptr -= Bhsize_wosize (wosize); \ if (young_ptr < young_limit){ \ young_ptr += Bhsize_wosize (wosize); \ Setup_for_gc; \ minor_collection (); \ Restore_after_gc; \ young_ptr -= Bhsize_wosize (wosize); \ } \ Hd_hp (young_ptr) = Make_header ((wosize), (tag), Caml_black); \ (result) = Val_hp (young_ptr); \ }while(0) #endif /*_COQ_CAML_GC_ */ coq-8.11.0/kernel/byterun/coq_fix_code.c0000644000175000017500000001620113612664533020024 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ /* Arnaud Spiwack: expanded the virtual machine with operators used for fast computation of bounded (31bits) integers */ #include #include #include #include #include #include #include #include #include #include "coq_instruct.h" #include "coq_fix_code.h" #ifdef THREADED_CODE char ** coq_instr_table; char * coq_instr_base; int arity[STOP+1]; void init_arity () { /* instruction with zero operand */ arity[ACC0]=arity[ACC1]=arity[ACC2]=arity[ACC3]=arity[ACC4]=arity[ACC5]= arity[ACC6]=arity[ACC7]=arity[PUSH]=arity[PUSHACC0]=arity[PUSHACC1]= arity[PUSHACC2]=arity[PUSHACC3]=arity[PUSHACC4]=arity[PUSHACC5]= arity[PUSHACC6]=arity[PUSHACC7]=arity[ENVACC1]=arity[ENVACC2]= arity[ENVACC3]=arity[ENVACC4]=arity[PUSHENVACC1]=arity[PUSHENVACC2]= arity[PUSHENVACC3]=arity[PUSHENVACC4]=arity[APPLY1]=arity[APPLY2]= arity[APPLY3]=arity[APPLY4]=arity[RESTART]=arity[OFFSETCLOSUREM2]= arity[OFFSETCLOSURE0]=arity[OFFSETCLOSURE2]=arity[PUSHOFFSETCLOSUREM2]= arity[PUSHOFFSETCLOSURE0]=arity[PUSHOFFSETCLOSURE2]= arity[GETFIELD0]=arity[GETFIELD1]=arity[SETFIELD0]=arity[SETFIELD1]= arity[CONST0]=arity[CONST1]=arity[CONST2]=arity[CONST3]= arity[PUSHCONST0]=arity[PUSHCONST1]=arity[PUSHCONST2]=arity[PUSHCONST3]= arity[ACCUMULATE]=arity[STOP]=arity[MAKEPROD]= arity[ADDINT63]=arity[SUBINT63]=arity[LTINT63]=arity[LEINT63]= arity[LTFLOAT]=arity[LEFLOAT]= arity[ISINT]=arity[AREINT2]=0; /* instruction with one operand */ arity[ACC]=arity[PUSHACC]=arity[POP]=arity[ENVACC]=arity[PUSHENVACC]= arity[PUSH_RETADDR]=arity[APPLY]=arity[APPTERM1]=arity[APPTERM2]= arity[APPTERM3]=arity[RETURN]=arity[GRAB]=arity[OFFSETCLOSURE]= arity[PUSHOFFSETCLOSURE]=arity[GETGLOBAL]=arity[PUSHGETGLOBAL]= arity[MAKEBLOCK1]=arity[MAKEBLOCK2]=arity[MAKEBLOCK3]=arity[MAKEBLOCK4]= arity[MAKEACCU]=arity[CONSTINT]=arity[PUSHCONSTINT]=arity[GRABREC]= arity[PUSHFIELDS]=arity[GETFIELD]=arity[SETFIELD]= arity[BRANCH]=arity[ENSURESTACKCAPACITY]= arity[CHECKADDINT63]=arity[CHECKADDCINT63]=arity[CHECKADDCARRYCINT63]= arity[CHECKSUBINT63]=arity[CHECKSUBCINT63]=arity[CHECKSUBCARRYCINT63]= arity[CHECKMULINT63]=arity[CHECKMULCINT63]= arity[CHECKDIVINT63]=arity[CHECKMODINT63]=arity[CHECKDIVEUCLINT63]= arity[CHECKDIV21INT63]= arity[CHECKLXORINT63]=arity[CHECKLORINT63]=arity[CHECKLANDINT63]= arity[CHECKLSLINT63]=arity[CHECKLSRINT63]=arity[CHECKADDMULDIVINT63]= arity[CHECKLSLINT63CONST1]=arity[CHECKLSRINT63CONST1]= arity[CHECKEQINT63]=arity[CHECKLTINT63]=arity[CHECKLEINT63]= arity[CHECKCOMPAREINT63]=arity[CHECKHEAD0INT63]=arity[CHECKTAIL0INT63]= arity[CHECKEQFLOAT]=arity[CHECKLTFLOAT]=arity[CHECKLEFLOAT]= arity[CHECKOPPFLOAT]=arity[CHECKABSFLOAT]=arity[CHECKCOMPAREFLOAT]= arity[CHECKCLASSIFYFLOAT]= arity[CHECKADDFLOAT]=arity[CHECKSUBFLOAT]=arity[CHECKMULFLOAT]= arity[CHECKDIVFLOAT]=arity[CHECKSQRTFLOAT]= arity[CHECKFLOATOFINT63]=arity[CHECKFLOATNORMFRMANTISSA]= arity[CHECKFRSHIFTEXP]=arity[CHECKLDSHIFTEXP]= arity[CHECKNEXTUPFLOAT]=arity[CHECKNEXTDOWNFLOAT]=1; /* instruction with two operands */ arity[APPTERM]=arity[MAKEBLOCK]=arity[CLOSURE]= arity[PROJ]=2; /* instruction with four operands */ arity[MAKESWITCHBLOCK]=4; /* instruction with arbitrary operands */ arity[CLOSUREREC]=arity[CLOSURECOFIX]=arity[SWITCH]=0; } #endif /* THREADED_CODE */ void * coq_stat_alloc (asize_t sz) { void * result = malloc (sz); if (result == NULL) raise_out_of_memory (); return result; } value coq_makeaccu (value i) { CAMLparam1(i); CAMLlocal1(res); code_t q = coq_stat_alloc(2 * sizeof(opcode_t)); res = caml_alloc_small(1, Abstract_tag); Code_val(res) = q; *q++ = VALINSTR(MAKEACCU); *q = (opcode_t)Int_val(i); CAMLreturn(res); } value coq_pushpop (value i) { CAMLparam1(i); CAMLlocal1(res); code_t q; res = caml_alloc_small(1, Abstract_tag); int n = Int_val(i); if (n == 0) { q = coq_stat_alloc(sizeof(opcode_t)); Code_val(res) = q; *q = VALINSTR(STOP); CAMLreturn(res); } else { q = coq_stat_alloc(3 * sizeof(opcode_t)); Code_val(res) = q; *q++ = VALINSTR(POP); *q++ = (opcode_t)n; *q = VALINSTR(STOP); CAMLreturn(res); } } value coq_is_accumulate_code(value code){ code_t q = Code_val(code); int res; res = Is_instruction(q,ACCUMULATE); return Val_bool(res); } #ifdef ARCH_BIG_ENDIAN #define Reverse_32(dst,src) { \ char * _p, * _q; \ char _a, _b; \ _p = (char *) (src); \ _q = (char *) (dst); \ _a = _p[0]; \ _b = _p[1]; \ _q[0] = _p[3]; \ _q[1] = _p[2]; \ _q[3] = _a; \ _q[2] = _b; \ } #define COPY32(dst,src) Reverse_32(dst,src) #else #define COPY32(dst,src) (*dst=*src) #endif /* ARCH_BIG_ENDIAN */ value coq_tcode_of_code (value code) { CAMLparam1 (code); CAMLlocal1 (res); code_t p, q; asize_t len = (asize_t) caml_string_length(code); res = caml_alloc_small(1, Abstract_tag); q = coq_stat_alloc(len); Code_val(res) = q; len /= sizeof(opcode_t); for (p = (code_t)code; p < (code_t)code + len; /*nothing*/) { opcode_t instr; COPY32(&instr,p); p++; if (instr < 0 || instr > STOP){ instr = STOP; }; *q++ = VALINSTR(instr); if (instr == SWITCH) { uint32_t i, sizes, const_size, block_size; COPY32(q,p); p++; sizes=*q++; const_size = sizes & 0xFFFFFF; block_size = sizes >> 24; sizes = const_size + block_size; for(i=0; i #include #include #include #include "coq_gc.h" #include "coq_instruct.h" #include "coq_fix_code.h" #include "coq_memory.h" #include "coq_interp.h" /* stack */ value * coq_stack_low; value * coq_stack_high; value * coq_stack_threshold; asize_t coq_max_stack_size = Coq_max_stack_size; /* global_data */ int drawinstr; /* interp state */ long coq_saved_sp_offset; value * coq_sp; /* Some predefined pointer code */ code_t accumulate; /* functions over global environment */ void coq_stat_free (void * blk) { free (blk); } value coq_static_alloc(value size) /* ML */ { return (value) coq_stat_alloc((asize_t) Long_val(size)); } value accumulate_code(value unit) /* ML */ { CAMLparam1(unit); CAMLlocal1(res); res = caml_alloc_small(1, Abstract_tag); Code_val(res) = accumulate; CAMLreturn(res); } static void (*coq_prev_scan_roots_hook) (scanning_action); static void coq_scan_roots(scanning_action action) { register value * i; /* Scan the stack */ for (i = coq_sp; i < coq_stack_high; i++) { #ifdef NO_NAKED_POINTERS /* The VM stack may contain C-allocated bytecode */ if (Is_block(*i) && !Is_in_heap_or_young(*i)) continue; #endif (*action) (*i, i); }; /* Hook */ if (coq_prev_scan_roots_hook != NULL) (*coq_prev_scan_roots_hook)(action); } void init_coq_stack() { coq_stack_low = (value *) coq_stat_alloc(Coq_stack_size); coq_stack_high = coq_stack_low + Coq_stack_size / sizeof (value); coq_stack_threshold = coq_stack_low + Coq_stack_threshold / sizeof(value); coq_max_stack_size = Coq_max_stack_size; } void init_coq_interpreter() { coq_sp = coq_stack_high; coq_interprete(NULL, Val_unit, Atom(0), Atom(0), Val_unit, 0); } static int coq_vm_initialized = 0; value init_coq_vm(value unit) /* ML */ { if (coq_vm_initialized == 1) { fprintf(stderr,"already open \n");fflush(stderr);} else { drawinstr=0; #ifdef THREADED_CODE init_arity(); #endif /* THREADED_CODE */ /* Allocate the table of global and the stack */ init_coq_stack(); /* Initialing the interpreter */ init_coq_interpreter(); /* Some predefined pointer code. * It is typically contained in accumulator blocks whose tag is 0 and thus * scanned by the GC, so make it look like an OCaml block. */ value accu_block = (value) coq_stat_alloc(2 * sizeof(value)); Hd_hp (accu_block) = Make_header (1, Abstract_tag, Caml_black); \ accumulate = (code_t) Val_hp(accu_block); *accumulate = VALINSTR(ACCUMULATE); /* Initialize GC */ if (coq_prev_scan_roots_hook == NULL) coq_prev_scan_roots_hook = scan_roots_hook; scan_roots_hook = coq_scan_roots; coq_vm_initialized = 1; } return Val_unit;; } /* [required_space] is a size in words */ void realloc_coq_stack(asize_t required_space) { asize_t size; value * new_low, * new_high, * new_sp; size = coq_stack_high - coq_stack_low; do { size *= 2; } while (size < coq_stack_high - coq_sp + required_space); new_low = (value *) coq_stat_alloc(size * sizeof(value)); new_high = new_low + size; #define shift(ptr) \ ((char *) new_high - ((char *) coq_stack_high - (char *) (ptr))) new_sp = (value *) shift(coq_sp); memmove((char *) new_sp, (char *) coq_sp, (coq_stack_high - coq_sp) * sizeof(value)); coq_stat_free(coq_stack_low); coq_stack_low = new_low; coq_stack_high = new_high; coq_stack_threshold = coq_stack_low + Coq_stack_threshold / sizeof(value); coq_sp = new_sp; #undef shift } value coq_set_drawinstr(value unit) { drawinstr = 1; return Val_unit; } coq-8.11.0/kernel/byterun/coq_interp.c0000644000175000017500000013170713612664533017556 0ustar treinentreinen/***********************************************************************/ /* */ /* Coq Compiler */ /* */ /* Benjamin Gregoire, projets Logical and Cristal */ /* INRIA Rocquencourt */ /* */ /* */ /***********************************************************************/ /* The bytecode interpreter */ /* Spiwack: expanded the virtual machine with operators used for fast computation of bounded (31bits) integers */ #include #include #include #include #include #include #include #include "coq_gc.h" #include "coq_instruct.h" #include "coq_fix_code.h" #include "coq_memory.h" #include "coq_values.h" #include "coq_float64.h" #ifdef ARCH_SIXTYFOUR #include "coq_uint63_native.h" #else #include "coq_uint63_emul.h" #endif /* Registers for the abstract machine: pc the code pointer sp the stack pointer (grows downward) accu the accumulator env heap-allocated environment trapsp pointer to the current trap frame extra_args number of extra arguments provided by the caller sp is a local copy of the global variable extern_sp. */ /* Instruction decoding */ #ifdef THREADED_CODE # define Instruct(name) coq_lbl_##name: # if defined(ARCH_SIXTYFOUR) && !defined(ARCH_CODE32) # define coq_Jumptbl_base ((char *) &&coq_lbl_ACC0) # else # define coq_Jumptbl_base ((char *) 0) # define coq_jumptbl_base ((char *) 0) # endif # ifdef DEBUG # define Next goto next_instr # else # define Next goto *(void *)(coq_jumptbl_base + *pc++) # endif #else # define Instruct(name) case name: # define Next break #endif /* #define _COQ_DEBUG_ */ #ifdef _COQ_DEBUG_ # define print_instr(s) /*if (drawinstr)*/ printf("%s\n",s) # define print_int(i) /*if (drawinstr)*/ printf("%d\n",i) # define print_lint(i) /*if (drawinstr)*/ printf("%ld\n",i) # else # define print_instr(s) # define print_int(i) # define print_lint(i) #endif #define CHECK_STACK(num_args) { \ if (sp - num_args < coq_stack_threshold) { \ coq_sp = sp; \ realloc_coq_stack(num_args + Coq_stack_threshold / sizeof(value)); \ sp = coq_sp; \ } \ } /* GC interface */ #define Setup_for_gc { sp -= 2; sp[0] = accu; sp[1] = coq_env; coq_sp = sp; } #define Restore_after_gc { accu = sp[0]; coq_env = sp[1]; sp += 2; } /* Register optimization. Some compilers underestimate the use of the local variables representing the abstract machine registers, and don't put them in hardware registers, which slows down the interpreter considerably. For GCC, Xavier Leroy have hand-assigned hardware registers for several architectures. */ #if defined(__GNUC__) && !defined(DEBUG) && !defined(__INTEL_COMPILER) \ && !defined(__llvm__) #ifdef __mips__ #define PC_REG asm("$16") #define SP_REG asm("$17") #define ACCU_REG asm("$18") #endif #ifdef __sparc__ #define PC_REG asm("%l0") #define SP_REG asm("%l1") #define ACCU_REG asm("%l2") #endif #ifdef __alpha__ #ifdef __CRAY__ #define PC_REG asm("r9") #define SP_REG asm("r10") #define ACCU_REG asm("r11") #define JUMPTBL_BASE_REG asm("r12") #else #define PC_REG asm("$9") #define SP_REG asm("$10") #define ACCU_REG asm("$11") #define JUMPTBL_BASE_REG asm("$12") #endif #endif #ifdef __i386__ #define PC_REG asm("%esi") #define SP_REG asm("%edi") #define ACCU_REG #endif #if defined(__ppc__) || defined(__ppc64__) #define PC_REG asm("26") #define SP_REG asm("27") #define ACCU_REG asm("28") #endif #ifdef __hppa__ #define PC_REG asm("%r18") #define SP_REG asm("%r17") #define ACCU_REG asm("%r16") #endif #ifdef __mc68000__ #define PC_REG asm("a5") #define SP_REG asm("a4") #define ACCU_REG asm("d7") #endif /* OCaml PR#4953: these specific registers not available in Thumb mode */ #if defined(__arm__) && !defined(__thumb__) #define PC_REG asm("r6") #define SP_REG asm("r8") #define ACCU_REG asm("r7") #endif #ifdef __ia64__ #define PC_REG asm("36") #define SP_REG asm("37") #define ACCU_REG asm("38") #define JUMPTBL_BASE_REG asm("39") #endif #ifdef __x86_64__ #define PC_REG asm("%r15") #define SP_REG asm("%r14") #define ACCU_REG asm("%r13") #endif #ifdef __aarch64__ #define PC_REG asm("%x19") #define SP_REG asm("%x20") #define ACCU_REG asm("%x21") #define JUMPTBL_BASE_REG asm("%x22") #endif #endif #define CheckPrimArgs(cond, apply_lbl) do{ \ if (cond) pc++; \ else{ \ *--sp=accu; \ accu = Field(coq_global_data, *pc++); \ goto apply_lbl; \ } \ }while(0) #define CheckInt1() CheckPrimArgs(Is_uint63(accu), apply1) #define CheckInt2() CheckPrimArgs(Is_uint63(accu) && Is_uint63(sp[0]), apply2) #define CheckInt3() CheckPrimArgs(Is_uint63(accu) && Is_uint63(sp[0]) \ && Is_uint63(sp[1]), apply3) #define CheckFloat1() CheckPrimArgs(Is_double(accu), apply1) #define CheckFloat2() CheckPrimArgs(Is_double(accu) && Is_double(sp[0]), apply2) #define AllocCarry(cond) Alloc_small(accu, 1, (cond)? coq_tag_C1 : coq_tag_C0) #define AllocPair() Alloc_small(accu, 2, coq_tag_pair) /* Beware: we cannot use caml_copy_double here as it doesn't use Alloc_small, hence doesn't protect the stack via Setup_for_gc/Restore_after_gc. */ #define Coq_copy_double(val) do{ \ double Coq_copy_double_f__ = (val); \ Alloc_small(accu, Double_wosize, Double_tag); \ Store_double_val(accu, Coq_copy_double_f__); \ }while(0); #define Swap_accu_sp do{ \ value swap_accu_sp_tmp__ = accu; \ accu = *sp; \ *sp = swap_accu_sp_tmp__; \ }while(0) #if OCAML_VERSION < 41000 /* For signal handling, we hijack some code from the caml runtime */ extern intnat volatile caml_signals_are_pending; extern intnat volatile caml_pending_signals[]; extern void caml_process_pending_signals(void); #endif /* The interpreter itself */ value coq_interprete (code_t coq_pc, value coq_accu, value coq_atom_tbl, value coq_global_data, value coq_env, long coq_extra_args) { /* coq_accu is not allocated on the OCaml heap */ CAMLparam2(coq_atom_tbl, coq_global_data); /*Declaration des variables */ #ifdef PC_REG register code_t pc PC_REG; register value * sp SP_REG; register value accu ACCU_REG; #else register code_t pc; register value * sp; register value accu; #endif #if defined(THREADED_CODE) && defined(ARCH_SIXTYFOUR) && !defined(ARCH_CODE32) #ifdef JUMPTBL_BASE_REG register char * coq_jumptbl_base JUMPTBL_BASE_REG; #else register char * coq_jumptbl_base; #endif #endif #ifdef THREADED_CODE static void * coq_jumptable[] = { # include "coq_jumptbl.h" }; #else opcode_t curr_instr; #endif print_instr("Enter Interpreter"); if (coq_pc == NULL) { /* Interpreter is initializing */ print_instr("Interpreter is initializing"); #ifdef THREADED_CODE coq_instr_table = (char **) coq_jumptable; coq_instr_base = coq_Jumptbl_base; #endif CAMLreturn(Val_unit); } #if defined(THREADED_CODE) && defined(ARCH_SIXTYFOUR) && !defined(ARCH_CODE32) coq_jumptbl_base = coq_Jumptbl_base; #endif /* Initialisation */ sp = coq_sp; pc = coq_pc; accu = coq_accu; CHECK_STACK(0); #ifdef THREADED_CODE goto *(void *)(coq_jumptbl_base + *pc++); /* Jump to the first instruction */ #else while(1) { curr_instr = *pc++; switch(curr_instr) { #endif /* Basic stack operations */ Instruct(ACC0){ print_instr("ACC0"); accu = sp[0]; Next; } Instruct(ACC1){ print_instr("ACC1"); accu = sp[1]; Next; } Instruct(ACC2){ print_instr("ACC2"); accu = sp[2]; Next; } Instruct(ACC3){ print_instr("ACC3"); accu = sp[3]; Next; } Instruct(ACC4){ print_instr("ACC4"); accu = sp[4]; Next; } Instruct(ACC5){ print_instr("ACC5"); accu = sp[5]; Next; } Instruct(ACC6){ print_instr("ACC6"); accu = sp[6]; Next; } Instruct(ACC7){ print_instr("ACC7"); accu = sp[7]; Next; } Instruct(PUSH){ print_instr("PUSH"); *--sp = accu; Next; } Instruct(PUSHACC0) { print_instr("PUSHACC0"); *--sp = accu; Next; } Instruct(PUSHACC1){ print_instr("PUSHACC1"); *--sp = accu; accu = sp[1]; Next; } Instruct(PUSHACC2){ print_instr("PUSHACC2"); *--sp = accu; accu = sp[2]; Next; } Instruct(PUSHACC3){ print_instr("PUSHACC3"); *--sp = accu; accu = sp[3]; Next; } Instruct(PUSHACC4){ print_instr("PUSHACC4"); *--sp = accu; accu = sp[4]; Next; } Instruct(PUSHACC5){ print_instr("PUSHACC5"); *--sp = accu; accu = sp[5]; Next; } Instruct(PUSHACC6){ print_instr("PUSHACC5"); *--sp = accu; accu = sp[6]; Next; } Instruct(PUSHACC7){ print_instr("PUSHACC7"); *--sp = accu; accu = sp[7]; Next; } Instruct(PUSHACC){ print_instr("PUSHACC"); *--sp = accu; } /* Fallthrough */ Instruct(ACC){ print_instr("ACC"); accu = sp[*pc++]; Next; } Instruct(POP){ print_instr("POP"); sp += *pc++; Next; } /* Access in heap-allocated environment */ Instruct(ENVACC1){ print_instr("ENVACC1"); accu = Field(coq_env, 1); Next; } Instruct(ENVACC2){ print_instr("ENVACC2"); accu = Field(coq_env, 2); Next; } Instruct(ENVACC3){ print_instr("ENVACC3"); accu = Field(coq_env, 3); Next; } Instruct(ENVACC4){ print_instr("ENVACC4"); accu = Field(coq_env, 4); Next; } Instruct(PUSHENVACC1){ print_instr("PUSHENVACC1"); *--sp = accu; accu = Field(coq_env, 1); Next; } Instruct(PUSHENVACC2){ print_instr("PUSHENVACC2"); *--sp = accu; accu = Field(coq_env, 2); Next; } Instruct(PUSHENVACC3){ print_instr("PUSHENVACC3"); *--sp = accu; accu = Field(coq_env, 3); Next; } Instruct(PUSHENVACC4){ print_instr("PUSHENVACC4"); *--sp = accu; accu = Field(coq_env, 4); Next; } Instruct(PUSHENVACC){ print_instr("PUSHENVACC"); *--sp = accu; } /* Fallthrough */ Instruct(ENVACC){ print_instr("ENVACC"); print_int(*pc); accu = Field(coq_env, *pc++); Next; } /* Function application */ Instruct(PUSH_RETADDR) { print_instr("PUSH_RETADDR"); sp -= 3; sp[0] = (value) (pc + *pc); sp[1] = coq_env; sp[2] = Val_long(coq_extra_args); coq_extra_args = 0; pc++; Next; } Instruct(APPLY) { print_instr("APPLY"); coq_extra_args = *pc - 1; pc = Code_val(accu); coq_env = accu; goto check_stack; } Instruct(APPLY1) { value arg1; apply1: print_instr("APPLY1"); arg1 = sp[0]; sp -= 3; sp[0] = arg1; sp[1] = (value)pc; sp[2] = coq_env; sp[3] = Val_long(coq_extra_args); print_instr("call stack="); print_lint(sp[1]); print_lint(sp[2]); print_lint(sp[3]); pc = Code_val(accu); coq_env = accu; coq_extra_args = 0; goto check_stack; } Instruct(APPLY2) { value arg1; value arg2; apply2: print_instr("APPLY2"); arg1 = sp[0]; arg2 = sp[1]; sp -= 3; sp[0] = arg1; sp[1] = arg2; sp[2] = (value)pc; sp[3] = coq_env; sp[4] = Val_long(coq_extra_args); pc = Code_val(accu); coq_env = accu; coq_extra_args = 1; goto check_stack; } Instruct(APPLY3) { value arg1; value arg2; value arg3; apply3: print_instr("APPLY3"); arg1 = sp[0]; arg2 = sp[1]; arg3 = sp[2]; sp -= 3; sp[0] = arg1; sp[1] = arg2; sp[2] = arg3; sp[3] = (value)pc; sp[4] = coq_env; sp[5] = Val_long(coq_extra_args); pc = Code_val(accu); coq_env = accu; coq_extra_args = 2; goto check_stack; } Instruct(APPLY4) { value arg1 = sp[0]; value arg2 = sp[1]; value arg3 = sp[2]; value arg4 = sp[3]; print_instr("APPLY4"); sp -= 3; sp[0] = arg1; sp[1] = arg2; sp[2] = arg3; sp[3] = arg4; sp[4] = (value)pc; sp[5] = coq_env; sp[6] = Val_long(coq_extra_args); pc = Code_val(accu); coq_env = accu; coq_extra_args = 3; goto check_stack; } /* Stack checks */ check_stack: print_instr("check_stack"); CHECK_STACK(0); /* We also check for signals */ #if OCAML_VERSION >= 41000 { value res = caml_process_pending_actions_exn(); if (Is_exception_result(res)) { /* If there is an asynchronous exception, we reset the vm */ coq_sp = coq_stack_high; caml_raise(Extract_exception(res)); } } #else if (caml_signals_are_pending) { /* If there's a Ctrl-C, we reset the vm */ intnat sigint = caml_pending_signals[SIGINT]; if (sigint) { coq_sp = coq_stack_high; } caml_process_pending_signals(); if (sigint) { caml_failwith("Coq VM: Fatal error: SIGINT signal detected " "but no exception was raised"); } } #endif Next; Instruct(ENSURESTACKCAPACITY) { print_instr("ENSURESTACKCAPACITY"); int size = *pc++; /* CHECK_STACK may trigger here a useless allocation because of the threshold, but check_stack: often does it anyway, so we prefer to factorize the code. */ CHECK_STACK(size); Next; } Instruct(APPTERM) { int nargs = *pc++; int slotsize = *pc; value * newsp; int i; print_instr("APPTERM"); /* Slide the nargs bottom words of the current frame to the top of the frame, and discard the remainder of the frame */ newsp = sp + slotsize - nargs; for (i = nargs - 1; i >= 0; i--) newsp[i] = sp[i]; sp = newsp; pc = Code_val(accu); coq_env = accu; coq_extra_args += nargs - 1; goto check_stack; } Instruct(APPTERM1) { value arg1 = sp[0]; print_instr("APPTERM1"); sp = sp + *pc - 1; sp[0] = arg1; pc = Code_val(accu); coq_env = accu; goto check_stack; } Instruct(APPTERM2) { value arg1 = sp[0]; value arg2 = sp[1]; print_instr("APPTERM2"); sp = sp + *pc - 2; sp[0] = arg1; sp[1] = arg2; pc = Code_val(accu); print_lint(accu); coq_env = accu; coq_extra_args += 1; goto check_stack; } Instruct(APPTERM3) { value arg1 = sp[0]; value arg2 = sp[1]; value arg3 = sp[2]; print_instr("APPTERM3"); sp = sp + *pc - 3; sp[0] = arg1; sp[1] = arg2; sp[2] = arg3; pc = Code_val(accu); coq_env = accu; coq_extra_args += 2; goto check_stack; } Instruct(RETURN) { print_instr("RETURN"); print_int(*pc); sp += *pc++; print_instr("stack="); print_lint(sp[0]); print_lint(sp[1]); print_lint(sp[2]); if (coq_extra_args > 0) { print_instr("extra args > 0"); print_lint(coq_extra_args); coq_extra_args--; pc = Code_val(accu); coq_env = accu; } else { print_instr("extra args = 0"); pc = (code_t)(sp[0]); coq_env = sp[1]; coq_extra_args = Long_val(sp[2]); sp += 3; } Next; } Instruct(RESTART) { int num_args = Wosize_val(coq_env) - 2; int i; print_instr("RESTART"); CHECK_STACK(num_args); sp -= num_args; for (i = 0; i < num_args; i++) sp[i] = Field(coq_env, i + 2); coq_env = Field(coq_env, 1); coq_extra_args += num_args; Next; } Instruct(GRAB) { int required = *pc++; print_instr("GRAB"); /* printf("GRAB %d\n",required); */ if (coq_extra_args >= required) { coq_extra_args -= required; } else { mlsize_t num_args, i; num_args = 1 + coq_extra_args; /* arg1 + extra args */ Alloc_small(accu, num_args + 2, Closure_tag); Field(accu, 1) = coq_env; for (i = 0; i < num_args; i++) Field(accu, i + 2) = sp[i]; Code_val(accu) = pc - 3; /* Point to the preceding RESTART instr. */ sp += num_args; pc = (code_t)(sp[0]); coq_env = sp[1]; coq_extra_args = Long_val(sp[2]); sp += 3; } Next; } Instruct(GRABREC) { int rec_pos = *pc++; /* commence a zero */ print_instr("GRABREC"); if (rec_pos <= coq_extra_args && !Is_accu(sp[rec_pos])) { pc++;/* On saute le Restart */ } else { if (coq_extra_args < rec_pos) { /* Partial application */ mlsize_t num_args, i; num_args = 1 + coq_extra_args; /* arg1 + extra args */ Alloc_small(accu, num_args + 2, Closure_tag); Field(accu, 1) = coq_env; for (i = 0; i < num_args; i++) Field(accu, i + 2) = sp[i]; Code_val(accu) = pc - 3; sp += num_args; pc = (code_t)(sp[0]); coq_env = sp[1]; coq_extra_args = Long_val(sp[2]); sp += 3; } else { /* The recursif argument is an accumulator */ mlsize_t num_args, i; /* Construction of fixpoint applied to its [rec_pos-1] first arguments */ Alloc_small(accu, rec_pos + 2, Closure_tag); Field(accu, 1) = coq_env; // We store the fixpoint in the first field for (i = 0; i < rec_pos; i++) Field(accu, i + 2) = sp[i]; // Storing args Code_val(accu) = pc; sp += rec_pos; *--sp = accu; /* Construction of the atom */ Alloc_small(accu, 2, ATOM_FIX_TAG); Field(accu,1) = sp[0]; Field(accu,0) = sp[1]; sp++; sp[0] = accu; /* Construction of the accumulator */ num_args = coq_extra_args - rec_pos; Alloc_small(accu, 2+num_args, Accu_tag); Code_val(accu) = accumulate; Field(accu,1) = sp[0]; sp++; for (i = 0; i < num_args;i++)Field(accu, i + 2) = sp[i]; sp += num_args; pc = (code_t)(sp[0]); coq_env = sp[1]; coq_extra_args = Long_val(sp[2]); sp += 3; } } Next; } Instruct(CLOSURE) { int nvars = *pc++; int i; print_instr("CLOSURE"); print_int(nvars); if (nvars > 0) *--sp = accu; Alloc_small(accu, 1 + nvars, Closure_tag); Code_val(accu) = pc + *pc; pc++; for (i = 0; i < nvars; i++) { print_lint(sp[i]); Field(accu, i + 1) = sp[i]; } sp += nvars; Next; } Instruct(CLOSUREREC) { int nfuncs = *pc++; int nvars = *pc++; int start = *pc++; int i; value * p; print_instr("CLOSUREREC"); if (nvars > 0) *--sp = accu; /* construction du vecteur de type */ Alloc_small(accu, nfuncs, Abstract_tag); for(i = 0; i < nfuncs; i++) { Field(accu,i) = (value)(pc+pc[i]); } pc += nfuncs; *--sp=accu; Alloc_small(accu, nfuncs * 2 + nvars, Closure_tag); Field(accu, nfuncs * 2 + nvars - 1) = *sp++; /* On remplie la partie pour les variables libres */ p = &Field(accu, nfuncs * 2 - 1); for (i = 0; i < nvars; i++) { *p++ = *sp++; } p = &Field(accu, 0); *p = (value) (pc + pc[0]); p++; for (i = 1; i < nfuncs; i++) { *p = Make_header(i * 2, Infix_tag, Caml_white); p++; /* color irrelevant. */ *p = (value) (pc + pc[i]); p++; } pc += nfuncs; accu = accu + 2 * start * sizeof(value); Next; } Instruct(CLOSURECOFIX){ int nfunc = *pc++; int nvars = *pc++; int start = *pc++; int i, j , size; value * p; print_instr("CLOSURECOFIX"); if (nvars > 0) *--sp = accu; /* construction du vecteur de type */ Alloc_small(accu, nfunc, Abstract_tag); for(i = 0; i < nfunc; i++) { Field(accu,i) = (value)(pc+pc[i]); } pc += nfunc; *--sp=accu; /* Creation des blocks accumulate */ for(i=0; i < nfunc; i++) { Alloc_small(accu, 2, Accu_tag); Code_val(accu) = accumulate; Field(accu,1) = Val_int(1); *--sp=accu; } /* creation des fonction cofix */ p = sp; size = nfunc + nvars + 2; for (i=0; i < nfunc; i++) { Alloc_small(accu, size, Closure_tag); Code_val(accu) = pc+pc[i]; for (j = 0; j < nfunc; j++) Field(accu, j+1) = p[j]; Field(accu, size - 1) = p[nfunc]; for (j = nfunc+1; j <= nfunc+nvars; j++) Field(accu, j) = p[j]; *--sp = accu; /* creation du block contenant le cofix */ Alloc_small(accu,1, ATOM_COFIX_TAG); Field(accu, 0) = sp[0]; *sp = accu; /* mise a jour du block accumulate */ caml_modify(&Field(p[i], 1),*sp); sp++; } pc += nfunc; accu = p[start]; sp = p + nfunc + 1 + nvars; print_instr("ici4"); Next; } Instruct(PUSHOFFSETCLOSURE) { print_instr("PUSHOFFSETCLOSURE"); *--sp = accu; } /* fallthrough */ Instruct(OFFSETCLOSURE) { print_instr("OFFSETCLOSURE"); accu = coq_env + *pc++ * sizeof(value); Next; } Instruct(PUSHOFFSETCLOSUREM2) { print_instr("PUSHOFFSETCLOSUREM2"); *--sp = accu; } /* fallthrough */ Instruct(OFFSETCLOSUREM2) { print_instr("OFFSETCLOSUREM2"); accu = coq_env - 2 * sizeof(value); Next; } Instruct(PUSHOFFSETCLOSURE0) { print_instr("PUSHOFFSETCLOSURE0"); *--sp = accu; }/* fallthrough */ Instruct(OFFSETCLOSURE0) { print_instr("OFFSETCLOSURE0"); accu = coq_env; Next; } Instruct(PUSHOFFSETCLOSURE2){ print_instr("PUSHOFFSETCLOSURE2"); *--sp = accu; /* fallthrough */ } Instruct(OFFSETCLOSURE2) { print_instr("OFFSETCLOSURE2"); accu = coq_env + 2 * sizeof(value); Next; } /* Access to global variables */ Instruct(PUSHGETGLOBAL) { print_instr("PUSH"); *--sp = accu; } /* Fallthrough */ Instruct(GETGLOBAL){ print_instr("GETGLOBAL"); print_int(*pc); accu = Field(coq_global_data, *pc); pc++; Next; } /* Allocation of blocks */ Instruct(MAKEBLOCK) { mlsize_t wosize = *pc++; tag_t tag = *pc++; mlsize_t i; value block; print_instr("MAKEBLOCK, tag="); Alloc_small(block, wosize, tag); Field(block, 0) = accu; for (i = 1; i < wosize; i++) Field(block, i) = *sp++; accu = block; Next; } Instruct(MAKEBLOCK1) { tag_t tag = *pc++; value block; print_instr("MAKEBLOCK1, tag="); print_int(tag); Alloc_small(block, 1, tag); Field(block, 0) = accu; accu = block; Next; } Instruct(MAKEBLOCK2) { tag_t tag = *pc++; value block; print_instr("MAKEBLOCK2, tag="); print_int(tag); Alloc_small(block, 2, tag); Field(block, 0) = accu; Field(block, 1) = sp[0]; sp += 1; accu = block; Next; } Instruct(MAKEBLOCK3) { tag_t tag = *pc++; value block; print_instr("MAKEBLOCK3, tag="); print_int(tag); Alloc_small(block, 3, tag); Field(block, 0) = accu; Field(block, 1) = sp[0]; Field(block, 2) = sp[1]; sp += 2; accu = block; Next; } Instruct(MAKEBLOCK4) { tag_t tag = *pc++; value block; print_instr("MAKEBLOCK4, tag="); print_int(tag); Alloc_small(block, 4, tag); Field(block, 0) = accu; Field(block, 1) = sp[0]; Field(block, 2) = sp[1]; Field(block, 3) = sp[2]; sp += 3; accu = block; Next; } /* Access to components of blocks */ Instruct(SWITCH) { uint32_t sizes = *pc++; print_instr("SWITCH"); print_int(sizes & 0xFFFFFF); if (Is_block(accu)) { long index = Tag_val(accu); print_instr("block"); print_lint(index); pc += pc[(sizes & 0xFFFFFF) + index]; } else { long index = Long_val(accu); print_instr("constant"); print_lint(index); pc += pc[index]; } Next; } Instruct(PUSHFIELDS){ int i; int size = *pc++; print_instr("PUSHFIELDS"); sp -= size; for(i=0;i y) { accu = coq_FGt; } else if(x == y) { accu = coq_FEq; } else { // nan value accu = coq_FNotComparable; } Next; } Instruct (CHECKCLASSIFYFLOAT) { double x; print_instr("CHECKCLASSIFYFLOAT"); CheckFloat1(); x = Double_val(accu); switch (fpclassify(x)) { case FP_NORMAL: accu = signbit(x) ? coq_NNormal : coq_PNormal; break; case FP_SUBNORMAL: accu = signbit(x) ? coq_NSubn : coq_PSubn; break; case FP_ZERO: accu = signbit(x) ? coq_NZero : coq_PZero; break; case FP_INFINITE: accu = signbit(x) ? coq_NInf : coq_PInf; break; default: /* FP_NAN */ accu = coq_NaN; break; } Next; } Instruct (CHECKADDFLOAT) { print_instr("CHECKADDFLOAT"); CheckFloat2(); Coq_copy_double(coq_fadd(Double_val(accu), Double_val(*sp++))); Next; } Instruct (CHECKSUBFLOAT) { print_instr("CHECKSUBFLOAT"); CheckFloat2(); Coq_copy_double(coq_fsub(Double_val(accu), Double_val(*sp++))); Next; } Instruct (CHECKMULFLOAT) { print_instr("CHECKMULFLOAT"); CheckFloat2(); Coq_copy_double(coq_fmul(Double_val(accu), Double_val(*sp++))); Next; } Instruct (CHECKDIVFLOAT) { print_instr("CHECKDIVFLOAT"); CheckFloat2(); Coq_copy_double(coq_fdiv(Double_val(accu), Double_val(*sp++))); Next; } Instruct (CHECKSQRTFLOAT) { print_instr("CHECKSQRTFLOAT"); CheckFloat1(); Coq_copy_double(coq_fsqrt(Double_val(accu))); Next; } Instruct (CHECKFLOATOFINT63) { print_instr("CHECKFLOATOFINT63"); CheckInt1(); Uint63_to_double(accu); Next; } Instruct (CHECKFLOATNORMFRMANTISSA) { double f; print_instr("CHECKFLOATNORMFRMANTISSA"); CheckFloat1(); f = fabs(Double_val(accu)); if (f >= 0.5 && f < 1) { Uint63_of_double(ldexp(f, DBL_MANT_DIG)); } else { Uint63_of_int(Val_int(0)); } Next; } Instruct (CHECKFRSHIFTEXP) { int exp; double f; print_instr("CHECKFRSHIFTEXP"); CheckFloat1(); /* frexp(infinity) incorrectly returns nan on mingw */ #if defined(__MINGW32__) || defined(__MINGW64__) if (fpclassify(Double_val(accu)) == FP_INFINITE) { f = Double_val(accu); } else #endif f = frexp(Double_val(accu), &exp); if (fpclassify(f) == FP_NORMAL) { exp += FLOAT_EXP_SHIFT; } else { exp = 0; } Coq_copy_double(f); *--sp = accu; #ifdef ARCH_SIXTYFOUR Alloc_small(accu, 2, coq_tag_pair); Field(accu, 1) = Val_int(exp); #else Uint63_of_int(Val_int(exp)); *--sp = accu; Alloc_small(accu, 2, coq_tag_pair); Field(accu, 1) = *sp++; #endif Field(accu, 0) = *sp++; Next; } Instruct (CHECKLDSHIFTEXP) { print_instr("CHECKLDSHIFTEXP"); CheckPrimArgs(Is_double(accu) && Is_uint63(sp[0]), apply2); Swap_accu_sp; Uint63_to_int_min(accu, Val_int(2 * FLOAT_EXP_SHIFT)); accu = Int_val(accu); Coq_copy_double(ldexp(Double_val(*sp++), accu - FLOAT_EXP_SHIFT)); Next; } Instruct (CHECKNEXTUPFLOAT) { print_instr("CHECKNEXTUPFLOAT"); CheckFloat1(); Coq_copy_double(coq_next_up(Double_val(accu))); Next; } Instruct (CHECKNEXTDOWNFLOAT) { print_instr("CHECKNEXTDOWNFLOAT"); CheckFloat1(); Coq_copy_double(coq_next_down(Double_val(accu))); Next; } /* Debugging and machine control */ Instruct(STOP){ print_instr("STOP"); coq_sp = sp; CAMLreturn(accu); } #ifndef THREADED_CODE default: /*fprintf(stderr, "%d\n", *pc);*/ caml_failwith("Coq VM: Fatal error: bad opcode"); } } #endif } value coq_push_ra(value code) { code_t tcode = Code_val(code); print_instr("push_ra"); coq_sp -= 3; coq_sp[0] = (value) tcode; coq_sp[1] = Val_unit; coq_sp[2] = Val_long(0); return Val_unit; } value coq_push_val(value v) { print_instr("push_val"); *--coq_sp = v; return Val_unit; } value coq_push_arguments(value args) { int nargs,i; value * sp = coq_sp; nargs = Wosize_val(args) - 2; CHECK_STACK(nargs); coq_sp -= nargs; print_instr("push_args");print_int(nargs); for(i = 0; i < nargs; i++) coq_sp[i] = Field(args, i+2); return Val_unit; } value coq_push_vstack(value stk, value max_stack_size) { int len,i; value * sp = coq_sp; len = Wosize_val(stk); CHECK_STACK(len); coq_sp -= len; print_instr("push_vstack");print_int(len); for(i = 0; i < len; i++) coq_sp[i] = Field(stk,i); sp = coq_sp; CHECK_STACK(uint_of_value(max_stack_size)); return Val_unit; } value coq_interprete_ml(value tcode, value a, value t, value g, value e, value ea) { // Registering the other arguments w.r.t. the OCaml GC is done by coq_interprete CAMLparam1(tcode); print_instr("coq_interprete"); CAMLreturn (coq_interprete(Code_val(tcode), a, t, g, e, Long_val(ea))); print_instr("end coq_interprete"); } value coq_interprete_byte(value* argv, int argn){ return coq_interprete_ml(argv[0], argv[1], argv[2], argv[3], argv[4], argv[5]); } value coq_eval_tcode (value tcode, value t, value g, value e) { return coq_interprete_ml(tcode, Val_unit, t, g, e, 0); } coq-8.11.0/kernel/environ.ml0000644000175000017500000006300013612664533015557 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* None | VKvalue v -> try Some (CEphemeron.get v) with CEphemeron.InvalidKey -> None let dummy_lazy_val () = ref VKnone let build_lazy_val vk key = vk := VKvalue (CEphemeron.create key) type named_context_val = { env_named_ctx : Constr.named_context; env_named_map : (Constr.named_declaration * lazy_val) Id.Map.t; } type rel_context_val = { env_rel_ctx : Constr.rel_context; env_rel_map : (Constr.rel_declaration * lazy_val) Range.t; } type env = { env_globals : Globals.t; env_named_context : named_context_val; (* section variables *) env_rel_context : rel_context_val; env_nb_rel : int; env_stratification : stratification; env_typing_flags : typing_flags; retroknowledge : Retroknowledge.retroknowledge; indirect_pterms : Opaqueproof.opaquetab; native_symbols : Nativevalues.symbols DPmap.t; } let empty_named_context_val = { env_named_ctx = []; env_named_map = Id.Map.empty; } let empty_rel_context_val = { env_rel_ctx = []; env_rel_map = Range.empty; } let empty_env = { env_globals = { Globals.constants = Cmap_env.empty ; inductives = Mindmap_env.empty ; modules = MPmap.empty ; modtypes = MPmap.empty }; env_named_context = empty_named_context_val; env_rel_context = empty_rel_context_val; env_nb_rel = 0; env_stratification = { env_universes = UGraph.initial_universes; env_sprop_allowed = false; env_universes_lbound = Univ.Level.set; env_engagement = PredicativeSet }; env_typing_flags = Declareops.safe_flags Conv_oracle.empty; retroknowledge = Retroknowledge.empty; indirect_pterms = Opaqueproof.empty_opaquetab; native_symbols = DPmap.empty; } (* Rel context *) let push_rel_context_val d ctx = { env_rel_ctx = Context.Rel.add d ctx.env_rel_ctx; env_rel_map = Range.cons (d, ref VKnone) ctx.env_rel_map; } let match_rel_context_val ctx = match ctx.env_rel_ctx with | [] -> None | decl :: rem -> let (_, lval) = Range.hd ctx.env_rel_map in let ctx = { env_rel_ctx = rem; env_rel_map = Range.tl ctx.env_rel_map } in Some (decl, lval, ctx) let push_rel d env = { env with env_rel_context = push_rel_context_val d env.env_rel_context; env_nb_rel = env.env_nb_rel + 1 } let lookup_rel n env = try fst (Range.get env.env_rel_context.env_rel_map (n - 1)) with Invalid_argument _ -> raise Not_found let lookup_rel_val n env = try snd (Range.get env.env_rel_context.env_rel_map (n - 1)) with Invalid_argument _ -> raise Not_found let rel_skipn n ctx = { env_rel_ctx = Util.List.skipn n ctx.env_rel_ctx; env_rel_map = Range.skipn n ctx.env_rel_map; } let env_of_rel n env = { env with env_rel_context = rel_skipn n env.env_rel_context; env_nb_rel = env.env_nb_rel - n } (* Named context *) let push_named_context_val_val d rval ctxt = (* assert (not (Id.Map.mem (NamedDecl.get_id d) ctxt.env_named_map)); *) { env_named_ctx = Context.Named.add d ctxt.env_named_ctx; env_named_map = Id.Map.add (NamedDecl.get_id d) (d, rval) ctxt.env_named_map; } let push_named_context_val d ctxt = push_named_context_val_val d (ref VKnone) ctxt let match_named_context_val c = match c.env_named_ctx with | [] -> None | decl :: ctx -> let (_, v) = Id.Map.find (NamedDecl.get_id decl) c.env_named_map in let map = Id.Map.remove (NamedDecl.get_id decl) c.env_named_map in let cval = { env_named_ctx = ctx; env_named_map = map } in Some (decl, v, cval) let map_named_val f ctxt = let open Context.Named.Declaration in let fold accu d = let d' = f d in let accu = if d == d' then accu else Id.Map.modify (get_id d) (fun _ (_, v) -> (d', v)) accu in (accu, d') in let map, ctx = List.fold_left_map fold ctxt.env_named_map ctxt.env_named_ctx in if map == ctxt.env_named_map then ctxt else { env_named_ctx = ctx; env_named_map = map } let push_named d env = {env with env_named_context = push_named_context_val d env.env_named_context} let lookup_named id env = fst (Id.Map.find id env.env_named_context.env_named_map) let lookup_named_val id env = snd(Id.Map.find id env.env_named_context.env_named_map) let lookup_named_ctxt id ctxt = fst (Id.Map.find id ctxt.env_named_map) let fold_constants f env acc = Cmap_env.fold (fun c (body,_) acc -> f c body acc) env.env_globals.Globals.constants acc let fold_inductives f env acc = Mindmap_env.fold (fun c (body,_) acc -> f c body acc) env.env_globals.Globals.inductives acc (* Global constants *) let lookup_constant_key kn env = Cmap_env.get kn env.env_globals.Globals.constants let lookup_constant kn env = fst (lookup_constant_key kn env) let mem_constant kn env = Cmap_env.mem kn env.env_globals.Globals.constants (* Mutual Inductives *) let lookup_mind_key kn env = Mindmap_env.get kn env.env_globals.Globals.inductives let lookup_mind kn env = fst (lookup_mind_key kn env) let mem_mind kn env = Mindmap_env.mem kn env.env_globals.Globals.inductives let mind_context env mind = let mib = lookup_mind mind env in Declareops.inductive_polymorphic_context mib let oracle env = env.env_typing_flags.conv_oracle let set_oracle env o = let env_typing_flags = { env.env_typing_flags with conv_oracle = o } in { env with env_typing_flags } let engagement env = env.env_stratification.env_engagement let typing_flags env = env.env_typing_flags let is_impredicative_set env = match engagement env with | ImpredicativeSet -> true | _ -> false let is_impredicative_sort env = function | Sorts.SProp | Sorts.Prop -> true | Sorts.Set -> is_impredicative_set env | Sorts.Type _ -> false let is_impredicative_univ env u = is_impredicative_sort env (Sorts.sort_of_univ u) let type_in_type env = not (typing_flags env).check_universes let deactivated_guard env = not (typing_flags env).check_guarded let indices_matter env = env.env_typing_flags.indices_matter let check_template env = env.env_typing_flags.check_template let universes env = env.env_stratification.env_universes let universes_lbound env = env.env_stratification.env_universes_lbound let set_universes_lbound env lbound = let env_stratification = { env.env_stratification with env_universes_lbound = lbound } in { env with env_stratification } let named_context env = env.env_named_context.env_named_ctx let named_context_val env = env.env_named_context let rel_context env = env.env_rel_context.env_rel_ctx let opaque_tables env = env.indirect_pterms let set_opaque_tables env indirect_pterms = { env with indirect_pterms } let empty_context env = match env.env_rel_context.env_rel_ctx, env.env_named_context.env_named_ctx with | [], [] -> true | _ -> false (* Rel context *) let evaluable_rel n env = is_local_def (lookup_rel n env) let nb_rel env = env.env_nb_rel let push_rel_context ctxt x = Context.Rel.fold_outside push_rel ctxt ~init:x let push_rec_types (lna,typarray,_) env = let ctxt = Array.map2_i (fun i na t -> LocalAssum (na, lift i t)) lna typarray in Array.fold_left (fun e assum -> push_rel assum e) env ctxt let fold_rel_context f env ~init = let rec fold_right env = match match_rel_context_val env.env_rel_context with | None -> init | Some (rd, _, rc) -> let env = { env with env_rel_context = rc; env_nb_rel = env.env_nb_rel - 1 } in f env rd (fold_right env) in fold_right env (* Named context *) let named_context_of_val c = c.env_named_ctx let ids_of_named_context_val c = Id.Map.domain c.env_named_map let empty_named_context = Context.Named.empty let push_named_context = List.fold_right push_named let val_of_named_context ctxt = List.fold_right push_named_context_val ctxt empty_named_context_val let eq_named_context_val c1 c2 = c1 == c2 || Context.Named.equal Constr.equal (named_context_of_val c1) (named_context_of_val c2) (* A local const is evaluable if it is defined *) let named_type id env = let open Context.Named.Declaration in get_type (lookup_named id env) let named_body id env = let open Context.Named.Declaration in get_value (lookup_named id env) let evaluable_named id env = match named_body id env with | Some _ -> true | _ -> false let reset_with_named_context ctxt env = { env with env_named_context = ctxt; env_rel_context = empty_rel_context_val; env_nb_rel = 0 } let reset_context = reset_with_named_context empty_named_context_val let pop_rel_context n env = let rec skip n ctx = if Int.equal n 0 then ctx else match match_rel_context_val ctx with | None -> invalid_arg "List.skipn" | Some (_, _, ctx) -> skip (pred n) ctx in let ctxt = env.env_rel_context in { env with env_rel_context = skip n ctxt; env_nb_rel = env.env_nb_rel - n } let fold_named_context f env ~init = let rec fold_right env = match match_named_context_val env.env_named_context with | None -> init | Some (d, _v, rem) -> let env = reset_with_named_context rem env in f env d (fold_right env) in fold_right env let fold_named_context_reverse f ~init env = Context.Named.fold_inside f ~init:init (named_context env) (* Universe constraints *) let map_universes f env = let s = env.env_stratification in { env with env_stratification = { s with env_universes = f s.env_universes } } let add_constraints c env = if Univ.Constraint.is_empty c then env else map_universes (UGraph.merge_constraints c) env let check_constraints c env = UGraph.check_constraints c env.env_stratification.env_universes let push_constraints_to_env (_,univs) env = add_constraints univs env let add_universes ~lbound ~strict ctx g = let g = Array.fold_left (fun g v -> UGraph.add_universe ~lbound ~strict v g) g (Univ.Instance.to_array (Univ.UContext.instance ctx)) in UGraph.merge_constraints (Univ.UContext.constraints ctx) g let push_context ?(strict=false) ctx env = map_universes (add_universes ~lbound:(universes_lbound env) ~strict ctx) env let add_universes_set ~lbound ~strict ctx g = let g = Univ.LSet.fold (* Be lenient, module typing reintroduces universes and constraints due to includes *) (fun v g -> try UGraph.add_universe ~lbound ~strict v g with UGraph.AlreadyDeclared -> g) (Univ.ContextSet.levels ctx) g in UGraph.merge_constraints (Univ.ContextSet.constraints ctx) g let push_context_set ?(strict=false) ctx env = map_universes (add_universes_set ~lbound:(universes_lbound env) ~strict ctx) env let push_subgraph (levels,csts) env = let lbound = universes_lbound env in let add_subgraph g = let newg = Univ.LSet.fold (fun v g -> UGraph.add_universe ~lbound ~strict:false v g) levels g in let newg = UGraph.merge_constraints csts newg in (if not (Univ.Constraint.is_empty csts) then let restricted = UGraph.constraints_for ~kept:(UGraph.domain g) newg in (if not (UGraph.check_constraints restricted g) then CErrors.anomaly Pp.(str "Local constraints imply new transitive constraints."))); newg in map_universes add_subgraph env let set_engagement c env = (* Unsafe *) { env with env_stratification = { env.env_stratification with env_engagement = c } } (* It's convenient to use [{flags with foo = bar}] so we're smart wrt to it. *) let same_flags { check_guarded; check_positive; check_universes; conv_oracle; indices_matter; share_reduction; enable_VM; enable_native_compiler; check_template; } alt = check_guarded == alt.check_guarded && check_positive == alt.check_positive && check_universes == alt.check_universes && conv_oracle == alt.conv_oracle && indices_matter == alt.indices_matter && share_reduction == alt.share_reduction && enable_VM == alt.enable_VM && enable_native_compiler == alt.enable_native_compiler && check_template == alt.check_template [@warning "+9"] let set_typing_flags c env = (* Unsafe *) if same_flags env.env_typing_flags c then env else { env with env_typing_flags = c } let make_sprop_cumulative = map_universes UGraph.make_sprop_cumulative let set_allow_sprop b env = { env with env_stratification = { env.env_stratification with env_sprop_allowed = b } } let sprop_allowed env = env.env_stratification.env_sprop_allowed (* Global constants *) let no_link_info = NotLinked let add_constant_key kn cb linkinfo env = let new_constants = Cmap_env.add kn (cb,(ref linkinfo, ref None)) env.env_globals.Globals.constants in let new_globals = { env.env_globals with Globals.constants = new_constants } in { env with env_globals = new_globals } let add_constant kn cb env = add_constant_key kn cb no_link_info env (* constant_type gives the type of a constant *) let constant_type env (kn,u) = let cb = lookup_constant kn env in let uctx = Declareops.constant_polymorphic_context cb in let csts = Univ.AUContext.instantiate u uctx in (subst_instance_constr u cb.const_type, csts) type const_evaluation_result = | NoBody | Opaque | IsPrimitive of CPrimitives.t exception NotEvaluableConst of const_evaluation_result let constant_value_and_type env (kn, u) = let cb = lookup_constant kn env in let uctx = Declareops.constant_polymorphic_context cb in let cst = Univ.AUContext.instantiate u uctx in let b' = match cb.const_body with | Def l_body -> Some (subst_instance_constr u (Mod_subst.force_constr l_body)) | OpaqueDef _ -> None | Undef _ | Primitive _ -> None in b', subst_instance_constr u cb.const_type, cst (* These functions should be called under the invariant that [env] already contains the constraints corresponding to the constant application. *) (* constant_type gives the type of a constant *) let constant_type_in env (kn,u) = let cb = lookup_constant kn env in subst_instance_constr u cb.const_type let constant_value_in env (kn,u) = let cb = lookup_constant kn env in match cb.const_body with | Def l_body -> let b = Mod_subst.force_constr l_body in subst_instance_constr u b | OpaqueDef _ -> raise (NotEvaluableConst Opaque) | Undef _ -> raise (NotEvaluableConst NoBody) | Primitive p -> raise (NotEvaluableConst (IsPrimitive p)) let constant_opt_value_in env cst = try Some (constant_value_in env cst) with NotEvaluableConst _ -> None (* A global const is evaluable if it is defined and not opaque *) let evaluable_constant kn env = let cb = lookup_constant kn env in match cb.const_body with | Def _ -> true | OpaqueDef _ -> false | Undef _ | Primitive _ -> false let is_primitive env c = let cb = lookup_constant c env in match cb.Declarations.const_body with | Declarations.Primitive _ -> true | _ -> false let polymorphic_constant cst env = Declareops.constant_is_polymorphic (lookup_constant cst env) let polymorphic_pconstant (cst,u) env = if Univ.Instance.is_empty u then false else polymorphic_constant cst env let type_in_type_constant cst env = not (lookup_constant cst env).const_typing_flags.check_universes let lookup_projection p env = let mind,i = Projection.inductive p in let mib = lookup_mind mind env in (if not (Int.equal mib.mind_nparams (Projection.npars p)) then anomaly ~label:"lookup_projection" Pp.(str "Bad number of parameters on projection.")); match mib.mind_record with | NotRecord | FakeRecord -> anomaly ~label:"lookup_projection" Pp.(str "not a projection") | PrimRecord infos -> let _,_,_,typs = infos.(i) in typs.(Projection.arg p) let get_projection env ind ~proj_arg = let mib = lookup_mind (fst ind) env in Declareops.inductive_make_projection ind mib ~proj_arg let get_projections env ind = let mib = lookup_mind (fst ind) env in Declareops.inductive_make_projections ind mib (* Mutual Inductives *) let polymorphic_ind (mind,_i) env = Declareops.inductive_is_polymorphic (lookup_mind mind env) let polymorphic_pind (ind,u) env = if Univ.Instance.is_empty u then false else polymorphic_ind ind env let type_in_type_ind (mind,_i) env = not (lookup_mind mind env).mind_typing_flags.check_universes let template_checked_ind (mind,_i) env = (lookup_mind mind env).mind_typing_flags.check_template let template_polymorphic_ind (mind,i) env = match (lookup_mind mind env).mind_packets.(i).mind_arity with | TemplateArity _ -> true | RegularArity _ -> false let template_polymorphic_variables (mind,i) env = match (lookup_mind mind env).mind_packets.(i).mind_arity with | TemplateArity { Declarations.template_param_levels = l; _ } -> List.map_filter (fun level -> level) l | RegularArity _ -> [] let template_polymorphic_pind (ind,u) env = if not (Univ.Instance.is_empty u) then false else template_polymorphic_ind ind env let add_mind_key kn (_mind, _ as mind_key) env = let new_inds = Mindmap_env.add kn mind_key env.env_globals.Globals.inductives in let new_globals = { env.env_globals with Globals.inductives = new_inds; } in { env with env_globals = new_globals } let add_mind kn mib env = let li = ref no_link_info in add_mind_key kn (mib, li) env (* Lookup of section variables *) let lookup_constant_variables c env = let cmap = lookup_constant c env in Context.Named.to_vars cmap.const_hyps let lookup_inductive_variables (kn,_i) env = let mis = lookup_mind kn env in Context.Named.to_vars mis.mind_hyps let lookup_constructor_variables (ind,_) env = lookup_inductive_variables ind env (* Universes *) let constant_context env c = let cb = lookup_constant c env in Declareops.constant_polymorphic_context cb let universes_of_global env r = let open GlobRef in match r with | VarRef _ -> Univ.AUContext.empty | ConstRef c -> constant_context env c | IndRef (mind,_) | ConstructRef ((mind,_),_) -> let mib = lookup_mind mind env in Declareops.inductive_polymorphic_context mib (* Returns the list of global variables in a term *) let vars_of_global env gr = let open GlobRef in match gr with | VarRef id -> Id.Set.singleton id | ConstRef kn -> lookup_constant_variables kn env | IndRef ind -> lookup_inductive_variables ind env | ConstructRef cstr -> lookup_constructor_variables cstr env let global_vars_set env constr = let rec filtrec acc c = match destRef c with | gr, _ -> Id.Set.union (vars_of_global env gr) acc | exception DestKO -> Constr.fold filtrec acc c in filtrec Id.Set.empty constr (* [keep_hyps env ids] keeps the part of the section context of [env] which contains the variables of the set [ids], and recursively the variables contained in the types of the needed variables. *) let really_needed env needed = let open! Context.Named.Declaration in Context.Named.fold_inside (fun need decl -> if Id.Set.mem (get_id decl) need then let globc = match decl with | LocalAssum _ -> Id.Set.empty | LocalDef (_,c,_) -> global_vars_set env c in Id.Set.union (global_vars_set env (get_type decl)) (Id.Set.union globc need) else need) ~init:needed (named_context env) let keep_hyps env needed = let open Context.Named.Declaration in let really_needed = really_needed env needed in Context.Named.fold_outside (fun d nsign -> if Id.Set.mem (get_id d) really_needed then Context.Named.add d nsign else nsign) (named_context env) ~init:empty_named_context (* Modules *) let add_modtype mtb env = let mp = mtb.mod_mp in let new_modtypes = MPmap.add mp mtb env.env_globals.Globals.modtypes in let new_globals = { env.env_globals with Globals.modtypes = new_modtypes } in { env with env_globals = new_globals } let shallow_add_module mb env = let mp = mb.mod_mp in let new_mods = MPmap.add mp mb env.env_globals.Globals.modules in let new_globals = { env.env_globals with Globals.modules = new_mods } in { env with env_globals = new_globals } let lookup_module mp env = MPmap.find mp env.env_globals.Globals.modules let lookup_modtype mp env = MPmap.find mp env.env_globals.Globals.modtypes (*s Judgments. *) type ('constr, 'types) punsafe_judgment = { uj_val : 'constr; uj_type : 'types } let on_judgment f j = { uj_val = f j.uj_val; uj_type = f j.uj_type } let on_judgment_value f j = { j with uj_val = f j.uj_val } let on_judgment_type f j = { j with uj_type = f j.uj_type } type unsafe_judgment = (constr, types) punsafe_judgment let make_judge v tj = { uj_val = v; uj_type = tj } let j_val j = j.uj_val let j_type j = j.uj_type type 'types punsafe_type_judgment = { utj_val : 'types; utj_type : Sorts.t } type unsafe_type_judgment = types punsafe_type_judgment exception Hyp_not_found let apply_to_hyp ctxt id f = let open Context.Named.Declaration in let rec aux rtail ctxt = match match_named_context_val ctxt with | Some (d, v, ctxt) -> if Id.equal (get_id d) id then push_named_context_val_val (f ctxt.env_named_ctx d rtail) v ctxt else let ctxt' = aux (d::rtail) ctxt in push_named_context_val_val d v ctxt' | None -> raise Hyp_not_found in aux [] ctxt (* To be used in Logic.clear_hyps *) let remove_hyps ids check_context check_value ctxt = let rec remove_hyps ctxt = match match_named_context_val ctxt with | None -> empty_named_context_val, false | Some (d, v, rctxt) -> let open Context.Named.Declaration in let (ans, seen) = remove_hyps rctxt in if Id.Set.mem (get_id d) ids then (ans, true) else if not seen then ctxt, false else let rctxt' = ans in let d' = check_context d in let v' = check_value v in if d == d' && v == v' && rctxt == rctxt' then ctxt, true else push_named_context_val_val d' v' rctxt', true in fst (remove_hyps ctxt) (* A general request *) let is_polymorphic env r = let open Names.GlobRef in match r with | VarRef _id -> false | ConstRef c -> polymorphic_constant c env | IndRef ind -> polymorphic_ind ind env | ConstructRef cstr -> polymorphic_ind (inductive_of_constructor cstr) env let is_template_polymorphic env r = let open Names.GlobRef in match r with | VarRef _id -> false | ConstRef _c -> false | IndRef ind -> template_polymorphic_ind ind env | ConstructRef cstr -> template_polymorphic_ind (inductive_of_constructor cstr) env let get_template_polymorphic_variables env r = let open Names.GlobRef in match r with | VarRef _id -> [] | ConstRef _c -> [] | IndRef ind -> template_polymorphic_variables ind env | ConstructRef cstr -> template_polymorphic_variables (inductive_of_constructor cstr) env let is_template_checked env r = let open Names.GlobRef in match r with | VarRef _id -> false | ConstRef _c -> false | IndRef ind -> template_checked_ind ind env | ConstructRef cstr -> template_checked_ind (inductive_of_constructor cstr) env let is_type_in_type env r = let open Names.GlobRef in match r with | VarRef _id -> false | ConstRef c -> type_in_type_constant c env | IndRef ind -> type_in_type_ind ind env | ConstructRef cstr -> type_in_type_ind (inductive_of_constructor cstr) env let set_retroknowledge env r = { env with retroknowledge = r } let set_native_symbols env native_symbols = { env with native_symbols } let add_native_symbols dir syms env = { env with native_symbols = DPmap.add dir syms env.native_symbols } coq-8.11.0/kernel/inferCumulativity.ml0000644000175000017500000002142713612664533017631 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Covariant so nontrivial *) LMap.update u (function | None -> None | Some CovariantI as x -> x | Some IrrelevantI -> Some CovariantI) variances let infer_generic_instance_eq variances u = Array.fold_left (fun variances u -> infer_level_eq u variances) variances (Instance.to_array u) let infer_cumulative_ind_instance cv_pb mind_variance variances u = Array.fold_left2 (fun variances varu u -> match cv_pb, varu with | _, Irrelevant -> variances | _, Invariant | CONV, Covariant -> infer_level_eq u variances | CUMUL, Covariant -> infer_level_leq u variances) variances mind_variance (Instance.to_array u) let infer_inductive_instance cv_pb env variances ind nargs u = let mind = Environ.lookup_mind (fst ind) env in match mind.mind_variance with | None -> infer_generic_instance_eq variances u | Some mind_variance -> if not (Int.equal (inductive_cumulativity_arguments (mind,snd ind)) nargs) then infer_generic_instance_eq variances u else infer_cumulative_ind_instance cv_pb mind_variance variances u let infer_constructor_instance_eq env variances ((mi,ind),ctor) nargs u = let mind = Environ.lookup_mind mi env in match mind.mind_variance with | None -> infer_generic_instance_eq variances u | Some _ -> if not (Int.equal (constructor_cumulativity_arguments (mind,ind,ctor)) nargs) then infer_generic_instance_eq variances u else variances (* constructors are convertible at common supertype *) let infer_sort cv_pb variances s = match cv_pb with | CONV -> LSet.fold infer_level_eq (Universe.levels (Sorts.univ_of_sort s)) variances | CUMUL -> LSet.fold infer_level_leq (Universe.levels (Sorts.univ_of_sort s)) variances let infer_table_key variances c = let open Names in match c with | ConstKey (_, u) -> infer_generic_instance_eq variances u | VarKey _ | RelKey _ -> variances let whd_stack (infos, tab) hd stk = CClosure.whd_stack infos tab hd stk let rec infer_fterm cv_pb infos variances hd stk = Control.check_for_interrupt (); let hd,stk = whd_stack infos hd stk in let open CClosure in match fterm_of hd with | FAtom a -> begin match kind a with | Sort s -> infer_sort cv_pb variances s | Meta _ -> infer_stack infos variances stk | _ -> assert false end | FEvar ((_,args),e) -> let variances = infer_stack infos variances stk in infer_vect infos variances (Array.map (mk_clos e) args) | FRel _ -> infer_stack infos variances stk | FInt _ -> infer_stack infos variances stk | FFloat _ -> infer_stack infos variances stk | FFlex fl -> let variances = infer_table_key variances fl in infer_stack infos variances stk | FProj (_,c) -> let variances = infer_fterm CONV infos variances c [] in infer_stack infos variances stk | FLambda _ -> let (_,ty,bd) = destFLambda mk_clos hd in let variances = infer_fterm CONV infos variances ty [] in infer_fterm CONV infos variances bd [] | FProd (_,dom,codom,e) -> let variances = infer_fterm CONV infos variances dom [] in infer_fterm cv_pb infos variances (mk_clos (Esubst.subs_lift e) codom) [] | FInd (ind, u) -> let variances = if Instance.is_empty u then variances else let nargs = stack_args_size stk in infer_inductive_instance cv_pb (info_env (fst infos)) variances ind nargs u in infer_stack infos variances stk | FConstruct (ctor,u) -> let variances = if Instance.is_empty u then variances else let nargs = stack_args_size stk in infer_constructor_instance_eq (info_env (fst infos)) variances ctor nargs u in infer_stack infos variances stk | FFix ((_,(_,tys,cl)),e) | FCoFix ((_,(_,tys,cl)),e) -> let n = Array.length cl in let variances = infer_vect infos variances (Array.map (mk_clos e) tys) in let le = Esubst.subs_liftn n e in let variances = infer_vect infos variances (Array.map (mk_clos le) cl) in infer_stack infos variances stk (* Removed by whnf *) | FLOCKED | FCaseT _ | FLetIn _ | FApp _ | FLIFT _ | FCLOS _ -> assert false and infer_stack infos variances (stk:CClosure.stack) = match stk with | [] -> variances | z :: stk -> let open CClosure in let variances = match z with | Zapp v -> infer_vect infos variances v | Zproj _ -> variances | Zfix (fx,a) -> let variances = infer_fterm CONV infos variances fx [] in infer_stack infos variances a | ZcaseT (_, p, br, e) -> let variances = infer_fterm CONV infos variances (mk_clos e p) [] in infer_vect infos variances (Array.map (mk_clos e) br) | Zshift _ -> variances | Zupdate _ -> variances | Zprimitive (_,_,rargs,kargs) -> let variances = List.fold_left (fun variances c -> infer_fterm CONV infos variances c []) variances rargs in let variances = List.fold_left (fun variances (_,c) -> infer_fterm CONV infos variances c []) variances kargs in variances in infer_stack infos variances stk and infer_vect infos variances v = Array.fold_left (fun variances c -> infer_fterm CONV infos variances c []) variances v let infer_term cv_pb env variances c = let open CClosure in let infos = (create_clos_infos all env, create_tab ()) in infer_fterm cv_pb infos variances (CClosure.inject c) [] let infer_arity_constructor is_arity env variances arcn = let infer_typ typ (env,variances) = match typ with | Context.Rel.Declaration.LocalAssum (_, typ') -> (Environ.push_rel typ env, infer_term CUMUL env variances typ') | Context.Rel.Declaration.LocalDef _ -> assert false in let typs, codom = Reduction.dest_prod env arcn in let env, variances = Context.Rel.fold_outside infer_typ typs ~init:(env, variances) in (* If we have Inductive foo@{i j} : ... -> Type@{i} := C : ... -> foo Type@{j} i is irrelevant, j is invariant. *) if not is_arity then infer_term CUMUL env variances codom else variances open Entries let infer_inductive_core env params entries uctx = let uarray = Instance.to_array @@ UContext.instance uctx in if Array.is_empty uarray then raise TrivialVariance; let env = Environ.push_context uctx env in let variances = Array.fold_left (fun variances u -> LMap.add u IrrelevantI variances) LMap.empty uarray in let env, _ = Typeops.check_context env params in let variances = List.fold_left (fun variances entry -> let variances = infer_arity_constructor true env variances entry.mind_entry_arity in List.fold_left (infer_arity_constructor false env) variances entry.mind_entry_lc) variances entries in Array.map (fun u -> match LMap.find u variances with | exception Not_found -> Invariant | IrrelevantI -> Irrelevant | CovariantI -> Covariant) uarray let infer_inductive env mie = let open Entries in let params = mie.mind_entry_params in let entries = mie.mind_entry_inds in let variances = match mie.mind_entry_variance with | None -> None | Some _ -> let uctx = match mie.mind_entry_universes with | Monomorphic_entry _ -> assert false | Polymorphic_entry (_,uctx) -> uctx in try Some (infer_inductive_core env params entries uctx) with TrivialVariance -> Some (Array.make (UContext.size uctx) Invariant) in { mie with mind_entry_variance = variances } let dummy_variance = let open Entries in function | Monomorphic_entry _ -> assert false | Polymorphic_entry (_,uctx) -> Array.make (UContext.size uctx) Irrelevant coq-8.11.0/kernel/float64.mli0000644000175000017500000000471513612664533015537 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* bool val is_infinity : t -> bool val is_neg_infinity : t -> bool val to_string : t -> string val of_string : string -> t val compile : t -> string val of_float : float -> t (** Return [true] for "-", [false] for "+". *) val sign : t -> bool val opp : t -> t val abs : t -> t type float_comparison = FEq | FLt | FGt | FNotComparable val eq : t -> t -> bool val lt : t -> t -> bool val le : t -> t -> bool (** The IEEE 754 float comparison. * NotComparable is returned if there is a NaN in the arguments *) val compare : t -> t -> float_comparison [@@ocaml.inline always] type float_class = | PNormal | NNormal | PSubn | NSubn | PZero | NZero | PInf | NInf | NaN val classify : t -> float_class [@@ocaml.inline always] val mul : t -> t -> t val add : t -> t -> t val sub : t -> t -> t val div : t -> t -> t val sqrt : t -> t (** Link with integers *) val of_int63 : Uint63.t -> t [@@ocaml.inline always] val normfr_mantissa : t -> Uint63.t [@@ocaml.inline always] (** Shifted exponent extraction *) val eshift : int val frshiftexp : t -> t * Uint63.t (* float remainder, shifted exponent *) [@@ocaml.inline always] val ldshiftexp : t -> Uint63.t -> t [@@ocaml.inline always] val next_up : t -> t val next_down : t -> t (** Return true if two floats are equal. * All NaN values are considered equal. *) val equal : t -> t -> bool [@@ocaml.inline always] val hash : t -> int (** Total order relation over float values. Behaves like [Pervasives.compare].*) val total_compare : t -> t -> int val is_float64 : Obj.t -> bool [@@ocaml.inline always] coq-8.11.0/kernel/indtypes.ml0000644000175000017500000006054713612664533015753 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* raise (InductiveError (NonPos (env,c',mkRel (kt+nparamsctxt)))) | LocalNotEnoughArgs kt -> raise (InductiveError (NotEnoughArgs (env,c',mkRel (kt+nparamsctxt)))) | LocalNotConstructor (paramsctxt,nargs)-> let nparams = Context.Rel.nhyps paramsctxt in raise (InductiveError (NotConstructor (env,id,c',mkRel (ntyp+nparamsctxt), nparams,nargs))) | LocalNonPar (n,i,l) -> raise (InductiveError (NonPar (env,c',n,mkRel i,mkRel (l+nparamsctxt)))) let failwith_non_pos n ntypes c = for k = n to n + ntypes - 1 do if not (noccurn k c) then raise (IllFormedInd (LocalNonPos (k-n+1))) done let failwith_non_pos_vect n ntypes v = Array.iter (failwith_non_pos n ntypes) v; anomaly ~label:"failwith_non_pos_vect" (Pp.str "some k in [n;n+ntypes-1] should occur.") let failwith_non_pos_list n ntypes l = List.iter (failwith_non_pos n ntypes) l; anomaly ~label:"failwith_non_pos_list" (Pp.str "some k in [n;n+ntypes-1] should occur.") (* Check the inductive type is called with the expected parameters *) (* [n] is the index of the last inductive type in [env] *) let check_correct_par (env,n,ntypes,_) paramdecls ind_index args = let nparams = Context.Rel.nhyps paramdecls in let args = Array.of_list args in if Array.length args < nparams then raise (IllFormedInd (LocalNotEnoughArgs ind_index)); let (params,realargs) = Array.chop nparams args in let nparamdecls = List.length paramdecls in let rec check param_index paramdecl_index = function | [] -> () | LocalDef _ :: paramdecls -> check param_index (paramdecl_index+1) paramdecls | _::paramdecls -> match kind (whd_all env params.(param_index)) with | Rel w when Int.equal w paramdecl_index -> check (param_index-1) (paramdecl_index+1) paramdecls | _ -> let paramdecl_index_in_env = paramdecl_index-n+nparamdecls+1 in let err = LocalNonPar (param_index+1, paramdecl_index_in_env, ind_index) in raise (IllFormedInd err) in check (nparams-1) (n-nparamdecls) paramdecls; if not (Array.for_all (noccur_between n ntypes) realargs) then failwith_non_pos_vect n ntypes realargs (* Computes the maximum number of recursive parameters: the first parameters which are constant in recursive arguments [n] is the current depth, [nmr] is the maximum number of possible recursive parameters *) let compute_rec_par (env,n,_,_) paramsctxt nmr largs = if Int.equal nmr 0 then 0 else (* start from 0, params will be in reverse order *) let (lpar,_) = List.chop nmr largs in let rec find k index = function ([],_) -> nmr | (_,[]) -> assert false (* |paramsctxt|>=nmr *) | (lp, LocalDef _ :: paramsctxt) -> find k (index-1) (lp,paramsctxt) | (p::lp,_::paramsctxt) -> ( match kind (whd_all env p) with | Rel w when Int.equal w index -> find (k+1) (index-1) (lp,paramsctxt) | _ -> k) in find 0 (n-1) (lpar,List.rev paramsctxt) (* [env] is the typing environment [n] is the dB of the last inductive type [ntypes] is the number of inductive types in the definition (i.e. range of inductives is [n; n+ntypes-1]) [lra] is the list of recursive tree of each variable *) let ienv_push_var (env, n, ntypes, lra) (x,a,ra) = (push_rel (LocalAssum (x,a)) env, n+1, ntypes, (Norec,ra)::lra) let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lrecparams) = let auxntyp = 1 in let specif = (lookup_mind_specif env mi, u) in let ty = type_of_inductive env specif in let env' = let r = (snd (fst specif)).mind_relevance in let anon = Context.make_annot Anonymous r in let decl = LocalAssum (anon, hnf_prod_applist env ty lrecparams) in push_rel decl env in let ra_env' = (Imbr mi,(Rtree.mk_rec_calls 1).(0)) :: List.map (fun (r,t) -> (r,Rtree.lift 1 t)) ra_env in (* New index of the inductive types *) let newidx = n + auxntyp in (env', newidx, ntypes, ra_env') let rec ienv_decompose_prod (env,_,_,_ as ienv) n c = if Int.equal n 0 then (ienv,c) else let c' = whd_all env c in match kind c' with Prod(na,a,b) -> let ienv' = ienv_push_var ienv (na,a,mk_norec) in ienv_decompose_prod ienv' (n-1) b | _ -> assert false let array_min nmr a = if Int.equal nmr 0 then 0 else Array.fold_left (fun k (nmri,_) -> min k nmri) nmr a (** [check_positivity_one ienv paramsctxt (mind,i) nnonrecargs lcnames indlc] checks the positivity of the [i]-th member of the mutually inductive definition [mind]. It returns an [Rtree.t] which represents the position of the recursive calls of inductive in [i] for use by the guard condition (terms at these positions are considered sub-terms) as well as the number of of non-uniform arguments (used to generate induction schemes, so a priori less relevant to the kernel). If [chkpos] is [false] then positivity is assumed, and [check_positivity_one] computes the subterms occurrences in a best-effort fashion. *) let check_positivity_one ~chkpos recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as ind) nnonrecargs lcnames indlc = let nparamsctxt = Context.Rel.length paramsctxt in let nmr = Context.Rel.nhyps paramsctxt in (** Positivity of one argument [c] of a constructor (i.e. the constructor [cn] has a type of the shape [… -> c … -> P], where, more generally, the arrows may be dependent). *) let rec check_pos (env, n, ntypes, ra_env as ienv) nmr c = let x,largs = decompose_app (whd_all env c) in match kind x with | Prod (na,b,d) -> let () = assert (List.is_empty largs) in (** If one of the inductives of the mutually inductive block occurs in the left-hand side of a product, then such an occurrence is a non-strictly-positive recursive call. Occurrences in the right-hand side of the product must be strictly positive.*) (match weaker_noccur_between env n ntypes b with | None when chkpos -> failwith_non_pos_list n ntypes [b] | None -> check_pos (ienv_push_var ienv (na, b, mk_norec)) nmr d | Some b -> check_pos (ienv_push_var ienv (na, b, mk_norec)) nmr d) | Rel k -> (try let (ra,rarg) = List.nth ra_env (k-1) in let largs = List.map (whd_all env) largs in let nmr1 = (match ra with Mrec _ -> compute_rec_par ienv paramsctxt nmr largs | _ -> nmr) in (** The case where one of the inductives of the mutually inductive block occurs as an argument of another is not known to be safe. So Coq rejects it. *) if chkpos && not (List.for_all (noccur_between n ntypes) largs) then failwith_non_pos_list n ntypes largs else (nmr1,rarg) with Failure _ | Invalid_argument _ -> (nmr,mk_norec)) | Ind ind_kn -> (** If one of the inductives of the mutually inductive block being defined appears in a parameter, then we have a nested inductive type. The positivity is then discharged to the [check_positive_nested] function. *) if List.for_all (noccur_between n ntypes) largs then (nmr,mk_norec) else check_positive_nested ienv nmr (ind_kn, largs) | _err -> (** If an inductive of the mutually inductive block appears in any other way, then the positivy check gives up. *) if not chkpos || (noccur_between n ntypes x && List.for_all (noccur_between n ntypes) largs) then (nmr,mk_norec) else failwith_non_pos_list n ntypes (x::largs) (** [check_positive_nested] handles the case of nested inductive calls, that is, when an inductive types from the mutually inductive block is called as an argument of an inductive types (for the moment, this inductive type must be a previously defined types, not one of the types of the mutually inductive block being defined). *) (* accesses to the environment are not factorised, but is it worth? *) and check_positive_nested (env,n,ntypes,_ra_env as ienv) nmr ((mi,u), largs) = let (mib,mip) = lookup_mind_specif env mi in let auxnrecpar = mib.mind_nparams_rec in let auxnnonrecpar = mib.mind_nparams - auxnrecpar in let (auxrecparams,auxnonrecargs) = try List.chop auxnrecpar largs with Failure _ -> raise (IllFormedInd (LocalNonPos n)) in (** Inductives of the inductive block being defined are only allowed to appear nested in the parameters of another inductive type. Not in the proper indices. *) if chkpos && not (List.for_all (noccur_between n ntypes) auxnonrecargs) then failwith_non_pos_list n ntypes auxnonrecargs; (* Nested mutual inductive types are not supported *) let auxntyp = mib.mind_ntypes in if not (Int.equal auxntyp 1) then raise (IllFormedInd (LocalNonPos n)); (* The nested inductive type with parameters removed *) let auxlcvect = abstract_mind_lc auxntyp auxnrecpar mip.mind_nf_lc in (* Extends the environment with a variable corresponding to the inductive def *) let (env',_,_,_ as ienv') = ienv_push_inductive ienv ((mi,u),auxrecparams) in (* Parameters expressed in env' *) let auxrecparams' = List.map (lift auxntyp) auxrecparams in let irecargs_nmr = (** Checks that the "nesting" inductive type is covariant in the relevant parameters. In other words, that the (nested) parameters which are instantiated with inductives of the mutually inductive block occur positively in the types of the nested constructors. *) Array.map (function c -> let c' = hnf_prod_applist env' c auxrecparams' in (* skip non-recursive parameters *) let (ienv',c') = ienv_decompose_prod ienv' auxnnonrecpar c' in check_constructors ienv' false nmr c') auxlcvect in let irecargs = Array.map snd irecargs_nmr and nmr' = array_min nmr irecargs_nmr in (nmr',(Rtree.mk_rec [|mk_paths (Imbr mi) irecargs|]).(0)) (** [check_constructors ienv check_head nmr c] checks the positivity condition in the type [c] of a constructor (i.e. that recursive calls to the inductives of the mutually inductive definition appear strictly positively in each of the arguments of the constructor, see also [check_pos]). If [check_head] is [true], then the type of the fully applied constructor (the "head" of the type [c]) is checked to be the right (properly applied) inductive type. *) and check_constructors ienv check_head nmr c = let rec check_constr_rec (env,n,ntypes,_ra_env as ienv) nmr lrec c = let x,largs = decompose_app (whd_all env c) in match kind x with | Prod (na,b,d) -> let () = assert (List.is_empty largs) in if not recursive && not (noccur_between n ntypes b) then raise (InductiveError Type_errors.BadEntry); let nmr',recarg = check_pos ienv nmr b in let ienv' = ienv_push_var ienv (na,b,mk_norec) in check_constr_rec ienv' nmr' (recarg::lrec) d | hd -> let () = if check_head then begin match hd with | Rel j when Int.equal j (n + ntypes - i - 1) -> check_correct_par ienv paramsctxt (ntypes - i) largs | _ -> raise (IllFormedInd (LocalNotConstructor(paramsctxt,nnonrecargs))) end else if chkpos && not (List.for_all (noccur_between n ntypes) largs) then failwith_non_pos_list n ntypes largs in (nmr, List.rev lrec) in check_constr_rec ienv nmr [] c in let irecargs_nmr = Array.map2 (fun id c -> let _,rawc = mind_extract_params nparamsctxt c in try check_constructors ienv true nmr rawc with IllFormedInd err -> explain_ind_err id (ntypes-i) env nparamsctxt c err) (Array.of_list lcnames) indlc in let irecargs = Array.map snd irecargs_nmr and nmr' = array_min nmr irecargs_nmr in (nmr', mk_paths (Mrec ind) irecargs) (** [check_positivity ~chkpos kn env_ar paramsctxt inds] checks that the mutually inductive block [inds] is strictly positive. If [chkpos] is [false] then positivity is assumed, and [check_positivity_one] computes the subterms occurrences in a best-effort fashion. *) let check_positivity ~chkpos kn names env_ar_par paramsctxt finite inds = let ntypes = Array.length inds in let recursive = finite != BiFinite in let rc = Array.mapi (fun j t -> (Mrec (kn,j),t)) (Rtree.mk_rec_calls ntypes) in let ra_env_ar = Array.rev_to_list rc in let nparamsctxt = Context.Rel.length paramsctxt in let nmr = Context.Rel.nhyps paramsctxt in let check_one i (_,lcnames) (nindices,lc) = let ra_env_ar_par = List.init nparamsctxt (fun _ -> (Norec,mk_norec)) @ ra_env_ar in let ienv = (env_ar_par, 1+nparamsctxt, ntypes, ra_env_ar_par) in check_positivity_one ~chkpos recursive ienv paramsctxt (kn,i) nindices lcnames lc in let irecargs_nmr = Array.map2_i check_one names inds in let irecargs = Array.map snd irecargs_nmr and nmr' = array_min nmr irecargs_nmr in (nmr',Rtree.mk_rec irecargs) (************************************************************************) (************************************************************************) (* Build the inductive packet *) let fold_arity f acc params arity indices = match arity with | RegularArity ar -> f acc ar.mind_user_arity | TemplateArity _ -> let fold_ctx acc ctx = List.fold_left (fun acc d -> Context.Rel.Declaration.fold_constr (fun c acc -> f acc c) d acc) acc ctx in fold_ctx (fold_ctx acc params) indices let fold_inductive_blocks f acc params inds = Array.fold_left (fun acc ((arity,lc),(indices,_),_) -> fold_arity f (Array.fold_left f acc lc) params arity indices) acc inds let used_section_variables env params inds = let fold l c = Id.Set.union (Environ.global_vars_set env c) l in let ids = fold_inductive_blocks fold Id.Set.empty params inds in keep_hyps env ids let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i)) let rel_list n m = Array.to_list (rel_vect n m) (** From a rel context describing the constructor arguments, build an expansion function. The term built is expecting to be substituted first by a substitution of the form [params, x : ind params] *) let compute_projections (kn, i as ind) mib = let pkt = mib.mind_packets.(i) in let u = Univ.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in let subst = List.init mib.mind_ntypes (fun i -> mkIndU ((kn, mib.mind_ntypes - i - 1), u)) in let (ctx, cty) = pkt.mind_nf_lc.(0) in let cty = it_mkProd_or_LetIn cty ctx in let rctx, _ = decompose_prod_assum (substl subst cty) in let ctx, paramslet = List.chop pkt.mind_consnrealdecls.(0) rctx in (** We build a substitution smashing the lets in the record parameters so that typechecking projections requires just a substitution and not matching with a parameter context. *) let paramsletsubst = (* [Ind inst] is typed in context [params-wo-let] *) let inst' = rel_list 0 mib.mind_nparams in (* {params-wo-let |- subst:params] *) let subst = subst_of_rel_context_instance paramslet inst' in (* {params-wo-let, x:Ind inst' |- subst':(params,x:Ind inst)] *) let subst = (* For the record parameter: *) mkRel 1 :: List.map (lift 1) subst in subst in let projections decl (i, j, labs, rs, pbs, letsubst) = match decl with | LocalDef (_na,c,_t) -> (* From [params, field1,..,fieldj |- c(params,field1,..,fieldj)] to [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)] *) let c = liftn 1 j c in (* From [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)] to [params-wo-let, x:I |- c(params,proj1 x,..,projj x)] *) let c2 = substl letsubst c in (* From [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj)] to [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj+1)] *) let letsubst = c2 :: letsubst in (i, j+1, labs, rs, pbs, letsubst) | LocalAssum (na,t) -> match na.Context.binder_name with | Name id -> let r = na.Context.binder_relevance in let lab = Label.of_id id in let kn = Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg:i lab in (* from [params, field1,..,fieldj |- t(params,field1,..,fieldj)] to [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj] *) let t = liftn 1 j t in (* from [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj)] to [params-wo-let, x:I |- t(params,proj1 x,..,projj x)] *) let projty = substl letsubst t in (* from [params, x:I, field1,..,fieldj |- t(field1,..,fieldj)] to [params, x:I |- t(proj1 x,..,projj x)] *) let fterm = mkProj (Projection.make kn false, mkRel 1) in (i + 1, j + 1, lab :: labs, r :: rs, projty :: pbs, fterm :: letsubst) | Anonymous -> assert false (* checked by indTyping *) in let (_, _, labs, rs, pbs, _letsubst) = List.fold_right projections ctx (0, 1, [], [], [], paramsletsubst) in Array.of_list (List.rev labs), Array.of_list (List.rev rs), Array.of_list (List.rev pbs) let build_inductive env names prv univs variance paramsctxt kn isrecord isfinite inds nmr recargs = let ntypes = Array.length inds in (* Compute the set of used section variables *) let hyps = used_section_variables env paramsctxt inds in let nparamargs = Context.Rel.nhyps paramsctxt in (* Check one inductive *) let build_one_packet (id,cnames) ((arity,lc),(indices,splayed_lc),kelim) recarg = (* Type of constructors in normal form *) let nf_lc = Array.map (fun (d, b) -> (d@paramsctxt, b)) splayed_lc in let consnrealdecls = Array.map (fun (d,_) -> Context.Rel.length d) splayed_lc in let consnrealargs = Array.map (fun (d,_) -> Context.Rel.nhyps d) splayed_lc in let mind_relevance = match arity with | RegularArity { mind_sort;_ } -> Sorts.relevance_of_sort mind_sort | TemplateArity _ -> Sorts.Relevant in (* Assigning VM tags to constructors *) let nconst, nblock = ref 0, ref 0 in let transf num = let arity = List.length (dest_subterms recarg).(num) in if Int.equal arity 0 then let p = (!nconst, 0) in incr nconst; p else let p = (!nblock + 1, arity) in incr nblock; p (* les tag des constructeur constant commence a 0, les tag des constructeur non constant a 1 (0 => accumulator) *) in let rtbl = Array.init (List.length cnames) transf in (* Build the inductive packet *) { mind_typename = id; mind_arity = arity; mind_arity_ctxt = indices @ paramsctxt; mind_nrealargs = Context.Rel.nhyps indices; mind_nrealdecls = Context.Rel.length indices; mind_kelim = kelim; mind_consnames = Array.of_list cnames; mind_consnrealdecls = consnrealdecls; mind_consnrealargs = consnrealargs; mind_user_lc = lc; mind_nf_lc = nf_lc; mind_recargs = recarg; mind_relevance; mind_nb_constant = !nconst; mind_nb_args = !nblock; mind_reloc_tbl = rtbl; } in let packets = Array.map3 build_one_packet names inds recargs in let mib = (* Build the mutual inductive *) { mind_record = NotRecord; mind_ntypes = ntypes; mind_finite = isfinite; mind_hyps = hyps; mind_nparams = nparamargs; mind_nparams_rec = nmr; mind_params_ctxt = paramsctxt; mind_packets = packets; mind_universes = univs; mind_variance = variance; mind_private = prv; mind_typing_flags = Environ.typing_flags env; } in let record_info = match isrecord with | Some (Some rid) -> (** The elimination criterion ensures that all projections can be defined. *) let map i id = let labs, rs, projs = compute_projections (kn, i) mib in (id, labs, rs, projs) in PrimRecord (Array.mapi map rid) | Some None -> FakeRecord | None -> NotRecord in { mib with mind_record = record_info } (************************************************************************) (************************************************************************) let check_inductive env kn mie = (* First type-check the inductive definition *) let (env_ar_par, univs, variance, record, paramsctxt, inds) = IndTyping.typecheck_inductive env mie in (* Then check positivity conditions *) let chkpos = (Environ.typing_flags env).check_positive in let names = Array.map_of_list (fun entry -> entry.mind_entry_typename, entry.mind_entry_consnames) mie.mind_entry_inds in let (nmr,recargs) = check_positivity ~chkpos kn names env_ar_par paramsctxt mie.mind_entry_finite (Array.map (fun ((_,lc),(indices,_),_) -> Context.Rel.nhyps indices,lc) inds) in (* Build the inductive packets *) build_inductive env names mie.mind_entry_private univs variance paramsctxt kn record mie.mind_entry_finite inds nmr recargs coq-8.11.0/kernel/declarations.ml0000644000175000017500000002763113612664533016561 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* module_implementation module_retroknowledge | ModTypeRK : unit module_retroknowledge (** Extra invariants : - No [MEwith] inside a [mod_expr] implementation : the 'with' syntax is only supported for module types - A module application is atomic, for instance ((M N) P) : * the head of [MEapply] can only be another [MEapply] or a [MEident] * the argument of [MEapply] is now directly forced to be a [ModPath.t]. *) coq-8.11.0/kernel/sorts.mli0000644000175000017500000000321513612664533015424 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t -> bool val compare : t -> t -> int val hash : t -> int val is_sprop : t -> bool val is_set : t -> bool val is_prop : t -> bool val is_small : t -> bool val family : t -> family val hcons : t -> t val family_compare : family -> family -> int val family_equal : family -> family -> bool val family_leq : family -> family -> bool val univ_of_sort : t -> Univ.Universe.t val sort_of_univ : Univ.Universe.t -> t val super : t -> t (** On binders: is this variable proof relevant *) type relevance = Relevant | Irrelevant val relevance_hash : relevance -> int val relevance_equal : relevance -> relevance -> bool val relevance_of_sort : t -> relevance val relevance_of_sort_family : family -> relevance val debug_print : t -> Pp.t val pr_sort_family : family -> Pp.t coq-8.11.0/kernel/conv_oracle.ml0000644000175000017500000000727513612664533016405 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* true | _ -> false let transparent = default let is_transparent = function | Level 0 -> true | _ -> false type oracle = { var_opacity : level Id.Map.t; cst_opacity : level Cmap.t; var_trstate : Id.Pred.t; cst_trstate : Cpred.t; } let empty = { var_opacity = Id.Map.empty; cst_opacity = Cmap.empty; var_trstate = Id.Pred.full; cst_trstate = Cpred.full; } let get_strategy { var_opacity; cst_opacity; _ } f = function | VarKey id -> (try Id.Map.find id var_opacity with Not_found -> default) | ConstKey c -> (try Cmap.find (f c) cst_opacity with Not_found -> default) | RelKey _ -> Expand let set_strategy ({ var_opacity; cst_opacity; _ } as oracle) k l = match k with | VarKey id -> let var_opacity = if is_default l then Id.Map.remove id var_opacity else Id.Map.add id l var_opacity in let var_trstate = match l with | Opaque -> Id.Pred.remove id oracle.var_trstate | _ -> Id.Pred.add id oracle.var_trstate in { oracle with var_opacity; var_trstate; } | ConstKey c -> let cst_opacity = if is_default l then Cmap.remove c cst_opacity else Cmap.add c l cst_opacity in let cst_trstate = match l with | Opaque -> Cpred.remove c oracle.cst_trstate | _ -> Cpred.add c oracle.cst_trstate in { oracle with cst_opacity; cst_trstate; } | RelKey _ -> CErrors.user_err Pp.(str "set_strategy: RelKey") let fold_strategy f { var_opacity; cst_opacity; _ } accu = let fvar id lvl accu = f (VarKey id) lvl accu in let fcst cst lvl accu = f (ConstKey cst) lvl accu in let accu = Id.Map.fold fvar var_opacity accu in Cmap.fold fcst cst_opacity accu let get_transp_state { var_trstate; cst_trstate; _ } = { TransparentState.tr_var = var_trstate; tr_cst = cst_trstate } let dep_order l2r k1 k2 = match k1, k2 with | RelKey _, RelKey _ -> l2r | RelKey _, (VarKey _ | ConstKey _) -> true | VarKey _, RelKey _ -> false | VarKey _, VarKey _ -> l2r | VarKey _, ConstKey _ -> true | ConstKey _, (RelKey _ | VarKey _) -> false | ConstKey _, ConstKey _ -> l2r (* Unfold the first constant only if it is "more transparent" than the second one. In case of tie, use the recommended default. *) let oracle_order f o l2r k1 k2 = match get_strategy o f k1, get_strategy o f k2 with | Expand, Expand -> dep_order l2r k1 k2 | Expand, (Opaque | Level _) -> true | (Opaque | Level _), Expand -> false | Opaque, Opaque -> dep_order l2r k1 k2 | Level _, Opaque -> true | Opaque, Level _ -> false | Level n1, Level n2 -> if Int.equal n1 n2 then dep_order l2r k1 k2 else n1 < n2 let get_strategy o = get_strategy o (fun x -> x) coq-8.11.0/kernel/transparentState.mli0000644000175000017500000000206613612664533017617 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool val is_transparent_variable : t -> Id.t -> bool val is_transparent_constant : t -> Constant.t -> bool coq-8.11.0/kernel/cPrimitives.ml0000644000175000017500000001442313612664533016402 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 1 | Int63tail0 -> 2 | Int63add -> 3 | Int63sub -> 4 | Int63mul -> 5 | Int63div -> 6 | Int63mod -> 7 | Int63lsr -> 8 | Int63lsl -> 9 | Int63land -> 10 | Int63lor -> 11 | Int63lxor -> 12 | Int63addc -> 13 | Int63subc -> 14 | Int63addCarryC -> 15 | Int63subCarryC -> 16 | Int63mulc -> 17 | Int63diveucl -> 18 | Int63div21 -> 19 | Int63addMulDiv -> 20 | Int63eq -> 21 | Int63lt -> 22 | Int63le -> 23 | Int63compare -> 24 | Float64opp -> 25 | Float64abs -> 26 | Float64compare -> 27 | Float64classify -> 28 | Float64add -> 29 | Float64sub -> 30 | Float64mul -> 31 | Float64div -> 32 | Float64sqrt -> 33 | Float64ofInt63 -> 34 | Float64normfr_mantissa -> 35 | Float64frshiftexp -> 36 | Float64ldshiftexp -> 37 | Float64next_up -> 38 | Float64next_down -> 39 | Float64eq -> 40 | Float64lt -> 41 | Float64le -> 42 (* Should match names in nativevalues.ml *) let to_string = function | Int63head0 -> "head0" | Int63tail0 -> "tail0" | Int63add -> "add" | Int63sub -> "sub" | Int63mul -> "mul" | Int63div -> "div" | Int63mod -> "rem" | Int63lsr -> "l_sr" | Int63lsl -> "l_sl" | Int63land -> "l_and" | Int63lor -> "l_or" | Int63lxor -> "l_xor" | Int63addc -> "addc" | Int63subc -> "subc" | Int63addCarryC -> "addCarryC" | Int63subCarryC -> "subCarryC" | Int63mulc -> "mulc" | Int63diveucl -> "diveucl" | Int63div21 -> "div21" | Int63addMulDiv -> "addMulDiv" | Int63eq -> "eq" | Int63lt -> "lt" | Int63le -> "le" | Int63compare -> "compare" | Float64opp -> "fopp" | Float64abs -> "fabs" | Float64eq -> "feq" | Float64lt -> "flt" | Float64le -> "fle" | Float64compare -> "fcompare" | Float64classify -> "fclassify" | Float64add -> "fadd" | Float64sub -> "fsub" | Float64mul -> "fmul" | Float64div -> "fdiv" | Float64sqrt -> "fsqrt" | Float64ofInt63 -> "float_of_int" | Float64normfr_mantissa -> "normfr_mantissa" | Float64frshiftexp -> "frshiftexp" | Float64ldshiftexp -> "ldshiftexp" | Float64next_up -> "next_up" | Float64next_down -> "next_down" type prim_type = | PT_int63 | PT_float64 type 'a prim_ind = | PIT_bool : unit prim_ind | PIT_carry : prim_type prim_ind | PIT_pair : (prim_type * prim_type) prim_ind | PIT_cmp : unit prim_ind | PIT_f_cmp : unit prim_ind | PIT_f_class : unit prim_ind type prim_ind_ex = PIE : 'a prim_ind -> prim_ind_ex type ind_or_type = | PITT_ind : 'a prim_ind * 'a -> ind_or_type | PITT_type : prim_type -> ind_or_type let types = let int_ty = PITT_type PT_int63 in let float_ty = PITT_type PT_float64 in function | Int63head0 | Int63tail0 -> [int_ty; int_ty] | Int63add | Int63sub | Int63mul | Int63div | Int63mod | Int63lsr | Int63lsl | Int63land | Int63lor | Int63lxor -> [int_ty; int_ty; int_ty] | Int63addc | Int63subc | Int63addCarryC | Int63subCarryC -> [int_ty; int_ty; PITT_ind (PIT_carry, PT_int63)] | Int63mulc | Int63diveucl -> [int_ty; int_ty; PITT_ind (PIT_pair, (PT_int63, PT_int63))] | Int63eq | Int63lt | Int63le -> [int_ty; int_ty; PITT_ind (PIT_bool, ())] | Int63compare -> [int_ty; int_ty; PITT_ind (PIT_cmp, ())] | Int63div21 -> [int_ty; int_ty; int_ty; PITT_ind (PIT_pair, (PT_int63, PT_int63))] | Int63addMulDiv -> [int_ty; int_ty; int_ty; int_ty] | Float64opp | Float64abs | Float64sqrt | Float64next_up | Float64next_down -> [float_ty; float_ty] | Float64ofInt63 -> [int_ty; float_ty] | Float64normfr_mantissa -> [float_ty; int_ty] | Float64frshiftexp -> [float_ty; PITT_ind (PIT_pair, (PT_float64, PT_int63))] | Float64eq | Float64lt | Float64le -> [float_ty; float_ty; PITT_ind (PIT_bool, ())] | Float64compare -> [float_ty; float_ty; PITT_ind (PIT_f_cmp, ())] | Float64classify -> [float_ty; PITT_ind (PIT_f_class, ())] | Float64add | Float64sub | Float64mul | Float64div -> [float_ty; float_ty; float_ty] | Float64ldshiftexp -> [float_ty; int_ty; float_ty] type arg_kind = | Kparam (* not needed for the evaluation of the primitive when it reduces *) | Kwhnf (* need to be reduced in whnf before reducing the primitive *) | Karg (* no need to be reduced in whnf. example: [v] in [Array.set t i v] *) type args_red = arg_kind list (* Invariant only argument of type int63, float or an inductive can have kind Kwhnf *) let arity t = List.length (types t) - 1 let kind t = let rec aux n = if n <= 0 then [] else Kwhnf :: aux (n - 1) in aux (arity t) (** Special Entries for Register **) type op_or_type = | OT_op of t | OT_type of prim_type let prim_ind_to_string (type a) (p : a prim_ind) = match p with | PIT_bool -> "bool" | PIT_carry -> "carry" | PIT_pair -> "pair" | PIT_cmp -> "cmp" | PIT_f_cmp -> "f_cmp" | PIT_f_class -> "f_class" let prim_type_to_string = function | PT_int63 -> "int63_type" | PT_float64 -> "float64_type" let op_or_type_to_string = function | OT_op op -> to_string op | OT_type t -> prim_type_to_string t coq-8.11.0/kernel/declareops.ml0000644000175000017500000003230713612664533016226 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let x' = f sub x in if x' == x then ar else RegularArity x' | TemplateArity x -> let x' = g sub x in if x' == x then ar else TemplateArity x' let map_decl_arity f g = function | RegularArity a -> RegularArity (f a) | TemplateArity a -> TemplateArity (g a) let hcons_template_arity ar = { template_param_levels = ar.template_param_levels; (* List.Smart.map (Option.Smart.map Univ.hcons_univ_level) ar.template_param_levels; *) template_level = Univ.hcons_univ ar.template_level } let universes_context = function | Monomorphic _ -> Univ.AUContext.empty | Polymorphic ctx -> ctx let abstract_universes = function | Entries.Monomorphic_entry ctx -> Univ.empty_level_subst, Monomorphic ctx | Entries.Polymorphic_entry (nas, ctx) -> let (inst, auctx) = Univ.abstract_universes nas ctx in let inst = Univ.make_instance_subst inst in (inst, Polymorphic auctx) (** {6 Constants } *) let constant_is_polymorphic cb = match cb.const_universes with | Monomorphic _ -> false | Polymorphic _ -> true let constant_has_body cb = match cb.const_body with | Undef _ | Primitive _ -> false | Def _ | OpaqueDef _ -> true let constant_polymorphic_context cb = universes_context cb.const_universes let is_opaque cb = match cb.const_body with | OpaqueDef _ -> true | Undef _ | Def _ | Primitive _ -> false (** {7 Constant substitutions } *) let subst_rel_declaration sub = RelDecl.map_constr (subst_mps sub) let subst_rel_context sub = List.Smart.map (subst_rel_declaration sub) let subst_const_type sub arity = if is_empty_subst sub then arity else subst_mps sub arity (** No need here to check for physical equality after substitution, at least for Def due to the delayed substitution [subst_constr_subst]. *) let subst_const_def sub def = match def with | Undef _ | Primitive _ -> def | Def c -> Def (subst_constr sub c) | OpaqueDef o -> OpaqueDef (Opaqueproof.subst_opaque sub o) let subst_const_body sub cb = assert (List.is_empty cb.const_hyps); (* we're outside sections *) if is_empty_subst sub then cb else let body' = subst_const_def sub cb.const_body in let type' = subst_const_type sub cb.const_type in if body' == cb.const_body && type' == cb.const_type then cb else { const_hyps = []; const_body = body'; const_type = type'; const_body_code = Option.map (Cemitcodes.subst_to_patch_subst sub) cb.const_body_code; const_universes = cb.const_universes; const_relevance = cb.const_relevance; const_inline_code = cb.const_inline_code; const_typing_flags = cb.const_typing_flags } (** {7 Hash-consing of constants } *) (** This hash-consing is currently quite partial : we only share internal fields (e.g. constr), and not the records themselves. But would it really bring substantial gains ? *) let hcons_rel_decl = RelDecl.map_name Names.Name.hcons %> RelDecl.map_value Constr.hcons %> RelDecl.map_type Constr.hcons let hcons_rel_context l = List.Smart.map hcons_rel_decl l let hcons_const_def = function | Undef inl -> Undef inl | Primitive p -> Primitive p | Def l_constr -> let constr = force_constr l_constr in Def (from_val (Constr.hcons constr)) | OpaqueDef _ as x -> x (* hashconsed when turned indirect *) let hcons_universes cbu = match cbu with | Monomorphic ctx -> Monomorphic (Univ.hcons_universe_context_set ctx) | Polymorphic ctx -> Polymorphic (Univ.hcons_abstract_universe_context ctx) let hcons_const_body cb = { cb with const_body = hcons_const_def cb.const_body; const_type = Constr.hcons cb.const_type; const_universes = hcons_universes cb.const_universes; } (** {6 Inductive types } *) let eq_recarg r1 r2 = match r1, r2 with | Norec, Norec -> true | Mrec i1, Mrec i2 -> Names.eq_ind i1 i2 | Imbr i1, Imbr i2 -> Names.eq_ind i1 i2 | _ -> false let subst_recarg sub r = match r with | Norec -> r | Mrec (kn,i) -> let kn' = subst_mind sub kn in if kn==kn' then r else Mrec (kn',i) | Imbr (kn,i) -> let kn' = subst_mind sub kn in if kn==kn' then r else Imbr (kn',i) let mk_norec = Rtree.mk_node Norec [||] let mk_paths r recargs = Rtree.mk_node r (Array.map (fun l -> Rtree.mk_node Norec (Array.of_list l)) recargs) let dest_recarg p = fst (Rtree.dest_node p) (* dest_subterms returns the sizes of each argument of each constructor of an inductive object of size [p]. This should never be done for Norec, because the number of sons does not correspond to the number of constructors. *) let dest_subterms p = let (ra,cstrs) = Rtree.dest_node p in assert (match ra with Norec -> false | _ -> true); Array.map (fun t -> Array.to_list (snd (Rtree.dest_node t))) cstrs let recarg_length p j = let (_,cstrs) = Rtree.dest_node p in Array.length (snd (Rtree.dest_node cstrs.(j-1))) let subst_wf_paths sub p = Rtree.Smart.map (subst_recarg sub) p (** {7 Substitution of inductive declarations } *) let subst_regular_ind_arity sub s = let uar' = subst_mps sub s.mind_user_arity in if uar' == s.mind_user_arity then s else { mind_user_arity = uar'; mind_sort = s.mind_sort } let subst_template_ind_arity _sub s = s (* FIXME records *) let subst_ind_arity = subst_decl_arity subst_regular_ind_arity subst_template_ind_arity let subst_mind_packet sub mbp = { mind_consnames = mbp.mind_consnames; mind_consnrealdecls = mbp.mind_consnrealdecls; mind_consnrealargs = mbp.mind_consnrealargs; mind_typename = mbp.mind_typename; mind_nf_lc = Array.Smart.map (fun (ctx, c) -> Context.Rel.map (subst_mps sub) ctx, subst_mps sub c) mbp.mind_nf_lc; mind_arity_ctxt = subst_rel_context sub mbp.mind_arity_ctxt; mind_arity = subst_ind_arity sub mbp.mind_arity; mind_user_lc = Array.Smart.map (subst_mps sub) mbp.mind_user_lc; mind_nrealargs = mbp.mind_nrealargs; mind_nrealdecls = mbp.mind_nrealdecls; mind_kelim = mbp.mind_kelim; mind_recargs = subst_wf_paths sub mbp.mind_recargs (*wf_paths*); mind_relevance = mbp.mind_relevance; mind_nb_constant = mbp.mind_nb_constant; mind_nb_args = mbp.mind_nb_args; mind_reloc_tbl = mbp.mind_reloc_tbl } let subst_mind_record sub r = match r with | NotRecord -> NotRecord | FakeRecord -> FakeRecord | PrimRecord infos -> let map (id, ps, rs, pb as info) = let pb' = Array.Smart.map (subst_mps sub) pb in if pb' == pb then info else (id, ps, rs, pb') in let infos' = Array.Smart.map map infos in if infos' == infos then r else PrimRecord infos' let subst_mind_body sub mib = { mind_record = subst_mind_record sub mib.mind_record ; mind_finite = mib.mind_finite ; mind_ntypes = mib.mind_ntypes ; mind_hyps = (match mib.mind_hyps with [] -> [] | _ -> assert false); mind_nparams = mib.mind_nparams; mind_nparams_rec = mib.mind_nparams_rec; mind_params_ctxt = Context.Rel.map (subst_mps sub) mib.mind_params_ctxt; mind_packets = Array.Smart.map (subst_mind_packet sub) mib.mind_packets ; mind_universes = mib.mind_universes; mind_variance = mib.mind_variance; mind_private = mib.mind_private; mind_typing_flags = mib.mind_typing_flags; } let inductive_polymorphic_context mib = universes_context mib.mind_universes let inductive_is_polymorphic mib = match mib.mind_universes with | Monomorphic _ -> false | Polymorphic _ctx -> true let inductive_is_cumulative mib = Option.has_some mib.mind_variance let inductive_make_projection ind mib ~proj_arg = match mib.mind_record with | NotRecord | FakeRecord -> None | PrimRecord infos -> let _, labs, _, _ = infos.(snd ind) in Some (Names.Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg labs.(proj_arg)) let inductive_make_projections ind mib = match mib.mind_record with | NotRecord | FakeRecord -> None | PrimRecord infos -> let _, labs, _, _ = infos.(snd ind) in let projs = Array.mapi (fun proj_arg lab -> Names.Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg lab) labs in Some projs let relevance_of_projection_repr mib p = let _mind,i = Names.Projection.Repr.inductive p in match mib.mind_record with | NotRecord | FakeRecord -> CErrors.anomaly ~label:"relevance_of_projection" Pp.(str "not a projection") | PrimRecord infos -> let _,_,rs,_ = infos.(i) in rs.(Names.Projection.Repr.arg p) (** {6 Hash-consing of inductive declarations } *) let hcons_regular_ind_arity a = { mind_user_arity = Constr.hcons a.mind_user_arity; mind_sort = Sorts.hcons a.mind_sort } (** Just as for constants, this hash-consing is quite partial *) let hcons_ind_arity = map_decl_arity hcons_regular_ind_arity hcons_template_arity (** Substitution of inductive declarations *) let hcons_mind_packet oib = let user = Array.Smart.map Constr.hcons oib.mind_user_lc in let map (ctx, c) = Context.Rel.map Constr.hcons ctx, Constr.hcons c in let nf = Array.Smart.map map oib.mind_nf_lc in { oib with mind_typename = Names.Id.hcons oib.mind_typename; mind_arity_ctxt = hcons_rel_context oib.mind_arity_ctxt; mind_arity = hcons_ind_arity oib.mind_arity; mind_consnames = Array.Smart.map Names.Id.hcons oib.mind_consnames; mind_user_lc = user; mind_nf_lc = nf } let hcons_mind mib = { mib with mind_packets = Array.Smart.map hcons_mind_packet mib.mind_packets; mind_params_ctxt = hcons_rel_context mib.mind_params_ctxt; mind_universes = hcons_universes mib.mind_universes } (** Hashconsing of modules *) let hcons_functorize hty he hself f = match f with | NoFunctor e -> let e' = he e in if e == e' then f else NoFunctor e' | MoreFunctor (mid, ty, nf) -> (** FIXME *) let mid' = mid in let ty' = hty ty in let nf' = hself nf in if mid == mid' && ty == ty' && nf == nf' then f else MoreFunctor (mid, ty', nf') let hcons_module_alg_expr me = me let rec hcons_structure_field_body sb = match sb with | SFBconst cb -> let cb' = hcons_const_body cb in if cb == cb' then sb else SFBconst cb' | SFBmind mib -> let mib' = hcons_mind mib in if mib == mib' then sb else SFBmind mib' | SFBmodule mb -> let mb' = hcons_module_body mb in if mb == mb' then sb else SFBmodule mb' | SFBmodtype mb -> let mb' = hcons_module_type mb in if mb == mb' then sb else SFBmodtype mb' and hcons_structure_body sb = (** FIXME *) let map (l, sfb as fb) = let l' = Names.Label.hcons l in let sfb' = hcons_structure_field_body sfb in if l == l' && sfb == sfb' then fb else (l', sfb') in List.Smart.map map sb and hcons_module_signature ms = hcons_functorize hcons_module_type hcons_structure_body hcons_module_signature ms and hcons_module_expression me = hcons_functorize hcons_module_type hcons_module_alg_expr hcons_module_expression me and hcons_module_implementation mip = match mip with | Abstract -> Abstract | Algebraic me -> let me' = hcons_module_expression me in if me == me' then mip else Algebraic me' | Struct ms -> let ms' = hcons_module_signature ms in if ms == ms' then mip else Struct ms | FullStruct -> FullStruct and hcons_generic_module_body : 'a. ('a -> 'a) -> 'a generic_module_body -> 'a generic_module_body = fun hcons_impl mb -> let mp' = mb.mod_mp in let expr' = hcons_impl mb.mod_expr in let type' = hcons_module_signature mb.mod_type in let type_alg' = mb.mod_type_alg in let constraints' = Univ.hcons_universe_context_set mb.mod_constraints in let delta' = mb.mod_delta in let retroknowledge' = mb.mod_retroknowledge in if mb.mod_mp == mp' && mb.mod_expr == expr' && mb.mod_type == type' && mb.mod_type_alg == type_alg' && mb.mod_constraints == constraints' && mb.mod_delta == delta' && mb.mod_retroknowledge == retroknowledge' then mb else { mod_mp = mp'; mod_expr = expr'; mod_type = type'; mod_type_alg = type_alg'; mod_constraints = constraints'; mod_delta = delta'; mod_retroknowledge = retroknowledge'; } and hcons_module_body mb = hcons_generic_module_body hcons_module_implementation mb and hcons_module_type mb = hcons_generic_module_body (fun () -> ()) mb coq-8.11.0/kernel/section.mli0000644000175000017500000000651213612664533015721 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool (** Checks whether there is no opened section *) val depth : 'a t -> int (** Number of nested sections (0 if no sections are open) *) (** {6 Manipulating sections} *) type section_entry = | SecDefinition of Constant.t | SecInductive of MutInd.t val open_section : custom:'a -> 'a t -> 'a t (** Open a new section with the provided universe polymorphic status. Sections can be nested, with the proviso that polymorphic sections cannot appear inside a monomorphic one. A custom data can be attached to this section, that will be returned by {!close_section}. *) val close_section : 'a t -> 'a t * section_entry list * ContextSet.t * 'a (** Close the current section and returns the entries defined inside, the set of global monomorphic constraints added in this section, and the custom data provided at the opening of the section. *) (** {6 Extending sections} *) val push_local : 'a t -> 'a t (** Extend the current section with a local definition (cf. push_named). *) val push_context : Name.t array * UContext.t -> 'a t -> 'a t (** Extend the current section with a local universe context. Assumes that the last opened section is polymorphic. *) val push_constraints : ContextSet.t -> 'a t -> 'a t (** Extend the current section with a global universe context. Assumes that the last opened section is monomorphic. *) val push_constant : poly:bool -> Constant.t -> 'a t -> 'a t (** Make the constant as having been defined in this section. *) val push_inductive : poly:bool -> MutInd.t -> 'a t -> 'a t (** Make the inductive block as having been defined in this section. *) (** {6 Retrieving section data} *) type abstr_info = private { abstr_ctx : Constr.named_context; (** Section variables of this prefix *) abstr_subst : Univ.Instance.t; (** Actual names of the abstracted variables *) abstr_uctx : Univ.AUContext.t; (** Universe quantification, same length as the substitution *) } (** Data needed to abstract over the section variable and universe hypotheses *) val empty_segment : abstr_info (** Nothing to abstract *) val segment_of_constant : Environ.env -> Constant.t -> 'a t -> abstr_info (** Section segment at the time of the constant declaration *) val segment_of_inductive : Environ.env -> MutInd.t -> 'a t -> abstr_info (** Section segment at the time of the inductive declaration *) val replacement_context : Environ.env -> 'a t -> Opaqueproof.work_list (** Section segments of all declarations from this section. *) val is_in_section : Environ.env -> GlobRef.t -> 'a t -> bool coq-8.11.0/kernel/nativelibrary.mli0000644000175000017500000000164213612664533017127 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* DirPath.t -> env -> module_signature -> global list * Nativevalues.symbols coq-8.11.0/kernel/cPrimitives.mli0000644000175000017500000000447313612664533016557 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t -> bool type arg_kind = | Kparam (* not needed for the elavuation of the primitive*) | Kwhnf (* need to be reduced in whnf before reducing the primitive *) | Karg (* no need to be reduced in whnf *) type args_red = arg_kind list val hash : t -> int val to_string : t -> string val arity : t -> int val kind : t -> args_red (** Special Entries for Register **) type prim_type = | PT_int63 | PT_float64 type 'a prim_ind = | PIT_bool : unit prim_ind | PIT_carry : prim_type prim_ind | PIT_pair : (prim_type * prim_type) prim_ind | PIT_cmp : unit prim_ind | PIT_f_cmp : unit prim_ind | PIT_f_class : unit prim_ind type prim_ind_ex = PIE : 'a prim_ind -> prim_ind_ex type op_or_type = | OT_op of t | OT_type of prim_type val prim_ind_to_string : 'a prim_ind -> string val op_or_type_to_string : op_or_type -> string type ind_or_type = | PITT_ind : 'a prim_ind * 'a -> ind_or_type | PITT_type : prim_type -> ind_or_type val types : t -> ind_or_type list coq-8.11.0/kernel/context.ml0000644000175000017500000004377113612664533015600 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* na (** Return the name bound by a given declaration. *) let get_name x = (get_annot x).binder_name (** Return [Some value] for local-declarations and [None] for local-assumptions. *) let get_value = function | LocalAssum _ -> None | LocalDef (_,v,_) -> Some v (** Return the type of the name bound by a given declaration. *) let get_type = function | LocalAssum (_,ty) | LocalDef (_,_,ty) -> ty let get_relevance x = (get_annot x).binder_relevance (** Set the name that is bound by a given declaration. *) let set_name na = function | LocalAssum (x,ty) -> LocalAssum ({x with binder_name=na}, ty) | LocalDef (x,v,ty) -> LocalDef ({x with binder_name=na}, v, ty) (** Set the type of the bound variable in a given declaration. *) let set_type ty = function | LocalAssum (na,_) -> LocalAssum (na, ty) | LocalDef (na,v,_) -> LocalDef (na, v, ty) (** Return [true] iff a given declaration is a local assumption. *) let is_local_assum = function | LocalAssum _ -> true | LocalDef _ -> false (** Return [true] iff a given declaration is a local definition. *) let is_local_def = function | LocalAssum _ -> false | LocalDef _ -> true (** Check whether any term in a given declaration satisfies a given predicate. *) let exists f = function | LocalAssum (_, ty) -> f ty | LocalDef (_, v, ty) -> f v || f ty (** Check whether all terms in a given declaration satisfy a given predicate. *) let for_all f = function | LocalAssum (_, ty) -> f ty | LocalDef (_, v, ty) -> f v && f ty (** Check whether the two given declarations are equal. *) let equal eq decl1 decl2 = match decl1, decl2 with | LocalAssum (n1,ty1), LocalAssum (n2, ty2) -> eq_annot Name.equal n1 n2 && eq ty1 ty2 | LocalDef (n1,v1,ty1), LocalDef (n2,v2,ty2) -> eq_annot Name.equal n1 n2 && eq v1 v2 && eq ty1 ty2 | _ -> false (** Map the name bound by a given declaration. *) let map_name f x = let na = get_name x in let na' = f na in if na == na' then x else set_name na' x (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) let map_value f = function | LocalAssum _ as decl -> decl | LocalDef (na, v, t) as decl -> let v' = f v in if v == v' then decl else LocalDef (na, v', t) (** Map the type of the name bound by a given declaration. *) let map_type f = function | LocalAssum (na, ty) as decl -> let ty' = f ty in if ty == ty' then decl else LocalAssum (na, ty') | LocalDef (na, v, ty) as decl -> let ty' = f ty in if ty == ty' then decl else LocalDef (na, v, ty') (** Map all terms in a given declaration. *) let map_constr f = function | LocalAssum (na, ty) as decl -> let ty' = f ty in if ty == ty' then decl else LocalAssum (na, ty') | LocalDef (na, v, ty) as decl -> let v' = f v in let ty' = f ty in if v == v' && ty == ty' then decl else LocalDef (na, v', ty') let map_constr_het f = function | LocalAssum (na, ty) -> let ty' = f ty in LocalAssum (na, ty') | LocalDef (na, v, ty) -> let v' = f v in let ty' = f ty in LocalDef (na, v', ty') (** Perform a given action on all terms in a given declaration. *) let iter_constr f = function | LocalAssum (_,ty) -> f ty | LocalDef (_,v,ty) -> f v; f ty (** Reduce all terms in a given declaration to a single value. *) let fold_constr f decl acc = match decl with | LocalAssum (_n,ty) -> f ty acc | LocalDef (_n,v,ty) -> f ty (f v acc) let to_tuple = function | LocalAssum (na, ty) -> na, None, ty | LocalDef (na, v, ty) -> na, Some v, ty let drop_body = function | LocalAssum _ as d -> d | LocalDef (na, _v, ty) -> LocalAssum (na, ty) end (** Rel-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list (** empty rel-context *) let empty = [] (** Return a new rel-context enriched by with a given inner-most declaration. *) let add d ctx = d :: ctx (** Return the number of {e local declarations} in a given context. *) let length = List.length (** [extended_rel_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ] with n = |Δ| and with the local definitions of [Γ] skipped in [args]. Example: for [x:T,y:=c,z:U] and [n]=2, it gives [Rel 5, Rel 3]. *) let nhyps ctx = let open Declaration in let rec nhyps acc = function | [] -> acc | LocalAssum _ :: hyps -> nhyps (succ acc) hyps | LocalDef _ :: hyps -> nhyps acc hyps in nhyps 0 ctx (** Return a declaration designated by a given de Bruijn index. @raise Not_found if the designated de Bruijn index is not present in the designated rel-context. *) let rec lookup n ctx = match n, ctx with | 1, decl :: _ -> decl | n, _ :: sign -> lookup (n-1) sign | _, [] -> raise Not_found (** Check whether given two rel-contexts are equal. *) let equal eq l = List.equal (fun c -> Declaration.equal eq c) l (** Map all terms in a given rel-context. *) let map f = List.Smart.map (Declaration.map_constr f) (** Perform a given action on every declaration in a given rel-context. *) let iter f = List.iter (Declaration.iter_constr f) (** Reduce all terms in a given rel-context to a single value. Innermost declarations are processed first. *) let fold_inside f ~init = List.fold_left f init (** Reduce all terms in a given rel-context to a single value. Outermost declarations are processed first. *) let fold_outside f l ~init = List.fold_right f l init (** Map a given rel-context to a list where each {e local assumption} is mapped to [true] and each {e local definition} is mapped to [false]. *) let to_tags l = let rec aux l = function | [] -> l | Declaration.LocalDef _ :: ctx -> aux (true::l) ctx | Declaration.LocalAssum _ :: ctx -> aux (false::l) ctx in aux [] l let drop_bodies l = List.Smart.map Declaration.drop_body l (** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ] with n = |Δ| and with the {e local definitions} of [Γ] skipped in [args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *) let to_extended_list mk n l = let rec reln l p = function | Declaration.LocalAssum _ :: hyps -> reln (mk (n+p) :: l) (p+1) hyps | Declaration.LocalDef _ :: hyps -> reln l (p+1) hyps | [] -> l in reln [] 1 l (** [extended_vect n Γ] does the same, returning instead an array. *) let to_extended_vect mk n hyps = Array.of_list (to_extended_list mk n hyps) end (** This module represents contexts that can capture non-anonymous variables. Individual declarations are then designated by the identifiers they bind. *) module Named = struct (** Representation of {e local declarations}. *) module Declaration = struct (** local declaration *) type ('constr, 'types) pt = | LocalAssum of Id.t binder_annot * 'types (** identifier, type *) | LocalDef of Id.t binder_annot * 'constr * 'types (** identifier, value, type *) let get_annot = function | LocalAssum (na,_) | LocalDef (na,_,_) -> na (** Return the identifier bound by a given declaration. *) let get_id x = (get_annot x).binder_name (** Return [Some value] for local-declarations and [None] for local-assumptions. *) let get_value = function | LocalAssum _ -> None | LocalDef (_,v,_) -> Some v (** Return the type of the name bound by a given declaration. *) let get_type = function | LocalAssum (_,ty) | LocalDef (_,_,ty) -> ty let get_relevance x = (get_annot x).binder_relevance (** Set the identifier that is bound by a given declaration. *) let set_id id = let set x = {x with binder_name = id} in function | LocalAssum (x,ty) -> LocalAssum (set x, ty) | LocalDef (x, v, ty) -> LocalDef (set x, v, ty) (** Set the type of the bound variable in a given declaration. *) let set_type ty = function | LocalAssum (id,_) -> LocalAssum (id, ty) | LocalDef (id,v,_) -> LocalDef (id, v, ty) (** Return [true] iff a given declaration is a local assumption. *) let is_local_assum = function | LocalAssum _ -> true | LocalDef _ -> false (** Return [true] iff a given declaration is a local definition. *) let is_local_def = function | LocalDef _ -> true | LocalAssum _ -> false (** Check whether any term in a given declaration satisfies a given predicate. *) let exists f = function | LocalAssum (_, ty) -> f ty | LocalDef (_, v, ty) -> f v || f ty (** Check whether all terms in a given declaration satisfy a given predicate. *) let for_all f = function | LocalAssum (_, ty) -> f ty | LocalDef (_, v, ty) -> f v && f ty (** Check whether the two given declarations are equal. *) let equal eq decl1 decl2 = match decl1, decl2 with | LocalAssum (id1, ty1), LocalAssum (id2, ty2) -> eq_annot Id.equal id1 id2 && eq ty1 ty2 | LocalDef (id1, v1, ty1), LocalDef (id2, v2, ty2) -> eq_annot Id.equal id1 id2 && eq v1 v2 && eq ty1 ty2 | _ -> false (** Map the identifier bound by a given declaration. *) let map_id f x = let id = get_id x in let id' = f id in if id == id' then x else set_id id' x (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) let map_value f = function | LocalAssum _ as decl -> decl | LocalDef (na, v, t) as decl -> let v' = f v in if v == v' then decl else LocalDef (na, v', t) (** Map the type of the name bound by a given declaration. *) let map_type f = function | LocalAssum (id, ty) as decl -> let ty' = f ty in if ty == ty' then decl else LocalAssum (id, ty') | LocalDef (id, v, ty) as decl -> let ty' = f ty in if ty == ty' then decl else LocalDef (id, v, ty') (** Map all terms in a given declaration. *) let map_constr f = function | LocalAssum (id, ty) as decl -> let ty' = f ty in if ty == ty' then decl else LocalAssum (id, ty') | LocalDef (id, v, ty) as decl -> let v' = f v in let ty' = f ty in if v == v' && ty == ty' then decl else LocalDef (id, v', ty') (** Perform a given action on all terms in a given declaration. *) let iter_constr f = function | LocalAssum (_, ty) -> f ty | LocalDef (_, v, ty) -> f v; f ty (** Reduce all terms in a given declaration to a single value. *) let fold_constr f decl a = match decl with | LocalAssum (_, ty) -> f ty a | LocalDef (_, v, ty) -> a |> f v |> f ty let to_tuple = function | LocalAssum (id, ty) -> id, None, ty | LocalDef (id, v, ty) -> id, Some v, ty let of_tuple = function | id, None, ty -> LocalAssum (id, ty) | id, Some v, ty -> LocalDef (id, v, ty) let drop_body = function | LocalAssum _ as d -> d | LocalDef (id, _v, ty) -> LocalAssum (id, ty) let of_rel_decl f = function | Rel.Declaration.LocalAssum (na,t) -> LocalAssum (map_annot f na, t) | Rel.Declaration.LocalDef (na,v,t) -> LocalDef (map_annot f na, v, t) let to_rel_decl = let name x = {binder_name=Name x.binder_name;binder_relevance=x.binder_relevance} in function | LocalAssum (id,t) -> Rel.Declaration.LocalAssum (name id, t) | LocalDef (id,v,t) -> Rel.Declaration.LocalDef (name id,v,t) end (** Named-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list (** empty named-context *) let empty = [] (** empty named-context *) let add d ctx = d :: ctx (** Return the number of {e local declarations} in a given named-context. *) let length = List.length (** Return a declaration designated by a given identifier @raise Not_found if the designated identifier is not present in the designated named-context. *) let rec lookup id = function | decl :: _ when Id.equal id (Declaration.get_id decl) -> decl | _ :: sign -> lookup id sign | [] -> raise Not_found (** Check whether given two named-contexts are equal. *) let equal eq l = List.equal (fun c -> Declaration.equal eq c) l (** Map all terms in a given named-context. *) let map f = List.Smart.map (Declaration.map_constr f) (** Perform a given action on every declaration in a given named-context. *) let iter f = List.iter (Declaration.iter_constr f) (** Reduce all terms in a given named-context to a single value. Innermost declarations are processed first. *) let fold_inside f ~init = List.fold_left f init (** Reduce all terms in a given named-context to a single value. Outermost declarations are processed first. *) let fold_outside f l ~init = List.fold_right f l init (** Return the set of all identifiers bound in a given named-context. *) let to_vars l = List.fold_left (fun accu decl -> Id.Set.add (Declaration.get_id decl) accu) Id.Set.empty l let drop_bodies l = List.Smart.map Declaration.drop_body l (** [instance_from_named_context Ω] builds an instance [args] such that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it gives [Var id1, Var id3]. All [idj] are supposed distinct. *) let to_instance mk l = let filter = function | Declaration.LocalAssum (id, _) -> Some (mk id.binder_name) | _ -> None in List.map_filter filter l end module Compacted = struct module Declaration = struct type ('constr, 'types) pt = | LocalAssum of Id.t binder_annot list * 'types | LocalDef of Id.t binder_annot list * 'constr * 'types let map_constr f = function | LocalAssum (ids, ty) as decl -> let ty' = f ty in if ty == ty' then decl else LocalAssum (ids, ty') | LocalDef (ids, c, ty) as decl -> let ty' = f ty in let c' = f c in if c == c' && ty == ty' then decl else LocalDef (ids,c',ty') let of_named_decl = function | Named.Declaration.LocalAssum (id,t) -> LocalAssum ([id],t) | Named.Declaration.LocalDef (id,v,t) -> LocalDef ([id],v,t) let to_named_context = function | LocalAssum (ids, t) -> List.map (fun id -> Named.Declaration.LocalAssum (id,t)) ids | LocalDef (ids, v, t) -> List.map (fun id -> Named.Declaration.LocalDef (id,v,t)) ids end type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list let fold f l ~init = List.fold_right f l init end coq-8.11.0/kernel/sorts.ml0000644000175000017500000000740413612664533015257 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* u | Set -> Universe.type0 | Prop -> Universe.type0m | SProp -> Universe.sprop let sort_of_univ u = if Universe.is_sprop u then sprop else if is_type0m_univ u then prop else if is_type0_univ u then set else Type u let compare s1 s2 = if s1 == s2 then 0 else match s1, s2 with | SProp, SProp -> 0 | SProp, _ -> -1 | _, SProp -> 1 | Prop, Prop -> 0 | Prop, _ -> -1 | Set, Prop -> 1 | Set, Set -> 0 | Set, _ -> -1 | Type u1, Type u2 -> Universe.compare u1 u2 | Type _, _ -> -1 let equal s1 s2 = Int.equal (compare s1 s2) 0 let super = function | SProp | Prop | Set -> Type (Universe.type1) | Type u -> Type (Universe.super u) let is_sprop = function | SProp -> true | Prop | Set | Type _ -> false let is_prop = function | Prop -> true | SProp | Set | Type _ -> false let is_set = function | Set -> true | SProp | Prop | Type _ -> false let is_small = function | SProp | Prop | Set -> true | Type _ -> false let family = function | SProp -> InSProp | Prop -> InProp | Set -> InSet | Type _ -> InType let family_compare a b = match a,b with | InSProp, InSProp -> 0 | InSProp, _ -> -1 | _, InSProp -> 1 | InProp, InProp -> 0 | InProp, _ -> -1 | _, InProp -> 1 | InSet, InSet -> 0 | InSet, _ -> -1 | _, InSet -> 1 | InType, InType -> 0 let family_equal = (==) let family_leq a b = family_compare a b <= 0 open Hashset.Combine let hash = function | SProp -> combinesmall 1 0 | Prop -> combinesmall 1 1 | Set -> combinesmall 1 2 | Type u -> let h = Univ.Universe.hash u in combinesmall 2 h module Hsorts = Hashcons.Make( struct type _t = t type t = _t type u = Universe.t -> Universe.t let hashcons huniv = function | Type u as c -> let u' = huniv u in if u' == u then c else Type u' | s -> s let eq s1 s2 = match (s1,s2) with | Prop, Prop | Set, Set -> true | (Type u1, Type u2) -> u1 == u2 |_ -> false let hash = hash end) let hcons = Hashcons.simple_hcons Hsorts.generate Hsorts.hcons hcons_univ (** On binders: is this variable proof relevant *) type relevance = Relevant | Irrelevant let relevance_equal r1 r2 = match r1,r2 with | Relevant, Relevant | Irrelevant, Irrelevant -> true | (Relevant | Irrelevant), _ -> false let relevance_of_sort_family = function | InSProp -> Irrelevant | _ -> Relevant let relevance_hash = function | Relevant -> 0 | Irrelevant -> 1 let relevance_of_sort = function | SProp -> Irrelevant | _ -> Relevant let debug_print = function | SProp -> Pp.(str "SProp") | Prop -> Pp.(str "Prop") | Set -> Pp.(str "Set") | Type u -> Pp.(str "Type(" ++ Univ.Universe.pr u ++ str ")") let pr_sort_family = function | InSProp -> Pp.(str "SProp") | InProp -> Pp.(str "Prop") | InSet -> Pp.(str "Set") | InType -> Pp.(str "Type") coq-8.11.0/kernel/genOpcodeFiles.ml0000644000175000017500000001141313612664533016766 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* (if isOcaml then pp_ocaml_comment else pp_c_comment) fmt Format.pp_print_string) "DO NOT EDIT: automatically generated by kernel/genOpcodeFiles.ml" let pp_with_commas fmt k = Array.iteri (fun n s -> Format.fprintf fmt " %a%s@." k s (if n + 1 < Array.length opcodes then "," else "") ) opcodes let pp_coq_instruct_h fmt = let line = Format.fprintf fmt "%s@." in pp_header false fmt; line "#pragma once"; line "enum instructions {"; pp_with_commas fmt Format.pp_print_string; line "};" let pp_coq_jumptbl_h fmt = pp_with_commas fmt (fun fmt -> Format.fprintf fmt "&&coq_lbl_%s") let pp_copcodes_ml fmt = pp_header true fmt; Array.iteri (fun n s -> Format.fprintf fmt "let op%s = %d@.@." s n ) opcodes let usage () = Format.eprintf "usage: %s [enum|jump|copml]@." Sys.argv.(0); exit 1 let main () = match Sys.argv.(1) with | "enum" -> pp_coq_instruct_h Format.std_formatter | "jump" -> pp_coq_jumptbl_h Format.std_formatter | "copml" -> pp_copcodes_ml Format.std_formatter | _ -> usage () | exception Invalid_argument _ -> usage () let () = main () coq-8.11.0/kernel/nativelambda.mli0000644000175000017500000000563013612664533016704 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr option; evars_metas : metavariable -> types } val empty_evars : evars val decompose_Llam : lambda -> Name.t Context.binder_annot array * lambda val decompose_Llam_Llet : lambda -> (Name.t Context.binder_annot * lambda option) array * lambda val is_lazy : constr -> bool val mk_lazy : lambda -> lambda val get_mind_prefix : env -> MutInd.t -> string val get_alias : env -> pconstant -> pconstant val lambda_of_constr : env -> evars -> Constr.constr -> lambda coq-8.11.0/kernel/mod_subst.mli0000644000175000017500000001352213612664533016253 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ModPath.t -> delta_resolver -> delta_resolver val add_kn_delta_resolver : KerName.t -> KerName.t -> delta_resolver -> delta_resolver val add_inline_delta_resolver : KerName.t -> (int * constr Univ.univ_abstracted option) -> delta_resolver -> delta_resolver val add_delta_resolver : delta_resolver -> delta_resolver -> delta_resolver (** Effect of a [delta_resolver] on a module path, on a kernel name *) val mp_of_delta : delta_resolver -> ModPath.t -> ModPath.t val kn_of_delta : delta_resolver -> KerName.t -> KerName.t (** Build a constant whose canonical part is obtained via a resolver *) val constant_of_delta_kn : delta_resolver -> KerName.t -> Constant.t (** Same, but a 2nd resolver is tried if the 1st one had no effect *) val constant_of_deltas_kn : delta_resolver -> delta_resolver -> KerName.t -> Constant.t (** Same for inductive names *) val mind_of_delta_kn : delta_resolver -> KerName.t -> MutInd.t val mind_of_deltas_kn : delta_resolver -> delta_resolver -> KerName.t -> MutInd.t (** Extract the set of inlined constant in the resolver *) val inline_of_delta : int option -> delta_resolver -> (int * KerName.t) list (** Does a [delta_resolver] contains a [mp], a constant, an inductive ? *) val mp_in_delta : ModPath.t -> delta_resolver -> bool val con_in_delta : Constant.t -> delta_resolver -> bool val mind_in_delta : MutInd.t -> delta_resolver -> bool (** {6 Substitution} *) type substitution val empty_subst : substitution val is_empty_subst : substitution -> bool (** add_* add [arg2/arg1]\{arg3\} to the substitution with no sequential composition. Most often this is not what you want. For sequential composition, try [join (map_mbid mp delta) subs] **) val add_mbid : MBId.t -> ModPath.t -> delta_resolver -> substitution -> substitution val add_mp : ModPath.t -> ModPath.t -> delta_resolver -> substitution -> substitution (** map_* create a new substitution [arg2/arg1]\{arg3\} *) val map_mbid : MBId.t -> ModPath.t -> delta_resolver -> substitution val map_mp : ModPath.t -> ModPath.t -> delta_resolver -> substitution (** sequential composition: [substitute (join sub1 sub2) t = substitute sub2 (substitute sub1 t)] *) val join : substitution -> substitution -> substitution (** Apply the substitution on the domain of the resolver *) val subst_dom_delta_resolver : substitution -> delta_resolver -> delta_resolver (** Apply the substitution on the codomain of the resolver *) val subst_codom_delta_resolver : substitution -> delta_resolver -> delta_resolver val subst_dom_codom_delta_resolver : substitution -> delta_resolver -> delta_resolver type 'a substituted val from_val : 'a -> 'a substituted val force : (substitution -> 'a -> 'a) -> 'a substituted -> 'a val subst_substituted : substitution -> 'a substituted -> 'a substituted (**/**) (* debugging *) val debug_string_of_subst : substitution -> string val debug_pr_subst : substitution -> Pp.t val debug_string_of_delta : delta_resolver -> string val debug_pr_delta : delta_resolver -> Pp.t (**/**) (** [subst_mp sub mp] guarantees that whenever the result of the substitution is structutally equal [mp], it is equal by pointers as well [==] *) val subst_mp : substitution -> ModPath.t -> ModPath.t val subst_mind : substitution -> MutInd.t -> MutInd.t val subst_ind : substitution -> inductive -> inductive val subst_pind : substitution -> pinductive -> pinductive val subst_kn : substitution -> KerName.t -> KerName.t val subst_con : substitution -> Constant.t -> Constant.t * constr Univ.univ_abstracted option val subst_pcon : substitution -> pconstant -> pconstant val subst_constant : substitution -> Constant.t -> Constant.t val subst_proj_repr : substitution -> Projection.Repr.t -> Projection.Repr.t val subst_proj : substitution -> Projection.t -> Projection.t val subst_retro_action : substitution -> Retroknowledge.action -> Retroknowledge.action (** Here the semantics is completely unclear. What does "Hint Unfold t" means when "t" is a parameter? Does the user mean "Unfold X.t" or does she mean "Unfold y" where X.t is later on instantiated with y? I choose the first interpretation (i.e. an evaluable reference is never expanded). *) val subst_evaluable_reference : substitution -> evaluable_global_reference -> evaluable_global_reference (** [replace_mp_in_con mp mp' con] replaces [mp] with [mp'] in [con] *) val replace_mp_in_kn : ModPath.t -> ModPath.t -> KerName.t -> KerName.t (** [subst_mps sub c] performs the substitution [sub] on all kernel names appearing in [c] *) val subst_mps : substitution -> constr -> constr (** [repr_substituted r] dumps the representation of a substituted: - [None, a] when r is a value - [Some s, a] when r is a delayed substitution [s] applied to [a] *) val repr_substituted : 'a substituted -> substitution list option * 'a val force_constr : constr substituted -> constr val subst_constr : substitution -> constr substituted -> constr substituted coq-8.11.0/kernel/type_errors.mli0000644000175000017500000001272013612664533016630 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int -> 'a val error_unbound_var : env -> variable -> 'a val error_not_type : env -> unsafe_judgment -> 'a val error_assumption : env -> unsafe_judgment -> 'a val error_reference_variables : env -> Id.t -> constr -> 'a val error_elim_arity : env -> pinductive -> constr -> unsafe_judgment -> (Sorts.family * Sorts.family * Sorts.family * arity_error) option -> 'a val error_case_not_inductive : env -> unsafe_judgment -> 'a val error_number_branches : env -> unsafe_judgment -> int -> 'a val error_ill_formed_branch : env -> constr -> pconstructor -> constr -> constr -> 'a val error_generalization : env -> Name.t * types -> unsafe_judgment -> 'a val error_actual_type : env -> unsafe_judgment -> types -> 'a val error_incorrect_primitive : env -> (CPrimitives.op_or_type,types) punsafe_judgment -> types -> 'a val error_cant_apply_not_functional : env -> unsafe_judgment -> unsafe_judgment array -> 'a val error_cant_apply_bad_type : env -> int * constr * constr -> unsafe_judgment -> unsafe_judgment array -> 'a val error_ill_formed_rec_body : env -> guard_error -> Name.t Context.binder_annot array -> int -> env -> unsafe_judgment array -> 'a val error_ill_typed_rec_body : env -> int -> Name.t Context.binder_annot array -> unsafe_judgment array -> types array -> 'a val error_elim_explain : Sorts.family -> Sorts.family -> arity_error val error_unsatisfied_constraints : env -> Univ.Constraint.t -> 'a val error_undeclared_universe : env -> Univ.Level.t -> 'a val error_disallowed_sprop : env -> 'a val error_bad_relevance : env -> 'a val map_pguard_error : ('c -> 'd) -> 'c pguard_error -> 'd pguard_error val map_ptype_error : ('c -> 'd) -> ('c, 'c) ptype_error -> ('d, 'd) ptype_error coq-8.11.0/kernel/retypeops.mli0000644000175000017500000000260613612664533016307 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constr.constr -> Sorts.relevance val relevance_of_fterm : Environ.env -> Sorts.relevance Range.t -> Esubst.lift -> CClosure.fconstr -> Sorts.relevance (** Helpers *) open Names val relevance_of_rel : Environ.env -> int -> Sorts.relevance val relevance_of_var : Environ.env -> Id.t -> Sorts.relevance val relevance_of_constant : Environ.env -> Constant.t -> Sorts.relevance val relevance_of_constructor : Environ.env -> constructor -> Sorts.relevance val relevance_of_projection : Environ.env -> Projection.t -> Sorts.relevance coq-8.11.0/kernel/declareops.mli0000644000175000017500000000635213612664533016400 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* AUContext.t val abstract_universes : Entries.universes_entry -> Univ.universe_level_subst * universes (** {6 Arities} *) val map_decl_arity : ('a -> 'c) -> ('b -> 'd) -> ('a, 'b) declaration_arity -> ('c, 'd) declaration_arity (** {6 Constants} *) val subst_const_body : substitution -> Opaqueproof.opaque constant_body -> Opaqueproof.opaque constant_body (** Is there a actual body in const_body ? *) val constant_has_body : 'a constant_body -> bool val constant_polymorphic_context : 'a constant_body -> AUContext.t (** Is the constant polymorphic? *) val constant_is_polymorphic : 'a constant_body -> bool (** Return the universe context, in case the definition is polymorphic, otherwise the context is empty. *) val is_opaque : 'a constant_body -> bool (** {6 Inductive types} *) val eq_recarg : recarg -> recarg -> bool val subst_recarg : substitution -> recarg -> recarg val mk_norec : wf_paths val mk_paths : recarg -> wf_paths list array -> wf_paths val dest_recarg : wf_paths -> recarg val dest_subterms : wf_paths -> wf_paths list array val recarg_length : wf_paths -> int -> int val subst_wf_paths : substitution -> wf_paths -> wf_paths val subst_mind_body : substitution -> mutual_inductive_body -> mutual_inductive_body val inductive_polymorphic_context : mutual_inductive_body -> AUContext.t (** Is the inductive polymorphic? *) val inductive_is_polymorphic : mutual_inductive_body -> bool (** Is the inductive cumulative? *) val inductive_is_cumulative : mutual_inductive_body -> bool val inductive_make_projection : Names.inductive -> mutual_inductive_body -> proj_arg:int -> Names.Projection.Repr.t option val inductive_make_projections : Names.inductive -> mutual_inductive_body -> Names.Projection.Repr.t array option val relevance_of_projection_repr : mutual_inductive_body -> Names.Projection.Repr.t -> Sorts.relevance (** {6 Kernel flags} *) (** A default, safe set of flags for kernel type-checking *) val safe_flags : Conv_oracle.oracle -> typing_flags (** {6 Hash-consing} *) (** Here, strictly speaking, we don't perform true hash-consing of the structure, but simply hash-cons all inner constr and other known elements *) val hcons_const_body : 'a constant_body -> 'a constant_body val hcons_mind : mutual_inductive_body -> mutual_inductive_body val hcons_module_body : module_body -> module_body val hcons_module_type : module_type_body -> module_type_body coq-8.11.0/kernel/retroknowledge.mli0000644000175000017500000000337113612664533017310 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* action | Register_type : CPrimitives.prim_type * Constant.t -> action coq-8.11.0/kernel/mod_subst.ml0000644000175000017500000004425413612664533016110 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool val add_mbi : MBId.t -> 'a -> 'a t -> 'a t val add_mp : ModPath.t -> 'a -> 'a t -> 'a t val find : ModPath.t -> 'a t -> 'a val join : 'a t -> 'a t -> 'a t val fold : (ModPath.t -> 'a -> 'b -> 'b) -> 'a t -> 'b -> 'b end = struct type 'a t = 'a MPmap.t let empty = MPmap.empty let is_empty m = MPmap.is_empty m let add_mbi mbi x m = MPmap.add (MPbound mbi) x m let add_mp mp x m = MPmap.add mp x m let find = MPmap.find let fold = MPmap.fold let join map1 map2 = fold add_mp map1 map2 end type substitution = (ModPath.t * delta_resolver) Umap.t let empty_subst = Umap.empty let is_empty_subst = Umap.is_empty (* *) let string_of_hint = function | Inline (_,Some _) -> "inline(Some _)" | Inline _ -> "inline()" | Equiv kn -> KerName.to_string kn let debug_string_of_delta resolve = let kn_to_string kn hint l = (KerName.to_string kn ^ "=>" ^ string_of_hint hint) :: l in let mp_to_string mp mp' l = (ModPath.to_string mp ^ "=>" ^ ModPath.to_string mp') :: l in let l = Deltamap.fold mp_to_string kn_to_string resolve [] in String.concat ", " (List.rev l) let list_contents sub = let one_pair (mp,reso) = (ModPath.to_string mp,debug_string_of_delta reso) in let mp_one_pair mp0 p l = (ModPath.to_string mp0, one_pair p)::l in Umap.fold mp_one_pair sub [] let debug_string_of_subst sub = let l = List.map (fun (s1,(s2,s3)) -> s1^"|->"^s2^"["^s3^"]") (list_contents sub) in "{" ^ String.concat "; " l ^ "}" let debug_pr_delta resolve = str (debug_string_of_delta resolve) let debug_pr_subst sub = let l = list_contents sub in let f (s1,(s2,s3)) = hov 2 (str s1 ++ spc () ++ str "|-> " ++ str s2 ++ spc () ++ str "[" ++ str s3 ++ str "]") in str "{" ++ hov 2 (prlist_with_sep pr_comma f l) ++ str "}" (* *) (** Extending a [delta_resolver] *) let add_inline_delta_resolver kn (lev,oc) = Deltamap.add_kn kn (Inline (lev,oc)) let add_kn_delta_resolver kn kn' = assert (Label.equal (KerName.label kn) (KerName.label kn')); Deltamap.add_kn kn (Equiv kn') let add_mp_delta_resolver mp1 mp2 = Deltamap.add_mp mp1 mp2 (** Extending a [substitution] without sequential composition *) let add_mbid mbid mp resolve s = Umap.add_mbi mbid (mp,resolve) s let add_mp mp1 mp2 resolve s = Umap.add_mp mp1 (mp2,resolve) s let map_mbid mbid mp resolve = add_mbid mbid mp resolve empty_subst let map_mp mp1 mp2 resolve = add_mp mp1 mp2 resolve empty_subst let mp_in_delta mp = Deltamap.mem_mp mp let kn_in_delta kn resolver = try match Deltamap.find_kn kn resolver with | Equiv _ -> true | Inline _ -> false with Not_found -> false let con_in_delta con resolver = kn_in_delta (Constant.user con) resolver let mind_in_delta mind resolver = kn_in_delta (MutInd.user mind) resolver let mp_of_delta resolve mp = try Deltamap.find_mp mp resolve with Not_found -> mp let find_prefix resolve mp = let rec sub_mp = function | MPdot(mp,l) as mp_sup -> (try Deltamap.find_mp mp_sup resolve with Not_found -> MPdot(sub_mp mp,l)) | p -> Deltamap.find_mp p resolve in try sub_mp mp with Not_found -> mp (** Applying a resolver to a kernel name *) exception Change_equiv_to_inline of (int * constr Univ.univ_abstracted) let solve_delta_kn resolve kn = try match Deltamap.find_kn kn resolve with | Equiv kn1 -> kn1 | Inline (lev, Some c) -> raise (Change_equiv_to_inline (lev,c)) | Inline (_, None) -> raise Not_found with Not_found -> let mp,l = KerName.repr kn in let new_mp = find_prefix resolve mp in if mp == new_mp then kn else KerName.make new_mp l let kn_of_delta resolve kn = try solve_delta_kn resolve kn with Change_equiv_to_inline _ -> kn (** Try a 1st resolver, and then a 2nd in case it had no effect *) let kn_of_deltas resolve1 resolve2 kn = let kn' = kn_of_delta resolve1 kn in if kn' == kn then kn_of_delta resolve2 kn else kn' let constant_of_delta_kn resolve kn = Constant.make kn (kn_of_delta resolve kn) let constant_of_deltas_kn resolve1 resolve2 kn = Constant.make kn (kn_of_deltas resolve1 resolve2 kn) let mind_of_delta_kn resolve kn = MutInd.make kn (kn_of_delta resolve kn) let mind_of_deltas_kn resolve1 resolve2 kn = MutInd.make kn (kn_of_deltas resolve1 resolve2 kn) let inline_of_delta inline resolver = match inline with | None -> [] | Some inl_lev -> let extract kn hint l = match hint with | Inline (lev,_) -> if lev <= inl_lev then (lev,kn)::l else l | _ -> l in Deltamap.fold_kn extract resolver [] let search_delta_inline resolve kn1 kn2 = let find kn = match Deltamap.find_kn kn resolve with | Inline (_,o) -> o | Equiv _ -> raise Not_found in try find kn1 with Not_found -> if kn1 == kn2 then None else try find kn2 with Not_found -> None let subst_mp0 sub mp = (* 's like subst *) let rec aux mp = match mp with | MPfile _ | MPbound _ -> Umap.find mp sub | MPdot (mp1,l) as mp2 -> begin try Umap.find mp2 sub with Not_found -> let mp1',resolve = aux mp1 in MPdot (mp1',l),resolve end in try Some (aux mp) with Not_found -> None let subst_mp sub mp = match subst_mp0 sub mp with None -> mp | Some (mp',_) -> mp' let subst_kn_delta sub kn = let mp,l = KerName.repr kn in match subst_mp0 sub mp with Some (mp',resolve) -> solve_delta_kn resolve (KerName.make mp' l) | None -> kn let subst_kn sub kn = let mp,l = KerName.repr kn in match subst_mp0 sub mp with Some (mp',_) -> (KerName.make mp' l) | None -> kn exception No_subst let subst_dual_mp sub mp1 mp2 = let o1 = subst_mp0 sub mp1 in let o2 = if mp1 == mp2 then o1 else subst_mp0 sub mp2 in match o1, o2 with | None, None -> raise No_subst | Some (mp1',resolve), None -> mp1', mp2, resolve, true | None, Some (mp2',resolve) -> mp1, mp2', resolve, false | Some (mp1',_), Some (mp2',resolve) -> mp1', mp2', resolve, false let progress f x ~orelse = let y = f x in if y != x then y else orelse let subst_mind sub mind = let mpu,l = MutInd.repr2 mind in let mpc = KerName.modpath (MutInd.canonical mind) in try let mpu,mpc,resolve,user = subst_dual_mp sub mpu mpc in let knu = KerName.make mpu l in let knc = if mpu == mpc then knu else KerName.make mpc l in let knc' = progress (kn_of_delta resolve) (if user then knu else knc) ~orelse:knc in MutInd.make knu knc' with No_subst -> mind let subst_ind sub (ind,i as indi) = let ind' = subst_mind sub ind in if ind' == ind then indi else ind',i let subst_pind sub (ind,u) = (subst_ind sub ind, u) let subst_con0 sub cst = let mpu,l = Constant.repr2 cst in let mpc = KerName.modpath (Constant.canonical cst) in let mpu,mpc,resolve,user = subst_dual_mp sub mpu mpc in let knu = KerName.make mpu l in let knc = if mpu == mpc then knu else KerName.make mpc l in match search_delta_inline resolve knu knc with | Some t -> (* In case of inlining, discard the canonical part (cf #2608) *) Constant.make1 knu, Some t | None -> let knc' = progress (kn_of_delta resolve) (if user then knu else knc) ~orelse:knc in let cst' = Constant.make knu knc' in cst', None let subst_con sub cst = try subst_con0 sub cst with No_subst -> cst, None let subst_pcon sub (con,u as pcon) = try let con', _can = subst_con0 sub con in con',u with No_subst -> pcon let subst_constant sub con = try fst (subst_con0 sub con) with No_subst -> con let subst_proj_repr sub p = Projection.Repr.map (subst_mind sub) p let subst_proj sub p = Projection.map (subst_mind sub) p let subst_retro_action subst action = let open Retroknowledge in match action with | Register_ind(prim,ind) -> let ind' = subst_ind subst ind in if ind == ind' then action else Register_ind(prim, ind') | Register_type(prim,c) -> let c' = subst_constant subst c in if c == c' then action else Register_type(prim, c') (* Here the semantics is completely unclear. What does "Hint Unfold t" means when "t" is a parameter? Does the user mean "Unfold X.t" or does she mean "Unfold y" where X.t is later on instantiated with y? I choose the first interpretation (i.e. an evaluable reference is never expanded). *) let subst_evaluable_reference subst = function | EvalVarRef id -> EvalVarRef id | EvalConstRef kn -> EvalConstRef (subst_constant subst kn) let rec map_kn f f' c = let func = map_kn f f' in match kind c with | Const kn -> (try f' kn with No_subst -> c) | Proj (p,t) -> let p' = Projection.map f p in let t' = func t in if p' == p && t' == t then c else mkProj (p', t') | Ind ((kn,i),u) -> let kn' = f kn in if kn'==kn then c else mkIndU ((kn',i),u) | Construct (((kn,i),j),u) -> let kn' = f kn in if kn'==kn then c else mkConstructU (((kn',i),j),u) | Case (ci,p,ct,l) -> let ci_ind = let (kn,i) = ci.ci_ind in let kn' = f kn in if kn'==kn then ci.ci_ind else kn',i in let p' = func p in let ct' = func ct in let l' = Array.Smart.map func l in if (ci.ci_ind==ci_ind && p'==p && l'==l && ct'==ct)then c else mkCase ({ci with ci_ind = ci_ind}, p',ct', l') | Cast (ct,k,t) -> let ct' = func ct in let t'= func t in if (t'==t && ct'==ct) then c else mkCast (ct', k, t') | Prod (na,t,ct) -> let ct' = func ct in let t'= func t in if (t'==t && ct'==ct) then c else mkProd (na, t', ct') | Lambda (na,t,ct) -> let ct' = func ct in let t'= func t in if (t'==t && ct'==ct) then c else mkLambda (na, t', ct') | LetIn (na,b,t,ct) -> let ct' = func ct in let t'= func t in let b'= func b in if (t'==t && ct'==ct && b==b') then c else mkLetIn (na, b', t', ct') | App (ct,l) -> let ct' = func ct in let l' = Array.Smart.map func l in if (ct'== ct && l'==l) then c else mkApp (ct',l') | Evar (e,l) -> let l' = Array.Smart.map func l in if (l'==l) then c else mkEvar (e,l') | Fix (ln,(lna,tl,bl)) -> let tl' = Array.Smart.map func tl in let bl' = Array.Smart.map func bl in if (bl == bl'&& tl == tl') then c else mkFix (ln,(lna,tl',bl')) | CoFix(ln,(lna,tl,bl)) -> let tl' = Array.Smart.map func tl in let bl' = Array.Smart.map func bl in if (bl == bl'&& tl == tl') then c else mkCoFix (ln,(lna,tl',bl')) | _ -> c let subst_mps sub c = let subst_pcon_term sub (con,u) = let con', can = subst_con0 sub con in match can with | None -> mkConstU (con',u) | Some t -> Vars.univ_instantiate_constr u t in if is_empty_subst sub then c else map_kn (subst_mind sub) (subst_pcon_term sub) c let rec replace_mp_in_mp mpfrom mpto mp = match mp with | _ when ModPath.equal mp mpfrom -> mpto | MPdot (mp1,l) -> let mp1' = replace_mp_in_mp mpfrom mpto mp1 in if mp1 == mp1' then mp else MPdot (mp1',l) | _ -> mp let replace_mp_in_kn mpfrom mpto kn = let mp,l = KerName.repr kn in let mp'' = replace_mp_in_mp mpfrom mpto mp in if mp==mp'' then kn else KerName.make mp'' l let rec mp_in_mp mp mp1 = match mp1 with | _ when ModPath.equal mp1 mp -> true | MPdot (mp2,_l) -> mp_in_mp mp mp2 | _ -> false let subset_prefixed_by mp resolver = let mp_prefix mkey mequ rslv = if mp_in_mp mp mkey then Deltamap.add_mp mkey mequ rslv else rslv in let kn_prefix kn hint rslv = match hint with | Inline _ -> rslv | Equiv _ -> if mp_in_mp mp (KerName.modpath kn) then Deltamap.add_kn kn hint rslv else rslv in Deltamap.fold mp_prefix kn_prefix resolver empty_delta_resolver let subst_dom_delta_resolver subst resolver = let mp_apply_subst mkey mequ rslv = Deltamap.add_mp (subst_mp subst mkey) mequ rslv in let kn_apply_subst kkey hint rslv = Deltamap.add_kn (subst_kn subst kkey) hint rslv in Deltamap.fold mp_apply_subst kn_apply_subst resolver empty_delta_resolver let subst_mp_delta sub mp mkey = match subst_mp0 sub mp with None -> empty_delta_resolver,mp | Some (mp',resolve) -> let mp1 = find_prefix resolve mp' in let resolve1 = subset_prefixed_by mp1 resolve in (subst_dom_delta_resolver (map_mp mp1 mkey empty_delta_resolver) resolve1),mp1 let gen_subst_delta_resolver dom subst resolver = let mp_apply_subst mkey mequ rslv = let mkey' = if dom then subst_mp subst mkey else mkey in let rslv',mequ' = subst_mp_delta subst mequ mkey in Deltamap.join rslv' (Deltamap.add_mp mkey' mequ' rslv) in let kn_apply_subst kkey hint rslv = let kkey' = if dom then subst_kn subst kkey else kkey in let hint' = match hint with | Equiv kequ -> (try Equiv (subst_kn_delta subst kequ) with Change_equiv_to_inline (lev,c) -> Inline (lev,Some c)) | Inline (lev,Some t) -> Inline (lev,Some (Univ.map_univ_abstracted (subst_mps subst) t)) | Inline (_,None) -> hint in Deltamap.add_kn kkey' hint' rslv in Deltamap.fold mp_apply_subst kn_apply_subst resolver empty_delta_resolver let subst_codom_delta_resolver = gen_subst_delta_resolver false let subst_dom_codom_delta_resolver = gen_subst_delta_resolver true let update_delta_resolver resolver1 resolver2 = let mp_apply_rslv mkey mequ rslv = Deltamap.add_mp mkey (find_prefix resolver2 mequ) rslv in let kn_apply_rslv kkey hint1 rslv = let hint = match hint1 with | Equiv kequ -> (try Equiv (solve_delta_kn resolver2 kequ) with Change_equiv_to_inline (lev,c) -> Inline (lev, Some c)) | Inline (_,Some _) -> hint1 | Inline (_,None) -> (try Deltamap.find_kn kkey resolver2 with Not_found -> hint1) in Deltamap.add_kn kkey hint rslv in Deltamap.fold mp_apply_rslv kn_apply_rslv resolver1 resolver2 let add_delta_resolver resolver1 resolver2 = if Deltamap.is_empty resolver2 then resolver1 else update_delta_resolver resolver1 resolver2 let substition_prefixed_by k mp subst = let mp_prefixmp kmp (mp_to,reso) sub = if mp_in_mp mp kmp && not (ModPath.equal mp kmp) then let new_key = replace_mp_in_mp mp k kmp in Umap.add_mp new_key (mp_to,reso) sub else sub in Umap.fold mp_prefixmp subst empty_subst let join subst1 subst2 = let apply_subst mpk add (mp,resolve) res = let mp',resolve' = match subst_mp0 subst2 mp with | None -> mp, None | Some (mp',resolve') -> mp', Some resolve' in let resolve'' = match resolve' with | Some res -> add_delta_resolver (subst_dom_codom_delta_resolver subst2 resolve) res | None -> subst_codom_delta_resolver subst2 resolve in let prefixed_subst = substition_prefixed_by mpk mp' subst2 in Umap.join prefixed_subst (add (mp',resolve'') res) in let mp_apply_subst mp = apply_subst mp (Umap.add_mp mp) in let subst = Umap.fold mp_apply_subst subst1 empty_subst in Umap.join subst2 subst type 'a substituted = { mutable subst_value : 'a; mutable subst_subst : substitution list; } let from_val x = { subst_value = x; subst_subst = []; } let force fsubst r = match r.subst_subst with | [] -> r.subst_value | s -> let subst = List.fold_left join empty_subst (List.rev s) in let x = fsubst subst r.subst_value in let () = r.subst_subst <- [] in let () = r.subst_value <- x in x let subst_substituted s r = { r with subst_subst = s :: r.subst_subst; } let force_constr = force subst_mps let subst_constr = subst_substituted (* debug *) let repr_substituted r = match r.subst_subst with | [] -> None, r.subst_value | s -> Some s, r.subst_value coq-8.11.0/kernel/modops.ml0000644000175000017500000005360213612664533015407 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* false |MoreFunctor _ -> true let destr_functor = function |NoFunctor _ -> error_not_a_functor () |MoreFunctor (mbid,ty,x) -> (mbid,ty,x) let destr_nofunctor = function |NoFunctor a -> a |MoreFunctor _ -> error_is_a_functor () let rec functor_smart_map fty f0 funct = match funct with |MoreFunctor (mbid,ty,e) -> let ty' = fty ty in let e' = functor_smart_map fty f0 e in if ty==ty' && e==e' then funct else MoreFunctor (mbid,ty',e') |NoFunctor a -> let a' = f0 a in if a==a' then funct else NoFunctor a' let rec functor_iter fty f0 = function |MoreFunctor (_mbid,ty,e) -> fty ty; functor_iter fty f0 e |NoFunctor a -> f0 a (** {6 Misc operations } *) let module_type_of_module mb = { mb with mod_expr = (); mod_type_alg = None; mod_retroknowledge = ModTypeRK; } let module_body_of_type mp mtb = { mtb with mod_expr = Abstract; mod_mp = mp; mod_retroknowledge = ModBodyRK []; } let check_modpath_equiv env mp1 mp2 = if ModPath.equal mp1 mp2 then () else let mp1' = mp_of_delta (lookup_module mp1 env).mod_delta mp1 in let mp2' = mp_of_delta (lookup_module mp2 env).mod_delta mp2 in if ModPath.equal mp1' mp2' then () else error_not_equal_modpaths mp1 mp2 let implem_smartmap fs fa impl = match impl with |Struct e -> let e' = fs e in if e==e' then impl else Struct e' |Algebraic a -> let a' = fa a in if a==a' then impl else Algebraic a' |Abstract|FullStruct -> impl let implem_iter fs fa impl = match impl with |Struct e -> fs e |Algebraic a -> fa a |Abstract|FullStruct -> () (** {6 Substitutions of modular structures } *) let id_delta x _y = x let subst_with_body sub = function |WithMod(id,mp) as orig -> let mp' = subst_mp sub mp in if mp==mp' then orig else WithMod(id,mp') |WithDef(id,(c,ctx)) as orig -> let c' = subst_mps sub c in if c==c' then orig else WithDef(id,(c',ctx)) let rec subst_structure sub do_delta sign = let subst_body ((l,body) as orig) = match body with |SFBconst cb -> let cb' = subst_const_body sub cb in if cb==cb' then orig else (l,SFBconst cb') |SFBmind mib -> let mib' = subst_mind_body sub mib in if mib==mib' then orig else (l,SFBmind mib') |SFBmodule mb -> let mb' = subst_module sub do_delta mb in if mb==mb' then orig else (l,SFBmodule mb') |SFBmodtype mtb -> let mtb' = subst_modtype sub do_delta mtb in if mtb==mtb' then orig else (l,SFBmodtype mtb') in List.Smart.map subst_body sign and subst_retro : type a. Mod_subst.substitution -> a module_retroknowledge -> a module_retroknowledge = fun subst retro -> match retro with | ModTypeRK as r -> r | ModBodyRK l as r -> let l' = List.Smart.map (subst_retro_action subst) l in if l == l' then r else ModBodyRK l and subst_body : 'a. _ -> _ -> (_ -> 'a -> 'a) -> _ -> 'a generic_module_body -> 'a generic_module_body = fun is_mod sub subst_impl do_delta mb -> let { mod_mp=mp; mod_expr=me; mod_type=ty; mod_type_alg=aty; mod_retroknowledge=retro; _ } = mb in let mp' = subst_mp sub mp in let sub = if ModPath.equal mp mp' then sub else if is_mod && not (is_functor ty) then sub else add_mp mp mp' empty_delta_resolver sub in let ty' = subst_signature sub do_delta ty in let me' = subst_impl sub me in let aty' = Option.Smart.map (subst_expression sub id_delta) aty in let retro' = subst_retro sub retro in let delta' = do_delta mb.mod_delta sub in if mp==mp' && me==me' && ty==ty' && aty==aty' && retro==retro' && delta'==mb.mod_delta then mb else { mb with mod_mp = mp'; mod_expr = me'; mod_type = ty'; mod_type_alg = aty'; mod_retroknowledge = retro'; mod_delta = delta'; } and subst_module sub do_delta mb = subst_body true sub subst_impl do_delta mb and subst_impl sub me = implem_smartmap (subst_signature sub id_delta) (subst_expression sub id_delta) me and subst_modtype sub do_delta mtb = subst_body false sub (fun _ () -> ()) do_delta mtb and subst_expr sub do_delta seb = match seb with |MEident mp -> let mp' = subst_mp sub mp in if mp==mp' then seb else MEident mp' |MEapply (meb1,mp2) -> let meb1' = subst_expr sub do_delta meb1 in let mp2' = subst_mp sub mp2 in if meb1==meb1' && mp2==mp2' then seb else MEapply(meb1',mp2') |MEwith (meb,wdb) -> let meb' = subst_expr sub do_delta meb in let wdb' = subst_with_body sub wdb in if meb==meb' && wdb==wdb' then seb else MEwith(meb',wdb') and subst_expression sub do_delta = functor_smart_map (subst_modtype sub do_delta) (subst_expr sub do_delta) and subst_signature sub do_delta = functor_smart_map (subst_modtype sub do_delta) (subst_structure sub do_delta) let do_delta_dom reso sub = subst_dom_delta_resolver sub reso let do_delta_codom reso sub = subst_codom_delta_resolver sub reso let do_delta_dom_codom reso sub = subst_dom_codom_delta_resolver sub reso let subst_signature subst = subst_signature subst do_delta_codom let subst_structure subst = subst_structure subst do_delta_codom (** {6 Adding a module in the environment } *) let add_retroknowledge r env = match r with | ModBodyRK l -> List.fold_left Primred.add_retroknowledge env l let rec add_structure mp sign resolver linkinfo env = let add_one env (l,elem) = match elem with |SFBconst cb -> let c = constant_of_delta_kn resolver (KerName.make mp l) in Environ.add_constant_key c cb linkinfo env |SFBmind mib -> let mind = mind_of_delta_kn resolver (KerName.make mp l) in let mib = if mib.mind_private != None then { mib with mind_private = Some true } else mib in Environ.add_mind_key mind (mib,ref linkinfo) env |SFBmodule mb -> add_module mb linkinfo env (* adds components as well *) |SFBmodtype mtb -> Environ.add_modtype mtb env in List.fold_left add_one env sign and add_module mb linkinfo env = let mp = mb.mod_mp in let env = Environ.shallow_add_module mb env in match mb.mod_type with |NoFunctor struc -> add_retroknowledge mb.mod_retroknowledge (add_structure mp struc mb.mod_delta linkinfo env) |MoreFunctor _ -> env let add_linked_module mb linkinfo env = add_module mb linkinfo env let add_structure mp sign resolver env = add_structure mp sign resolver no_link_info env let add_module mb env = add_module mb no_link_info env let add_module_type mp mtb env = add_module (module_body_of_type mp mtb) env (** {6 Strengtening } *) let strengthen_const mp_from l cb resolver = match cb.const_body with |Def _ -> cb |_ -> let kn = KerName.make mp_from l in let con = constant_of_delta_kn resolver kn in let u = Univ.make_abstract_instance (Declareops.constant_polymorphic_context cb) in { cb with const_body = Def (Mod_subst.from_val (mkConstU (con,u))); const_body_code = Some (Cemitcodes.from_val (Cbytegen.compile_alias con)) } let rec strengthen_mod mp_from mp_to mb = if mp_in_delta mb.mod_mp mb.mod_delta then mb else match mb.mod_type with |NoFunctor struc -> let reso,struc' = strengthen_sig mp_from struc mp_to mb.mod_delta in { mb with mod_expr = Algebraic (NoFunctor (MEident mp_to)); mod_type = NoFunctor struc'; mod_delta = add_mp_delta_resolver mp_from mp_to (add_delta_resolver mb.mod_delta reso) } |MoreFunctor _ -> mb and strengthen_sig mp_from struc mp_to reso = match struc with |[] -> empty_delta_resolver,[] |(l,SFBconst cb) :: rest -> let item' = l,SFBconst (strengthen_const mp_from l cb reso) in let reso',rest' = strengthen_sig mp_from rest mp_to reso in reso',item'::rest' |(_,SFBmind _ as item):: rest -> let reso',rest' = strengthen_sig mp_from rest mp_to reso in reso',item::rest' |(l,SFBmodule mb) :: rest -> let mp_from' = MPdot (mp_from,l) in let mp_to' = MPdot(mp_to,l) in let mb' = strengthen_mod mp_from' mp_to' mb in let item' = l,SFBmodule mb' in let reso',rest' = strengthen_sig mp_from rest mp_to reso in add_delta_resolver reso' mb.mod_delta, item':: rest' |(_l,SFBmodtype _mty as item) :: rest -> let reso',rest' = strengthen_sig mp_from rest mp_to reso in reso',item::rest' let strengthen mtb mp = (* Has mtb already been strengthened ? *) if mp_in_delta mtb.mod_mp mtb.mod_delta then mtb else match mtb.mod_type with |NoFunctor struc -> let reso',struc' = strengthen_sig mtb.mod_mp struc mp mtb.mod_delta in { mtb with mod_type = NoFunctor struc'; mod_delta = add_delta_resolver mtb.mod_delta (add_mp_delta_resolver mtb.mod_mp mp reso') } |MoreFunctor _ -> mtb let inline_delta_resolver env inl mp mbid mtb delta = let constants = inline_of_delta inl mtb.mod_delta in let rec make_inline delta = function | [] -> delta | (lev,kn)::r -> let kn = replace_mp_in_kn (MPbound mbid) mp kn in let con = constant_of_delta_kn delta kn in try let constant = lookup_constant con env in let l = make_inline delta r in match constant.const_body with | Undef _ | OpaqueDef _ | Primitive _ -> l | Def body -> let constr = Mod_subst.force_constr body in let ctx = Declareops.constant_polymorphic_context constant in let constr = Univ.{univ_abstracted_value=constr; univ_abstracted_binder=ctx} in add_inline_delta_resolver kn (lev, Some constr) l with Not_found -> error_no_such_label_sub (Constant.label con) (ModPath.to_string (Constant.modpath con)) in make_inline delta constants let rec strengthen_and_subst_mod mb subst mp_from mp_to = match mb.mod_type with |NoFunctor struc -> let mb_is_an_alias = mp_in_delta mb.mod_mp mb.mod_delta in if mb_is_an_alias then subst_module subst do_delta_dom mb else let reso',struc' = strengthen_and_subst_struct struc subst mp_from mp_to false false mb.mod_delta in { mb with mod_mp = mp_to; mod_expr = Algebraic (NoFunctor (MEident mp_from)); mod_type = NoFunctor struc'; mod_delta = add_mp_delta_resolver mp_to mp_from reso' } |MoreFunctor _ -> let subst = add_mp mb.mod_mp mp_to empty_delta_resolver subst in subst_module subst do_delta_dom mb and strengthen_and_subst_struct str subst mp_from mp_to alias incl reso = match str with | [] -> empty_delta_resolver,[] | (l,SFBconst cb) as item :: rest -> let cb' = subst_const_body subst cb in let cb' = if alias then cb' else strengthen_const mp_from l cb' reso in let item' = if cb' == cb then item else (l, SFBconst cb') in let reso',rest' = strengthen_and_subst_struct rest subst mp_from mp_to alias incl reso in let str' = if rest' == rest && item' == item then str else item' :: rest' in if incl then (* If we are performing an inclusion we need to add the fact that the constant mp_to.l is \Delta-equivalent to reso(mp_from.l) *) let kn_from = KerName.make mp_from l in let kn_to = KerName.make mp_to l in let old_name = kn_of_delta reso kn_from in add_kn_delta_resolver kn_to old_name reso', str' else (* In this case the fact that the constant mp_to.l is \Delta-equivalent to resolver(mp_from.l) is already known because reso' contains mp_to maps to reso(mp_from) *) reso', str' | (l,SFBmind mib) as item :: rest -> let mib' = subst_mind_body subst mib in let item' = if mib' == mib then item else (l, SFBmind mib') in let reso',rest' = strengthen_and_subst_struct rest subst mp_from mp_to alias incl reso in let str' = if rest' == rest && item' == item then str else item' :: rest' in (* Same as constant *) if incl then let kn_from = KerName.make mp_from l in let kn_to = KerName.make mp_to l in let old_name = kn_of_delta reso kn_from in add_kn_delta_resolver kn_to old_name reso', str' else reso', str' | (l,SFBmodule mb) as item :: rest -> let mp_from' = MPdot (mp_from,l) in let mp_to' = MPdot (mp_to,l) in let mb' = if alias then subst_module subst do_delta_dom mb else strengthen_and_subst_mod mb subst mp_from' mp_to' in let item' = if mb' == mb then item else (l, SFBmodule mb') in let reso',rest' = strengthen_and_subst_struct rest subst mp_from mp_to alias incl reso in let str' = if rest' == rest && item' == item then str else item' :: rest' in (* if mb is a functor we should not derive new equivalences on names, hence we add the fact that the functor can only be equivalent to itself. If we adopt an applicative semantic for functor this should be changed.*) if is_functor mb'.mod_type then add_mp_delta_resolver mp_to' mp_to' reso', str' else add_delta_resolver reso' mb'.mod_delta, str' | (l,SFBmodtype mty) as item :: rest -> let mp_from' = MPdot (mp_from,l) in let mp_to' = MPdot(mp_to,l) in let subst' = add_mp mp_from' mp_to' empty_delta_resolver subst in let mty' = subst_modtype subst' (fun resolver _ -> subst_dom_codom_delta_resolver subst' resolver) mty in let item' = if mty' == mty then item else (l, SFBmodtype mty') in let reso',rest' = strengthen_and_subst_struct rest subst mp_from mp_to alias incl reso in let str' = if rest' == rest && item' == item then str else item' :: rest' in add_mp_delta_resolver mp_to' mp_to' reso', str' (** Let P be a module path when we write: "Module M:=P." or "Module M. Include P. End M." We need to perform two operations to compute the body of M. - The first one is applying the substitution {P <- M} on the type of P - The second one is strengthening. *) let strengthen_and_subst_mb mb mp include_b = match mb.mod_type with |NoFunctor struc -> let mb_is_an_alias = mp_in_delta mb.mod_mp mb.mod_delta in (* if mb.mod_mp is an alias then the strengthening is useless (i.e. it is already done)*) let mp_alias = mp_of_delta mb.mod_delta mb.mod_mp in let subst_resolver = map_mp mb.mod_mp mp empty_delta_resolver in let new_resolver = add_mp_delta_resolver mp mp_alias (subst_dom_delta_resolver subst_resolver mb.mod_delta) in let subst = map_mp mb.mod_mp mp new_resolver in let reso',struc' = strengthen_and_subst_struct struc subst mb.mod_mp mp mb_is_an_alias include_b mb.mod_delta in { mb with mod_mp = mp; mod_type = NoFunctor struc'; mod_expr = Algebraic (NoFunctor (MEident mb.mod_mp)); mod_delta = if include_b then reso' else add_delta_resolver new_resolver reso' } |MoreFunctor _ -> let subst = map_mp mb.mod_mp mp empty_delta_resolver in subst_module subst do_delta_dom_codom mb let subst_modtype_and_resolver mtb mp = let subst = map_mp mtb.mod_mp mp empty_delta_resolver in let new_delta = subst_dom_codom_delta_resolver subst mtb.mod_delta in let full_subst = map_mp mtb.mod_mp mp new_delta in subst_modtype full_subst do_delta_dom_codom mtb (** {6 Cleaning a module expression from bounded parts } For instance: functor(X:T)->struct module M:=X end) becomes: functor(X:T)->struct module M:= end) *) let rec is_bounded_expr l = function | MEident (MPbound mbid) -> MBIset.mem mbid l | MEapply (fexpr,mp) -> is_bounded_expr l (MEident mp) || is_bounded_expr l fexpr | _ -> false let rec clean_module_body l mb = let impl, typ = mb.mod_expr, mb.mod_type in let typ' = clean_signature l typ in let impl' = match impl with | Algebraic (NoFunctor m) when is_bounded_expr l m -> FullStruct | _ -> implem_smartmap (clean_signature l) (clean_expression l) impl in if typ==typ' && impl==impl' then mb else { mb with mod_type=typ'; mod_expr=impl' } and clean_module_type l mb = let (), typ = mb.mod_expr, mb.mod_type in let typ' = clean_signature l typ in if typ==typ' then mb else { mb with mod_type=typ' } and clean_field l field = match field with |(lab,SFBmodule mb) -> let mb' = clean_module_body l mb in if mb==mb' then field else (lab,SFBmodule mb') |_ -> field and clean_structure l = List.Smart.map (clean_field l) and clean_signature l = functor_smart_map (clean_module_type l) (clean_structure l) and clean_expression l = functor_smart_map (clean_module_type l) (fun me -> me) let rec collect_mbid l sign = match sign with |MoreFunctor (mbid,ty,m) -> let m' = collect_mbid (MBIset.add mbid l) m in if m==m' then sign else MoreFunctor (mbid,ty,m') |NoFunctor struc -> let struc' = clean_structure l struc in if struc==struc' then sign else NoFunctor struc' let clean_bounded_mod_expr sign = if is_functor sign then collect_mbid MBIset.empty sign else sign (** {6 Stm machinery } *) let join_constant_body except otab cb = match cb.const_body with | OpaqueDef o -> Opaqueproof.join_opaque ~except otab o | _ -> () let join_structure except otab s = let rec join_module : 'a. 'a generic_module_body -> unit = fun mb -> Option.iter join_expression mb.mod_type_alg; join_signature mb.mod_type and join_field (_l,body) = match body with |SFBconst sb -> join_constant_body except otab sb |SFBmind _ -> () |SFBmodule m -> implem_iter join_signature join_expression m.mod_expr; join_module m |SFBmodtype m -> join_module m and join_structure struc = List.iter join_field struc and join_signature sign = functor_iter join_module join_structure sign and join_expression me = functor_iter join_module (fun _ -> ()) me in join_structure s coq-8.11.0/kernel/cemitcodes.mli0000644000175000017500000000277513612664533016403 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* patches -> (reloc_info -> int) -> Vmvalues.tcode type to_patch = emitcodes * patches * fv type body_code = | BCdefined of to_patch | BCalias of Constant.t | BCconstant type to_patch_substituted val from_val : body_code -> to_patch_substituted val force : to_patch_substituted -> body_code val subst_to_patch_subst : Mod_subst.substitution -> to_patch_substituted -> to_patch_substituted val repr_body_code : to_patch_substituted -> Mod_subst.substitution list option * body_code val to_memory : bytecodes * bytecodes * fv -> to_patch (** init code, fun code, fv *) coq-8.11.0/kernel/uint63.mli0000644000175000017500000000453213612664533015405 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t val to_int2 : t -> int * int (* msb, lsb *) val of_int64 : Int64.t -> t (* val of_uint : int -> t *) (** [int_min n m] returns the minimum of [n] and [m], [m] must be in [0, 2^30-1]. *) val to_int_min : t -> int -> int (* conversion to float *) val of_float : float -> t val to_float : t -> float val hash : t -> int (* conversion to a string *) val to_string : t -> string val of_string : string -> t val compile : t -> string (* constants *) val zero : t val one : t (* logical operations *) val l_sl : t -> t -> t val l_sr : t -> t -> t val l_and : t -> t -> t val l_xor : t -> t -> t val l_or : t -> t -> t (* Arithmetic operations *) val add : t -> t -> t val sub : t -> t -> t val mul : t -> t -> t val div : t -> t -> t val rem : t -> t -> t val diveucl : t -> t -> t * t (* Specific arithmetic operations *) val mulc : t -> t -> t * t val addmuldiv : t -> t -> t -> t (** [div21 xh xl y] returns [q % 2^63, r] s.t. [xh * 2^63 + xl = q * y + r] and [r < y]. When [y] is [0], returns [0, 0]. *) val div21 : t -> t -> t -> t * t (* comparison *) val lt : t -> t -> bool val equal : t -> t -> bool val le : t -> t -> bool val compare : t -> t -> int (* head and tail *) val head0 : t -> t val tail0 : t -> t val is_uint63 : Obj.t -> bool (* Arithmetic with explicit carries *) (* Analog of Numbers.Abstract.Cyclic.carry *) type 'a carry = C0 of 'a | C1 of 'a val addc : t -> t -> t carry val addcarryc : t -> t -> t carry val subc : t -> t -> t carry val subcarryc : t -> t -> t carry coq-8.11.0/kernel/constr.mli0000644000175000017500000006206613612664533015573 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* = ci_cstr_nargs.(i) *) type case_info = { ci_ind : inductive; (* inductive type to which belongs the value that is being matched *) ci_npar : int; (* number of parameters of the above inductive type *) ci_cstr_ndecls : int array; (* For each constructor, the corresponding integer determines the number of values that can be bound in a match-construct. NOTE: parameters of the inductive type are therefore excluded from the count *) ci_cstr_nargs : int array; (* for each constructor, the corresponding integers determines the number of values that can be applied to the constructor, in addition to the parameters of the related inductive type NOTE: "lets" are therefore excluded from the count NOTE: parameters of the inductive type are also excluded from the count *) ci_relevance : Sorts.relevance; (* relevance of the predicate (not of the inductive!) *) ci_pp_info : case_printing (* not interpreted by the kernel *) } (** {6 The type of constructions } *) type t type constr = t (** [types] is the same as [constr] but is intended to be used for documentation to indicate that such or such function specifically works with {e types} (i.e. terms of type a sort). (Rem:plurial form since [type] is a reserved ML keyword) *) type types = constr (** {5 Functions for dealing with constr terms. } The following functions are intended to simplify and to uniform the manipulation of terms. Some of these functions may be overlapped with previous ones. *) (** {6 Term constructors. } *) (** Constructs a de Bruijn index (DB indices begin at 1) *) val mkRel : int -> constr (** Constructs a Variable *) val mkVar : Id.t -> constr (** Constructs a machine integer *) val mkInt : Uint63.t -> constr (** Constructs a machine float number *) val mkFloat : Float64.t -> constr (** Constructs an patvar named "?n" *) val mkMeta : metavariable -> constr (** Constructs an existential variable *) type existential = Evar.t * constr array val mkEvar : existential -> constr (** Construct a sort *) val mkSort : Sorts.t -> types val mkSProp : types val mkProp : types val mkSet : types val mkType : Univ.Universe.t -> types (** This defines the strategy to use for verifiying a Cast *) type cast_kind = VMcast | NATIVEcast | DEFAULTcast | REVERTcast (** Constructs the term [t1::t2], i.e. the term t{_ 1} casted with the type t{_ 2} (that means t2 is declared as the type of t1). *) val mkCast : constr * cast_kind * constr -> constr (** Constructs the product [(x:t1)t2] *) val mkProd : Name.t Context.binder_annot * types * types -> types (** Constructs the abstraction \[x:t{_ 1}\]t{_ 2} *) val mkLambda : Name.t Context.binder_annot * types * constr -> constr (** Constructs the product [let x = t1 : t2 in t3] *) val mkLetIn : Name.t Context.binder_annot * constr * types * constr -> constr (** [mkApp (f, [|t1; ...; tN|]] constructs the application {%html:(f t1 ... tn)%} {%latex:$(f~t_1\dots f_n)$%}. *) val mkApp : constr * constr array -> constr val map_puniverses : ('a -> 'b) -> 'a Univ.puniverses -> 'b Univ.puniverses (** Constructs a Constant.t *) val mkConst : Constant.t -> constr val mkConstU : pconstant -> constr (** Constructs a projection application *) val mkProj : (Projection.t * constr) -> constr (** Inductive types *) (** Constructs the ith (co)inductive type of the block named kn *) val mkInd : inductive -> constr val mkIndU : pinductive -> constr (** Constructs the jth constructor of the ith (co)inductive type of the block named kn. *) val mkConstruct : constructor -> constr val mkConstructU : pconstructor -> constr val mkConstructUi : pinductive * int -> constr (** Make a constant, inductive, constructor or variable. *) val mkRef : GlobRef.t Univ.puniverses -> constr (** Constructs a destructor of inductive type. [mkCase ci p c ac] stand for match [c] as [x] in [I args] return [p] with [ac] presented as describe in [ci]. [p] structure is [fun args x -> "return clause"] [ac]{^ ith} element is ith constructor case presented as {e lambda construct_args (without params). case_term } *) val mkCase : case_info * constr * constr * constr array -> constr (** If [recindxs = [|i1,...in|]] [funnames = [|f1,.....fn|]] [typarray = [|t1,...tn|]] [bodies = [|b1,.....bn|]] then [mkFix ((recindxs,i), funnames, typarray, bodies) ] constructs the {% $ %}i{% $ %}th function of the block (counting from 0) [Fixpoint f1 [ctx1] = b1 with f2 [ctx2] = b2 ... with fn [ctxn] = bn.] where the length of the {% $ %}j{% $ %}th context is {% $ %}ij{% $ %}. *) type ('constr, 'types) prec_declaration = Name.t Context.binder_annot array * 'types array * 'constr array type ('constr, 'types) pfixpoint = (int array * int) * ('constr, 'types) prec_declaration (* The array of [int]'s tells for each component of the array of mutual fixpoints the number of lambdas to skip before finding the recursive argument (e.g., value is 2 in "fix f (x:A) (y:=t) (z:B) (v:=u) (w:I) {struct w}"), telling to skip x and z and that w is the recursive argument); The second component [int] tells which component of the block is returned *) type ('constr, 'types) pcofixpoint = int * ('constr, 'types) prec_declaration (* The component [int] tells which component of the block of cofixpoint is returned *) type rec_declaration = (constr, types) prec_declaration type fixpoint = (constr, types) pfixpoint val mkFix : fixpoint -> constr (** If [funnames = [|f1,.....fn|]] [typarray = [|t1,...tn|]] [bodies = [b1,.....bn]] then [mkCoFix (i, (funnames, typarray, bodies))] constructs the ith function of the block [CoFixpoint f1 = b1 with f2 = b2 ... with fn = bn.] *) type cofixpoint = (constr, types) pcofixpoint val mkCoFix : cofixpoint -> constr (** {6 Concrete type for making pattern-matching. } *) (** [constr array] is an instance matching definitional [named_context] in the same order (i.e. last argument first) *) type 'constr pexistential = Evar.t * 'constr array type ('constr, 'types, 'sort, 'univs) kind_of_term = | Rel of int (** Gallina-variable introduced by [forall], [fun], [let-in], [fix], or [cofix]. *) | Var of Id.t (** Gallina-variable that was introduced by Vernacular-command that extends the local context of the currently open section (i.e. [Variable] or [Let]). *) | Meta of metavariable | Evar of 'constr pexistential | Sort of 'sort | Cast of 'constr * cast_kind * 'types | Prod of Name.t Context.binder_annot * 'types * 'types (** Concrete syntax ["forall A:B,C"] is represented as [Prod (A,B,C)]. *) | Lambda of Name.t Context.binder_annot * 'types * 'constr (** Concrete syntax ["fun A:B => C"] is represented as [Lambda (A,B,C)]. *) | LetIn of Name.t Context.binder_annot * 'constr * 'types * 'constr (** Concrete syntax ["let A:C := B in D"] is represented as [LetIn (A,B,C,D)]. *) | App of 'constr * 'constr array (** Concrete syntax ["(F P1 P2 ... Pn)"] is represented as [App (F, [|P1; P2; ...; Pn|])]. The {!mkApp} constructor also enforces the following invariant: - [F] itself is not {!App} - and [[|P1;..;Pn|]] is not empty. *) | Const of (Constant.t * 'univs) (** Gallina-variable that was introduced by Vernacular-command that extends the global environment (i.e. [Parameter], or [Axiom], or [Definition], or [Theorem] etc.) *) | Ind of (inductive * 'univs) (** A name of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *) | Construct of (constructor * 'univs) (** A constructor of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *) | Case of case_info * 'constr * 'constr * 'constr array | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint | Proj of Projection.t * 'constr | Int of Uint63.t | Float of Float64.t (** User view of [constr]. For [App], it is ensured there is at least one argument and the function is not itself an applicative term *) val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term val of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constr val kind_nocast_gen : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) val kind_nocast : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term (** {6 Simple case analysis} *) val isRel : constr -> bool val isRelN : int -> constr -> bool val isVar : constr -> bool val isVarId : Id.t -> constr -> bool val isInd : constr -> bool val isEvar : constr -> bool val isMeta : constr -> bool val isEvar_or_Meta : constr -> bool val isSort : constr -> bool val isCast : constr -> bool val isApp : constr -> bool val isLambda : constr -> bool val isLetIn : constr -> bool val isProd : constr -> bool val isConst : constr -> bool val isConstruct : constr -> bool val isFix : constr -> bool val isCoFix : constr -> bool val isCase : constr -> bool val isProj : constr -> bool val is_Prop : constr -> bool val is_Set : constr -> bool val isprop : constr -> bool val is_Type : constr -> bool val iskind : constr -> bool val is_small : Sorts.t -> bool (** {6 Term destructors } *) (** Destructor operations are partial functions and @raise DestKO if the term has not the expected form. *) exception DestKO (** Destructs a de Bruijn index *) val destRel : constr -> int (** Destructs an existential variable *) val destMeta : constr -> metavariable (** Destructs a variable *) val destVar : constr -> Id.t (** Destructs a sort. [is_Prop] recognizes the sort [Prop], whether [isprop] recognizes both [Prop] and [Set]. *) val destSort : constr -> Sorts.t (** Destructs a casted term *) val destCast : constr -> constr * cast_kind * constr (** Destructs the product {% $ %}(x:t_1)t_2{% $ %} *) val destProd : types -> Name.t Context.binder_annot * types * types (** Destructs the abstraction {% $ %}[x:t_1]t_2{% $ %} *) val destLambda : constr -> Name.t Context.binder_annot * types * constr (** Destructs the let {% $ %}[x:=b:t_1]t_2{% $ %} *) val destLetIn : constr -> Name.t Context.binder_annot * constr * types * constr (** Destructs an application *) val destApp : constr -> constr * constr array (** Decompose any term as an applicative term; the list of args can be empty *) val decompose_app : constr -> constr * constr list (** Same as [decompose_app], but returns an array. *) val decompose_appvect : constr -> constr * constr array (** Destructs a constant *) val destConst : constr -> Constant.t Univ.puniverses (** Destructs an existential variable *) val destEvar : constr -> existential (** Destructs a (co)inductive type *) val destInd : constr -> inductive Univ.puniverses (** Destructs a constructor *) val destConstruct : constr -> constructor Univ.puniverses (** Destructs a [match c as x in I args return P with ... | Ci(...yij...) => ti | ... end] (or [let (..y1i..) := c as x in I args return P in t1], or [if c then t1 else t2]) @return [(info,c,fun args x => P,[|...|fun yij => ti| ...|])] where [info] is pretty-printing information *) val destCase : constr -> case_info * constr * constr * constr array (** Destructs a projection *) val destProj : constr -> Projection.t * constr (** Destructs the {% $ %}i{% $ %}th function of the block [Fixpoint f{_ 1} ctx{_ 1} = b{_ 1} with f{_ 2} ctx{_ 2} = b{_ 2} ... with f{_ n} ctx{_ n} = b{_ n}], where the length of the {% $ %}j{% $ %}th context is {% $ %}ij{% $ %}. *) val destFix : constr -> fixpoint val destCoFix : constr -> cofixpoint val destRef : constr -> GlobRef.t Univ.puniverses (** {6 Equality} *) (** [equal a b] is true if [a] equals [b] modulo alpha, casts, and application grouping *) val equal : constr -> constr -> bool (** [eq_constr_univs u a b] is [true] if [a] equals [b] modulo alpha, casts, application grouping and the universe equalities in [u]. *) val eq_constr_univs : constr UGraph.check_function (** [leq_constr_univs u a b] is [true] if [a] is convertible to [b] modulo alpha, casts, application grouping and the universe inequalities in [u]. *) val leq_constr_univs : constr UGraph.check_function (** [eq_constr_univs u a b] is [true] if [a] equals [b] modulo alpha, casts, application grouping and the universe equalities in [u]. *) val eq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrained (** [leq_constr_univs u a b] is [true] if [a] is convertible to [b] modulo alpha, casts, application grouping and the universe inequalities in [u]. *) val leq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrained (** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts, application grouping and ignoring universe instances. *) val eq_constr_nounivs : constr -> constr -> bool (** Total ordering compatible with [equal] *) val compare : constr -> constr -> int (** {6 Extension of Context with declarations on constr} *) type rel_declaration = (constr, types) Context.Rel.Declaration.pt type named_declaration = (constr, types) Context.Named.Declaration.pt type compacted_declaration = (constr, types) Context.Compacted.Declaration.pt type rel_context = rel_declaration list type named_context = named_declaration list type compacted_context = compacted_declaration list (** {6 Relocation and substitution } *) (** [exliftn el c] lifts [c] with lifting [el] *) val exliftn : Esubst.lift -> constr -> constr (** [liftn n k c] lifts by [n] indexes above or equal to [k] in [c] *) val liftn : int -> int -> constr -> constr (** [lift n c] lifts by [n] the positive indexes in [c] *) val lift : int -> constr -> constr (** {6 Functionals working on expressions canonically abstracted over a local context (possibly with let-ins)} *) (** [map_under_context f l c] maps [f] on the immediate subterms of a term abstracted over a context of length [n] (local definitions are counted) *) val map_under_context : (constr -> constr) -> int -> constr -> constr (** [map_branches f br] maps [f] on the immediate subterms of an array of "match" branches [br] in canonical eta-let-expanded form; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of each branch as well as the body of each branch *) val map_branches : (constr -> constr) -> case_info -> constr array -> constr array (** [map_return_predicate f p] maps [f] on the immediate subterms of a return predicate of a "match" in canonical eta-let-expanded form; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of each branch as well as the body of the predicate *) val map_return_predicate : (constr -> constr) -> case_info -> constr -> constr (** [map_under_context_with_binders g f n l c] maps [f] on the immediate subterms of a term abstracted over a context of length [n] (local definitions are counted); it preserves sharing; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to [n]) at each binder traversal *) val map_under_context_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> int -> constr -> constr (** [map_branches_with_binders f br] maps [f] on the immediate subterms of an array of "match" branches [br] in canonical eta-let-expanded form; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically adds 1 to [n]) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of the branch as well as the body of the branch *) val map_branches_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_info -> constr array -> constr array (** [map_return_predicate_with_binders f p] maps [f] on the immediate subterms of a return predicate of a "match" in canonical eta-let-expanded form; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically adds 1 to [n]) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified; it preserves sharing; the immediate subterms are the types and possibly terms occurring in the context of each branch as well as the body of the predicate *) val map_return_predicate_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_info -> constr -> constr (** [map_under_context_with_full_binders g f n l c] is similar to [map_under_context_with_binders] except that [g] takes also a full binder as argument and that only the number of binders (and not their signature) is required *) val map_under_context_with_full_binders : ((constr, constr) Context.Rel.Declaration.pt -> 'a -> 'a) -> ('a -> constr -> constr) -> 'a -> int -> constr -> constr (** [map_branches_with_full_binders g f l br] is equivalent to [map_branches_with_binders] but using [map_under_context_with_full_binders] *) val map_branches_with_full_binders : ((constr, constr) Context.Rel.Declaration.pt -> 'a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_info -> constr array -> constr array (** [map_return_predicate_with_full_binders g f l p] is equivalent to [map_return_predicate_with_binders] but using [map_under_context_with_full_binders] *) val map_return_predicate_with_full_binders : ((constr, constr) Context.Rel.Declaration.pt -> 'a -> 'a) -> ('a -> constr -> constr) -> 'a -> case_info -> constr -> constr (** {6 Functionals working on the immediate subterm of a construction } *) (** [fold f acc c] folds [f] on the immediate subterms of [c] starting from [acc] and proceeding from left to right according to the usual representation of the constructions; it is not recursive *) val fold : ('a -> constr -> 'a) -> 'a -> constr -> 'a val fold_with_full_binders : (rel_declaration -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b (** [map f c] maps [f] on the immediate subterms of [c]; it is not recursive and the order with which subterms are processed is not specified *) val map : (constr -> constr) -> constr -> constr (** [map_user_view f c] maps [f] on the immediate subterms of [c]; it differs from [map f c] in that the typing context and body of the return predicate and of the branches of a [match] are considered as immediate subterm of a [match] *) val map_user_view : (constr -> constr) -> constr -> constr (** Like {!map}, but also has an additional accumulator. *) val fold_map : ('a -> constr -> 'a * constr) -> 'a -> constr -> 'a * constr (** [map_with_binders g f n c] maps [f n] on the immediate subterms of [c]; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to [n]) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified *) val map_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr (** [iter f c] iters [f] on the immediate subterms of [c]; it is not recursive and the order with which subterms are processed is not specified *) val iter : (constr -> unit) -> constr -> unit (** [iter_with_binders g f n c] iters [f n] on the immediate subterms of [c]; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to [n]) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified *) val iter_with_binders : ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit (** [iter_with_binders g f n c] iters [f n] on the immediate subterms of [c]; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to [n]) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified *) val fold_constr_with_binders : ('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b type 'constr constr_compare_fn = int -> 'constr -> 'constr -> bool (** [compare_head f c1 c2] compare [c1] and [c2] using [f] to compare the immediate subterms of [c1] of [c2] if needed; Cast's, binders name and Cases annotations are not taken into account *) val compare_head : constr constr_compare_fn -> constr constr_compare_fn (** Convert a global reference applied to 2 instances. The int says how many arguments are given (as we can only use cumulativity for fully applied inductives/constructors) .*) type 'univs instance_compare_fn = GlobRef.t -> int -> 'univs -> 'univs -> bool (** [compare_head_gen u s f c1 c2] compare [c1] and [c2] using [f] to compare the immediate subterms of [c1] of [c2] if needed, [u] to compare universe instances, [s] to compare sorts; Cast's, binders name and Cases annotations are not taken into account *) val compare_head_gen : Univ.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> constr constr_compare_fn -> constr constr_compare_fn val compare_head_gen_leq_with : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> 'univs instance_compare_fn -> ('sort -> 'sort -> bool) -> 'v constr_compare_fn -> 'v constr_compare_fn -> 'v constr_compare_fn (** [compare_head_gen_with k1 k2 u s f c1 c2] compares [c1] and [c2] like [compare_head_gen u s f c1 c2], except that [k1] (resp. [k2]) is used,rather than {!kind}, to expose the immediate subterms of [c1] (resp. [c2]). *) val compare_head_gen_with : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> 'univs instance_compare_fn -> ('sort -> 'sort -> bool) -> 'v constr_compare_fn -> 'v constr_compare_fn (** [compare_head_gen_leq u s f fle c1 c2] compare [c1] and [c2] using [f] to compare the immediate subterms of [c1] of [c2] for conversion, [fle] for cumulativity, [u] to compare universe instances (the first boolean tells if they belong to a Constant.t), [s] to compare sorts for for subtyping; Cast's, binders name and Cases annotations are not taken into account *) val compare_head_gen_leq : Univ.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> constr constr_compare_fn -> constr constr_compare_fn -> constr constr_compare_fn (** {6 Hashconsing} *) val hash : constr -> int val case_info_hash : case_info -> int (*********************************************************************) val hcons : constr -> constr val debug_print : constr -> Pp.t val debug_print_fix : ('a -> Pp.t) -> ('a, 'a) pfixpoint -> Pp.t coq-8.11.0/kernel/univ.ml0000644000175000017500000010236413612664533015067 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* true | Prop, Prop -> true | Set, Set -> true | Level l, Level l' -> UGlobal.equal l l' | Var n, Var n' -> Int.equal n n' | _ -> false let compare u v = match u, v with | SProp, SProp -> 0 | SProp, _ -> -1 | _, SProp -> 1 | Prop,Prop -> 0 | Prop, _ -> -1 | _, Prop -> 1 | Set, Set -> 0 | Set, _ -> -1 | _, Set -> 1 | Level (dp1, i1), Level (dp2, i2) -> if i1 < i2 then -1 else if i1 > i2 then 1 else DirPath.compare dp1 dp2 | Level _, _ -> -1 | _, Level _ -> 1 | Var n, Var m -> Int.compare n m let hequal x y = x == y || match x, y with | SProp, SProp -> true | Prop, Prop -> true | Set, Set -> true | Level (n,d), Level (n',d') -> n == n' && d == d' | Var n, Var n' -> n == n' | _ -> false let hcons = function | SProp as x -> x | Prop as x -> x | Set as x -> x | Level (d,n) as x -> let d' = Names.DirPath.hcons d in if d' == d then x else Level (d',n) | Var _n as x -> x open Hashset.Combine let hash = function | SProp -> combinesmall 1 0 | Prop -> combinesmall 1 1 | Set -> combinesmall 1 2 | Var n -> combinesmall 2 n | Level (d, n) -> combinesmall 3 (combine n (Names.DirPath.hash d)) end module Level = struct module UGlobal = RawLevel.UGlobal type raw_level = RawLevel.t = | SProp | Prop | Set | Level of UGlobal.t | Var of int (** Embed levels with their hash value *) type t = { hash : int; data : RawLevel.t } let equal x y = x == y || Int.equal x.hash y.hash && RawLevel.equal x.data y.data let hash x = x.hash let data x = x.data (** Hashcons on levels + their hash *) module Self = struct type nonrec t = t type u = unit let eq x y = x.hash == y.hash && RawLevel.hequal x.data y.data let hash x = x.hash let hashcons () x = let data' = RawLevel.hcons x.data in if x.data == data' then x else { x with data = data' } end let hcons = let module H = Hashcons.Make(Self) in Hashcons.simple_hcons H.generate H.hcons () let make l = hcons { hash = RawLevel.hash l; data = l } let set = make Set let prop = make Prop let sprop = make SProp let is_small x = match data x with | Level _ -> false | Var _ -> false | SProp -> true | Prop -> true | Set -> true let is_prop x = match data x with | Prop -> true | _ -> false let is_set x = match data x with | Set -> true | _ -> false let is_sprop x = match data x with | SProp -> true | _ -> false let compare u v = if u == v then 0 else RawLevel.compare (data u) (data v) let to_string x = match data x with | SProp -> "SProp" | Prop -> "Prop" | Set -> "Set" | Level (d,n) -> Names.DirPath.to_string d^"."^string_of_int n | Var n -> "Var(" ^ string_of_int n ^ ")" let pr u = str (to_string u) let apart u v = match data u, data v with | SProp, _ | _, SProp | Prop, Set | Set, Prop -> true | _ -> false let vars = Array.init 20 (fun i -> make (Var i)) let var n = if n < 20 then vars.(n) else make (Var n) let var_index u = match data u with | Var n -> Some n | _ -> None let make qid = make (Level qid) let name u = match data u with | Level (d, n) -> Some (d, n) | _ -> None end (** Level maps *) module LMap = struct module M = HMap.Make (Level) include M let lunion l r = union (fun _k l _r -> Some l) l r let subst_union l r = union (fun _k l r -> match l, r with | Some _, _ -> Some l | None, None -> Some l | _, _ -> Some r) l r let diff ext orig = fold (fun u v acc -> if mem u orig then acc else add u v acc) ext empty let pr f m = h 0 (prlist_with_sep fnl (fun (u, v) -> Level.pr u ++ f v) (bindings m)) end module LSet = struct include LMap.Set let pr prl s = str"{" ++ prlist_with_sep spc prl (elements s) ++ str"}" let of_array l = Array.fold_left (fun acc x -> add x acc) empty l end type 'a universe_map = 'a LMap.t type universe_level = Level.t type universe_level_subst_fn = universe_level -> universe_level type universe_set = LSet.t (* An algebraic universe [universe] is either a universe variable [Level.t] or a formal universe known to be greater than some universe variables and strictly greater than some (other) universe variables Universes variables denote universes initially present in the term to type-check and non variable algebraic universes denote the universes inferred while type-checking: it is either the successor of a universe present in the initial term to type-check or the maximum of two algebraic universes *) module Universe = struct (* Invariants: non empty, sorted and without duplicates *) module Expr = struct type t = Level.t * int (* Hashing of expressions *) module ExprHash = struct type t = Level.t * int type u = Level.t -> Level.t let hashcons hdir (b,n as x) = let b' = hdir b in if b' == b then x else (b',n) let eq l1 l2 = l1 == l2 || match l1,l2 with | (b,n), (b',n') -> b == b' && n == n' let hash (x, n) = n + Level.hash x end module H = Hashcons.Make(ExprHash) let hcons = Hashcons.simple_hcons H.generate H.hcons Level.hcons let make l = (l, 0) let compare u v = if u == v then 0 else let (x, n) = u and (x', n') = v in if Int.equal n n' then Level.compare x x' else n - n' let sprop = hcons (Level.sprop, 0) let prop = hcons (Level.prop, 0) let set = hcons (Level.set, 0) let type1 = hcons (Level.set, 1) let is_small = function | (l,0) -> Level.is_small l | _ -> false let equal x y = x == y || (let (u,n) = x and (v,n') = y in Int.equal n n' && Level.equal u v) let hash = ExprHash.hash let leq (u,n) (v,n') = let cmp = Level.compare u v in if Int.equal cmp 0 then n <= n' else if n <= n' then (Level.is_prop u && not (Level.is_sprop v)) else false let successor (u,n as e) = if is_small e then type1 else (u, n + 1) let addn k (u,n as x) = if k = 0 then x else if Level.is_small u then (Level.set,n+k) else (u,n+k) type super_result = SuperSame of bool (* The level expressions are in cumulativity relation. boolean indicates if left is smaller than right? *) | SuperDiff of int (* The level expressions are unrelated, the comparison result is canonical *) (** [super u v] compares two level expressions, returning [SuperSame] if they refer to the same level at potentially different increments or [SuperDiff] if they are different. The booleans indicate if the left expression is "smaller" than the right one in both cases. *) let super (u,n) (v,n') = let cmp = Level.compare u v in if Int.equal cmp 0 then SuperSame (n < n') else let open RawLevel in match Level.data u, n, Level.data v, n' with | SProp, _, SProp, _ | Prop, _, Prop, _ -> SuperSame (n < n') | SProp, 0, Prop, 0 -> SuperSame true | Prop, 0, SProp, 0 -> SuperSame false | (SProp | Prop), 0, _, _ -> SuperSame true | _, _, (SProp | Prop), 0 -> SuperSame false | _, _, _, _ -> SuperDiff cmp let to_string (v, n) = if Int.equal n 0 then Level.to_string v else Level.to_string v ^ "+" ^ string_of_int n let pr x = str(to_string x) let pr_with f (v, n) = if Int.equal n 0 then f v else f v ++ str"+" ++ int n let is_level = function | (_v, 0) -> true | _ -> false let level = function | (v,0) -> Some v | _ -> None let get_level (v,_n) = v let map f (v, n as x) = let v' = f v in if v' == v then x else if Level.is_prop v' && n != 0 then (Level.set, n) else (v', n) end type t = Expr.t list let tip l = [l] let cons x l = x :: l let rec hash = function | [] -> 0 | e :: l -> Hashset.Combine.combinesmall (Expr.ExprHash.hash e) (hash l) let equal x y = x == y || List.equal Expr.equal x y let compare x y = if x == y then 0 else List.compare Expr.compare x y module Huniv = Hashcons.Hlist(Expr) let hcons = Hashcons.recursive_hcons Huniv.generate Huniv.hcons Expr.hcons let make l = tip (Expr.make l) let tip x = tip x let pr l = match l with | [u] -> Expr.pr u | _ -> str "max(" ++ hov 0 (prlist_with_sep pr_comma Expr.pr l) ++ str ")" let pr_with f l = match l with | [u] -> Expr.pr_with f u | _ -> str "max(" ++ hov 0 (prlist_with_sep pr_comma (Expr.pr_with f) l) ++ str ")" let is_level l = match l with | [l] -> Expr.is_level l | _ -> false let rec is_levels l = match l with | l :: r -> Expr.is_level l && is_levels r | [] -> true let level l = match l with | [l] -> Expr.level l | _ -> None let levels l = List.fold_left (fun acc x -> LSet.add (Expr.get_level x) acc) LSet.empty l let is_small u = match u with | [l] -> Expr.is_small l | _ -> false let sprop = tip Expr.sprop (* The lower predicative level of the hierarchy that contains (impredicative) Prop and singleton inductive types *) let type0m = tip Expr.prop (* The level of sets *) let type0 = tip Expr.set (* When typing [Prop] and [Set], there is no constraint on the level, hence the definition of [type1_univ], the type of [Prop] *) let type1 = tip Expr.type1 let is_sprop x = equal sprop x let is_type0m x = equal type0m x let is_type0 x = equal type0 x (* Returns the formal universe that lies just above the universe variable u. Used to type the sort u. *) let super l = if is_small l then type1 else List.Smart.map (fun x -> Expr.successor x) l let addn n l = List.Smart.map (fun x -> Expr.addn n x) l let rec merge_univs l1 l2 = match l1, l2 with | [], _ -> l2 | _, [] -> l1 | h1 :: t1, h2 :: t2 -> let open Expr in (match super h1 h2 with | SuperSame true (* h1 < h2 *) -> merge_univs t1 l2 | SuperSame false -> merge_univs l1 t2 | SuperDiff c -> if c <= 0 (* h1 < h2 is name order *) then cons h1 (merge_univs t1 l2) else cons h2 (merge_univs l1 t2)) let sort u = let rec aux a l = match l with | b :: l' -> let open Expr in (match super a b with | SuperSame false -> aux a l' | SuperSame true -> l | SuperDiff c -> if c <= 0 then cons a l else cons b (aux a l')) | [] -> cons a l in List.fold_right (fun a acc -> aux a acc) u [] (* Returns the formal universe that is greater than the universes u and v. Used to type the products. *) let sup x y = merge_univs x y let empty = [] let exists = List.exists let for_all = List.for_all let smart_map = List.Smart.map let map = List.map end type universe = Universe.t (* The level of predicative Set *) let type0m_univ = Universe.type0m let type0_univ = Universe.type0 let type1_univ = Universe.type1 let is_type0m_univ = Universe.is_type0m let is_type0_univ = Universe.is_type0 let is_univ_variable l = Universe.level l != None let is_small_univ = Universe.is_small let pr_uni = Universe.pr let sup = Universe.sup let super = Universe.super open Universe let universe_level = Universe.level type constraint_type = AcyclicGraph.constraint_type = Lt | Le | Eq type explanation = (constraint_type * Level.t) list let constraint_type_ord c1 c2 = match c1, c2 with | Lt, Lt -> 0 | Lt, _ -> -1 | Le, Lt -> 1 | Le, Le -> 0 | Le, Eq -> -1 | Eq, Eq -> 0 | Eq, _ -> 1 (* Universe inconsistency: error raised when trying to enforce a relation that would create a cycle in the graph of universes. *) type univ_inconsistency = constraint_type * universe * universe * explanation Lazy.t option exception UniverseInconsistency of univ_inconsistency let error_inconsistency o u v p = raise (UniverseInconsistency (o,make u,make v,p)) (* Constraints and sets of constraints. *) type univ_constraint = Level.t * constraint_type * Level.t let pr_constraint_type op = let op_str = match op with | Lt -> " < " | Le -> " <= " | Eq -> " = " in str op_str module UConstraintOrd = struct type t = univ_constraint let compare (u,c,v) (u',c',v') = let i = constraint_type_ord c c' in if not (Int.equal i 0) then i else let i' = Level.compare u u' in if not (Int.equal i' 0) then i' else Level.compare v v' end module Constraint = struct module S = Set.Make(UConstraintOrd) include S let pr prl c = v 0 (prlist_with_sep spc (fun (u1,op,u2) -> hov 0 (prl u1 ++ pr_constraint_type op ++ prl u2)) (elements c)) end let empty_constraint = Constraint.empty let union_constraint = Constraint.union let eq_constraint = Constraint.equal type constraints = Constraint.t module Hconstraint = Hashcons.Make( struct type t = univ_constraint type u = universe_level -> universe_level let hashcons hul (l1,k,l2) = (hul l1, k, hul l2) let eq (l1,k,l2) (l1',k',l2') = l1 == l1' && k == k' && l2 == l2' let hash = Hashtbl.hash end) module Hconstraints = Hashcons.Make( struct type t = constraints type u = univ_constraint -> univ_constraint let hashcons huc s = Constraint.fold (fun x -> Constraint.add (huc x)) s Constraint.empty let eq s s' = List.for_all2eq (==) (Constraint.elements s) (Constraint.elements s') let hash = Hashtbl.hash end) let hcons_constraint = Hashcons.simple_hcons Hconstraint.generate Hconstraint.hcons Level.hcons let hcons_constraints = Hashcons.simple_hcons Hconstraints.generate Hconstraints.hcons hcons_constraint (** A value with universe constraints. *) type 'a constrained = 'a * constraints let constraints_of (_, cst) = cst (** Constraint functions. *) type 'a constraint_function = 'a -> 'a -> constraints -> constraints let enforce_eq_level u v c = (* We discard trivial constraints like u=u *) if Level.equal u v then c else if Level.apart u v then error_inconsistency Eq u v None else Constraint.add (u,Eq,v) c let enforce_eq u v c = match Universe.level u, Universe.level v with | Some u, Some v -> enforce_eq_level u v c | _ -> anomaly (Pp.str "A universe comparison can only happen between variables.") let check_univ_eq u v = Universe.equal u v let enforce_eq u v c = if check_univ_eq u v then c else enforce_eq u v c let constraint_add_leq v u c = (* We just discard trivial constraints like u<=u *) if Expr.equal v u then c else match v, u with | (x,n), (y,m) -> let j = m - n in if j = -1 (* n = m+1, v+1 <= u <-> v < u *) then Constraint.add (x,Lt,y) c else if j <= -1 (* n = m+k, v+k <= u <-> v+(k-1) < u *) then if Level.equal x y then (* u+(k+1) <= u *) raise (UniverseInconsistency (Le, Universe.tip v, Universe.tip u, None)) else anomaly (Pp.str"Unable to handle arbitrary u+k <= v constraints.") else if j = 0 then Constraint.add (x,Le,y) c else (* j >= 1 *) (* m = n + k, u <= v+k *) if Level.equal x y then c (* u <= u+k, trivial *) else if Level.is_small x then c (* Prop,Set <= u+S k, trivial *) else Constraint.add (x,Le,y) c (* u <= v implies u <= v+k *) let check_univ_leq_one u v = Universe.exists (Expr.leq u) v let check_univ_leq u v = Universe.for_all (fun u -> check_univ_leq_one u v) u let enforce_leq u v c = match is_sprop u, is_sprop v with | true, true -> c | true, false | false, true -> raise (UniverseInconsistency (Le, u, v, None)) | false, false -> List.fold_left (fun c v -> (List.fold_left (fun c u -> constraint_add_leq u v c) c u)) c v let enforce_leq u v c = if check_univ_leq u v then c else enforce_leq u v c let enforce_leq_level u v c = if Level.equal u v then c else Constraint.add (u,Le,v) c (* Miscellaneous functions to remove or test local univ assumed to occur in a universe *) let univ_level_mem u v = List.exists (fun (l, n) -> Int.equal n 0 && Level.equal u l) v let univ_level_rem u v min = match Universe.level v with | Some u' -> if Level.equal u u' then min else v | None -> List.filter (fun (l, n) -> not (Int.equal n 0 && Level.equal u l)) v (* Is u mentioned in v (or equals to v) ? *) (**********************************************************************) (** Universe polymorphism *) (**********************************************************************) (** A universe level substitution, note that no algebraic universes are involved *) type universe_level_subst = universe_level universe_map (** A full substitution might involve algebraic universes *) type universe_subst = universe universe_map module Variance = struct (** A universe position in the instance given to a cumulative inductive can be the following. Note there is no Contravariant case because [forall x : A, B <= forall x : A', B'] requires [A = A'] as opposed to [A' <= A]. *) type t = Irrelevant | Covariant | Invariant let sup x y = match x, y with | Irrelevant, s | s, Irrelevant -> s | Invariant, _ | _, Invariant -> Invariant | Covariant, Covariant -> Covariant let check_subtype x y = match x, y with | (Irrelevant | Covariant | Invariant), Irrelevant -> true | Irrelevant, Covariant -> false | (Covariant | Invariant), Covariant -> true | (Irrelevant | Covariant), Invariant -> false | Invariant, Invariant -> true let pr = function | Irrelevant -> str "*" | Covariant -> str "+" | Invariant -> str "=" let leq_constraint csts variance u u' = match variance with | Irrelevant -> csts | Covariant -> enforce_leq_level u u' csts | Invariant -> enforce_eq_level u u' csts let eq_constraint csts variance u u' = match variance with | Irrelevant -> csts | Covariant | Invariant -> enforce_eq_level u u' csts let leq_constraints variance u u' csts = let len = Array.length u in assert (len = Array.length u' && len = Array.length variance); Array.fold_left3 leq_constraint csts variance u u' let eq_constraints variance u u' csts = let len = Array.length u in assert (len = Array.length u' && len = Array.length variance); Array.fold_left3 eq_constraint csts variance u u' end module Instance : sig type t = Level.t array val empty : t val is_empty : t -> bool val of_array : Level.t array -> t val to_array : t -> Level.t array val append : t -> t -> t val equal : t -> t -> bool val length : t -> int val hcons : t -> t val hash : t -> int val share : t -> t * int val subst_fn : universe_level_subst_fn -> t -> t val pr : (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t val levels : t -> LSet.t end = struct type t = Level.t array let empty : t = [||] module HInstancestruct = struct type nonrec t = t type u = Level.t -> Level.t let hashcons huniv a = let len = Array.length a in if Int.equal len 0 then empty else begin for i = 0 to len - 1 do let x = Array.unsafe_get a i in let x' = huniv x in if x == x' then () else Array.unsafe_set a i x' done; a end let eq t1 t2 = t1 == t2 || (Int.equal (Array.length t1) (Array.length t2) && let rec aux i = (Int.equal i (Array.length t1)) || (t1.(i) == t2.(i) && aux (i + 1)) in aux 0) let hash a = let accu = ref 0 in for i = 0 to Array.length a - 1 do let l = Array.unsafe_get a i in let h = Level.hash l in accu := Hashset.Combine.combine !accu h; done; (* [h] must be positive. *) let h = !accu land 0x3FFFFFFF in h end module HInstance = Hashcons.Make(HInstancestruct) let hcons = Hashcons.simple_hcons HInstance.generate HInstance.hcons Level.hcons let hash = HInstancestruct.hash let share a = (hcons a, hash a) let empty = hcons [||] let is_empty x = Int.equal (Array.length x) 0 let append x y = if Array.length x = 0 then y else if Array.length y = 0 then x else Array.append x y let of_array a = assert(Array.for_all (fun x -> not (Level.is_prop x || Level.is_sprop x)) a); a let to_array a = a let length a = Array.length a let subst_fn fn t = let t' = CArray.Smart.map fn t in if t' == t then t else of_array t' let levels x = LSet.of_array x let pr prl ?variance = let ppu i u = let v = Option.map (fun v -> v.(i)) variance in pr_opt_no_spc Variance.pr v ++ prl u in prvecti_with_sep spc ppu let equal t u = t == u || (Array.is_empty t && Array.is_empty u) || (CArray.for_all2 Level.equal t u (* Necessary as universe instances might come from different modules and unmarshalling doesn't preserve sharing *)) end let enforce_eq_instances x y = let ax = Instance.to_array x and ay = Instance.to_array y in if Array.length ax != Array.length ay then anomaly (Pp.(++) (Pp.str "Invalid argument: enforce_eq_instances called with") (Pp.str " instances of different lengths.")); CArray.fold_right2 enforce_eq_level ax ay let enforce_eq_variance_instances = Variance.eq_constraints let enforce_leq_variance_instances = Variance.leq_constraints let subst_instance_level s l = match l.Level.data with | Level.Var n -> s.(n) | _ -> l let subst_instance_instance s i = Array.Smart.map (fun l -> subst_instance_level s l) i let subst_instance_universe s u = let f x = Universe.Expr.map (fun u -> subst_instance_level s u) x in let u' = Universe.smart_map f u in if u == u' then u else Universe.sort u' let subst_instance_constraint s (u,d,v as c) = let u' = subst_instance_level s u in let v' = subst_instance_level s v in if u' == u && v' == v then c else (u',d,v') let subst_instance_constraints s csts = Constraint.fold (fun c csts -> Constraint.add (subst_instance_constraint s c) csts) csts Constraint.empty type 'a puniverses = 'a * Instance.t let out_punivs (x, _y) = x let in_punivs x = (x, Instance.empty) let eq_puniverses f (x, u) (y, u') = f x y && Instance.equal u u' (** A context of universe levels with universe constraints, representing local universe variables and constraints *) module UContext = struct type t = Instance.t constrained let make x = x (** Universe contexts (variables as a list) *) let empty = (Instance.empty, Constraint.empty) let is_empty (univs, cst) = Instance.is_empty univs && Constraint.is_empty cst let pr prl ?variance (univs, cst as ctx) = if is_empty ctx then mt() else h 0 (Instance.pr prl ?variance univs ++ str " |= ") ++ h 0 (v 0 (Constraint.pr prl cst)) let hcons (univs, cst) = (Instance.hcons univs, hcons_constraints cst) let instance (univs, _cst) = univs let constraints (_univs, cst) = cst let union (univs, cst) (univs', cst') = Instance.append univs univs', Constraint.union cst cst' let dest x = x let size (x,_) = Instance.length x end type universe_context = UContext.t let hcons_universe_context = UContext.hcons module AUContext = struct type t = Names.Name.t array constrained let repr (inst, cst) = (Array.init (Array.length inst) (fun i -> Level.var i), cst) let pr f ?variance ctx = UContext.pr f ?variance (repr ctx) let instantiate inst (u, cst) = assert (Array.length u = Array.length inst); subst_instance_constraints inst cst let names (nas, _) = nas let hcons (univs, cst) = (Array.map Names.Name.hcons univs, hcons_constraints cst) let empty = ([||], Constraint.empty) let is_empty (nas, cst) = Array.is_empty nas && Constraint.is_empty cst let union (nas, cst) (nas', cst') = (Array.append nas nas', Constraint.union cst cst') let size (nas, _) = Array.length nas end type 'a univ_abstracted = { univ_abstracted_value : 'a; univ_abstracted_binder : AUContext.t; } let map_univ_abstracted f {univ_abstracted_value;univ_abstracted_binder} = let univ_abstracted_value = f univ_abstracted_value in {univ_abstracted_value;univ_abstracted_binder} let hcons_abstract_universe_context = AUContext.hcons (** A set of universes with universe constraints. We linearize the set to a list after typechecking. Beware, representation could change. *) module ContextSet = struct type t = universe_set constrained let empty = (LSet.empty, Constraint.empty) let is_empty (univs, cst) = LSet.is_empty univs && Constraint.is_empty cst let equal (univs, cst as x) (univs', cst' as y) = x == y || (LSet.equal univs univs' && Constraint.equal cst cst') let of_set s = (s, Constraint.empty) let singleton l = of_set (LSet.singleton l) let of_instance i = of_set (Instance.levels i) let union (univs, cst as x) (univs', cst' as y) = if x == y then x else LSet.union univs univs', Constraint.union cst cst' let append (univs, cst) (univs', cst') = let univs = LSet.fold LSet.add univs univs' in let cst = Constraint.fold Constraint.add cst cst' in (univs, cst) let diff (univs, cst) (univs', cst') = LSet.diff univs univs', Constraint.diff cst cst' let add_universe u (univs, cst) = LSet.add u univs, cst let add_constraints cst' (univs, cst) = univs, Constraint.union cst cst' let add_instance inst (univs, cst) = let v = Instance.to_array inst in let fold accu u = LSet.add u accu in let univs = Array.fold_left fold univs v in (univs, cst) let sort_levels a = Array.sort Level.compare a; a let to_context (ctx, cst) = (Instance.of_array (sort_levels (Array.of_list (LSet.elements ctx))), cst) let of_context (ctx, cst) = (Instance.levels ctx, cst) let pr prl (univs, cst as ctx) = if is_empty ctx then mt() else h 0 (LSet.pr prl univs ++ str " |= ") ++ h 0 (v 0 (Constraint.pr prl cst)) let constraints (_univs, cst) = cst let levels (univs, _cst) = univs let size (univs,_) = LSet.cardinal univs end type universe_context_set = ContextSet.t (** A value in a universe context (resp. context set). *) type 'a in_universe_context = 'a * universe_context type 'a in_universe_context_set = 'a * universe_context_set let extend_in_context_set (a, ctx) ctx' = (a, ContextSet.union ctx ctx') (** Substitutions. *) let empty_subst = LMap.empty let is_empty_subst = LMap.is_empty let empty_level_subst = LMap.empty let is_empty_level_subst = LMap.is_empty (** Substitution functions *) (** With level to level substitutions. *) let subst_univs_level_level subst l = try LMap.find l subst with Not_found -> l let subst_univs_level_universe subst u = let f x = Universe.Expr.map (fun u -> subst_univs_level_level subst u) x in let u' = Universe.smart_map f u in if u == u' then u else Universe.sort u' let subst_univs_level_instance subst i = let i' = Instance.subst_fn (subst_univs_level_level subst) i in if i == i' then i else i' let subst_univs_level_constraint subst (u,d,v) = let u' = subst_univs_level_level subst u and v' = subst_univs_level_level subst v in if d != Lt && Level.equal u' v' then None else Some (u',d,v') let subst_univs_level_constraints subst csts = Constraint.fold (fun c -> Option.fold_right Constraint.add (subst_univs_level_constraint subst c)) csts Constraint.empty let subst_univs_level_abstract_universe_context subst (inst, csts) = inst, subst_univs_level_constraints subst csts (** With level to universe substitutions. *) type universe_subst_fn = universe_level -> universe let make_subst subst = fun l -> LMap.find l subst let subst_univs_expr_opt fn (l,n) = Universe.addn n (fn l) let subst_univs_universe fn ul = let subst, nosubst = List.fold_right (fun u (subst,nosubst) -> try let a' = subst_univs_expr_opt fn u in (a' :: subst, nosubst) with Not_found -> (subst, u :: nosubst)) ul ([], []) in if CList.is_empty subst then ul else let substs = List.fold_left Universe.merge_univs Universe.empty subst in List.fold_left (fun acc u -> Universe.merge_univs acc (Universe.tip u)) substs nosubst let make_instance_subst i = let arr = Instance.to_array i in Array.fold_left_i (fun i acc l -> LMap.add l (Level.var i) acc) LMap.empty arr let make_inverse_instance_subst i = let arr = Instance.to_array i in Array.fold_left_i (fun i acc l -> LMap.add (Level.var i) l acc) LMap.empty arr let make_abstract_instance (ctx, _) = Array.init (Array.length ctx) (fun i -> Level.var i) let abstract_universes nas ctx = let instance = UContext.instance ctx in let () = assert (Int.equal (Array.length nas) (Instance.length instance)) in let subst = make_instance_subst instance in let cstrs = subst_univs_level_constraints subst (UContext.constraints ctx) in let ctx = (nas, cstrs) in instance, ctx let rec compact_univ s vars i u = match u with | [] -> (s, List.rev vars) | (lvl, _) :: u -> match Level.var_index lvl with | Some k when not (LMap.mem lvl s) -> let lvl' = Level.var i in compact_univ (LMap.add lvl lvl' s) (k :: vars) (i+1) u | _ -> compact_univ s vars i u let compact_univ u = let (s, s') = compact_univ LMap.empty [] 0 u in (subst_univs_level_universe s u, s') (** Pretty-printing *) let pr_constraints prl = Constraint.pr prl let pr_universe_context = UContext.pr let pr_abstract_universe_context = AUContext.pr let pr_universe_context_set = ContextSet.pr let pr_universe_subst = LMap.pr (fun u -> str" := " ++ Universe.pr u ++ spc ()) let pr_universe_level_subst = LMap.pr (fun u -> str" := " ++ Level.pr u ++ spc ()) module Huniverse_set = Hashcons.Make( struct type t = universe_set type u = universe_level -> universe_level let hashcons huc s = LSet.fold (fun x -> LSet.add (huc x)) s LSet.empty let eq s s' = LSet.equal s s' let hash = Hashtbl.hash end) let hcons_universe_set = Hashcons.simple_hcons Huniverse_set.generate Huniverse_set.hcons Level.hcons let hcons_universe_context_set (v, c) = (hcons_universe_set v, hcons_constraints c) let hcons_univ x = Universe.hcons x let explain_universe_inconsistency prl (o,u,v,p : univ_inconsistency) = let pr_uni = Universe.pr_with prl in let pr_rel = function | Eq -> str"=" | Lt -> str"<" | Le -> str"<=" in let reason = match p with | None -> mt() | Some p -> let p = Lazy.force p in if p = [] then mt () else str " because" ++ spc() ++ pr_uni v ++ prlist (fun (r,v) -> spc() ++ pr_rel r ++ str" " ++ prl v) p ++ (if Universe.equal (Universe.make (snd (List.last p))) u then mt() else (spc() ++ str "= " ++ pr_uni u)) in str "Cannot enforce" ++ spc() ++ pr_uni u ++ spc() ++ pr_rel o ++ spc() ++ pr_uni v ++ reason coq-8.11.0/kernel/safe_typing.ml0000644000175000017500000014530413612664533016417 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* String.equal d1 d2 | Dvivo (d1,e1), Dvivo (d2,e2) -> String.equal d1 d2 && String.equal e1 e2 | Dvo_or_vi _, Dvivo _ -> false type library_info = DirPath.t * vodigest (** Functor and funsig parameters, most recent first *) type module_parameters = (MBId.t * module_type_body) list (** Part of the safe_env at a section opening time to be backtracked *) type section_data = { rev_env : Environ.env; rev_univ : Univ.ContextSet.t; rev_objlabels : Label.Set.t; } type safe_environment = { env : Environ.env; sections : section_data Section.t; modpath : ModPath.t; modvariant : modvariant; modresolver : Mod_subst.delta_resolver; paramresolver : Mod_subst.delta_resolver; revstruct : structure_body; modlabels : Label.Set.t; objlabels : Label.Set.t; univ : Univ.ContextSet.t; future_cst : Univ.ContextSet.t Future.computation list; engagement : engagement option; required : vodigest DPmap.t; loads : (ModPath.t * module_body) list; local_retroknowledge : Retroknowledge.action list; } and modvariant = | NONE | LIBRARY | SIG of module_parameters * safe_environment (** saved env *) | STRUCT of module_parameters * safe_environment (** saved env *) let rec library_dp_of_senv senv = match senv.modvariant with | NONE | LIBRARY -> ModPath.dp senv.modpath | SIG(_,senv) -> library_dp_of_senv senv | STRUCT(_,senv) -> library_dp_of_senv senv let empty_environment = { env = Environ.empty_env; modpath = ModPath.initial; modvariant = NONE; modresolver = Mod_subst.empty_delta_resolver; paramresolver = Mod_subst.empty_delta_resolver; revstruct = []; modlabels = Label.Set.empty; objlabels = Label.Set.empty; sections = Section.empty; future_cst = []; univ = Univ.ContextSet.empty; engagement = None; required = DPmap.empty; loads = []; local_retroknowledge = []; } let is_initial senv = match senv.revstruct, senv.modvariant with | [], NONE -> ModPath.equal senv.modpath ModPath.initial | _ -> false let sections_are_opened senv = not (Section.is_empty senv.sections) let delta_of_senv senv = senv.modresolver,senv.paramresolver let constant_of_delta_kn_senv senv kn = Mod_subst.constant_of_deltas_kn senv.paramresolver senv.modresolver kn let mind_of_delta_kn_senv senv kn = Mod_subst.mind_of_deltas_kn senv.paramresolver senv.modresolver kn (** The safe_environment state monad *) type safe_transformer0 = safe_environment -> safe_environment type 'a safe_transformer = safe_environment -> 'a * safe_environment (** {6 Engagement } *) let set_engagement_opt env = function | Some c -> Environ.set_engagement c env | None -> env let set_engagement c senv = { senv with env = Environ.set_engagement c senv.env; engagement = Some c } let set_typing_flags c senv = let env = Environ.set_typing_flags c senv.env in if env == senv.env then senv else { senv with env } let set_check_guarded b senv = let flags = Environ.typing_flags senv.env in set_typing_flags { flags with check_guarded = b } senv let set_check_positive b senv = let flags = Environ.typing_flags senv.env in set_typing_flags { flags with check_positive = b } senv let set_check_universes b senv = let flags = Environ.typing_flags senv.env in set_typing_flags { flags with check_universes = b } senv let set_indices_matter indices_matter senv = set_typing_flags { (Environ.typing_flags senv.env) with indices_matter } senv let set_share_reduction b senv = let flags = Environ.typing_flags senv.env in set_typing_flags { flags with share_reduction = b } senv let set_VM b senv = let flags = Environ.typing_flags senv.env in set_typing_flags { flags with enable_VM = b } senv let set_native_compiler b senv = let flags = Environ.typing_flags senv.env in set_typing_flags { flags with enable_native_compiler = b } senv let make_sprop_cumulative senv = { senv with env = Environ.make_sprop_cumulative senv.env } let set_allow_sprop b senv = { senv with env = Environ.set_allow_sprop b senv.env } (** Check that the engagement [c] expected by a library matches the current (initial) one *) let check_engagement env expected_impredicative_set = let impredicative_set = Environ.engagement env in begin match impredicative_set, expected_impredicative_set with | PredicativeSet, ImpredicativeSet -> CErrors.user_err Pp.(str "Needs option -impredicative-set.") | _ -> () end (** {6 Stm machinery } *) type side_effect = { from_env : Declarations.structure_body CEphemeron.key; seff_constant : Constant.t; seff_body : Constr.t Declarations.constant_body; } module SideEffects : sig type t val repr : t -> side_effect list val empty : t val add : side_effect -> t -> t val concat : t -> t -> t end = struct module SeffOrd = struct type t = side_effect let compare e1 e2 = Constant.CanOrd.compare e1.seff_constant e2.seff_constant end module SeffSet = Set.Make(SeffOrd) type t = { seff : side_effect list; elts : SeffSet.t } (** Invariant: [seff] is a permutation of the elements of [elts] *) let repr eff = eff.seff let empty = { seff = []; elts = SeffSet.empty } let add x es = if SeffSet.mem x es.elts then es else { seff = x :: es.seff; elts = SeffSet.add x es.elts } let concat xes yes = List.fold_right add xes.seff yes end type private_constants = SideEffects.t let side_effects_of_private_constants l = List.rev (SideEffects.repr l) (* Only used to push in an Environ.env. *) let lift_constant c = let body = match c.const_body with | OpaqueDef _ -> Undef None | Def _ | Undef _ | Primitive _ as body -> body in { c with const_body = body } let push_private_constants env eff = let eff = side_effects_of_private_constants eff in let add_if_undefined env eff = if Environ.mem_constant eff.seff_constant env then env else Environ.add_constant eff.seff_constant (lift_constant eff.seff_body) env in List.fold_left add_if_undefined env eff let empty_private_constants = SideEffects.empty let concat_private = SideEffects.concat let universes_of_private eff = let fold acc eff = match eff.seff_body.const_universes with | Monomorphic ctx -> Univ.ContextSet.union ctx acc | Polymorphic _ -> acc in List.fold_left fold Univ.ContextSet.empty (side_effects_of_private_constants eff) let env_of_safe_env senv = senv.env let env_of_senv = env_of_safe_env let sections_of_safe_env senv = senv.sections type constraints_addition = | Now of Univ.ContextSet.t | Later of Univ.ContextSet.t Future.computation let push_context_set poly cst senv = if Univ.ContextSet.is_empty cst then senv else let sections = if Section.is_empty senv.sections then senv.sections else Section.push_constraints cst senv.sections in { senv with env = Environ.push_context_set ~strict:(not poly) cst senv.env; univ = Univ.ContextSet.union cst senv.univ; sections } let add_constraints cst senv = match cst with | Later fc -> {senv with future_cst = fc :: senv.future_cst} | Now cst -> push_context_set false cst senv let add_constraints_list cst senv = List.fold_left (fun acc c -> add_constraints c acc) senv cst let is_curmod_library senv = match senv.modvariant with LIBRARY -> true | _ -> false let join_safe_environment ?(except=Future.UUIDSet.empty) e = Modops.join_structure except (Environ.opaque_tables e.env) e.revstruct; List.fold_left (fun e fc -> if Future.UUIDSet.mem (Future.uuid fc) except then e else add_constraints (Now (Future.join fc)) e) {e with future_cst = []} e.future_cst let is_joined_environment e = List.is_empty e.future_cst (** {6 Various checks } *) let exists_modlabel l senv = Label.Set.mem l senv.modlabels let exists_objlabel l senv = Label.Set.mem l senv.objlabels let check_modlabel l senv = if exists_modlabel l senv then Modops.error_existing_label l let check_objlabel l senv = if exists_objlabel l senv then Modops.error_existing_label l let check_objlabels ls senv = Label.Set.iter (fun l -> check_objlabel l senv) ls (** Are we closing the right module / modtype ? No user error here, since the opening/ending coherence is now verified in [vernac_end_segment] *) let check_current_label lab = function | MPdot (_,l) -> assert (Label.equal lab l) | _ -> assert false let check_struct = function | STRUCT (params,oldsenv) -> params, oldsenv | NONE | LIBRARY | SIG _ -> assert false let check_sig = function | SIG (params,oldsenv) -> params, oldsenv | NONE | LIBRARY | STRUCT _ -> assert false let check_current_library dir senv = match senv.modvariant with | LIBRARY -> assert (ModPath.equal senv.modpath (MPfile dir)) | NONE | STRUCT _ | SIG _ -> assert false (* cf Lib.end_compilation *) (** When operating on modules, we're normally outside sections *) let check_empty_context senv = assert (Environ.empty_context senv.env && Section.is_empty senv.sections) (** When adding a parameter to the current module/modtype, it must have been freshly started *) let check_empty_struct senv = assert (List.is_empty senv.revstruct && List.is_empty senv.loads) (** When starting a library, the current environment should be initial i.e. only composed of Require's *) let check_initial senv = assert (is_initial senv) (** When loading a library, its dependencies should be already there, with the correct digests. *) let check_required current_libs needed = let check (id,required) = try let actual = DPmap.find id current_libs in if not(digest_match ~actual ~required) then CErrors.user_err Pp.(pr_sequence str ["Inconsistent assumptions over module"; DirPath.to_string id; "."]) with Not_found -> CErrors.user_err Pp.(pr_sequence str ["Reference to unknown module"; DirPath.to_string id; "."]) in Array.iter check needed (** {6 Insertion of section variables} *) (** They are now typed before being added to the environment. Same as push_named, but check that the variable is not already there. Should *not* be done in Environ because tactics add temporary hypothesis many many times, and the check performed here would cost too much. *) let safe_push_named d env = let id = NamedDecl.get_id d in let _ = try let _ = Environ.lookup_named id env in CErrors.user_err Pp.(pr_sequence str ["Identifier"; Id.to_string id; "already defined."]) with Not_found -> () in Environ.push_named d env let push_named_def (id,de) senv = let sections = Section.push_local senv.sections in let c, r, typ = Term_typing.translate_local_def senv.env id de in let x = Context.make_annot id r in let env'' = safe_push_named (LocalDef (x, c, typ)) senv.env in { senv with sections; env = env'' } let push_named_assum (x,t) senv = let sections = Section.push_local senv.sections in let t, r = Term_typing.translate_local_assum senv.env t in let x = Context.make_annot x r in let env'' = safe_push_named (LocalAssum (x,t)) senv.env in { senv with sections; env = env'' } let push_section_context (nas, ctx) senv = let sections = Section.push_context (nas, ctx) senv.sections in let senv = { senv with sections } in let ctx = Univ.ContextSet.of_context ctx in (* We check that the universes are fresh. FIXME: This should be done implicitly, but we have to work around the API. *) let () = assert (Univ.LSet.for_all (fun u -> not (Univ.LSet.mem u (fst senv.univ))) (fst ctx)) in { senv with env = Environ.push_context_set ~strict:false ctx senv.env; univ = Univ.ContextSet.union ctx senv.univ } (** {6 Insertion of new declarations to current environment } *) let labels_of_mib mib = let add,get = let labels = ref Label.Set.empty in (fun id -> labels := Label.Set.add (Label.of_id id) !labels), (fun () -> !labels) in let visit_mip mip = add mip.mind_typename; Array.iter add mip.mind_consnames in Array.iter visit_mip mib.mind_packets; get () let globalize_constant_universes cb = match cb.const_universes with | Monomorphic cstrs -> (* Constraints hidden in the opaque body are added by [add_constant_aux] *) [cstrs] | Polymorphic _ -> [] let globalize_mind_universes mb = match mb.mind_universes with | Monomorphic ctx -> [ctx] | Polymorphic _ -> [] let constraints_of_sfb sfb = match sfb with | SFBconst cb -> globalize_constant_universes cb | SFBmind mib -> globalize_mind_universes mib | SFBmodtype mtb -> [mtb.mod_constraints] | SFBmodule mb -> [mb.mod_constraints] let add_retroknowledge pttc senv = { senv with env = Primred.add_retroknowledge senv.env pttc; local_retroknowledge = pttc::senv.local_retroknowledge } (** A generic function for adding a new field in a same environment. It also performs the corresponding [add_constraints]. *) type generic_name = | C of Constant.t | I of MutInd.t | M (** name already known, cf the mod_mp field *) | MT (** name already known, cf the mod_mp field *) let add_field ?(is_include=false) ((l,sfb) as field) gn senv = let mlabs,olabs = match sfb with | SFBmind mib -> let l = labels_of_mib mib in check_objlabels l senv; (Label.Set.empty,l) | SFBconst _ -> check_objlabel l senv; (Label.Set.empty, Label.Set.singleton l) | SFBmodule _ | SFBmodtype _ -> check_modlabel l senv; (Label.Set.singleton l, Label.Set.empty) in let senv = if is_include then (* Universes and constraints were added when the included module was defined eg in [Include F X.] (one of the trickier versions of Include) the constraints on the fields are exactly those of the fields of F which was defined separately. *) senv else (* Delayed constraints from opaque body are added by [add_constant_aux] *) let cst = constraints_of_sfb sfb in List.fold_left (fun senv cst -> push_context_set false cst senv) senv cst in let env' = match sfb, gn with | SFBconst cb, C con -> Environ.add_constant con cb senv.env | SFBmind mib, I mind -> Environ.add_mind mind mib senv.env | SFBmodtype mtb, MT -> Environ.add_modtype mtb senv.env | SFBmodule mb, M -> Modops.add_module mb senv.env | _ -> assert false in let sections = match sfb, gn with | SFBconst cb, C con -> let poly = Declareops.constant_is_polymorphic cb in Section.push_constant ~poly con senv.sections | SFBmind mib, I mind -> let poly = Declareops.inductive_is_polymorphic mib in Section.push_inductive ~poly mind senv.sections | _, (M | MT) -> senv.sections | _ -> assert false in { senv with env = env'; sections; revstruct = field :: senv.revstruct; modlabels = Label.Set.union mlabs senv.modlabels; objlabels = Label.Set.union olabs senv.objlabels } (** Applying a certain function to the resolver of a safe environment *) let update_resolver f senv = { senv with modresolver = f senv.modresolver } type global_declaration = | ConstantEntry : Entries.constant_entry -> global_declaration | OpaqueEntry : private_constants Entries.const_entry_body Entries.opaque_entry -> global_declaration type exported_private_constant = Constant.t let add_constant_aux senv (kn, cb) = let l = Constant.label kn in (* This is the only place where we hashcons the contents of a constant body *) let cb = if sections_are_opened senv then cb else Declareops.hcons_const_body cb in let senv' = add_field (l,SFBconst cb) (C kn) senv in let senv'' = match cb.const_body with | Undef (Some lev) -> update_resolver (Mod_subst.add_inline_delta_resolver (Constant.user kn) (lev,None)) senv' | _ -> senv' in senv'' let mk_pure_proof c = (c, Univ.ContextSet.empty), SideEffects.empty let inline_side_effects env body side_eff = let open Constr in (** First step: remove the constants that are still in the environment *) let filter e = let cb = (e.seff_constant, e.seff_body) in if Environ.mem_constant e.seff_constant env then None else Some (cb, e.from_env) in (* CAVEAT: we assure that most recent effects come first *) let side_eff = List.map_filter filter (SideEffects.repr side_eff) in let sigs = List.rev_map (fun (_, mb) -> mb) side_eff in let side_eff = List.fold_left (fun accu (cb, _) -> cb :: accu) [] side_eff in let side_eff = List.rev side_eff in (** Most recent side-effects first in side_eff *) if List.is_empty side_eff then (body, Univ.ContextSet.empty, sigs) else (** Second step: compute the lifts and substitutions to apply *) let cname c r = Context.make_annot (Name (Label.to_id (Constant.label c))) r in let fold (subst, var, ctx, args) (c, cb) = let (b, opaque) = match cb.const_body with | Def b -> (Mod_subst.force_constr b, false) | OpaqueDef b -> (b, true) | _ -> assert false in match cb.const_universes with | Monomorphic univs -> (** Abstract over the term at the top of the proof *) let ty = cb.const_type in let subst = Cmap_env.add c (Inr var) subst in let ctx = Univ.ContextSet.union ctx univs in (subst, var + 1, ctx, (cname c cb.const_relevance, b, ty, opaque) :: args) | Polymorphic _ -> (** Inline the term to emulate universe polymorphism *) let subst = Cmap_env.add c (Inl b) subst in (subst, var, ctx, args) in let (subst, len, ctx, args) = List.fold_left fold (Cmap_env.empty, 1, Univ.ContextSet.empty, []) side_eff in (** Third step: inline the definitions *) let rec subst_const i k t = match Constr.kind t with | Const (c, u) -> let data = try Some (Cmap_env.find c subst) with Not_found -> None in begin match data with | None -> t | Some (Inl b) -> (** [b] is closed but may refer to other constants *) subst_const i k (Vars.subst_instance_constr u b) | Some (Inr n) -> mkRel (k + n - i) end | Rel n -> (** Lift free rel variables *) if n <= k then t else mkRel (n + len - i - 1) | _ -> Constr.map_with_binders ((+) 1) (fun k t -> subst_const i k t) k t in let map_args i (na, b, ty, opaque) = (** Both the type and the body may mention other constants *) let ty = subst_const (len - i - 1) 0 ty in let b = subst_const (len - i - 1) 0 b in (na, b, ty, opaque) in let args = List.mapi map_args args in let body = subst_const 0 0 body in let fold_arg (na, b, ty, opaque) accu = if opaque then mkApp (mkLambda (na, ty, accu), [|b|]) else mkLetIn (na, b, ty, accu) in let body = List.fold_right fold_arg args body in (body, ctx, sigs) let inline_private_constants env ((body, ctx), side_eff) = let body, ctx',_ = inline_side_effects env body side_eff in let ctx' = Univ.ContextSet.union ctx ctx' in (body, ctx') let is_suffix l suf = match l with | [] -> false | _ :: l -> l == suf (* Given the list of signatures of side effects, checks if they match. * I.e. if they are ordered descendants of the current revstruct. Returns the number of effects that can be trusted. *) let check_signatures curmb sl = let is_direct_ancestor accu mb = match accu with | None -> None | Some (n, curmb) -> try let mb = CEphemeron.get mb in if is_suffix mb curmb then Some (n + 1, mb) else None with CEphemeron.InvalidKey -> None in let sl = List.fold_left is_direct_ancestor (Some (0, curmb)) sl in match sl with | None -> 0 | Some (n, _) -> n type side_effect_declaration = | DefinitionEff : Entries.definition_entry -> side_effect_declaration | OpaqueEff : Constr.constr Entries.opaque_entry -> side_effect_declaration let constant_entry_of_side_effect eff = let cb = eff.seff_body in let open Entries in let univs = match cb.const_universes with | Monomorphic uctx -> Monomorphic_entry uctx | Polymorphic auctx -> Polymorphic_entry (Univ.AUContext.names auctx, Univ.AUContext.repr auctx) in let p = match cb.const_body with | OpaqueDef b -> b | Def b -> Mod_subst.force_constr b | _ -> assert false in if Declareops.is_opaque cb then OpaqueEff { opaque_entry_body = p; opaque_entry_secctx = Context.Named.to_vars cb.const_hyps; opaque_entry_feedback = None; opaque_entry_type = cb.const_type; opaque_entry_universes = univs; } else DefinitionEff { const_entry_body = p; const_entry_secctx = Some (Context.Named.to_vars cb.const_hyps); const_entry_feedback = None; const_entry_type = Some cb.const_type; const_entry_universes = univs; const_entry_inline_code = cb.const_inline_code } let export_eff eff = (eff.seff_constant, eff.seff_body) let is_empty_private = function | Opaqueproof.PrivateMonomorphic ctx -> Univ.ContextSet.is_empty ctx | Opaqueproof.PrivatePolymorphic (_, ctx) -> Univ.ContextSet.is_empty ctx let empty_private univs = match univs with | Monomorphic _ -> Opaqueproof.PrivateMonomorphic Univ.ContextSet.empty | Polymorphic auctx -> Opaqueproof.PrivatePolymorphic (Univ.AUContext.size auctx, Univ.ContextSet.empty) (* Special function to call when the body of an opaque definition is provided. It performs the type-checking of the body immediately. *) let translate_direct_opaque env kn ce = let cb, ctx = Term_typing.translate_opaque env kn ce in let body = ce.Entries.opaque_entry_body, Univ.ContextSet.empty in let handle _env c () = (c, Univ.ContextSet.empty, 0) in let (c, u) = Term_typing.check_delayed handle ctx (body, ()) in (* No constraints can be generated, we set it empty everywhere *) let () = assert (is_empty_private u) in { cb with const_body = OpaqueDef c } let export_side_effects mb env (b_ctx, eff) = let not_exists e = not (Environ.mem_constant e.seff_constant env) in let aux (acc,sl) e = if not (not_exists e) then acc, sl else e :: acc, e.from_env :: sl in let seff, signatures = List.fold_left aux ([],[]) (SideEffects.repr eff) in let trusted = check_signatures mb signatures in let push_seff env eff = let { seff_constant = kn; seff_body = cb ; _ } = eff in let env = Environ.add_constant kn (lift_constant cb) env in match cb.const_universes with | Polymorphic _ -> env | Monomorphic ctx -> Environ.push_context_set ~strict:true ctx env in let rec translate_seff sl seff acc env = match seff with | [] -> List.rev acc, b_ctx | eff :: rest -> if Int.equal sl 0 then let env, cb = let kn = eff.seff_constant in let ce = constant_entry_of_side_effect eff in let open Entries in let cb = match ce with | DefinitionEff ce -> Term_typing.translate_constant env kn (DefinitionEntry ce) | OpaqueEff ce -> translate_direct_opaque env kn ce in let eff = { eff with seff_body = cb } in (push_seff env eff, export_eff eff) in translate_seff 0 rest (cb :: acc) env else let env = push_seff env eff in let ecb = export_eff eff in translate_seff (sl - 1) rest (ecb :: acc) env in translate_seff trusted seff [] env let push_opaque_proof pf senv = let o, otab = Opaqueproof.create (library_dp_of_senv senv) pf (Environ.opaque_tables senv.env) in let senv = { senv with env = Environ.set_opaque_tables senv.env otab } in senv, o let export_private_constants ce senv = let exported, ce = export_side_effects senv.revstruct senv.env ce in let map senv (kn, c) = match c.const_body with | OpaqueDef p -> let local = empty_private c.const_universes in let senv, o = push_opaque_proof (Future.from_val (p, local)) senv in senv, (kn, { c with const_body = OpaqueDef o }) | Def _ | Undef _ | Primitive _ as body -> senv, (kn, { c with const_body = body }) in let senv, bodies = List.fold_left_map map senv exported in let exported = List.map (fun (kn, _) -> kn) exported in (* No delayed constants to declare *) let senv = List.fold_left add_constant_aux senv bodies in (ce, exported), senv let add_constant l decl senv = let kn = Constant.make2 senv.modpath l in let cb = match decl with | OpaqueEntry ce -> let handle env body eff = let body, uctx, signatures = inline_side_effects env body eff in let trusted = check_signatures senv.revstruct signatures in body, uctx, trusted in let cb, ctx = Term_typing.translate_opaque senv.env kn ce in let map pf = Term_typing.check_delayed handle ctx pf in let pf = Future.chain ce.Entries.opaque_entry_body map in { cb with const_body = OpaqueDef pf } | ConstantEntry ce -> Term_typing.translate_constant senv.env kn ce in let senv = let senv, cb, delayed_cst = match cb.const_body with | OpaqueDef fc -> let senv, o = push_opaque_proof fc senv in let delayed_cst = if not (Declareops.constant_is_polymorphic cb) then let map (_, u) = match u with | Opaqueproof.PrivateMonomorphic ctx -> ctx | Opaqueproof.PrivatePolymorphic _ -> assert false in let fc = Future.chain fc map in match Future.peek_val fc with | None -> [Later fc] | Some c -> [Now c] else [] in senv, { cb with const_body = OpaqueDef o }, delayed_cst | Undef _ | Def _ | Primitive _ as body -> senv, { cb with const_body = body }, [] in let senv = add_constant_aux senv (kn, cb) in add_constraints_list delayed_cst senv in let senv = match decl with | ConstantEntry (Entries.PrimitiveEntry { Entries.prim_entry_content = CPrimitives.OT_type t; _ }) -> if sections_are_opened senv then CErrors.anomaly (Pp.str "Primitive type not allowed in sections"); add_retroknowledge (Retroknowledge.Register_type(t,kn)) senv | _ -> senv in kn, senv let add_private_constant l decl senv : (Constant.t * private_constants) * safe_environment = let kn = Constant.make2 senv.modpath l in let cb = match decl with | OpaqueEff ce -> translate_direct_opaque senv.env kn ce | DefinitionEff ce -> Term_typing.translate_constant senv.env kn (Entries.DefinitionEntry ce) in let senv, dcb = match cb.const_body with | Def _ as const_body -> senv, { cb with const_body } | OpaqueDef c -> let local = empty_private cb.const_universes in let senv, o = push_opaque_proof (Future.from_val (c, local)) senv in senv, { cb with const_body = OpaqueDef o } | Undef _ | Primitive _ -> assert false in let senv = add_constant_aux senv (kn, dcb) in let eff = let from_env = CEphemeron.create senv.revstruct in let eff = { from_env = from_env; seff_constant = kn; seff_body = cb; } in SideEffects.add eff empty_private_constants in (kn, eff), senv (** Insertion of inductive types *) let check_mind mie lab = let open Entries in match mie.mind_entry_inds with | [] -> assert false (* empty inductive entry *) | oie::_ -> (* The label and the first inductive type name should match *) assert (Id.equal (Label.to_id lab) oie.mind_entry_typename) let add_mind l mie senv = let () = check_mind mie l in let kn = MutInd.make2 senv.modpath l in let mib = Indtypes.check_inductive senv.env kn mie in let mib = match mib.mind_hyps with [] -> Declareops.hcons_mind mib | _ -> mib in kn, add_field (l,SFBmind mib) (I kn) senv (** Insertion of module types *) let add_modtype l params_mte inl senv = let mp = MPdot(senv.modpath, l) in let mtb = Mod_typing.translate_modtype senv.env mp inl params_mte in let mtb = Declareops.hcons_module_type mtb in let senv' = add_field (l,SFBmodtype mtb) MT senv in mp, senv' (** full_add_module adds module with universes and constraints *) let full_add_module mb senv = let senv = add_constraints (Now mb.mod_constraints) senv in let dp = ModPath.dp mb.mod_mp in let linkinfo = Nativecode.link_info_of_dirpath dp in { senv with env = Modops.add_linked_module mb linkinfo senv.env } let full_add_module_type mp mt senv = let senv = add_constraints (Now mt.mod_constraints) senv in { senv with env = Modops.add_module_type mp mt senv.env } (** Insertion of modules *) let add_module l me inl senv = let mp = MPdot(senv.modpath, l) in let mb = Mod_typing.translate_module senv.env mp inl me in let mb = Declareops.hcons_module_body mb in let senv' = add_field (l,SFBmodule mb) M senv in let senv'' = if Modops.is_functor mb.mod_type then senv' else update_resolver (Mod_subst.add_delta_resolver mb.mod_delta) senv' in (mp,mb.mod_delta),senv'' (** {6 Interactive sections *) let open_section senv = let custom = { rev_env = senv.env; rev_univ = senv.univ; rev_objlabels = senv.objlabels; } in let sections = Section.open_section ~custom senv.sections in { senv with sections } let close_section senv = let open Section in let sections0 = senv.sections in let env0 = senv.env in (* First phase: revert the declarations added in the section *) let sections, entries, cstrs, revert = Section.close_section sections0 in let rec pop_revstruct accu entries revstruct = match entries, revstruct with | [], revstruct -> accu, revstruct | _ :: _, [] -> CErrors.anomaly (Pp.str "Unmatched section data") | entry :: entries, (lbl, leaf) :: revstruct -> let data = match entry, leaf with | SecDefinition kn, SFBconst cb -> let () = assert (Label.equal lbl (Constant.label kn)) in `Definition (kn, cb) | SecInductive ind, SFBmind mib -> let () = assert (Label.equal lbl (MutInd.label ind)) in `Inductive (ind, mib) | (SecDefinition _ | SecInductive _), (SFBconst _ | SFBmind _) -> CErrors.anomaly (Pp.str "Section content mismatch") | (SecDefinition _ | SecInductive _), (SFBmodule _ | SFBmodtype _) -> CErrors.anomaly (Pp.str "Module inside a section") in pop_revstruct (data :: accu) entries revstruct in let redo, revstruct = pop_revstruct [] entries senv.revstruct in (* Don't revert the delayed constraints. If some delayed constraints were forced inside the section, they have been turned into global monomorphic that are going to be replayed. Those that are not forced are not readded by {!add_constant_aux}. *) let { rev_env = env; rev_univ = univ; rev_objlabels = objlabels } = revert in (* Do not revert the opaque table, the discharged opaque constants are referring to it. *) let env = Environ.set_opaque_tables env (Environ.opaque_tables senv.env) in let senv = { senv with env; revstruct; sections; univ; objlabels; } in (* Second phase: replay the discharged section contents *) let senv = add_constraints (Now cstrs) senv in let modlist = Section.replacement_context env0 sections0 in let cooking_info seg = let { abstr_ctx; abstr_subst; abstr_uctx } = seg in let abstract = (abstr_ctx, abstr_subst, abstr_uctx) in { Opaqueproof.modlist; abstract } in let fold senv = function | `Definition (kn, cb) -> let info = cooking_info (Section.segment_of_constant env0 kn sections0) in let r = { Cooking.from = cb; info } in let cb = Term_typing.translate_recipe senv.env kn r in (* Delayed constants are already in the global environment *) add_constant_aux senv (kn, cb) | `Inductive (ind, mib) -> let info = cooking_info (Section.segment_of_inductive env0 ind sections0) in let mie = Cooking.cook_inductive info mib in let mie = InferCumulativity.infer_inductive senv.env mie in let _, senv = add_mind (MutInd.label ind) mie senv in senv in List.fold_left fold senv redo (** {6 Starting / ending interactive modules and module types } *) let start_module l senv = let () = check_modlabel l senv in let () = check_empty_context senv in let mp = MPdot(senv.modpath, l) in mp, { empty_environment with env = senv.env; modpath = mp; modvariant = STRUCT ([],senv); required = senv.required } let start_modtype l senv = let () = check_modlabel l senv in let () = check_empty_context senv in let mp = MPdot(senv.modpath, l) in mp, { empty_environment with env = senv.env; modpath = mp; modvariant = SIG ([], senv); required = senv.required } (** Adding parameters to the current module or module type. This module should have been freshly started. *) let add_module_parameter mbid mte inl senv = let () = check_empty_struct senv in let mp = MPbound mbid in let mtb = Mod_typing.translate_modtype senv.env mp inl ([],mte) in let senv = full_add_module_type mp mtb senv in let new_variant = match senv.modvariant with | STRUCT (params,oldenv) -> STRUCT ((mbid,mtb) :: params, oldenv) | SIG (params,oldenv) -> SIG ((mbid,mtb) :: params, oldenv) | _ -> assert false in let new_paramresolver = if Modops.is_functor mtb.mod_type then senv.paramresolver else Mod_subst.add_delta_resolver mtb.mod_delta senv.paramresolver in mtb.mod_delta, { senv with modvariant = new_variant; paramresolver = new_paramresolver } let functorize params init = List.fold_left (fun e (mbid,mt) -> MoreFunctor(mbid,mt,e)) init params let propagate_loads senv = List.fold_left (fun env (_,mb) -> full_add_module mb env) senv (List.rev senv.loads) (** Build the module body of the current module, taking in account a possible return type (_:T) *) let functorize_module params mb = let f x = functorize params x in { mb with mod_expr = Modops.implem_smartmap f f mb.mod_expr; mod_type = f mb.mod_type; mod_type_alg = Option.map f mb.mod_type_alg } let build_module_body params restype senv = let struc = NoFunctor (List.rev senv.revstruct) in let restype' = Option.map (fun (ty,inl) -> (([],ty),inl)) restype in let mb = Mod_typing.finalize_module senv.env senv.modpath (struc,None,senv.modresolver,senv.univ) restype' in let mb' = functorize_module params mb in { mb' with mod_retroknowledge = ModBodyRK senv.local_retroknowledge } (** Returning back to the old pre-interactive-module environment, with one extra component and some updated fields (constraints, required, etc) *) let allow_delayed_constants = ref false let propagate_senv newdef newenv newresolver senv oldsenv = let now_cst, later_cst = List.partition Future.is_val senv.future_cst in (* This asserts that after Paral-ITP, standard vo compilation is behaving * exctly as before: the same universe constraints are added to modules *) if not !allow_delayed_constants && later_cst <> [] then CErrors.anomaly ~label:"safe_typing" Pp.(str "True Future.t were created for opaque constants even if -async-proofs is off"); { oldsenv with env = newenv; modresolver = newresolver; revstruct = newdef::oldsenv.revstruct; modlabels = Label.Set.add (fst newdef) oldsenv.modlabels; univ = List.fold_left (fun acc cst -> Univ.ContextSet.union acc (Future.force cst)) (Univ.ContextSet.union senv.univ oldsenv.univ) now_cst; future_cst = later_cst @ oldsenv.future_cst; (* engagement is propagated to the upper level *) engagement = senv.engagement; required = senv.required; loads = senv.loads@oldsenv.loads; local_retroknowledge = senv.local_retroknowledge@oldsenv.local_retroknowledge; } let end_module l restype senv = let mp = senv.modpath in let params, oldsenv = check_struct senv.modvariant in let () = check_current_label l mp in let () = check_empty_context senv in let mbids = List.rev_map fst params in let mb = build_module_body params restype senv in let newenv = Environ.set_opaque_tables oldsenv.env (Environ.opaque_tables senv.env) in let newenv = Environ.set_native_symbols newenv senv.env.Environ.native_symbols in let newenv = set_engagement_opt newenv senv.engagement in let senv'= propagate_loads { senv with env = newenv; univ = Univ.ContextSet.union senv.univ mb.mod_constraints} in let newenv = Environ.push_context_set ~strict:true mb.mod_constraints senv'.env in let newenv = Modops.add_module mb newenv in let newresolver = if Modops.is_functor mb.mod_type then oldsenv.modresolver else Mod_subst.add_delta_resolver mb.mod_delta oldsenv.modresolver in (mp,mbids,mb.mod_delta), propagate_senv (l,SFBmodule mb) newenv newresolver senv' oldsenv let build_mtb mp sign cst delta = { mod_mp = mp; mod_expr = (); mod_type = sign; mod_type_alg = None; mod_constraints = cst; mod_delta = delta; mod_retroknowledge = ModTypeRK } let end_modtype l senv = let mp = senv.modpath in let params, oldsenv = check_sig senv.modvariant in let () = check_current_label l mp in let () = check_empty_context senv in let mbids = List.rev_map fst params in let newenv = Environ.set_opaque_tables oldsenv.env (Environ.opaque_tables senv.env) in let newenv = Environ.set_native_symbols newenv senv.env.Environ.native_symbols in let newenv = Environ.push_context_set ~strict:true senv.univ newenv in let newenv = set_engagement_opt newenv senv.engagement in let senv' = propagate_loads {senv with env=newenv} in let auto_tb = functorize params (NoFunctor (List.rev senv.revstruct)) in let mtb = build_mtb mp auto_tb senv'.univ senv.modresolver in let newenv = Environ.add_modtype mtb senv'.env in let newresolver = oldsenv.modresolver in (mp,mbids), propagate_senv (l,SFBmodtype mtb) newenv newresolver senv' oldsenv (** {6 Inclusion of module or module type } *) let add_include me is_module inl senv = let open Mod_typing in let mp_sup = senv.modpath in let sign,(),resolver,cst = translate_mse_incl is_module senv.env mp_sup inl me in let senv = add_constraints (Now cst) senv in (* Include Self support *) let rec compute_sign sign mb resolver senv = match sign with | MoreFunctor(mbid,mtb,str) -> let cst_sub = Subtyping.check_subtypes senv.env mb mtb in let senv = add_constraints (Now (Univ.ContextSet.add_constraints cst_sub Univ.ContextSet.empty)) senv in let mpsup_delta = Modops.inline_delta_resolver senv.env inl mp_sup mbid mtb mb.mod_delta in let subst = Mod_subst.map_mbid mbid mp_sup mpsup_delta in let resolver = Mod_subst.subst_codom_delta_resolver subst resolver in compute_sign (Modops.subst_signature subst str) mb resolver senv | NoFunctor str -> resolver,str,senv in let resolver,str,senv = let struc = NoFunctor (List.rev senv.revstruct) in let mtb = build_mtb mp_sup struc Univ.ContextSet.empty senv.modresolver in compute_sign sign mtb resolver senv in let senv = update_resolver (Mod_subst.add_delta_resolver resolver) senv in let add senv ((l,elem) as field) = let new_name = match elem with | SFBconst _ -> C (Mod_subst.constant_of_delta_kn resolver (KerName.make mp_sup l)) | SFBmind _ -> I (Mod_subst.mind_of_delta_kn resolver (KerName.make mp_sup l)) | SFBmodule _ -> M | SFBmodtype _ -> MT in add_field ~is_include:true field new_name senv in resolver, List.fold_left add senv str (** {6 Libraries, i.e. compiled modules } *) type compiled_library = { comp_name : DirPath.t; comp_mod : module_body; comp_deps : library_info array; comp_enga : engagement; comp_natsymbs : Nativevalues.symbols } let module_of_library lib = lib.comp_mod type native_library = Nativecode.global list (** FIXME: MS: remove?*) let current_modpath senv = senv.modpath let current_dirpath senv = Names.ModPath.dp (current_modpath senv) let start_library dir senv = check_initial senv; assert (not (DirPath.is_empty dir)); let mp = MPfile dir in mp, { empty_environment with env = senv.env; modpath = mp; modvariant = LIBRARY; required = senv.required } let export ?except ~output_native_objects senv dir = let senv = try join_safe_environment ?except senv with e -> let e = CErrors.push e in CErrors.user_err ~hdr:"export" (CErrors.iprint e) in assert(senv.future_cst = []); let () = check_current_library dir senv in let mp = senv.modpath in let str = NoFunctor (List.rev senv.revstruct) in let mb = { mod_mp = mp; mod_expr = FullStruct; mod_type = str; mod_type_alg = None; mod_constraints = senv.univ; mod_delta = senv.modresolver; mod_retroknowledge = ModBodyRK senv.local_retroknowledge } in let ast, symbols = if output_native_objects then Nativelibrary.dump_library mp dir senv.env str else [], Nativevalues.empty_symbols in let lib = { comp_name = dir; comp_mod = mb; comp_deps = Array.of_list (DPmap.bindings senv.required); comp_enga = Environ.engagement senv.env; comp_natsymbs = symbols } in mp, lib, ast (* cst are the constraints that were computed by the vi2vo step and hence are * not part of the mb.mod_constraints field (but morally should be) *) let import lib cst vodigest senv = check_required senv.required lib.comp_deps; check_engagement senv.env lib.comp_enga; if DirPath.equal (ModPath.dp senv.modpath) lib.comp_name then CErrors.user_err ~hdr:"Safe_typing.import" (Pp.strbrk "Cannot load a library with the same name as the current one."); let mp = MPfile lib.comp_name in let mb = lib.comp_mod in let env = Environ.push_context_set ~strict:true (Univ.ContextSet.union mb.mod_constraints cst) senv.env in let env = let linkinfo = Nativecode.link_info_of_dirpath lib.comp_name in Modops.add_linked_module mb linkinfo env in let env = Environ.add_native_symbols lib.comp_name lib.comp_natsymbs env in mp, { senv with env; modresolver = Mod_subst.add_delta_resolver mb.mod_delta senv.modresolver; required = DPmap.add lib.comp_name vodigest senv.required; loads = (mp,mb)::senv.loads; } (** {6 Safe typing } *) type judgment = Environ.unsafe_judgment let j_val j = j.Environ.uj_val let j_type j = j.Environ.uj_type let typing senv = Typeops.infer (env_of_senv senv) (** {6 Retroknowledge / native compiler } *) let register_inline kn senv = let open Environ in if not (evaluable_constant kn senv.env) then CErrors.user_err Pp.(str "Register inline: an evaluable constant is expected"); let env = senv.env in let cb = lookup_constant kn env in let cb = {cb with const_inline_code = true} in let env = add_constant kn cb env in { senv with env} let check_register_ind (type t) ind (r : t CPrimitives.prim_ind) env = let (mb,ob as spec) = Inductive.lookup_mind_specif env ind in let check_if b msg = if not b then CErrors.user_err ~hdr:"check_register_ind" msg in check_if (Int.equal (Array.length mb.mind_packets) 1) Pp.(str "A non mutual inductive is expected"); let is_monomorphic = function Monomorphic _ -> true | Polymorphic _ -> false in check_if (is_monomorphic mb.mind_universes) Pp.(str "A universe monomorphic inductive type is expected"); check_if (not @@ Inductive.is_private spec) Pp.(str "A non-private inductive type is expected"); let check_nparams n = check_if (Int.equal mb.mind_nparams n) Pp.(str "An inductive type with " ++ int n ++ str " parameters is expected") in let check_nconstr n = check_if (Int.equal (Array.length ob.mind_consnames) n) Pp.(str "an inductive type with " ++ int n ++ str " constructors is expected") in let check_name pos s = check_if (Id.equal ob.mind_consnames.(pos) (Id.of_string s)) Pp.(str"the " ++ int (pos + 1) ++ str "th constructor does not have the expected name: " ++ str s) in let check_type pos t = check_if (Constr.equal t ob.mind_user_lc.(pos)) Pp.(str"the " ++ int (pos + 1) ++ str "th constructor does not have the expected type") in let check_type_cte pos = check_type pos (Constr.mkRel 1) in match r with | CPrimitives.PIT_bool -> check_nparams 0; check_nconstr 2; check_name 0 "true"; check_type_cte 0; check_name 1 "false"; check_type_cte 1 | CPrimitives.PIT_carry -> check_nparams 1; check_nconstr 2; let test_type pos = let c = ob.mind_user_lc.(pos) in let s = Pp.(str"the " ++ int (pos + 1) ++ str "th constructor does not have the expected type") in check_if (Constr.isProd c) s; let (_,d,cd) = Constr.destProd c in check_if (Constr.is_Type d) s; check_if (Constr.equal (mkProd (Context.anonR,mkRel 1, mkApp (mkRel 3,[|mkRel 2|]))) cd) s in check_name 0 "C0"; test_type 0; check_name 1 "C1"; test_type 1; | CPrimitives.PIT_pair -> check_nparams 2; check_nconstr 1; check_name 0 "pair"; let c = ob.mind_user_lc.(0) in let s = Pp.str "the constructor does not have the expected type" in begin match Term.decompose_prod c with | ([_,b;_,a;_,_B;_,_A], codom) -> check_if (is_Type _A) s; check_if (is_Type _B) s; check_if (Constr.equal a (mkRel 2)) s; check_if (Constr.equal b (mkRel 2)) s; check_if (Constr.equal codom (mkApp (mkRel 5,[|mkRel 4; mkRel 3|]))) s | _ -> check_if false s end | CPrimitives.PIT_cmp -> check_nparams 0; check_nconstr 3; check_name 0 "Eq"; check_type_cte 0; check_name 1 "Lt"; check_type_cte 1; check_name 2 "Gt"; check_type_cte 2 | CPrimitives.PIT_f_cmp -> check_nconstr 4; check_name 0 "FEq"; check_type_cte 0; check_name 1 "FLt"; check_type_cte 1; check_name 2 "FGt"; check_type_cte 2; check_name 3 "FNotComparable"; check_type_cte 3 | CPrimitives.PIT_f_class -> check_nconstr 9; check_name 0 "PNormal"; check_type_cte 0; check_name 1 "NNormal"; check_type_cte 1; check_name 2 "PSubn"; check_type_cte 2; check_name 3 "NSubn"; check_type_cte 3; check_name 4 "PZero"; check_type_cte 4; check_name 5 "NZero"; check_type_cte 5; check_name 6 "PInf"; check_type_cte 6; check_name 7 "NInf"; check_type_cte 7; check_name 8 "NaN"; check_type_cte 8 let register_inductive ind prim senv = check_register_ind ind prim senv.env; let action = Retroknowledge.Register_ind(prim,ind) in add_retroknowledge action senv let add_constraints c = add_constraints (Now (Univ.ContextSet.add_constraints c Univ.ContextSet.empty)) (* NB: The next old comment probably refers to [propagate_loads] above. When a Require is done inside a module, we'll redo this require at the upper level after the module is ended, and so on. This is probably not a big deal anyway, since these Require's inside modules should be pretty rare. Maybe someday we could brutally forbid this tricky "feature"... *) (* we have an inefficiency: Since loaded files are added to the environment every time a module is closed, their components are calculated many times. This could be avoided in several ways: 1 - for each file create a dummy environment containing only this file's components, merge this environment with the global environment, and store for the future (instead of just its type) 2 - create "persistent modules" environment table in Environ add put loaded by side-effect once and for all (like it is done in OCaml). Would this be correct with respect to undo's and stuff ? *) let set_strategy k l e = { e with env = (Environ.set_oracle e.env (Conv_oracle.set_strategy (Environ.oracle e.env) k l)) } coq-8.11.0/kernel/uGraph.mli0000644000175000017500000000700113612664533015475 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t (** Don't use this in the kernel, it makes the system incomplete. *) type 'a check_function = t -> 'a -> 'a -> bool val check_leq : Universe.t check_function val check_eq : Universe.t check_function val check_eq_level : Level.t check_function (** The initial graph of universes: Prop < Set *) val initial_universes : t (** Check equality of instances w.r.t. a universe graph *) val check_eq_instances : Instance.t check_function (** {6 ... } *) (** Merge of constraints in a universes graph. The function [merge_constraints] merges a set of constraints in a given universes graph. It raises the exception [UniverseInconsistency] if the constraints are not satisfiable. *) val enforce_constraint : univ_constraint -> t -> t val merge_constraints : Constraint.t -> t -> t val check_constraint : t -> univ_constraint -> bool val check_constraints : Constraint.t -> t -> bool (** Picks an arbitrary set of constraints sufficient to ensure [u <= v]. *) val enforce_leq_alg : Universe.t -> Universe.t -> t -> Constraint.t * t (** Adds a universe to the graph, ensuring it is >= or > Set. @raise AlreadyDeclared if the level is already declared in the graph. *) exception AlreadyDeclared val add_universe : Level.t -> lbound:Level.t -> strict:bool -> t -> t (** Add a universe without (Prop,Set) <= u *) val add_universe_unconstrained : Level.t -> t -> t (** Check that the universe levels are declared. Otherwise @raise UndeclaredLevel l for the first undeclared level found. *) exception UndeclaredLevel of Univ.Level.t val check_declared_universes : t -> Univ.LSet.t -> unit (** {6 Pretty-printing of universes. } *) val pr_universes : (Level.t -> Pp.t) -> t -> Pp.t (** The empty graph of universes *) val empty_universes : t val sort_universes : t -> t (** [constraints_of_universes g] returns [csts] and [partition] where [csts] are the non-Eq constraints and [partition] is the partition of the universes into equivalence classes. *) val constraints_of_universes : t -> Constraint.t * LSet.t list val choose : (Level.t -> bool) -> t -> Level.t -> Level.t option (** [choose p g u] picks a universe verifying [p] and equal to [u] in [g]. *) (** [constraints_for ~kept g] returns the constraints about the universes [kept] in [g] up to transitivity. eg if [g] is [a <= b <= c] then [constraints_for ~kept:{a, c} g] is [a <= c]. *) val constraints_for : kept:LSet.t -> t -> Constraint.t val domain : t -> LSet.t (** Known universes *) val check_subtype : lbound:Level.t -> AUContext.t check_function (** [check_subtype univ ctx1 ctx2] checks whether [ctx2] is an instance of [ctx1]. *) (** {6 Dumping to a file } *) val dump_universes : (constraint_type -> Level.t -> Level.t -> unit) -> t -> unit (** {6 Debugging} *) val check_universes_invariants : t -> unit coq-8.11.0/kernel/cooking.mli0000644000175000017500000000300013612664533015673 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Opaqueproof.opaque result val cook_constr : Opaqueproof.cooking_info list -> (constr * unit Opaqueproof.delayed_universes) -> (constr * unit Opaqueproof.delayed_universes) val cook_inductive : Opaqueproof.cooking_info -> mutual_inductive_body -> Entries.mutual_inductive_entry (** {6 Utility functions used in module [Discharge]. } *) val expmod_constr : Opaqueproof.work_list -> constr -> constr coq-8.11.0/kernel/cooking.ml0000644000175000017500000003167113612664533015541 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constant.SyntacticOrd.equal c1 c2 | IndRef i1, IndRef i2 -> eq_syntactic_ind i1 i2 | ConstructRef c1, ConstructRef c2 -> eq_syntactic_constructor c1 c2 | _ -> false open Hashset.Combine let hash = function | ConstRef c -> combinesmall 1 (Constant.SyntacticOrd.hash c) | IndRef i -> combinesmall 2 (ind_syntactic_hash i) | ConstructRef c -> combinesmall 3 (constructor_syntactic_hash c) end module RefTable = Hashtbl.Make(RefHash) let instantiate_my_gr gr u = match gr with | ConstRef c -> mkConstU (c, u) | IndRef i -> mkIndU (i, u) | ConstructRef c -> mkConstructU (c, u) let share cache r (cstl,knl) = try RefTable.find cache r with Not_found -> let (u,l) = match r with | IndRef (kn,_i) -> Mindmap.find kn knl | ConstructRef ((kn,_i),_j) -> Mindmap.find kn knl | ConstRef cst -> Cmap.find cst cstl in let c = (u, Array.map mkVar l) in RefTable.add cache r c; c let share_univs cache r u l = let (u', args) = share cache r l in mkApp (instantiate_my_gr r (Instance.append u' u), args) let update_case_info cache ci modlist = try let (_u,l) = share cache (IndRef ci.ci_ind) modlist in { ci with ci_npar = ci.ci_npar + Array.length l } with Not_found -> ci let is_empty_modlist (cm, mm) = Cmap.is_empty cm && Mindmap.is_empty mm let expmod_constr cache modlist c = let share_univs = share_univs cache in let update_case_info = update_case_info cache in let rec substrec c = match kind c with | Case (ci,p,t,br) -> Constr.map substrec (mkCase (update_case_info ci modlist,p,t,br)) | Ind (ind,u) -> (try share_univs (IndRef ind) u modlist with | Not_found -> Constr.map substrec c) | Construct (cstr,u) -> (try share_univs (ConstructRef cstr) u modlist with | Not_found -> Constr.map substrec c) | Const (cst,u) -> (try share_univs (ConstRef cst) u modlist with | Not_found -> Constr.map substrec c) | Proj (p, c') -> let map cst npars = let _, newpars = Mindmap.find cst (snd modlist) in (cst, npars + Array.length newpars) in let p' = try Projection.map_npars map p with Not_found -> p in let c'' = substrec c' in if p == p' && c' == c'' then c else mkProj (p', c'') | _ -> Constr.map substrec c in if is_empty_modlist modlist then c else substrec c (** Transforms a named context into a rel context. Also returns the list of variables [id1 ... idn] that need to be replaced by [Rel 1 ... Rel n] to abstract a term that lived in that context. *) let abstract_context hyps = let fold decl (ctx, subst) = let id, decl = match decl with | NamedDecl.LocalDef (id, b, t) -> let b = Vars.subst_vars subst b in let t = Vars.subst_vars subst t in id, RelDecl.LocalDef (map_annot Name.mk_name id, b, t) | NamedDecl.LocalAssum (id, t) -> let t = Vars.subst_vars subst t in id, RelDecl.LocalAssum (map_annot Name.mk_name id, t) in (decl :: ctx, id.binder_name :: subst) in Context.Named.fold_outside fold hyps ~init:([], []) let abstract_constant_type t (hyps, subst) = let t = Vars.subst_vars subst t in List.fold_left (fun c d -> mkProd_wo_LetIn d c) t hyps let abstract_constant_body c (hyps, subst) = let c = Vars.subst_vars subst c in it_mkLambda_or_LetIn c hyps type recipe = { from : Opaqueproof.opaque constant_body; info : Opaqueproof.cooking_info } type inline = bool type 'opaque result = { cook_body : (constr Mod_subst.substituted, 'opaque) constant_def; cook_type : types; cook_universes : universes; cook_relevance : Sorts.relevance; cook_inline : inline; cook_context : Id.Set.t option; } let expmod_constr_subst cache modlist subst c = let subst = Univ.make_instance_subst subst in let c = expmod_constr cache modlist c in Vars.subst_univs_level_constr subst c let discharge_abstract_universe_context subst abs_ctx auctx = (** Given a named instance [subst := u₀ ... uₙ₋₁] together with an abstract context [auctx0 := 0 ... n - 1 |= C{0, ..., n - 1}] of the same length, and another abstract context relative to the former context [auctx := 0 ... m - 1 |= C'{u₀, ..., uₙ₋₁, 0, ..., m - 1}], construct the lifted abstract universe context [0 ... n - 1 n ... n + m - 1 |= C{0, ... n - 1} ∪ C'{0, ..., n - 1, n, ..., n + m - 1} ] together with the instance [u₀ ... uₙ₋₁ Var(0) ... Var (m - 1)]. *) if (Univ.Instance.is_empty subst) then (** Still need to take the union for the constraints between globals *) subst, (AUContext.union abs_ctx auctx) else let open Univ in let ainst = make_abstract_instance auctx in let subst = Instance.append subst ainst in let substf = make_instance_subst subst in let auctx = Univ.subst_univs_level_abstract_universe_context substf auctx in subst, (AUContext.union abs_ctx auctx) let lift_univs cb subst auctx0 = match cb.const_universes with | Monomorphic ctx -> assert (AUContext.is_empty auctx0); subst, (Monomorphic ctx) | Polymorphic auctx -> let subst, auctx = discharge_abstract_universe_context subst auctx0 auctx in subst, (Polymorphic auctx) let cook_constr { Opaqueproof.modlist ; abstract } (c, priv) = let cache = RefTable.create 13 in let abstract, usubst, abs_ctx = abstract in let usubst, priv = match priv with | Opaqueproof.PrivateMonomorphic () -> let () = assert (AUContext.is_empty abs_ctx) in let () = assert (Instance.is_empty usubst) in usubst, priv | Opaqueproof.PrivatePolymorphic (univs, ctx) -> let ainst = Instance.of_array (Array.init univs Level.var) in let usubst = Instance.append usubst ainst in let ctx = on_snd (Univ.subst_univs_level_constraints (Univ.make_instance_subst usubst)) ctx in let univs = univs + AUContext.size abs_ctx in usubst, Opaqueproof.PrivatePolymorphic (univs, ctx) in let expmod = expmod_constr_subst cache modlist usubst in let hyps = Context.Named.map expmod abstract in let hyps = abstract_context hyps in let c = abstract_constant_body (expmod c) hyps in (c, priv) let cook_constr infos c = let fold info c = cook_constr info c in List.fold_right fold infos c let cook_constant { from = cb; info } = let { Opaqueproof.modlist; abstract } = info in let cache = RefTable.create 13 in let abstract, usubst, abs_ctx = abstract in let usubst, univs = lift_univs cb usubst abs_ctx in let expmod = expmod_constr_subst cache modlist usubst in let hyps0 = Context.Named.map expmod abstract in let hyps = abstract_context hyps0 in let map c = abstract_constant_body (expmod c) hyps in let body = match cb.const_body with | Undef _ as x -> x | Def cs -> Def (Mod_subst.from_val (map (Mod_subst.force_constr cs))) | OpaqueDef o -> OpaqueDef (Opaqueproof.discharge_opaque info o) | Primitive _ -> CErrors.anomaly (Pp.str "Primitives cannot be cooked") in let const_hyps = Id.Set.diff (Context.Named.to_vars cb.const_hyps) (Context.Named.to_vars hyps0) in let typ = abstract_constant_type (expmod cb.const_type) hyps in { cook_body = body; cook_type = typ; cook_universes = univs; cook_relevance = cb.const_relevance; cook_inline = cb.const_inline_code; cook_context = Some const_hyps; } (* let cook_constant_key = CProfile.declare_profile "cook_constant" *) (* let cook_constant = CProfile.profile2 cook_constant_key cook_constant *) (********************************) (* Discharging mutual inductive *) (* Replace Var(y1)..Var(yq):C1..Cq |- Ij:Bj Var(y1)..Var(yq):C1..Cq; I1..Ip:B1..Bp |- ci : Ti by |- Ij: (y1..yq:C1..Cq)Bj I1..Ip:(B1 y1..yq)..(Bp y1..yq) |- ci : (y1..yq:C1..Cq)Ti[Ij:=(Ij y1..yq)] *) let it_mkNamedProd_wo_LetIn b d = List.fold_left (fun c d -> mkNamedProd_wo_LetIn d c) b d let abstract_inductive decls nparamdecls inds = let open Entries in let ntyp = List.length inds in let ndecls = Context.Named.length decls in let args = Context.Named.to_instance mkVar (List.rev decls) in let args = Array.of_list args in let subs = List.init ntyp (fun k -> lift ndecls (mkApp(mkRel (k+1),args))) in let inds' = List.map (function (tname,arity,template,cnames,lc) -> let lc' = List.map (Vars.substl subs) lc in let lc'' = List.map (fun b -> it_mkNamedProd_wo_LetIn b decls) lc' in let arity' = it_mkNamedProd_wo_LetIn arity decls in (tname,arity',template,cnames,lc'')) inds in let nparamdecls' = nparamdecls + Array.length args in (* To be sure to be the same as before, should probably be moved to cook_inductive *) let params' = let (_,arity,_,_,_) = List.hd inds' in let (params,_) = decompose_prod_n_assum nparamdecls' arity in params in let ind'' = List.map (fun (a,arity,template,c,lc) -> let _, short_arity = decompose_prod_n_assum nparamdecls' arity in let shortlc = List.map (fun c -> snd (decompose_prod_n_assum nparamdecls' c)) lc in { mind_entry_typename = a; mind_entry_arity = short_arity; mind_entry_template = template; mind_entry_consnames = c; mind_entry_lc = shortlc }) inds' in (params',ind'') let refresh_polymorphic_type_of_inductive (_,mip) = match mip.mind_arity with | RegularArity s -> s.mind_user_arity, false | TemplateArity ar -> let ctx = List.rev mip.mind_arity_ctxt in mkArity (List.rev ctx, Sorts.sort_of_univ ar.template_level), true let dummy_variance = let open Entries in function | Monomorphic_entry _ -> assert false | Polymorphic_entry (_,uctx) -> Array.make (Univ.UContext.size uctx) Univ.Variance.Irrelevant let cook_inductive { Opaqueproof.modlist; abstract } mib = let open Entries in let (section_decls, subst, abs_uctx) = abstract in let nparamdecls = Context.Rel.length mib.mind_params_ctxt in let subst, ind_univs = match mib.mind_universes with | Monomorphic ctx -> Univ.empty_level_subst, Monomorphic_entry ctx | Polymorphic auctx -> let subst, auctx = discharge_abstract_universe_context subst abs_uctx auctx in let subst = Univ.make_instance_subst subst in let nas = Univ.AUContext.names auctx in let auctx = Univ.AUContext.repr auctx in subst, Polymorphic_entry (nas, auctx) in let variance = match mib.mind_variance with | None -> None | Some _ -> Some (dummy_variance ind_univs) in let cache = RefTable.create 13 in let discharge c = Vars.subst_univs_level_constr subst (expmod_constr cache modlist c) in let inds = Array.map_to_list (fun mip -> let ty, template = refresh_polymorphic_type_of_inductive (mib,mip) in let arity = discharge ty in let lc = Array.map discharge mip.mind_user_lc in (mip.mind_typename, arity, template, Array.to_list mip.mind_consnames, Array.to_list lc)) mib.mind_packets in let section_decls' = Context.Named.map discharge section_decls in let (params',inds') = abstract_inductive section_decls' nparamdecls inds in let record = match mib.mind_record with | PrimRecord info -> Some (Some (Array.map (fun (x,_,_,_) -> x) info)) | FakeRecord -> Some None | NotRecord -> None in { mind_entry_record = record; mind_entry_finite = mib.mind_finite; mind_entry_params = params'; mind_entry_inds = inds'; mind_entry_private = mib.mind_private; mind_entry_variance = variance; mind_entry_universes = ind_univs } let expmod_constr modlist c = expmod_constr (RefTable.create 13) modlist c coq-8.11.0/kernel/vm.mli0000644000175000017500000000264613612664533014703 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit val reduce_fix : int -> vfix -> vfun array * values array (** bodies , types *) val reduce_cofix : int -> vcofix -> values array * values array (** bodies , types *) val type_of_switch : vswitch -> values val branch_of_switch : int -> vswitch -> (int * values) array val reduce_fun : int -> vfun -> values (** [decompose_vfun2 k f1 f2] takes two functions [f1] and [f2] at current DeBruijn level [k], with [n] lambdas in common, returns [n] and the reduced bodies under those lambdas. *) val decompose_vfun2 : int -> vfun -> vfun -> int * values * values (** Apply a value *) val apply_whd : int -> whd -> values coq-8.11.0/kernel/reduction.mli0000644000175000017500000001263213612664533016251 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr -> constr val whd_all : env -> constr -> constr val whd_allnolet : env -> constr -> constr val whd_betaiota : env -> constr -> constr val nf_betaiota : env -> constr -> constr (*********************************************************************** s conversion functions *) exception NotConvertible type 'a kernel_conversion_function = env -> 'a -> 'a -> unit type 'a extended_conversion_function = ?l2r:bool -> ?reds:TransparentState.t -> env -> ?evars:((existential->constr option) * UGraph.t) -> 'a -> 'a -> unit type conv_pb = CONV | CUMUL type 'a universe_compare = { (* Might raise NotConvertible *) compare_sorts : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a; compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a; compare_cumul_instances : conv_pb -> Univ.Variance.t array -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a } type 'a universe_state = 'a * 'a universe_compare type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a -> 'b type 'a infer_conversion_function = env -> UGraph.t -> 'a -> 'a -> Univ.Constraint.t val get_cumulativity_constraints : conv_pb -> Univ.Variance.t array -> Univ.Instance.t -> Univ.Instance.t -> Univ.Constraint.t val inductive_cumulativity_arguments : (Declarations.mutual_inductive_body * int) -> int val constructor_cumulativity_arguments : (Declarations.mutual_inductive_body * int * int) -> int val sort_cmp_universes : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a * 'a universe_compare -> 'a * 'a universe_compare (* [flex] should be true for constants, false for inductive types and constructors. *) val convert_instances : flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a * 'a universe_compare -> 'a * 'a universe_compare (** These two never raise UnivInconsistency, inferred_universes just gathers the constraints. *) val checked_universes : UGraph.t universe_compare val inferred_universes : (UGraph.t * Univ.Constraint.t) universe_compare (** These two functions can only raise NotConvertible *) val conv : constr extended_conversion_function val conv_leq : types extended_conversion_function (** These conversion functions are used by module subtyping, which needs to infer universe constraints inside the kernel *) val infer_conv : ?l2r:bool -> ?evars:(existential->constr option) -> ?ts:TransparentState.t -> constr infer_conversion_function val infer_conv_leq : ?l2r:bool -> ?evars:(existential->constr option) -> ?ts:TransparentState.t -> types infer_conversion_function (** Depending on the universe state functions, this might raise [UniverseInconsistency] in addition to [NotConvertible] (for better error messages). *) val generic_conv : conv_pb -> l2r:bool -> (existential->constr option) -> TransparentState.t -> (constr,'a) generic_conversion_function val default_conv : conv_pb -> ?l2r:bool -> types kernel_conversion_function val default_conv_leq : ?l2r:bool -> types kernel_conversion_function (************************************************************************) (** Builds an application node, reducing beta redexes it may produce. *) val beta_applist : constr -> constr list -> constr (** Builds an application node, reducing beta redexes it may produce. *) val beta_appvect : constr -> constr array -> constr (** Builds an application node, reducing beta redexe it may produce. *) val beta_app : constr -> constr -> constr (** Pseudo-reduction rule Prod(x,A,B) a --> B[x\a] *) val hnf_prod_applist : env -> types -> constr list -> types (** In [hnf_prod_applist_assum n c args], [c] is supposed to (whd-)reduce to the form [∀Γ.t] with [Γ] of length [n] and possibly with let-ins; it returns [t] with the assumptions of [Γ] instantiated by [args] and the local definitions of [Γ] expanded. *) val hnf_prod_applist_assum : env -> int -> types -> constr list -> types (** Compatibility alias for Term.lambda_appvect_assum *) val betazeta_appvect : int -> constr -> constr array -> constr (*********************************************************************** s Recognizing products and arities modulo reduction *) val dest_prod : env -> types -> Constr.rel_context * types val dest_prod_assum : env -> types -> Constr.rel_context * types val dest_lam : env -> constr -> Constr.rel_context * constr val dest_lam_assum : env -> constr -> Constr.rel_context * constr exception NotArity val dest_arity : env -> types -> Term.arity (* raises NotArity if not an arity *) val is_arity : env -> types -> bool val eta_expand : env -> constr -> types -> constr coq-8.11.0/kernel/nativevalues.ml0000644000175000017500000004237113612664533016615 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t type accumulator (* = t (* a block [0:code;atom;arguments] *) *) type tag = int type arity = int type reloc_table = (tag * arity) array type annot_sw = { asw_ind : inductive; asw_ci : case_info; asw_reloc : reloc_table; asw_finite : bool; asw_prefix : string } (* We compare only what is relevant for generation of ml code *) let eq_annot_sw asw1 asw2 = eq_ind asw1.asw_ind asw2.asw_ind && String.equal asw1.asw_prefix asw2.asw_prefix open Hashset.Combine let hash_annot_sw asw = combine (ind_hash asw.asw_ind) (String.hash asw.asw_prefix) type sort_annot = string * int type rec_pos = int array let eq_rec_pos = Array.equal Int.equal type atom = | Arel of int | Aconstant of pconstant | Aind of pinductive | Asort of Sorts.t | Avar of Id.t | Acase of annot_sw * accumulator * t * (t -> t) | Afix of t array * t array * rec_pos * int (* types, bodies, rec_pos, pos *) | Acofix of t array * t array * int * t | Acofixe of t array * t array * int * t | Aprod of Name.t * t * (t -> t) | Ameta of metavariable * t | Aevar of Evar.t * t array | Aproj of (inductive * int) * accumulator type symbol = | SymbValue of t | SymbSort of Sorts.t | SymbName of Name.t | SymbConst of Constant.t | SymbMatch of annot_sw | SymbInd of inductive | SymbMeta of metavariable | SymbEvar of Evar.t | SymbLevel of Univ.Level.t | SymbProj of (inductive * int) type symbols = symbol array let empty_symbols = [| |] let accumulate_tag = 0 (** Unique pointer used to drive the accumulator function *) let ret_accu = Obj.repr (ref ()) type accu_val = { mutable acc_atm : atom; acc_arg : Obj.t list } let mk_accu (a : atom) : t = let rec accumulate data x = if x == ret_accu then Obj.repr data else let data = { data with acc_arg = x :: data.acc_arg } in let ans = Obj.repr (accumulate data) in let () = Obj.set_tag ans accumulate_tag [@ocaml.alert "--deprecated"] in ans in let acc = { acc_atm = a; acc_arg = [] } in let ans = Obj.repr (accumulate acc) in (** FIXME: use another representation for accumulators, this causes naked pointers. *) let () = Obj.set_tag ans accumulate_tag [@ocaml.alert "--deprecated"] in (Obj.obj ans : t) let get_accu (k : accumulator) = (Obj.magic k : Obj.t -> accu_val) ret_accu let mk_rel_accu i = mk_accu (Arel i) let rel_tbl_size = 100 let rel_tbl = Array.init rel_tbl_size mk_rel_accu let mk_rel_accu i = if i < rel_tbl_size then rel_tbl.(i) else mk_rel_accu i let mk_rels_accu lvl len = Array.init len (fun i -> mk_rel_accu (lvl + i)) let napply (f:t) (args: t array) = Array.fold_left (fun f a -> f a) f args let mk_constant_accu kn u = mk_accu (Aconstant (kn,Univ.Instance.of_array u)) let mk_ind_accu ind u = mk_accu (Aind (ind,Univ.Instance.of_array u)) let mk_sort_accu s u = let open Sorts in match s with | SProp | Prop | Set -> mk_accu (Asort s) | Type s -> let u = Univ.Instance.of_array u in let s = Sorts.sort_of_univ (Univ.subst_instance_universe u s) in mk_accu (Asort s) let mk_var_accu id = mk_accu (Avar id) let mk_sw_accu annot c p ac = mk_accu (Acase(annot,c,p,ac)) let mk_prod_accu s dom codom = mk_accu (Aprod (s,dom,codom)) let mk_meta_accu mv ty = mk_accu (Ameta (mv,ty)) let mk_evar_accu ev args = mk_accu (Aevar (ev, args)) let mk_proj_accu kn c = mk_accu (Aproj (kn,c)) let atom_of_accu (k:accumulator) = (get_accu k).acc_atm let set_atom_of_accu (k:accumulator) (a:atom) = (get_accu k).acc_atm <- a let accu_nargs (k:accumulator) = List.length (get_accu k).acc_arg let args_of_accu (k:accumulator) = let acc = (get_accu k).acc_arg in (Obj.magic (Array.of_list acc) : t array) let mk_fix_accu rec_pos pos types bodies = mk_accu (Afix(types,bodies,rec_pos, pos)) let mk_cofix_accu pos types norm = mk_accu (Acofix(types,norm,pos,(Obj.magic 0 : t))) let upd_cofix (cofix :t) (cofix_fun : t) = let atom = atom_of_accu (Obj.magic cofix) in match atom with | Acofix (typ,norm,pos,_) -> set_atom_of_accu (Obj.magic cofix) (Acofix(typ,norm,pos,cofix_fun)) | _ -> assert false let force_cofix (cofix : t) = let accu = (Obj.magic cofix : accumulator) in let atom = atom_of_accu accu in match atom with | Acofix(typ,norm,pos,f) -> let args = args_of_accu accu in let f = Array.fold_right (fun arg f -> f arg) args f in let v = f (Obj.magic ()) in set_atom_of_accu accu (Acofixe(typ,norm,pos,v)); v | Acofixe(_,_,_,v) -> v | _ -> cofix let mk_const tag = Obj.magic tag let mk_block tag args = let nargs = Array.length args in let r = Obj.new_block tag nargs in for i = 0 to nargs - 1 do Obj.set_field r i (Obj.magic args.(i)) done; (Obj.magic r : t) (* Two instances of dummy_value should not be pointer equal, otherwise comparing them as terms would succeed *) let dummy_value : unit -> t = fun () _ -> anomaly ~label:"native" (Pp.str "Evaluation failed.") let cast_accu v = (Obj.magic v:accumulator) [@@ocaml.inline always] let mk_int (x : int) = (Obj.magic x : t) [@@ocaml.inline always] (* Coq's booleans are reversed... *) let mk_bool (b : bool) = (Obj.magic (not b) : t) [@@ocaml.inline always] let mk_uint (x : Uint63.t) = (Obj.magic x : t) [@@ocaml.inline always] let mk_float (x : Float64.t) = (Obj.magic x : t) [@@ocaml.inline always] type block let block_size (b:block) = Obj.size (Obj.magic b) let block_field (b:block) i = (Obj.magic (Obj.field (Obj.magic b) i) : t) let block_tag (b:block) = Obj.tag (Obj.magic b) type kind_of_value = | Vaccu of accumulator | Vfun of (t -> t) | Vconst of int | Vint64 of int64 | Vfloat64 of float | Vblock of block let kind_of_value (v:t) = let o = Obj.repr v in if Obj.is_int o then Vconst (Obj.magic v) else if Obj.tag o == Obj.double_tag then Vfloat64 (Obj.magic v) else let tag = Obj.tag o in if Int.equal tag accumulate_tag then Vaccu (Obj.magic v) else if Int.equal tag Obj.custom_tag then Vint64 (Obj.magic v) else if Int.equal tag Obj.double_tag then Vfloat64 (Obj.magic v) else if (tag < Obj.lazy_tag) then Vblock (Obj.magic v) else (* assert (tag = Obj.closure_tag || tag = Obj.infix_tag); or ??? what is 1002*) Vfun v (** Support for machine integers *) let is_int (x:t) = let o = Obj.repr x in Obj.is_int o || Int.equal (Obj.tag o) Obj.custom_tag [@@ocaml.inline always] let val_to_int (x:t) = (Obj.magic x : int) [@@ocaml.inline always] let to_uint (x:t) = (Obj.magic x : Uint63.t) [@@ocaml.inline always] let no_check_head0 x = mk_uint (Uint63.head0 (to_uint x)) [@@ocaml.inline always] let head0 accu x = if is_int x then no_check_head0 x else accu x let no_check_tail0 x = mk_uint (Uint63.tail0 (to_uint x)) [@@ocaml.inline always] let tail0 accu x = if is_int x then no_check_tail0 x else accu x let no_check_add x y = mk_uint (Uint63.add (to_uint x) (to_uint y)) [@@ocaml.inline always] let add accu x y = if is_int x && is_int y then no_check_add x y else accu x y let no_check_sub x y = mk_uint (Uint63.sub (to_uint x) (to_uint y)) [@@ocaml.inline always] let sub accu x y = if is_int x && is_int y then no_check_sub x y else accu x y let no_check_mul x y = mk_uint (Uint63.mul (to_uint x) (to_uint y)) [@@ocaml.inline always] let mul accu x y = if is_int x && is_int y then no_check_mul x y else accu x y let no_check_div x y = mk_uint (Uint63.div (to_uint x) (to_uint y)) [@@ocaml.inline always] let div accu x y = if is_int x && is_int y then no_check_div x y else accu x y let no_check_rem x y = mk_uint (Uint63.rem (to_uint x) (to_uint y)) [@@ocaml.inline always] let rem accu x y = if is_int x && is_int y then no_check_rem x y else accu x y let no_check_l_sr x y = mk_uint (Uint63.l_sr (to_uint x) (to_uint y)) [@@ocaml.inline always] let l_sr accu x y = if is_int x && is_int y then no_check_l_sr x y else accu x y let no_check_l_sl x y = mk_uint (Uint63.l_sl (to_uint x) (to_uint y)) [@@ocaml.inline always] let l_sl accu x y = if is_int x && is_int y then no_check_l_sl x y else accu x y let no_check_l_and x y = mk_uint (Uint63.l_and (to_uint x) (to_uint y)) [@@ocaml.inline always] let l_and accu x y = if is_int x && is_int y then no_check_l_and x y else accu x y let no_check_l_xor x y = mk_uint (Uint63.l_xor (to_uint x) (to_uint y)) [@@ocaml.inline always] let l_xor accu x y = if is_int x && is_int y then no_check_l_xor x y else accu x y let no_check_l_or x y = mk_uint (Uint63.l_or (to_uint x) (to_uint y)) [@@ocaml.inline always] let l_or accu x y = if is_int x && is_int y then no_check_l_or x y else accu x y [@@@ocaml.warning "-37"] type coq_carry = | Caccu of t | C0 of t | C1 of t type coq_pair = | Paccu of t | PPair of t * t let mkCarry b i = if b then (Obj.magic (C1(mk_uint i)):t) else (Obj.magic (C0(mk_uint i)):t) let no_check_addc x y = let s = Uint63.add (to_uint x) (to_uint y) in mkCarry (Uint63.lt s (to_uint x)) s [@@ocaml.inline always] let addc accu x y = if is_int x && is_int y then no_check_addc x y else accu x y let no_check_subc x y = let s = Uint63.sub (to_uint x) (to_uint y) in mkCarry (Uint63.lt (to_uint x) (to_uint y)) s [@@ocaml.inline always] let subc accu x y = if is_int x && is_int y then no_check_subc x y else accu x y let no_check_addCarryC x y = let s = Uint63.add (Uint63.add (to_uint x) (to_uint y)) (Uint63.of_int 1) in mkCarry (Uint63.le s (to_uint x)) s [@@ocaml.inline always] let addCarryC accu x y = if is_int x && is_int y then no_check_addCarryC x y else accu x y let no_check_subCarryC x y = let s = Uint63.sub (Uint63.sub (to_uint x) (to_uint y)) (Uint63.of_int 1) in mkCarry (Uint63.le (to_uint x) (to_uint y)) s [@@ocaml.inline always] let subCarryC accu x y = if is_int x && is_int y then no_check_subCarryC x y else accu x y let of_pair (x, y) = (Obj.magic (PPair(mk_uint x, mk_uint y)):t) [@@ocaml.inline always] let no_check_mulc x y = of_pair (Uint63.mulc (to_uint x) (to_uint y)) [@@ocaml.inline always] let mulc accu x y = if is_int x && is_int y then no_check_mulc x y else accu x y let no_check_diveucl x y = let i1, i2 = to_uint x, to_uint y in of_pair(Uint63.div i1 i2, Uint63.rem i1 i2) [@@ocaml.inline always] let diveucl accu x y = if is_int x && is_int y then no_check_diveucl x y else accu x y let no_check_div21 x y z = let i1, i2, i3 = to_uint x, to_uint y, to_uint z in of_pair (Uint63.div21 i1 i2 i3) [@@ocaml.inline always] let div21 accu x y z = if is_int x && is_int y && is_int z then no_check_div21 x y z else accu x y z let no_check_addMulDiv x y z = let p, i, j = to_uint x, to_uint y, to_uint z in mk_uint (Uint63.addmuldiv p i j) [@@ocaml.inline always] let addMulDiv accu x y z = if is_int x && is_int y && is_int z then no_check_addMulDiv x y z else accu x y z [@@@ocaml.warning "-34"] type coq_bool = | Baccu of t | Btrue | Bfalse type coq_cmp = | CmpAccu of t | CmpEq | CmpLt | CmpGt let no_check_eq x y = mk_bool (Uint63.equal (to_uint x) (to_uint y)) [@@ocaml.inline always] let eq accu x y = if is_int x && is_int y then no_check_eq x y else accu x y let no_check_lt x y = mk_bool (Uint63.lt (to_uint x) (to_uint y)) [@@ocaml.inline always] let lt accu x y = if is_int x && is_int y then no_check_lt x y else accu x y let no_check_le x y = mk_bool (Uint63.le (to_uint x) (to_uint y)) [@@ocaml.inline always] let le accu x y = if is_int x && is_int y then no_check_le x y else accu x y let no_check_compare x y = match Uint63.compare (to_uint x) (to_uint y) with | x when x < 0 -> (Obj.magic CmpLt:t) | 0 -> (Obj.magic CmpEq:t) | _ -> (Obj.magic CmpGt:t) let compare accu x y = if is_int x && is_int y then no_check_compare x y else accu x y let print x = Printf.fprintf stderr "%s" (Uint63.to_string (to_uint x)); flush stderr; x (** Support for machine floating point values *) external is_float : t -> bool = "coq_is_double" [@@noalloc] let to_float (x:t) = (Obj.magic x : Float64.t) [@@ocaml.inline always] let no_check_fopp x = mk_float (Float64.opp (to_float x)) [@@ocaml.inline always] let fopp accu x = if is_float x then no_check_fopp x else accu x let no_check_fabs x = mk_float (Float64.abs (to_float x)) [@@ocaml.inline always] let fabs accu x = if is_float x then no_check_fabs x else accu x let no_check_feq x y = mk_bool (Float64.eq (to_float x) (to_float y)) let feq accu x y = if is_float x && is_float y then no_check_feq x y else accu x y let no_check_flt x y = mk_bool (Float64.lt (to_float x) (to_float y)) let flt accu x y = if is_float x && is_float y then no_check_flt x y else accu x y let no_check_fle x y = mk_bool (Float64.le (to_float x) (to_float y)) let fle accu x y = if is_float x && is_float y then no_check_fle x y else accu x y type coq_fcmp = | CFcmpAccu of t | CFcmpEq | CFcmpLt | CFcmpGt | CFcmpNotComparable let no_check_fcompare x y = let c = Float64.compare (to_float x) (to_float y) in (Obj.magic c:t) [@@ocaml.inline always] let fcompare accu x y = if is_float x && is_float y then no_check_fcompare x y else accu x y type coq_fclass = | CFclassAccu of t | CFclassPNormal | CFclassNNormal | CFclassPSubn | CFclassNSubn | CFclassPZero | CFclassNZero | CFclassPInf | CFclassNInf | CFclassNaN let no_check_fclassify x = let c = Float64.classify (to_float x) in (Obj.magic c:t) [@@ocaml.inline always] let fclassify accu x = if is_float x then no_check_fclassify x else accu x let no_check_fadd x y = mk_float (Float64.add (to_float x) (to_float y)) [@@ocaml.inline always] let fadd accu x y = if is_float x && is_float y then no_check_fadd x y else accu x y let no_check_fsub x y = mk_float (Float64.sub (to_float x) (to_float y)) [@@ocaml.inline always] let fsub accu x y = if is_float x && is_float y then no_check_fsub x y else accu x y let no_check_fmul x y = mk_float (Float64.mul (to_float x) (to_float y)) [@@ocaml.inline always] let fmul accu x y = if is_float x && is_float y then no_check_fmul x y else accu x y let no_check_fdiv x y = mk_float (Float64.div (to_float x) (to_float y)) [@@ocaml.inline always] let fdiv accu x y = if is_float x && is_float y then no_check_fdiv x y else accu x y let no_check_fsqrt x = mk_float (Float64.sqrt (to_float x)) [@@ocaml.inline always] let fsqrt accu x = if is_float x then no_check_fsqrt x else accu x let no_check_float_of_int x = mk_float (Float64.of_int63 (to_uint x)) [@@ocaml.inline always] let float_of_int accu x = if is_int x then no_check_float_of_int x else accu x let no_check_normfr_mantissa x = mk_uint (Float64.normfr_mantissa (to_float x)) [@@ocaml.inline always] let normfr_mantissa accu x = if is_float x then no_check_normfr_mantissa x else accu x let no_check_frshiftexp x = let f, e = Float64.frshiftexp (to_float x) in (Obj.magic (PPair(mk_float f, mk_uint e)):t) [@@ocaml.inline always] let frshiftexp accu x = if is_float x then no_check_frshiftexp x else accu x let no_check_ldshiftexp x e = mk_float (Float64.ldshiftexp (to_float x) (to_uint e)) [@@ocaml.inline always] let ldshiftexp accu x e = if is_float x && is_int e then no_check_ldshiftexp x e else accu x e let no_check_next_up x = mk_float (Float64.next_up (to_float x)) [@@ocaml.inline always] let next_up accu x = if is_float x then no_check_next_up x else accu x let no_check_next_down x = mk_float (Float64.next_down (to_float x)) [@@ocaml.inline always] let next_down accu x = if is_float x then no_check_next_down x else accu x let hobcnv = Array.init 256 (fun i -> Printf.sprintf "%02x" i) let bohcnv = Array.init 256 (fun i -> i - (if 0x30 <= i then 0x30 else 0) - (if 0x41 <= i then 0x7 else 0) - (if 0x61 <= i then 0x20 else 0)) let hex_of_bin ch = hobcnv.(int_of_char ch) let bin_of_hex s = char_of_int (bohcnv.(int_of_char s.[0]) * 16 + bohcnv.(int_of_char s.[1])) let str_encode expr = let mshl_expr = Marshal.to_string expr [] in let payload = Buffer.create (String.length mshl_expr * 2) in String.iter (fun c -> Buffer.add_string payload (hex_of_bin c)) mshl_expr; Buffer.contents payload let str_decode s = let mshl_expr_len = String.length s / 2 in let mshl_expr = Buffer.create mshl_expr_len in let buf = Bytes.create 2 in for i = 0 to mshl_expr_len - 1 do Bytes.blit_string s (2*i) buf 0 2; Buffer.add_char mshl_expr (bin_of_hex (Bytes.to_string buf)) done; Marshal.from_bytes (Buffer.to_bytes mshl_expr) 0 coq-8.11.0/kernel/typeops.mli0000644000175000017500000001155113612664533015757 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr -> unsafe_judgment val infer_type : env -> types -> unsafe_type_judgment val check_context : env -> Constr.rel_context -> env * Constr.rel_context (** {6 Basic operations of the typing machine. } *) (** If [j] is the judgement {% $ %}c:t{% $ %}, then [assumption_of_judgement env j] returns the type {% $ %}c{% $ %}, checking that {% $ %}t{% $ %} is a sort. *) val assumption_of_judgment : env -> unsafe_judgment -> Sorts.relevance val type_judgment : env -> unsafe_judgment -> unsafe_type_judgment val check_binder_annot : Sorts.t -> Name.t Context.binder_annot -> Name.t Context.binder_annot (** {6 Type of sorts. } *) val type1 : types val type_of_sort : Sorts.t -> types val judge_of_prop : unsafe_judgment val judge_of_set : unsafe_judgment val judge_of_type : Universe.t -> unsafe_judgment (** {6 Type of a bound variable. } *) val type_of_relative : env -> int -> types val judge_of_relative : env -> int -> unsafe_judgment (** {6 Type of variables } *) val type_of_variable : env -> variable -> types val judge_of_variable : env -> variable -> unsafe_judgment (** {6 type of a constant } *) val type_of_constant_in : env -> pconstant -> types val judge_of_constant : env -> pconstant -> unsafe_judgment (** {6 type of an applied projection } *) val judge_of_projection : env -> Projection.t -> unsafe_judgment -> unsafe_judgment (** {6 Type of application. } *) val judge_of_apply : env -> unsafe_judgment -> unsafe_judgment array -> unsafe_judgment (** {6 Type of an abstraction. } *) (* val judge_of_abstraction : *) (* env -> Name.t -> unsafe_type_judgment -> unsafe_judgment *) (* -> unsafe_judgment *) (** {6 Type of a product. } *) val sort_of_product : env -> Sorts.t -> Sorts.t -> Sorts.t val type_of_product : env -> Name.t Context.binder_annot -> Sorts.t -> Sorts.t -> types (* val judge_of_product : *) (* env -> Name.t -> unsafe_type_judgment -> unsafe_type_judgment *) (* -> unsafe_judgment *) (** s Type of a let in. *) (* val judge_of_letin : *) (* env -> Name.t -> unsafe_judgment -> unsafe_type_judgment -> unsafe_judgment *) (* -> unsafe_judgment *) (** {6 Type of a cast. } *) val judge_of_cast : env -> unsafe_judgment -> cast_kind -> unsafe_type_judgment -> unsafe_judgment (** {6 Inductive types. } *) val judge_of_inductive : env -> inductive puniverses -> unsafe_judgment val judge_of_constructor : env -> constructor puniverses -> unsafe_judgment (** {6 Type of Cases. } *) val judge_of_case : env -> case_info -> unsafe_judgment -> unsafe_judgment -> unsafe_judgment array -> unsafe_judgment (** {6 Type of global references. } *) val type_of_global_in_context : env -> GlobRef.t -> types * Univ.AUContext.t (** Returns the type of the global reference, by creating a fresh instance of polymorphic references and computing their instantiated universe context. The type should not be used without pushing it's universe context in the environmnent of usage. For non-universe-polymorphic constants, it does not matter. *) (** {6 Miscellaneous. } *) (** Check that hyps are included in env and fails with error otherwise *) val check_hyps_inclusion : env -> ?evars:((existential->constr option) * UGraph.t) -> ('a -> constr) -> 'a -> Constr.named_context -> unit val check_primitive_type : env -> CPrimitives.op_or_type -> types -> unit (** Types for primitives *) val type_of_int : env -> types val judge_of_int : env -> Uint63.t -> unsafe_judgment val type_of_float : env -> types val judge_of_float : env -> Float64.t -> unsafe_judgment val type_of_prim_type : env -> CPrimitives.prim_type -> types val type_of_prim : env -> CPrimitives.t -> types val warn_bad_relevance_name : string (** Allow the checker to make this warning into an error. *) coq-8.11.0/kernel/nativeconv.ml0000644000175000017500000001666313612664533016270 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* let v = mk_rel_accu lvl in conv_val env CONV (lvl+1) (f1 v) (f2 v) cu | Vfun _f1, _ -> conv_val env CONV lvl v1 (fun x -> v2 x) cu | _, Vfun _f2 -> conv_val env CONV lvl (fun x -> v1 x) v2 cu | Vaccu k1, Vaccu k2 -> conv_accu env pb lvl k1 k2 cu | Vconst i1, Vconst i2 -> if Int.equal i1 i2 then cu else raise NotConvertible | Vint64 i1, Vint64 i2 -> if Int64.equal i1 i2 then cu else raise NotConvertible | Vfloat64 f1, Vfloat64 f2 -> if Float64.(equal (of_float f1) (of_float f2)) then cu else raise NotConvertible | Vblock b1, Vblock b2 -> let n1 = block_size b1 in let n2 = block_size b2 in if not (Int.equal (block_tag b1) (block_tag b2)) || not (Int.equal n1 n2) then raise NotConvertible; let rec aux lvl max b1 b2 i cu = if Int.equal i max then conv_val env CONV lvl (block_field b1 i) (block_field b2 i) cu else let cu = conv_val env CONV lvl (block_field b1 i) (block_field b2 i) cu in aux lvl max b1 b2 (i+1) cu in aux lvl (n1-1) b1 b2 0 cu | (Vaccu _ | Vconst _ | Vint64 _ | Vfloat64 _ | Vblock _), _ -> raise NotConvertible and conv_accu env pb lvl k1 k2 cu = let n1 = accu_nargs k1 in let n2 = accu_nargs k2 in if not (Int.equal n1 n2) then raise NotConvertible; if Int.equal n1 0 then conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu else let cu = conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu in Array.fold_right2 (conv_val env CONV lvl) (args_of_accu k1) (args_of_accu k2) cu and conv_atom env pb lvl a1 a2 cu = if a1 == a2 then cu else match a1, a2 with | Ameta (m1,_), Ameta (m2,_) -> if Int.equal m1 m2 then cu else raise NotConvertible | Aevar (ev1, args1), Aevar (ev2, args2) -> if Evar.equal ev1 ev2 then Array.fold_right2 (conv_val env CONV lvl) args1 args2 cu else raise NotConvertible | Arel i1, Arel i2 -> if Int.equal i1 i2 then cu else raise NotConvertible | Aind (ind1,u1), Aind (ind2,u2) -> if eq_ind ind1 ind2 then convert_instances ~flex:false u1 u2 cu else raise NotConvertible | Aconstant (c1,u1), Aconstant (c2,u2) -> if Constant.equal c1 c2 then convert_instances ~flex:true u1 u2 cu else raise NotConvertible | Asort s1, Asort s2 -> sort_cmp_universes env pb s1 s2 cu | Avar id1, Avar id2 -> if Id.equal id1 id2 then cu else raise NotConvertible | Acase(a1,ac1,p1,bs1), Acase(a2,ac2,p2,bs2) -> if not (eq_ind a1.asw_ind a2.asw_ind) then raise NotConvertible; let cu = conv_accu env CONV lvl ac1 ac2 cu in let tbl = a1.asw_reloc in let len = Array.length tbl in if Int.equal len 0 then conv_val env CONV lvl p1 p2 cu else begin let cu = conv_val env CONV lvl p1 p2 cu in let max = len - 1 in let rec aux i cu = let tag,arity = tbl.(i) in let ci = if Int.equal arity 0 then mk_const tag else mk_block tag (mk_rels_accu lvl arity) in let bi1 = bs1 ci and bi2 = bs2 ci in if Int.equal i max then conv_val env CONV (lvl + arity) bi1 bi2 cu else aux (i+1) (conv_val env CONV (lvl + arity) bi1 bi2 cu) in aux 0 cu end | Afix(t1,f1,rp1,s1), Afix(t2,f2,rp2,s2) -> if not (Int.equal s1 s2) || not (Array.equal Int.equal rp1 rp2) then raise NotConvertible; if f1 == f2 then cu else conv_fix env lvl t1 f1 t2 f2 cu | (Acofix(t1,f1,s1,_) | Acofixe(t1,f1,s1,_)), (Acofix(t2,f2,s2,_) | Acofixe(t2,f2,s2,_)) -> if not (Int.equal s1 s2) then raise NotConvertible; if f1 == f2 then cu else if not (Int.equal (Array.length f1) (Array.length f2)) then raise NotConvertible else conv_fix env lvl t1 f1 t2 f2 cu | Aprod(_,d1,_c1), Aprod(_,d2,_c2) -> let cu = conv_val env CONV lvl d1 d2 cu in let v = mk_rel_accu lvl in conv_val env pb (lvl + 1) (d1 v) (d2 v) cu | Aproj((ind1, i1), ac1), Aproj((ind2, i2), ac2) -> if not (eq_ind ind1 ind2 && Int.equal i1 i2) then raise NotConvertible else conv_accu env CONV lvl ac1 ac2 cu | Arel _, _ | Aind _, _ | Aconstant _, _ | Asort _, _ | Avar _, _ | Acase _, _ | Afix _, _ | Acofix _, _ | Acofixe _, _ | Aprod _, _ | Aproj _, _ | Ameta _, _ | Aevar _, _ -> raise NotConvertible (* Precondition length t1 = length f1 = length f2 = length t2 *) and conv_fix env lvl t1 f1 t2 f2 cu = let len = Array.length f1 in let max = len - 1 in let fargs = mk_rels_accu lvl len in let flvl = lvl + len in let rec aux i cu = let cu = conv_val env CONV lvl t1.(i) t2.(i) cu in let fi1 = napply f1.(i) fargs in let fi2 = napply f2.(i) fargs in if Int.equal i max then conv_val env CONV flvl fi1 fi2 cu else aux (i+1) (conv_val env CONV flvl fi1 fi2 cu) in aux 0 cu let warn_no_native_compiler = let open Pp in CWarnings.create ~name:"native-compiler-disabled" ~category:"native-compiler" (fun () -> strbrk "Native compiler is disabled," ++ strbrk " falling back to VM conversion test.") let native_conv_gen pb sigma env univs t1 t2 = if not (typing_flags env).Declarations.enable_native_compiler then begin warn_no_native_compiler (); Vconv.vm_conv_gen pb env univs t1 t2 end else let ml_filename, prefix = get_ml_filename () in let code, upds = mk_conv_code env sigma prefix t1 t2 in let fn = compile ml_filename code ~profile:false in if !Flags.debug then Feedback.msg_debug (Pp.str "Running test..."); let t0 = Sys.time () in call_linker env ~fatal:true ~prefix fn (Some upds); let t1 = Sys.time () in let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in if !Flags.debug then Feedback.msg_debug (Pp.str time_info); (* TODO change 0 when we can have de Bruijn *) fst (conv_val env pb 0 !rt1 !rt2 univs) (* Wrapper for [native_conv] above *) let native_conv cv_pb sigma env t1 t2 = let univs = Environ.universes env in let b = if cv_pb = CUMUL then Constr.leq_constr_univs univs t1 t2 else Constr.eq_constr_univs univs t1 t2 in if not b then let univs = (univs, checked_universes) in let t1 = Term.it_mkLambda_or_LetIn t1 (Environ.rel_context env) in let t2 = Term.it_mkLambda_or_LetIn t2 (Environ.rel_context env) in let _ = native_conv_gen cv_pb sigma env univs t1 t2 in () coq-8.11.0/kernel/vmvalues.ml0000644000175000017500000005737613612664533015764 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit = "init_coq_vm" let _ = init_vm () (******************************************************) (* Abstract data types and utility functions **********) (******************************************************) (* The representation of values relies on this assertion *) let _ = assert (Int.equal Obj.first_non_constant_constructor_tag 0) (* Values of the abstract machine *) type values type structured_values = values let val_of_obj v = ((Obj.obj v):values) let crazy_val = (val_of_obj (Obj.repr 0)) type tag = int let accu_tag = 0 let type_atom_tag = 2 let max_atom_tag = 2 let proj_tag = 3 let fix_app_tag = 4 let switch_tag = 5 let cofix_tag = 6 let cofix_evaluated_tag = 7 (** Structured constants are constants whose construction is done once. Their occurrences share the same value modulo kernel name substitutions (for functor application). Structured values have the additional property that no substitution will need to be performed, so their runtime value can directly be shared without reallocating a more structured representation. *) type structured_constant = | Const_sort of Sorts.t | Const_ind of inductive | Const_b0 of tag | Const_univ_level of Univ.Level.t | Const_val of structured_values | Const_uint of Uint63.t | Const_float of Float64.t type reloc_table = (tag * int) array type annot_switch = {ci : case_info; rtbl : reloc_table; tailcall : bool; max_stack_size : int} let rec eq_structured_values v1 v2 = v1 == v2 || let o1 = Obj.repr v1 in let o2 = Obj.repr v2 in if Obj.is_int o1 && Obj.is_int o2 then o1 == o2 else let t1 = Obj.tag o1 in let t2 = Obj.tag o2 in if Int.equal t1 t2 && Int.equal (Obj.size o1) (Obj.size o2) then if Int.equal t1 Obj.custom_tag then Int64.equal (Obj.magic v1 : int64) (Obj.magic v2 : int64) else if Int.equal t1 Obj.double_tag then Float64.(equal (of_float (Obj.magic v1)) (of_float (Obj.magic v2))) else begin assert (t1 <= Obj.last_non_constant_constructor_tag && t2 <= Obj.last_non_constant_constructor_tag); let i = ref 0 in while (!i < Obj.size o1 && eq_structured_values (Obj.magic (Obj.field o1 !i) : structured_values) (Obj.magic (Obj.field o2 !i) : structured_values)) do incr i done; !i >= Obj.size o1 end else false let hash_structured_values (v : structured_values) = (* We may want a better hash function here *) Hashtbl.hash v let eq_structured_constant c1 c2 = match c1, c2 with | Const_sort s1, Const_sort s2 -> Sorts.equal s1 s2 | Const_sort _, _ -> false | Const_ind i1, Const_ind i2 -> eq_ind i1 i2 | Const_ind _, _ -> false | Const_b0 t1, Const_b0 t2 -> Int.equal t1 t2 | Const_b0 _, _ -> false | Const_univ_level l1 , Const_univ_level l2 -> Univ.Level.equal l1 l2 | Const_univ_level _ , _ -> false | Const_val v1, Const_val v2 -> eq_structured_values v1 v2 | Const_val _, _ -> false | Const_uint i1, Const_uint i2 -> Uint63.equal i1 i2 | Const_uint _, _ -> false | Const_float f1, Const_float f2 -> Float64.equal f1 f2 | Const_float _, _ -> false let hash_structured_constant c = let open Hashset.Combine in match c with | Const_sort s -> combinesmall 1 (Sorts.hash s) | Const_ind i -> combinesmall 2 (ind_hash i) | Const_b0 t -> combinesmall 3 (Int.hash t) | Const_univ_level l -> combinesmall 4 (Univ.Level.hash l) | Const_val v -> combinesmall 5 (hash_structured_values v) | Const_uint i -> combinesmall 6 (Uint63.hash i) | Const_float f -> combinesmall 7 (Float64.hash f) let eq_annot_switch asw1 asw2 = let eq_ci ci1 ci2 = eq_ind ci1.ci_ind ci2.ci_ind && Int.equal ci1.ci_npar ci2.ci_npar && CArray.equal Int.equal ci1.ci_cstr_ndecls ci2.ci_cstr_ndecls in let eq_rlc (i1, j1) (i2, j2) = Int.equal i1 i2 && Int.equal j1 j2 in eq_ci asw1.ci asw2.ci && CArray.equal eq_rlc asw1.rtbl asw2.rtbl && (asw1.tailcall : bool) == asw2.tailcall let hash_annot_switch asw = let open Hashset.Combine in let h1 = Constr.case_info_hash asw.ci in let h2 = Array.fold_left (fun h (t, i) -> combine3 h t i) 0 asw.rtbl in let h3 = if asw.tailcall then 1 else 0 in combine3 h1 h2 h3 let pp_sort s = let open Sorts in match s with | SProp -> Pp.str "SProp" | Prop -> Pp.str "Prop" | Set -> Pp.str "Set" | Type u -> Pp.(str "Type@{" ++ Univ.pr_uni u ++ str "}") let pp_struct_const = function | Const_sort s -> pp_sort s | Const_ind (mind, i) -> Pp.(MutInd.print mind ++ str"#" ++ int i) | Const_b0 i -> Pp.int i | Const_univ_level l -> Univ.Level.pr l | Const_val _ -> Pp.str "(value)" | Const_uint i -> Pp.str (Uint63.to_string i) | Const_float f -> Pp.str (Float64.to_string f) (* Abstract data *) type vprod type vfun type vfix type vcofix type vblock type arguments let fun_val v = (Obj.magic v : values) let fix_val v = (Obj.magic v : values) let cofix_upd_val v = (Obj.magic v : values) type vm_env type vm_global let fun_env v = (Obj.magic v : vm_env) let fix_env v = (Obj.magic v : vm_env) let cofix_env v = (Obj.magic v : vm_env) let cofix_upd_env v = (Obj.magic v : vm_env) type vstack = values array let fun_of_val v = (Obj.magic v : vfun) let vm_global (v : values array) = (Obj.magic v : vm_global) (*******************************************) (* Machine code *** ************************) (*******************************************) type tcode (** A block whose first field is a C-allocated VM bytecode, encoded as char*. This is compatible with the representation of the Coq VM closures. *) type tcode_array external mkAccuCode : int -> tcode = "coq_makeaccu" external offset_tcode : tcode -> int -> tcode = "coq_offset_tcode" let fun_code v = (Obj.magic v : tcode) let fix_code = fun_code let cofix_upd_code = fun_code type vswitch = { sw_type_code : tcode; sw_code : tcode; sw_annot : annot_switch; sw_stk : vstack; sw_env : vm_env } (* Representation of values *) (* + Products : *) (* - vprod = 0_[ dom | codom] *) (* dom : values, codom : vfun *) (* *) (* + Functions have two representations : *) (* - unapplied fun : vf = Ct_[ C | fv1 | ... | fvn] *) (* C:tcode, fvi : values *) (* Remark : a function and its environment is the same value. *) (* - partially applied fun : Ct_[Restart:C| vf | arg1 | ... argn] *) (* *) (* + Fixpoints : *) (* - Ct_[C1|Infix_t|C2|...|Infix_t|Cn|fv1|...|fvn] *) (* One single block to represent all of the fixpoints, each fixpoint *) (* is the pointer to the field holding the pointer to its code, and *) (* the infix tag is used to know where the block starts. *) (* - Partial application follows the scheme of partially applied *) (* functions. Note: only fixpoints not having been applied to its *) (* recursive argument are coded this way. When the rec. arg. is *) (* applied, either it's a constructor and the fix reduces, or it's *) (* and the fix is coded as an accumulator. *) (* *) (* + Cofixpoints : see cbytegen.ml *) (* *) (* + vblock's encode (non constant) constructors as in Ocaml, but *) (* starting from 0 up. tag 0 ( = accu_tag) is reserved for *) (* accumulators. *) (* *) (* + vm_env is the type of the machine environments (i.e. a function or *) (* a fixpoint) *) (* *) (* + Accumulators : At_[accumulate| accu | arg1 | ... | argn ] *) (* - representation of [accu] : tag_[....] *) (* -- tag <= 3 : encoding atom type (sorts, free vars, etc.) *) (* -- 10_[accu|proj name] : a projection blocked by an accu *) (* -- 11_[accu|fix_app] : a fixpoint blocked by an accu *) (* -- 12_[accu|vswitch] : a match blocked by an accu *) (* -- 13_[fcofix] : a cofix function *) (* -- 14_[fcofix|val] : a cofix function, val represent the value *) (* of the function applied to arg1 ... argn *) (* The [arguments] type, which is abstracted as an array, represents : *) (* tag[ _ | _ |v1|... | vn] *) (* Generally the first field is a code pointer. *) (* Do not edit this type without editing C code, especially "coq_values.h" *) type id_key = | ConstKey of Constant.t | VarKey of Id.t | RelKey of Int.t | EvarKey of Evar.t let eq_id_key (k1 : id_key) (k2 : id_key) = match k1, k2 with | ConstKey c1, ConstKey c2 -> Constant.equal c1 c2 | VarKey id1, VarKey id2 -> Id.equal id1 id2 | RelKey n1, RelKey n2 -> Int.equal n1 n2 | EvarKey evk1, EvarKey evk2 -> Evar.equal evk1 evk2 | _ -> false type atom = | Aid of id_key | Aind of inductive | Asort of Sorts.t (* Zippers *) type zipper = | Zapp of arguments | Zfix of vfix*arguments (* Possibly empty *) | Zswitch of vswitch | Zproj of Projection.Repr.t (* name of the projection *) type stack = zipper list type to_update = values type whd = | Vprod of vprod | Vfun of vfun | Vfix of vfix * arguments option | Vcofix of vcofix * to_update * arguments option | Vconstr_const of int | Vconstr_block of vblock | Vint64 of int64 | Vfloat64 of float | Vatom_stk of atom * stack | Vuniv_level of Univ.Level.t (* Functions over arguments *) let nargs : arguments -> int = fun args -> (Obj.size (Obj.repr args)) - 2 let arg args i = if 0 <= i && i < (nargs args) then val_of_obj (Obj.field (Obj.repr args) (i+2)) else invalid_arg ("Vm.arg size = "^(string_of_int (nargs args))^ " acces "^(string_of_int i)) (*************************************************) (* Destructors ***********************************) (*************************************************) let uni_lvl_val (v : values) : Univ.Level.t = let whd = Obj.magic v in match whd with | Vuniv_level lvl -> lvl | _ -> let pr = let open Pp in match whd with | Vprod _ -> str "Vprod" | Vfun _ -> str "Vfun" | Vfix _ -> str "Vfix" | Vcofix _ -> str "Vcofix" | Vconstr_const _i -> str "Vconstr_const" | Vconstr_block _b -> str "Vconstr_block" | Vint64 _ -> str "Vint64" | Vfloat64 _ -> str "Vfloat64" | Vatom_stk (_a,_stk) -> str "Vatom_stk" | Vuniv_level _ -> assert false in CErrors.anomaly Pp.( strbrk "Parsing virtual machine value expected universe level, got " ++ pr ++ str ".") let rec whd_accu a stk = let stk = if Int.equal (Obj.size a) 2 then stk else Zapp (Obj.obj a) :: stk in let at = Obj.field a 1 in match Obj.tag at with | i when Int.equal i type_atom_tag -> begin match stk with | [] -> Vatom_stk(Obj.magic at, stk) | [Zapp args] -> let args = Array.init (nargs args) (arg args) in let s = Obj.obj (Obj.field at 0) in begin match s with | Sorts.Type u -> let inst = Instance.of_array (Array.map uni_lvl_val args) in let u = Univ.subst_instance_universe inst u in Vatom_stk (Asort (Sorts.sort_of_univ u), []) | _ -> assert false end | _ -> assert false end | i when i <= max_atom_tag -> Vatom_stk(Obj.magic at, stk) | i when Int.equal i proj_tag -> let zproj = Zproj (Obj.obj (Obj.field at 0)) in whd_accu (Obj.field at 1) (zproj :: stk) | i when Int.equal i fix_app_tag -> let fa = Obj.field at 1 in let zfix = Zfix (Obj.obj (Obj.field fa 1), Obj.obj fa) in whd_accu (Obj.field at 0) (zfix :: stk) | i when Int.equal i switch_tag -> let zswitch = Zswitch (Obj.obj (Obj.field at 1)) in whd_accu (Obj.field at 0) (zswitch :: stk) | i when Int.equal i cofix_tag -> let vcfx = Obj.obj (Obj.field at 0) in let to_up = Obj.obj a in begin match stk with | [] -> Vcofix(vcfx, to_up, None) | [Zapp args] -> Vcofix(vcfx, to_up, Some args) | _ -> assert false end | i when Int.equal i cofix_evaluated_tag -> let vcofix = Obj.obj (Obj.field at 0) in let res = Obj.obj a in begin match stk with | [] -> Vcofix(vcofix, res, None) | [Zapp args] -> Vcofix(vcofix, res, Some args) | _ -> assert false end | i when Int.equal i Obj.custom_tag -> Vint64 (Obj.magic i) | i when Int.equal i Obj.double_tag -> Vfloat64 (Obj.magic i) | tg -> CErrors.anomaly Pp.(strbrk "Failed to parse VM value. Tag = " ++ int tg ++ str ".") external kind_of_closure : Obj.t -> int = "coq_kind_of_closure" external is_accumulate : tcode -> bool = "coq_is_accumulate_code" external int_tcode : tcode -> int -> int = "coq_int_tcode" external accumulate : unit -> tcode = "accumulate_code" external set_bytecode_field : Obj.t -> int -> tcode -> unit = "coq_set_bytecode_field" let accumulate = accumulate () let whd_val : values -> whd = fun v -> let o = Obj.repr v in if Obj.is_int o then Vconstr_const (Obj.obj o) else let tag = Obj.tag o in if tag = accu_tag then if is_accumulate (fun_code o) then whd_accu o [] else Vprod(Obj.obj o) else if tag = Obj.closure_tag || tag = Obj.infix_tag then (match kind_of_closure o with | 0 -> Vfun(Obj.obj o) | 1 -> Vfix(Obj.obj o, None) | 2 -> Vfix(Obj.obj (Obj.field o 1), Some (Obj.obj o)) | 3 -> Vatom_stk(Aid(RelKey(int_tcode (fun_code o) 1)), []) | _ -> CErrors.anomaly ~label:"Vm.whd " (Pp.str "kind_of_closure does not work.")) else if Int.equal tag Obj.custom_tag then Vint64 (Obj.magic v) else if Int.equal tag Obj.double_tag then Vfloat64 (Obj.magic v) else Vconstr_block(Obj.obj o) (**********************************************) (* Constructors *******************************) (**********************************************) let obj_of_atom : atom -> Obj.t = fun a -> let res = Obj.new_block accu_tag 2 in set_bytecode_field res 0 accumulate; Obj.set_field res 1 (Obj.repr a); res (* obj_of_str_const : structured_constant -> Obj.t *) let obj_of_str_const str = match str with | Const_sort s -> obj_of_atom (Asort s) | Const_ind ind -> obj_of_atom (Aind ind) | Const_b0 tag -> Obj.repr tag | Const_univ_level l -> Obj.repr (Vuniv_level l) | Const_val v -> Obj.repr v | Const_uint i -> Obj.repr i | Const_float f -> Obj.repr f let val_of_block tag (args : structured_values array) = let nargs = Array.length args in let r = Obj.new_block tag nargs in for i = 0 to nargs - 1 do Obj.set_field r i (Obj.repr args.(i)) done; (Obj.magic r : structured_values) let val_of_obj o = ((Obj.obj o) : values) let val_of_str_const str = val_of_obj (obj_of_str_const str) let val_of_atom a = val_of_obj (obj_of_atom a) let val_of_int i = (Obj.magic i : values) let val_of_uint i = (Obj.magic i : values) let atom_of_proj kn v = let r = Obj.new_block proj_tag 2 in Obj.set_field r 0 (Obj.repr kn); Obj.set_field r 1 (Obj.repr v); ((Obj.obj r) : atom) let val_of_proj kn v = val_of_atom (atom_of_proj kn v) module IdKeyHash = struct type t = id_key let equal = eq_id_key open Hashset.Combine let hash : t -> tag = function | ConstKey c -> combinesmall 1 (Constant.hash c) | VarKey id -> combinesmall 2 (Id.hash id) | RelKey i -> combinesmall 3 (Int.hash i) | EvarKey evk -> combinesmall 4 (Evar.hash evk) end module KeyTable = Hashtbl.Make(IdKeyHash) let idkey_tbl = KeyTable.create 31 let val_of_idkey key = try KeyTable.find idkey_tbl key with Not_found -> let v = val_of_atom (Aid key) in KeyTable.add idkey_tbl key v; v let val_of_rel k = val_of_idkey (RelKey k) let val_of_named id = val_of_idkey (VarKey id) let val_of_constant c = val_of_idkey (ConstKey c) let val_of_evar evk = val_of_idkey (EvarKey evk) external val_of_annot_switch : annot_switch -> values = "%identity" external val_of_proj_name : Projection.Repr.t -> values = "%identity" (*************************************************) (** Operations manipulating data types ***********) (*************************************************) (* Functions over products *) let dom : vprod -> values = fun p -> val_of_obj (Obj.field (Obj.repr p) 0) let codom : vprod -> vfun = fun p -> (Obj.obj (Obj.field (Obj.repr p) 1)) (* Functions over vfun *) external closure_arity : vfun -> int = "coq_closure_arity" (* Functions over fixpoint *) external offset : Obj.t -> int = "coq_offset" external offset_closure : Obj.t -> int -> Obj.t = "coq_offset_closure" external offset_closure_fix : vfix -> int -> vm_env = "coq_offset_closure" external tcode_array : tcode_array -> tcode array = "coq_tcode_array" let first o = (offset_closure o (offset o)) let first_fix (v:vfix) = (Obj.magic (first (Obj.repr v)) : vfix) let last o = (Obj.field o (Obj.size o - 1)) let fix_types (v:vfix) = tcode_array (Obj.magic (last (Obj.repr v)) : tcode_array) let cofix_types (v:vcofix) = tcode_array (Obj.magic (last (Obj.repr v)) : tcode_array) let current_fix vf = - (offset (Obj.repr vf) / 2) let unsafe_fb_code fb i = let off = (2 * i) * (Sys.word_size / 8) in Obj.obj (Obj.add_offset (Obj.repr fb) (Int32.of_int off)) let unsafe_rec_arg fb i = int_tcode (unsafe_fb_code fb i) 1 let rec_args vf = let fb = first (Obj.repr vf) in let size = Obj.size (last fb) in Array.init size (unsafe_rec_arg fb) exception FALSE let check_fix f1 f2 = let i1, i2 = current_fix f1, current_fix f2 in (* Checking starting point *) if i1 = i2 then let fb1,fb2 = first (Obj.repr f1), first (Obj.repr f2) in let n = Obj.size (last fb1) in (* Checking number of definitions *) if n = Obj.size (last fb2) then (* Checking recursive arguments *) try for i = 0 to n - 1 do if unsafe_rec_arg fb1 i <> unsafe_rec_arg fb2 i then raise FALSE done; true with FALSE -> false else false else false let atom_rel : atom array ref = let init i = Aid (RelKey i) in ref (Array.init 40 init) let get_atom_rel () = !atom_rel let realloc_atom_rel n = let n = min (2 * n + 0x100) Sys.max_array_length in let init i = Aid (RelKey i) in let ans = Array.init n init in atom_rel := ans let relaccu_tbl = let len = Array.length !atom_rel in ref (Array.init len mkAccuCode) let relaccu_code i = let len = Array.length !relaccu_tbl in if i < len then !relaccu_tbl.(i) else begin realloc_atom_rel i; let nl = Array.length !atom_rel in relaccu_tbl := Array.init nl (fun j -> if j < len then !relaccu_tbl.(j) else mkAccuCode j); !relaccu_tbl.(i) end let mk_fix_body k ndef fb = let e = Obj.dup (Obj.repr fb) in for i = 0 to ndef - 1 do set_bytecode_field e (2 * i) (relaccu_code (k + i)) done; let fix_body i = let c = offset_tcode (unsafe_fb_code fb i) 2 in let res = Obj.new_block Obj.closure_tag 2 in set_bytecode_field res 0 c; Obj.set_field res 1 (offset_closure e (2*i)); ((Obj.obj res) : vfun) in Array.init ndef fix_body (* Functions over vcofix *) let get_fcofix vcf i = match whd_val (Obj.obj (Obj.field (Obj.repr vcf) (i+1))) with | Vcofix(vcfi, _, _) -> vcfi | _ -> assert false let current_cofix vcf = let ndef = Obj.size (last (Obj.repr vcf)) in let rec find_cofix pos = if pos < ndef then if get_fcofix vcf pos == vcf then pos else find_cofix (pos+1) else raise Not_found in try find_cofix 0 with Not_found -> assert false let check_cofix vcf1 vcf2 = (current_cofix vcf1 = current_cofix vcf2) && (Obj.size (last (Obj.repr vcf1)) = Obj.size (last (Obj.repr vcf2))) let mk_cofix_body apply_varray k ndef vcf = let e = Obj.dup (Obj.repr vcf) in for i = 0 to ndef - 1 do Obj.set_field e (i+1) (Obj.repr (val_of_rel (k+i))) done; let cofix_body i = let vcfi = get_fcofix vcf i in let c = Obj.field (Obj.repr vcfi) 0 in Obj.set_field e 0 c; let atom = Obj.new_block cofix_tag 1 in let self = Obj.new_block accu_tag 2 in set_bytecode_field self 0 accumulate; Obj.set_field self 1 (Obj.repr atom); apply_varray (Obj.obj e) [|Obj.obj self|] in Array.init ndef cofix_body (* Functions over vblock *) let btag : vblock -> int = fun b -> Obj.tag (Obj.repr b) let bsize : vblock -> int = fun b -> Obj.size (Obj.repr b) let bfield b i = if 0 <= i && i < (bsize b) then val_of_obj (Obj.field (Obj.repr b) i) else invalid_arg "Vm.bfield" (* Functions over vswitch *) let check_switch sw1 sw2 = sw1.sw_annot.rtbl = sw2.sw_annot.rtbl let branch_arg k (tag,arity) = if Int.equal arity 0 then ((Obj.magic tag):values) else let b, ofs = if tag < Obj.last_non_constant_constructor_tag then Obj.new_block tag arity, 0 else let b = Obj.new_block Obj.last_non_constant_constructor_tag (arity+1) in Obj.set_field b 0 (Obj.repr (tag-Obj.last_non_constant_constructor_tag)); b,1 in for i = ofs to ofs + arity - 1 do Obj.set_field b i (Obj.repr (val_of_rel (k+i))) done; val_of_obj b (* Printing *) let rec pr_atom a = Pp.(match a with | Aid c -> str "Aid(" ++ (match c with | ConstKey c -> Constant.print c | RelKey i -> str "#" ++ int i | _ -> str "...") ++ str ")" | Aind (mi,i) -> str "Aind(" ++ MutInd.print mi ++ str "#" ++ int i ++ str ")" | Asort _ -> str "Asort(") and pr_whd w = Pp.(match w with | Vprod _ -> str "Vprod" | Vfun _ -> str "Vfun" | Vfix _ -> str "Vfix" | Vcofix _ -> str "Vcofix" | Vconstr_const i -> str "Vconstr_const(" ++ int i ++ str ")" | Vconstr_block _b -> str "Vconstr_block" | Vint64 i -> i |> Format.sprintf "Vint64(%LiL)" |> str | Vfloat64 f -> str "Vfloat64(" ++ str (Float64.(to_string (of_float f))) ++ str ")" | Vatom_stk (a,stk) -> str "Vatom_stk(" ++ pr_atom a ++ str ", " ++ pr_stack stk ++ str ")" | Vuniv_level _ -> assert false) and pr_stack stk = Pp.(match stk with | [] -> str "[]" | s :: stk -> pr_zipper s ++ str " :: " ++ pr_stack stk) and pr_zipper z = Pp.(match z with | Zapp args -> str "Zapp(len = " ++ int (nargs args) ++ str ")" | Zfix (_f,args) -> str "Zfix(..., len=" ++ int (nargs args) ++ str ")" | Zswitch _s -> str "Zswitch(...)" | Zproj c -> str "Zproj(" ++ Projection.Repr.print c ++ str ")") coq-8.11.0/kernel/nativevalues.mli0000644000175000017500000001716113612664533016765 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t type accumulator type tag = int type arity = int type reloc_table = (tag * arity) array type annot_sw = { asw_ind : inductive; asw_ci : case_info; asw_reloc : reloc_table; asw_finite : bool; asw_prefix : string } val eq_annot_sw : annot_sw -> annot_sw -> bool val hash_annot_sw : annot_sw -> int type sort_annot = string * int type rec_pos = int array val eq_rec_pos : rec_pos -> rec_pos -> bool type atom = | Arel of int | Aconstant of pconstant | Aind of pinductive | Asort of Sorts.t | Avar of Id.t | Acase of annot_sw * accumulator * t * (t -> t) | Afix of t array * t array * rec_pos * int | Acofix of t array * t array * int * t | Acofixe of t array * t array * int * t | Aprod of Name.t * t * (t -> t) | Ameta of metavariable * t | Aevar of Evar.t * t array (* arguments *) | Aproj of (inductive * int) * accumulator type symbol = | SymbValue of t | SymbSort of Sorts.t | SymbName of Name.t | SymbConst of Constant.t | SymbMatch of annot_sw | SymbInd of inductive | SymbMeta of metavariable | SymbEvar of Evar.t | SymbLevel of Univ.Level.t | SymbProj of (inductive * int) type symbols = symbol array val empty_symbols : symbols (* Constructors *) val mk_accu : atom -> t val mk_rel_accu : int -> t val mk_rels_accu : int -> int -> t array val mk_constant_accu : Constant.t -> Univ.Level.t array -> t val mk_ind_accu : inductive -> Univ.Level.t array -> t val mk_sort_accu : Sorts.t -> Univ.Level.t array -> t val mk_var_accu : Id.t -> t val mk_sw_accu : annot_sw -> accumulator -> t -> (t -> t) val mk_prod_accu : Name.t -> t -> t -> t val mk_fix_accu : rec_pos -> int -> t array -> t array -> t val mk_cofix_accu : int -> t array -> t array -> t val mk_meta_accu : metavariable -> t val mk_evar_accu : Evar.t -> t array -> t val mk_proj_accu : (inductive * int) -> accumulator -> t val upd_cofix : t -> t -> unit val force_cofix : t -> t val mk_const : tag -> t val mk_block : tag -> t array -> t val mk_bool : bool -> t [@@ocaml.inline always] val mk_int : int -> t [@@ocaml.inline always] val mk_uint : Uint63.t -> t [@@ocaml.inline always] val mk_float : Float64.t -> t [@@ocaml.inline always] val napply : t -> t array -> t (* Functions over accumulators *) val dummy_value : unit -> t val atom_of_accu : accumulator -> atom val args_of_accu : accumulator -> t array val accu_nargs : accumulator -> int val cast_accu : t -> accumulator [@@ocaml.inline always] (* Functions over block: i.e constructors *) type block val block_size : block -> int val block_field : block -> int -> t val block_tag : block -> int (* kind_of_value *) type kind_of_value = | Vaccu of accumulator | Vfun of (t -> t) | Vconst of int | Vint64 of int64 | Vfloat64 of float | Vblock of block val kind_of_value : t -> kind_of_value val str_encode : 'a -> string val str_decode : string -> 'a (** Support for machine integers *) val val_to_int : t -> int val is_int : t -> bool [@@ocaml.inline always] (* function with check *) val head0 : t -> t -> t val tail0 : t -> t -> t val add : t -> t -> t -> t val sub : t -> t -> t -> t val mul : t -> t -> t -> t val div : t -> t -> t -> t val rem : t -> t -> t -> t val l_sr : t -> t -> t -> t val l_sl : t -> t -> t -> t val l_and : t -> t -> t -> t val l_xor : t -> t -> t -> t val l_or : t -> t -> t -> t val addc : t -> t -> t -> t val subc : t -> t -> t -> t val addCarryC : t -> t -> t -> t val subCarryC : t -> t -> t -> t val mulc : t -> t -> t -> t val diveucl : t -> t -> t -> t val div21 : t -> t -> t -> t -> t val addMulDiv : t -> t -> t -> t -> t val eq : t -> t -> t -> t val lt : t -> t -> t -> t val le : t -> t -> t -> t val compare : t -> t -> t -> t val print : t -> t (* Function without check *) val no_check_head0 : t -> t [@@ocaml.inline always] val no_check_tail0 : t -> t [@@ocaml.inline always] val no_check_add : t -> t -> t [@@ocaml.inline always] val no_check_sub : t -> t -> t [@@ocaml.inline always] val no_check_mul : t -> t -> t [@@ocaml.inline always] val no_check_div : t -> t -> t [@@ocaml.inline always] val no_check_rem : t -> t -> t [@@ocaml.inline always] val no_check_l_sr : t -> t -> t [@@ocaml.inline always] val no_check_l_sl : t -> t -> t [@@ocaml.inline always] val no_check_l_and : t -> t -> t [@@ocaml.inline always] val no_check_l_xor : t -> t -> t [@@ocaml.inline always] val no_check_l_or : t -> t -> t [@@ocaml.inline always] val no_check_addc : t -> t -> t [@@ocaml.inline always] val no_check_subc : t -> t -> t [@@ocaml.inline always] val no_check_addCarryC : t -> t -> t [@@ocaml.inline always] val no_check_subCarryC : t -> t -> t [@@ocaml.inline always] val no_check_mulc : t -> t -> t [@@ocaml.inline always] val no_check_diveucl : t -> t -> t [@@ocaml.inline always] val no_check_div21 : t -> t -> t -> t [@@ocaml.inline always] val no_check_addMulDiv : t -> t -> t -> t [@@ocaml.inline always] val no_check_eq : t -> t -> t [@@ocaml.inline always] val no_check_lt : t -> t -> t [@@ocaml.inline always] val no_check_le : t -> t -> t [@@ocaml.inline always] val no_check_compare : t -> t -> t (** Support for machine floating point values *) val is_float : t -> bool [@@ocaml.inline always] val fopp : t -> t -> t val fabs : t -> t -> t val feq : t -> t -> t -> t val flt : t -> t -> t -> t val fle : t -> t -> t -> t val fcompare : t -> t -> t -> t val fclassify : t -> t -> t val fadd : t -> t -> t -> t val fsub : t -> t -> t -> t val fmul : t -> t -> t -> t val fdiv : t -> t -> t -> t val fsqrt : t -> t -> t val float_of_int : t -> t -> t val normfr_mantissa : t -> t -> t val frshiftexp : t -> t -> t val ldshiftexp : t -> t -> t -> t val next_up : t -> t -> t val next_down : t -> t -> t (* Function without check *) val no_check_fopp : t -> t [@@ocaml.inline always] val no_check_fabs : t -> t [@@ocaml.inline always] val no_check_feq : t -> t -> t [@@ocaml.inline always] val no_check_flt : t -> t -> t [@@ocaml.inline always] val no_check_fle : t -> t -> t [@@ocaml.inline always] val no_check_fcompare : t -> t -> t [@@ocaml.inline always] val no_check_fclassify : t -> t [@@ocaml.inline always] val no_check_fadd : t -> t -> t [@@ocaml.inline always] val no_check_fsub : t -> t -> t [@@ocaml.inline always] val no_check_fmul : t -> t -> t [@@ocaml.inline always] val no_check_fdiv : t -> t -> t [@@ocaml.inline always] val no_check_fsqrt : t -> t [@@ocaml.inline always] val no_check_float_of_int : t -> t [@@ocaml.inline always] val no_check_normfr_mantissa : t -> t [@@ocaml.inline always] val no_check_frshiftexp : t -> t [@@ocaml.inline always] val no_check_ldshiftexp : t -> t -> t [@@ocaml.inline always] val no_check_next_up : t -> t [@@ocaml.inline always] val no_check_next_down : t -> t [@@ocaml.inline always] coq-8.11.0/kernel/esubst.mli0000644000175000017500000000655413612664533015570 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a subs val subs_cons: 'a array * 'a subs -> 'a subs val subs_shft: int * 'a subs -> 'a subs val subs_lift: 'a subs -> 'a subs val subs_liftn: int -> 'a subs -> 'a subs (** [subs_shift_cons(k,s,[|t1..tn|])] builds (^k s).t1..tn *) val subs_shift_cons: int * 'a subs * 'a array -> 'a subs (** [expand_rel k subs] expands de Bruijn [k] in the explicit substitution [subs]. The result is either (Inl(lams,v)) when the variable is substituted by value [v] under lams binders (i.e. v *has* to be shifted by lams), or (Inr (k',p)) when the variable k is just relocated as k'; p is None if the variable points inside subs and Some(k) if the variable points k bindings beyond subs (cf argument of ESID). *) val expand_rel: int -> 'a subs -> (int * 'a, int * int option) Util.union (** Tests whether a substitution behaves like the identity *) val is_subs_id: 'a subs -> bool (** Composition of substitutions: [comp mk_clos s1 s2] computes a substitution equivalent to applying s2 then s1. Argument mk_clos is used when a closure has to be created, i.e. when s1 is applied on an element of s2. *) val comp : ('a subs * 'a -> 'a) -> 'a subs -> 'a subs -> 'a subs (** {6 Compact representation } *) (** Compact representation of explicit relocations - [ELSHFT(l,n)] == lift of [n], then apply [lift l]. - [ELLFT(n,l)] == apply [l] to de Bruijn > [n] i.e under n binders. Invariant ensured by the private flag: no lift contains two consecutive [ELSHFT] nor two consecutive [ELLFT]. *) type lift = private | ELID | ELSHFT of lift * int | ELLFT of int * lift val el_id : lift val el_shft : int -> lift -> lift val el_liftn : int -> lift -> lift val el_lift : lift -> lift val reloc_rel : int -> lift -> int val is_lift_id : lift -> bool (** Lift applied to substitution: [lift_subst mk_clos el s] computes a substitution equivalent to applying el then s. Argument mk_clos is used when a closure has to be created, i.e. when el is applied on an element of s. *) val lift_subst : (lift -> 'a -> 'b) -> lift -> 'a subs -> 'b subs coq-8.11.0/kernel/vm.ml0000644000175000017500000001375613612664533014536 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* unit = "coq_set_drawinstr" external mkPopStopCode : int -> tcode = "coq_pushpop" let popstop_tbl = ref (Array.init 30 mkPopStopCode) let popstop_code i = let len = Array.length !popstop_tbl in if i < len then !popstop_tbl.(i) else begin popstop_tbl := Array.init (i+10) (fun j -> if j < len then !popstop_tbl.(j) else mkPopStopCode j); !popstop_tbl.(i) end let stop = popstop_code 0 (************************************************) (* Abstract machine *****************************) (************************************************) (* gestion de la pile *) external push_ra : tcode -> unit = "coq_push_ra" external push_val : values -> unit = "coq_push_val" external push_arguments : arguments -> unit = "coq_push_arguments" external push_vstack : vstack -> int -> unit = "coq_push_vstack" (* interpreteur *) external coq_interprete : tcode -> values -> atom array -> vm_global -> vm_env -> int -> values = "coq_interprete_byte" "coq_interprete_ml" let interprete code v env k = coq_interprete code v (get_atom_rel ()) (Csymtable.get_global_data ()) env k (* Functions over arguments *) (* Apply a value to arguments contained in [vargs] *) let apply_arguments vf vargs = let n = nargs vargs in if Int.equal n 0 then fun_val vf else begin push_ra stop; push_arguments vargs; interprete (fun_code vf) (fun_val vf) (fun_env vf) (n - 1) end (* Apply value [vf] to an array of argument values [varray] *) let apply_varray vf varray = let n = Array.length varray in if Int.equal n 0 then fun_val vf else begin push_ra stop; (* The fun code of [vf] will make sure we have enough stack, so we put 0 here. *) push_vstack varray 0; interprete (fun_code vf) (fun_val vf) (fun_env vf) (n - 1) end let mkrel_vstack k arity = let max = k + arity - 1 in Array.init arity (fun i -> val_of_rel (max - i)) let reduce_fun k vf = let vargs = mkrel_vstack k 1 in apply_varray vf vargs let decompose_vfun2 k vf1 vf2 = let arity = min (closure_arity vf1) (closure_arity vf2) in assert (0 < arity && arity < Sys.max_array_length); let vargs = mkrel_vstack k arity in let v1 = apply_varray vf1 vargs in let v2 = apply_varray vf2 vargs in arity, v1, v2 (* Functions over vfix *) let reduce_fix k vf = let fb = first_fix vf in (* computing types *) let fc_typ = fix_types fb in let ndef = Array.length fc_typ in let et = offset_closure_fix fb (2*(ndef - 1)) in let ftyp = Array.map (fun c -> interprete c crazy_val et 0) fc_typ in (* Construction of the environment of fix bodies *) (mk_fix_body k ndef fb, ftyp) let reduce_cofix k vcf = let fc_typ = cofix_types vcf in let ndef = Array.length fc_typ in let ftyp = (* Evaluate types *) Array.map (fun c -> interprete c crazy_val (cofix_env vcf) 0) fc_typ in (* Construction of the environment of cofix bodies *) (mk_cofix_body apply_varray k ndef vcf, ftyp) let type_of_switch sw = (* The fun code of types will make sure we have enough stack, so we put 0 here. *) push_vstack sw.sw_stk 0; interprete sw.sw_type_code crazy_val sw.sw_env 0 let apply_switch sw arg = let tc = sw.sw_annot.tailcall in if tc then (push_ra stop;push_vstack sw.sw_stk sw.sw_annot.max_stack_size) else (push_vstack sw.sw_stk sw.sw_annot.max_stack_size; push_ra (popstop_code (Array.length sw.sw_stk))); interprete sw.sw_code arg sw.sw_env 0 let branch_of_switch k sw = let eval_branch (_,arity as ta) = let arg = branch_arg k ta in let v = apply_switch sw arg in (arity, v) in Array.map eval_branch sw.sw_annot.rtbl (* Apply the term represented by a under stack stk to argument v *) (* t = a stk --> t v *) let rec apply_stack a stk v = match stk with | [] -> apply_varray (fun_of_val a) [|v|] | Zapp args :: stk -> apply_stack (apply_arguments (fun_of_val a) args) stk v | Zproj kn :: stk -> apply_stack (val_of_proj kn a) stk v | Zfix(f,args) :: stk -> let a,stk = match stk with | Zapp args' :: stk -> push_ra stop; push_arguments args'; push_val a; push_arguments args; let a = interprete (fix_code f) (fix_val f) (fix_env f) (nargs args+ nargs args') in a, stk | _ -> push_ra stop; push_val a; push_arguments args; let a = interprete (fix_code f) (fix_val f) (fix_env f) (nargs args) in a, stk in apply_stack a stk v | Zswitch sw :: stk -> apply_stack (apply_switch sw a) stk v let apply_whd k whd = let v = val_of_rel k in match whd with | Vprod _ | Vconstr_const _ | Vconstr_block _ | Vint64 _ | Vfloat64 _ -> assert false | Vfun f -> reduce_fun k f | Vfix(f, None) -> push_ra stop; push_val v; interprete (fix_code f) (fix_val f) (fix_env f) 0 | Vfix(f, Some args) -> push_ra stop; push_val v; push_arguments args; interprete (fix_code f) (fix_val f) (fix_env f) (nargs args) | Vcofix(_,to_up,_) -> push_ra stop; push_val v; interprete (cofix_upd_code to_up) (cofix_upd_val to_up) (cofix_upd_env to_up) 0 | Vatom_stk(a,stk) -> apply_stack (val_of_atom a) stk v | Vuniv_level _lvl -> assert false coq-8.11.0/kernel/names.mli0000644000175000017500000004644413612664533015370 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t -> bool (** Equality over identifiers. *) val compare : t -> t -> int (** Comparison over identifiers. *) val hash : t -> int (** Hash over identifiers. *) val is_valid : string -> bool (** Check that a string may be converted to an identifier. *) val of_bytes : bytes -> t val of_string : string -> t (** Converts a string into an identifier. @raise UserError if the string is invalid as an identifier. *) val of_string_soft : string -> t (** Same as {!of_string} except that any string made of supported UTF-8 characters is accepted. @raise UserError if the string is invalid as an UTF-8 string. *) val to_string : t -> string (** Converts a identifier into an string. *) val print : t -> Pp.t (** Pretty-printer. *) module Set : Set.S with type elt = t (** Finite sets of identifiers. *) module Map : Map.ExtS with type key = t and module Set := Set (** Finite maps of identifiers. *) module Pred : Predicate.S with type elt = t (** Predicates over identifiers. *) module List : List.MonoS with type elt = t (** Operations over lists of identifiers. *) val hcons : t -> t (** Hashconsing of identifiers. *) end (** Representation and operations on identifiers that are allowed to be anonymous (i.e. "_" in concrete syntax). *) module Name : sig type t = Anonymous (** anonymous identifier *) | Name of Id.t (** non-anonymous identifier *) val mk_name : Id.t -> t (** constructor *) val is_anonymous : t -> bool (** Return [true] iff a given name is [Anonymous]. *) val is_name : t -> bool (** Return [true] iff a given name is [Name _]. *) val compare : t -> t -> int (** Comparison over names. *) val equal : t -> t -> bool (** Equality over names. *) val hash : t -> int (** Hash over names. *) val hcons : t -> t (** Hashconsing over names. *) val print : t -> Pp.t (** Pretty-printer (print "_" for [Anonymous]. *) end (** {6 Type aliases} *) type name = Name.t = Anonymous | Name of Id.t [@@ocaml.deprecated "Use Name.t"] type variable = Id.t type module_ident = Id.t module ModIdset : Set.S with type elt = module_ident module ModIdmap : Map.ExtS with type key = module_ident and module Set := ModIdset (** {6 Directory paths = section names paths } *) module DirPath : sig type t (** Type of directory paths. Essentially a list of module identifiers. The order is reversed to improve sharing. E.g. A.B.C is ["C";"B";"A"] *) val equal : t -> t -> bool (** Equality over directory paths. *) val compare : t -> t -> int (** Comparison over directory paths. *) val hash : t -> int (** Hash over directory paths. *) val make : module_ident list -> t (** Create a directory path. (The list must be reversed). *) val repr : t -> module_ident list (** Represent a directory path. (The result list is reversed). *) val empty : t (** The empty directory path. *) val is_empty : t -> bool (** Test whether a directory path is empty. *) val initial : t (** Initial "seed" of the unique identifier generator *) val hcons : t -> t (** Hashconsing of directory paths. *) val to_string : t -> string (** Print non-empty directory paths as ["coq_root.module.submodule"] *) val print : t -> Pp.t end module DPset : Set.S with type elt = DirPath.t module DPmap : Map.ExtS with type key = DirPath.t and module Set := DPset (** {6 Names of structure elements } *) module Label : sig type t (** Type of labels *) val equal : t -> t -> bool (** Equality over labels *) val compare : t -> t -> int (** Comparison over labels. *) val hash : t -> int (** Hash over labels. *) val make : string -> t (** Create a label out of a string. *) val of_id : Id.t -> t (** Conversion from an identifier. *) val to_id : t -> Id.t (** Conversion to an identifier. *) val to_string : t -> string (** Conversion to string. *) val print : t -> Pp.t (** Pretty-printer. *) module Set : Set.S with type elt = t module Map : Map.ExtS with type key = t and module Set := Set val hcons : t -> t end (** {6 Unique names for bound modules} *) module MBId : sig type t (** Unique names for bound modules. Each call to [make] constructs a fresh unique identifier. *) val equal : t -> t -> bool (** Equality over unique bound names. *) val compare : t -> t -> int (** Comparison over unique bound names. *) val hash : t -> int (** Hash over unique bound names. *) val make : DirPath.t -> Id.t -> t (** The first argument is a file name, to prevent conflict between different files. *) val repr : t -> int * Id.t * DirPath.t (** Reverse of [make]. *) val to_id : t -> Id.t (** Return the identifier contained in the argument. *) val to_string : t -> string (** Encode as a string (not to be used for user-facing messages). *) val debug_to_string : t -> string (** Same as [to_string], but outputs extra information related to debug. *) end module MBIset : Set.S with type elt = MBId.t module MBImap : Map.ExtS with type key = MBId.t and module Set := MBIset (** {6 The module part of the kernel name } *) module ModPath : sig type t = | MPfile of DirPath.t | MPbound of MBId.t | MPdot of t * Label.t val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int val is_bound : t -> bool val initial : t (** Name of the toplevel structure ([= MPfile initial_dir]) *) val dp : t -> DirPath.t val to_string : t -> string (** Encode as a string (not to be used for user-facing messages). *) val debug_to_string : t -> string (** Same as [to_string], but outputs extra information related to debug. *) end module MPset : Set.S with type elt = ModPath.t module MPmap : Map.ExtS with type key = ModPath.t and module Set := MPset (** {6 The absolute names of objects seen by kernel } *) module KerName : sig type t (** Constructor and destructor *) val make : ModPath.t -> Label.t -> t val repr : t -> ModPath.t * Label.t (** Projections *) val modpath : t -> ModPath.t val label : t -> Label.t val to_string : t -> string (** Encode as a string (not to be used for user-facing messages). *) val print : t -> Pp.t (** Print internal representation (not to be used for user-facing messages). *) val debug_to_string : t -> string (** Same as [to_string], but outputs extra information related to debug. *) val debug_print : t -> Pp.t (** Same as [print], but outputs extra information related to debug. *) (** Comparisons *) val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module KNset : CSig.SetS with type elt = KerName.t module KNpred : Predicate.S with type elt = KerName.t module KNmap : Map.ExtS with type key = KerName.t and module Set := KNset (** {6 Constant Names } *) module Constant: sig type t (** Constructors *) val make : KerName.t -> KerName.t -> t (** Builds a constant name from a user and a canonical kernel name. *) val make1 : KerName.t -> t (** Special case of [make] where the user name is canonical. *) val make2 : ModPath.t -> Label.t -> t (** Shortcut for [(make1 (KerName.make2 ...))] *) (** Projections *) val user : t -> KerName.t val canonical : t -> KerName.t val repr2 : t -> ModPath.t * Label.t (** Shortcut for [KerName.repr (user ...)] *) val modpath : t -> ModPath.t (** Shortcut for [KerName.modpath (user ...)] *) val label : t -> Label.t (** Shortcut for [KerName.label (user ...)] *) (** Comparisons *) module CanOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module UserOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module SyntacticOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end val equal : t -> t -> bool (** Default comparison, alias for [CanOrd.equal] *) val hash : t -> int (** Hashing function *) val change_label : t -> Label.t -> t (** Builds a new constant name with a different label *) (** Displaying *) val to_string : t -> string (** Encode as a string (not to be used for user-facing messages). *) val print : t -> Pp.t (** Print internal representation (not to be used for user-facing messages). *) val debug_to_string : t -> string (** Same as [to_string], but outputs extra information related to debug. *) val debug_print : t -> Pp.t (** Same as [print], but outputs extra information related to debug. *) end (** The [*_env] modules consider an order on user part of names the others consider an order on canonical part of names*) module Cpred : Predicate.S with type elt = Constant.t module Cset : CSig.SetS with type elt = Constant.t module Cset_env : CSig.SetS with type elt = Constant.t module Cmap : Map.ExtS with type key = Constant.t and module Set := Cset (** A map whose keys are constants (values of the {!Constant.t} type). Keys are ordered wrt. "canonical form" of the constant. *) module Cmap_env : Map.ExtS with type key = Constant.t and module Set := Cset_env (** A map whose keys are constants (values of the {!Constant.t} type). Keys are ordered wrt. "user form" of the constant. *) (** {6 Inductive names} *) module MutInd : sig type t (** Constructors *) val make : KerName.t -> KerName.t -> t (** Builds a mutual inductive name from a user and a canonical kernel name. *) val make1 : KerName.t -> t (** Special case of [make] where the user name is canonical. *) val make2 : ModPath.t -> Label.t -> t (** Shortcut for [(make1 (KerName.make2 ...))] *) (** Projections *) val user : t -> KerName.t val canonical : t -> KerName.t val repr2 : t -> ModPath.t * Label.t (** Shortcut for [KerName.repr (user ...)] *) val modpath : t -> ModPath.t (** Shortcut for [KerName.modpath (user ...)] *) val label : t -> Label.t (** Shortcut for [KerName.label (user ...)] *) (** Comparisons *) module CanOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module UserOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module SyntacticOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end val equal : t -> t -> bool (** Default comparison, alias for [CanOrd.equal] *) val hash : t -> int (** Displaying *) val to_string : t -> string (** Encode as a string (not to be used for user-facing messages). *) val print : t -> Pp.t (** Print internal representation (not to be used for user-facing messages). *) val debug_to_string : t -> string (** Same as [to_string], but outputs extra information related to debug. *) val debug_print : t -> Pp.t (** Same as [print], but outputs extra information related to debug. *) end module Mindset : CSig.SetS with type elt = MutInd.t module Mindmap : Map.ExtS with type key = MutInd.t and module Set := Mindset module Mindmap_env : CMap.ExtS with type key = MutInd.t (** Designation of a (particular) inductive type. *) type inductive = MutInd.t (* the name of the inductive type *) * int (* the position of this inductive type within the block of mutually-recursive inductive types. BEWARE: indexing starts from 0. *) (** Designation of a (particular) constructor of a (particular) inductive type. *) type constructor = inductive (* designates the inductive type *) * int (* the index of the constructor BEWARE: indexing starts from 1. *) module Indset : CSet.S with type elt = inductive module Constrset : CSet.S with type elt = constructor module Indset_env : CSet.S with type elt = inductive module Constrset_env : CSet.S with type elt = constructor module Indmap : CMap.ExtS with type key = inductive and module Set := Indset module Constrmap : CMap.ExtS with type key = constructor and module Set := Constrset module Indmap_env : CMap.ExtS with type key = inductive and module Set := Indset_env module Constrmap_env : CMap.ExtS with type key = constructor and module Set := Constrset_env val ind_modpath : inductive -> ModPath.t val constr_modpath : constructor -> ModPath.t val ith_mutual_inductive : inductive -> int -> inductive val ith_constructor_of_inductive : inductive -> int -> constructor val inductive_of_constructor : constructor -> inductive val index_of_constructor : constructor -> int val eq_ind : inductive -> inductive -> bool val eq_user_ind : inductive -> inductive -> bool val eq_syntactic_ind : inductive -> inductive -> bool val ind_ord : inductive -> inductive -> int val ind_hash : inductive -> int val ind_user_ord : inductive -> inductive -> int val ind_user_hash : inductive -> int val ind_syntactic_ord : inductive -> inductive -> int val ind_syntactic_hash : inductive -> int val eq_constructor : constructor -> constructor -> bool val eq_user_constructor : constructor -> constructor -> bool val eq_syntactic_constructor : constructor -> constructor -> bool val constructor_ord : constructor -> constructor -> int val constructor_hash : constructor -> int val constructor_user_ord : constructor -> constructor -> int val constructor_user_hash : constructor -> int val constructor_syntactic_ord : constructor -> constructor -> int val constructor_syntactic_hash : constructor -> int (** {6 Hash-consing } *) val hcons_con : Constant.t -> Constant.t val hcons_mind : MutInd.t -> MutInd.t val hcons_ind : inductive -> inductive val hcons_construct : constructor -> constructor (******) type 'a tableKey = | ConstKey of 'a | VarKey of Id.t | RelKey of Int.t type inv_rel_key = int (** index in the [rel_context] part of environment starting by the end, {e inverse} of de Bruijn indice *) val eq_table_key : ('a -> 'a -> bool) -> 'a tableKey -> 'a tableKey -> bool val eq_constant_key : Constant.t -> Constant.t -> bool (** equalities on constant and inductive names (for the checker) *) val eq_ind_chk : inductive -> inductive -> bool (** {5 Module paths} *) type module_path = ModPath.t = | MPfile of DirPath.t | MPbound of MBId.t | MPdot of ModPath.t * Label.t [@@ocaml.deprecated "Alias type"] module Projection : sig module Repr : sig type t val make : inductive -> proj_npars:int -> proj_arg:int -> Label.t -> t module SyntacticOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module CanOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module UserOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end val constant : t -> Constant.t (** Don't use this if you don't have to. *) val inductive : t -> inductive val mind : t -> MutInd.t val npars : t -> int val arg : t -> int val label : t -> Label.t val equal : t -> t -> bool val hash : t -> int val compare : t -> t -> int val map : (MutInd.t -> MutInd.t) -> t -> t val map_npars : (MutInd.t -> int -> MutInd.t * int) -> t -> t val to_string : t -> string (** Encode as a string (not to be used for user-facing messages). *) val print : t -> Pp.t (** Print internal representation (not to be used for user-facing messages). *) end type t (* = Repr.t * bool *) val make : Repr.t -> bool -> t val repr : t -> Repr.t module SyntacticOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module CanOrd : sig val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end val constant : t -> Constant.t val mind : t -> MutInd.t val inductive : t -> inductive val npars : t -> int val arg : t -> int val label : t -> Label.t val unfolded : t -> bool val unfold : t -> t val equal : t -> t -> bool val hash : t -> int val hcons : t -> t (** Hashconsing of projections. *) val repr_equal : t -> t -> bool (** Ignoring the unfolding boolean. *) val compare : t -> t -> int val map : (MutInd.t -> MutInd.t) -> t -> t val map_npars : (MutInd.t -> int -> MutInd.t * int) -> t -> t val to_string : t -> string (** Encode as a string (not to be used for user-facing messages). *) val print : t -> Pp.t (** Print internal representation (not to be used for user-facing messages). *) end (** {6 Global reference is a kernel side type for all references together } *) (* XXX: Should we define GlobRefCan GlobRefUser? *) module GlobRef : sig type t = | VarRef of variable (** A reference to the section-context. *) | ConstRef of Constant.t (** A reference to the environment. *) | IndRef of inductive (** A reference to an inductive type. *) | ConstructRef of constructor (** A reference to a constructor of an inductive type. *) val equal : t -> t -> bool module Ordered : sig type nonrec t = t val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module Ordered_env : sig type nonrec t = t val compare : t -> t -> int val equal : t -> t -> bool val hash : t -> int end module Set_env : CSig.SetS with type elt = t module Map_env : Map.ExtS with type key = t and module Set := Set_env module Set : CSig.SetS with type elt = t module Map : Map.ExtS with type key = t and module Set := Set end (** Better to have it here that in Closure, since required in grammar.cma *) (* XXX: Move to a module *) type evaluable_global_reference = | EvalVarRef of Id.t | EvalConstRef of Constant.t val eq_egr : evaluable_global_reference -> evaluable_global_reference -> bool (** Located identifiers and objects with syntax. *) type lident = Id.t CAst.t type lname = Name.t CAst.t type lstring = string CAst.t coq-8.11.0/kernel/context.mli0000644000175000017500000003227213612664533015743 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a -> bool) -> 'a binder_annot -> 'a binder_annot -> bool val hash_annot : ('a -> int) -> 'a binder_annot -> int val map_annot : ('a -> 'b) -> 'a binder_annot -> 'b binder_annot val make_annot : 'a -> Sorts.relevance -> 'a binder_annot val binder_name : 'a binder_annot -> 'a val binder_relevance : 'a binder_annot -> Sorts.relevance val annotR : 'a -> 'a binder_annot (** Always Relevant *) val nameR : Id.t -> Name.t binder_annot (** Relevant + Name *) val anonR : Name.t binder_annot (** Relevant + Anonymous *) (** Representation of contexts that can capture anonymous as well as non-anonymous variables. Individual declarations are then designated by de Bruijn indexes. *) module Rel : sig module Declaration : sig (* local declaration *) type ('constr, 'types) pt = | LocalAssum of Name.t binder_annot * 'types (** name, type *) | LocalDef of Name.t binder_annot * 'constr * 'types (** name, value, type *) val get_annot : _ pt -> Name.t binder_annot (** Return the name bound by a given declaration. *) val get_name : ('c, 't) pt -> Name.t (** Return [Some value] for local-declarations and [None] for local-assumptions. *) val get_value : ('c, 't) pt -> 'c option (** Return the type of the name bound by a given declaration. *) val get_type : ('c, 't) pt -> 't val get_relevance : ('c, 't) pt -> Sorts.relevance (** Set the name that is bound by a given declaration. *) val set_name : Name.t -> ('c, 't) pt -> ('c, 't) pt (** Set the type of the bound variable in a given declaration. *) val set_type : 't -> ('c, 't) pt -> ('c, 't) pt (** Return [true] iff a given declaration is a local assumption. *) val is_local_assum : ('c, 't) pt -> bool (** Return [true] iff a given declaration is a local definition. *) val is_local_def : ('c, 't) pt -> bool (** Check whether any term in a given declaration satisfies a given predicate. *) val exists : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether all terms in a given declaration satisfy a given predicate. *) val for_all : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether the two given declarations are equal. *) val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Map the name bound by a given declaration. *) val map_name : (Name.t -> Name.t) -> ('c, 't) pt -> ('c, 't) pt (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt (** Map the type of the name bound by a given declaration. *) val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt (** Map all terms in a given declaration. *) val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Map all terms, with an heterogeneous function. *) val map_constr_het : ('a -> 'b) -> ('a, 'a) pt -> ('b, 'b) pt (** Perform a given action on all terms in a given declaration. *) val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given declaration to a single value. *) val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a val to_tuple : ('c, 't) pt -> Name.t binder_annot * 'c option * 't (** Turn [LocalDef] into [LocalAssum], identity otherwise. *) val drop_body : ('c, 't) pt -> ('c, 't) pt end (** Rel-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list (** empty rel-context *) val empty : ('c, 't) pt (** Return a new rel-context enriched by with a given inner-most declaration. *) val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt (** Return the number of {e local declarations} in a given context. *) val length : ('c, 't) pt -> int (** Check whether given two rel-contexts are equal. *) val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Return the number of {e local assumptions} in a given rel-context. *) val nhyps : ('c, 't) pt -> int (** Return a declaration designated by a given de Bruijn index. @raise Not_found if the designated de Bruijn index outside the range. *) val lookup : int -> ('c, 't) pt -> ('c, 't) Declaration.pt (** Map all terms in a given rel-context. *) val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Perform a given action on every declaration in a given rel-context. *) val iter : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given rel-context to a single value. Innermost declarations are processed first. *) val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a (** Reduce all terms in a given rel-context to a single value. Outermost declarations are processed first. *) val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a (** Map a given rel-context to a list where each {e local assumption} is mapped to [true] and each {e local definition} is mapped to [false]. *) val to_tags : ('c, 't) pt -> bool list (** Turn all [LocalDef] into [LocalAssum], leave [LocalAssum] unchanged. *) val drop_bodies : ('c, 't) pt -> ('c, 't) pt (** [extended_list mk n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ] with n = |Δ| and with the {e local definitions} of [Γ] skipped in [args] where [mk] is used to build the corresponding variables. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [mk 5, mk 3]. *) val to_extended_list : (int -> 'r) -> int -> ('c, 't) pt -> 'r list (** [extended_vect n Γ] does the same, returning instead an array. *) val to_extended_vect : (int -> 'r) -> int -> ('c, 't) pt -> 'r array end (** This module represents contexts that can capture non-anonymous variables. Individual declarations are then designated by the identifiers they bind. *) module Named : sig (** Representation of {e local declarations}. *) module Declaration : sig type ('constr, 'types) pt = | LocalAssum of Id.t binder_annot * 'types (** identifier, type *) | LocalDef of Id.t binder_annot * 'constr * 'types (** identifier, value, type *) val get_annot : _ pt -> Id.t binder_annot (** Return the identifier bound by a given declaration. *) val get_id : ('c, 't) pt -> Id.t (** Return [Some value] for local-declarations and [None] for local-assumptions. *) val get_value : ('c, 't) pt -> 'c option (** Return the type of the name bound by a given declaration. *) val get_type : ('c, 't) pt -> 't val get_relevance : ('c, 't) pt -> Sorts.relevance (** Set the identifier that is bound by a given declaration. *) val set_id : Id.t -> ('c, 't) pt -> ('c, 't) pt (** Set the type of the bound variable in a given declaration. *) val set_type : 't -> ('c, 't) pt -> ('c, 't) pt (** Return [true] iff a given declaration is a local assumption. *) val is_local_assum : ('c, 't) pt -> bool (** Return [true] iff a given declaration is a local definition. *) val is_local_def : ('c, 't) pt -> bool (** Check whether any term in a given declaration satisfies a given predicate. *) val exists : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether all terms in a given declaration satisfy a given predicate. *) val for_all : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether the two given declarations are equal. *) val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Map the identifier bound by a given declaration. *) val map_id : (Id.t -> Id.t) -> ('c, 't) pt -> ('c, 't) pt (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt (** Map the type of the name bound by a given declaration. *) val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt (** Map all terms in a given declaration. *) val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Perform a given action on all terms in a given declaration. *) val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given declaration to a single value. *) val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a val to_tuple : ('c, 't) pt -> Id.t binder_annot * 'c option * 't val of_tuple : Id.t binder_annot * 'c option * 't -> ('c, 't) pt (** Turn [LocalDef] into [LocalAssum], identity otherwise. *) val drop_body : ('c, 't) pt -> ('c, 't) pt (** Convert [Rel.Declaration.t] value to the corresponding [Named.Declaration.t] value. The function provided as the first parameter determines how to translate "names" to "ids". *) val of_rel_decl : (Name.t -> Id.t) -> ('c, 't) Rel.Declaration.pt -> ('c, 't) pt (** Convert [Named.Declaration.t] value to the corresponding [Rel.Declaration.t] value. *) (* TODO: Move this function to [Rel.Declaration] module and rename it to [of_named]. *) val to_rel_decl : ('c, 't) pt -> ('c, 't) Rel.Declaration.pt end (** Named-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list (** empty named-context *) val empty : ('c, 't) pt (** Return a new named-context enriched by with a given inner-most declaration. *) val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt (** Return the number of {e local declarations} in a given named-context. *) val length : ('c, 't) pt -> int (** Return a declaration designated by an identifier of the variable bound in that declaration. @raise Not_found if the designated identifier is not bound in a given named-context. *) val lookup : Id.t -> ('c, 't) pt -> ('c, 't) Declaration.pt (** Check whether given two named-contexts are equal. *) val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Map all terms in a given named-context. *) val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Perform a given action on every declaration in a given named-context. *) val iter : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given named-context to a single value. Innermost declarations are processed first. *) val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a (** Reduce all terms in a given named-context to a single value. Outermost declarations are processed first. *) val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a (** Return the set of all identifiers bound in a given named-context. *) val to_vars : ('c, 't) pt -> Id.Set.t (** Turn all [LocalDef] into [LocalAssum], leave [LocalAssum] unchanged. *) val drop_bodies : ('c, 't) pt -> ('c, 't) pt (** [to_instance Ω] builds an instance [args] such that [Ω ⊢ args:Ω] where [Ω] is a named-context and with the local definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it gives [Var id1, Var id3]. All [idj] are supposed distinct. *) val to_instance : (Id.t -> 'r) -> ('c, 't) pt -> 'r list end module Compacted : sig module Declaration : sig type ('constr, 'types) pt = | LocalAssum of Id.t binder_annot list * 'types | LocalDef of Id.t binder_annot list * 'constr * 'types val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt val of_named_decl : ('c, 't) Named.Declaration.pt -> ('c, 't) pt val to_named_context : ('c, 't) pt -> ('c, 't) Named.pt end type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list val fold : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a end coq-8.11.0/kernel/section.ml0000644000175000017500000001616413612664533015554 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Level.equal u u') (Instance.to_array (UContext.instance uctx)) in List.exists check s let push_constraints uctx s = let on_sec sec = if sec.has_poly_univs && Constraint.exists (fun (l,_,r) -> is_polymorphic_univ l s || is_polymorphic_univ r s) (snd uctx) then CErrors.user_err Pp.(str "Cannot add monomorphic constraints which refer to section polymorphic universes."); let uctx' = sec.sec_mono_universes in let sec_mono_universes = (ContextSet.union uctx uctx') in { sec with sec_mono_universes } in on_last_section on_sec s let open_section ~custom sections = let sec = { sec_context = 0; sec_mono_universes = ContextSet.empty; sec_poly_universes = ([||], UContext.empty); has_poly_univs = has_poly_univs sections; sec_entries = []; sec_data = (Cmap.empty, Mindmap.empty); sec_custom = custom; } in sec :: sections let close_section sections = match sections with | sec :: sections -> sections, sec.sec_entries, sec.sec_mono_universes, sec.sec_custom | [] -> CErrors.user_err (Pp.str "No opened section.") let make_decl_univs (nas,uctx) = abstract_universes nas uctx let push_global ~poly e s = if is_empty s then s else if has_poly_univs s && not poly then CErrors.user_err Pp.(str "Cannot add a universe monomorphic declaration when \ section polymorphic universes are present.") else let on_sec sec = { sec with sec_entries = e :: sec.sec_entries; sec_data = add_emap e (make_decl_univs sec.sec_poly_universes) sec.sec_data; } in on_last_section on_sec s let push_constant ~poly con s = push_global ~poly (SecDefinition con) s let push_inductive ~poly ind s = push_global ~poly (SecInductive ind) s type abstr_info = { abstr_ctx : Constr.named_context; abstr_subst : Instance.t; abstr_uctx : AUContext.t; } let empty_segment = { abstr_ctx = []; abstr_subst = Instance.empty; abstr_uctx = AUContext.empty; } let extract_hyps sec vars used = (* Keep the section-local segment of variables *) let vars = List.firstn sec.sec_context vars in (* Only keep the part that is used by the declaration *) List.filter (fun d -> Id.Set.mem (NamedDecl.get_id d) used) vars let section_segment_of_entry vars e hyps sections = (* [vars] are the named hypotheses, [hyps] the subset that is declared by the global *) let with_sec s = match s with | None -> CErrors.user_err (Pp.str "No opened section.") | Some sec -> let hyps = extract_hyps sec vars hyps in let inst, auctx = find_emap e sec.sec_data in { abstr_ctx = hyps; abstr_subst = inst; abstr_uctx = auctx; } in with_last_section with_sec sections let segment_of_constant env con s = let body = Environ.lookup_constant con env in let vars = Environ.named_context env in let used = Context.Named.to_vars body.Declarations.const_hyps in section_segment_of_entry vars (SecDefinition con) used s let segment_of_inductive env mind s = let mib = Environ.lookup_mind mind env in let vars = Environ.named_context env in let used = Context.Named.to_vars mib.Declarations.mind_hyps in section_segment_of_entry vars (SecInductive mind) used s let instance_from_variable_context = List.rev %> List.filter NamedDecl.is_local_assum %> List.map NamedDecl.get_id %> Array.of_list let extract_worklist info = let args = instance_from_variable_context info.abstr_ctx in info.abstr_subst, args let replacement_context env s = let with_sec sec = match sec with | None -> CErrors.user_err (Pp.str "No opened section.") | Some sec -> let cmap, imap = sec.sec_data in let cmap = Cmap.mapi (fun con _ -> extract_worklist @@ segment_of_constant env con s) cmap in let imap = Mindmap.mapi (fun ind _ -> extract_worklist @@ segment_of_inductive env ind s) imap in (cmap, imap) in with_last_section with_sec s let is_in_section env gr s = let with_sec sec = match sec with | None -> false | Some sec -> let open GlobRef in match gr with | VarRef id -> let vars = List.firstn sec.sec_context (Environ.named_context env) in List.exists (fun decl -> Id.equal id (NamedDecl.get_id decl)) vars | ConstRef con -> Cmap.mem con (fst sec.sec_data) | IndRef (ind, _) | ConstructRef ((ind, _), _) -> Mindmap.mem ind (snd sec.sec_data) in with_last_section with_sec s coq-8.11.0/kernel/retroknowledge.ml0000644000175000017500000000425113612664533017135 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* action | Register_type : CPrimitives.prim_type * Constant.t -> action coq-8.11.0/kernel/cbytecodes.mli0000644000175000017500000000602613612664533016401 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t val reset_label_counter : unit -> unit end type instruction = | Klabel of Label.t | Kacc of int (** accu = sp[n] *) | Kenvacc of int (** accu = coq_env[n] *) | Koffsetclosure of int (** accu = &coq_env[n] *) | Kpush (** sp = accu :: sp *) | Kpop of int (** sp = skipn n sp *) | Kpush_retaddr of Label.t (** sp = pc :: coq_env :: coq_extra_args :: sp ; coq_extra_args = 0 *) | Kapply of int (** number of arguments (arguments on top of stack) *) | Kappterm of int * int (** number of arguments, slot size *) | Kreturn of int (** slot size *) | Kjump | Krestart | Kgrab of int (** number of arguments *) | Kgrabrec of int (** rec arg *) | Kclosure of Label.t * int (** label, number of free variables *) | Kclosurerec of int * int * Label.t array * Label.t array (** nb fv, init, lbl types, lbl bodies *) | Kclosurecofix of int * int * Label.t array * Label.t array (** nb fv, init, lbl types, lbl bodies *) | Kgetglobal of Constant.t | Kconst of structured_constant | Kmakeblock of (* size: *) int * tag (** allocate an ocaml block. Index 0 ** is accu, all others are popped from ** the top of the stack *) | Kmakeprod | Kmakeswitchblock of Label.t * Label.t * annot_switch * int | Kswitch of Label.t array * Label.t array (** consts,blocks *) | Kpushfields of int | Kfield of int (** accu = accu[n] *) | Ksetfield of int (** accu[n] = sp[0] ; sp = pop sp *) | Kstop | Ksequence of bytecodes * bytecodes | Kproj of Projection.Repr.t | Kensurestackcapacity of int | Kbranch of Label.t (** jump to label, is it needed ? *) | Kprim of CPrimitives.t * pconstant option | Kareint of int and bytecodes = instruction list val pp_bytecodes : bytecodes -> Pp.t type fv_elem = FVnamed of Id.t | FVrel of int | FVuniv_var of int | FVevar of Evar.t type fv = fv_elem array val pp_fv_elem : fv_elem -> Pp.t coq-8.11.0/kernel/transparentState.ml0000644000175000017500000000230013612664533017435 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Entries.mutual_inductive_entry -> Entries.mutual_inductive_entry val dummy_variance : Entries.universes_entry -> Univ.Variance.t array coq-8.11.0/kernel/subtyping.ml0000644000175000017500000003520213612664533016126 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Label.Map.add (Label.of_id id) (IndConstr((ip,i+1), mib)) map) oib.mind_consnames map in Label.Map.add (Label.of_id oib.mind_typename) (IndType (ip, mib)) map in Array.fold_right_i add_mip_nameobjects mib.mind_packets map (* creates (namedobject/namedmodule) map for the whole signature *) type labmap = { objs : namedobject Label.Map.t; mods : namedmodule Label.Map.t } let empty_labmap = { objs = Label.Map.empty; mods = Label.Map.empty } let get_obj mp map l = try Label.Map.find l map.objs with Not_found -> error_no_such_label_sub l (ModPath.to_string mp) let get_mod mp map l = try Label.Map.find l map.mods with Not_found -> error_no_such_label_sub l (ModPath.to_string mp) let make_labmap mp list = let add_one (l,e) map = match e with | SFBconst cb -> { map with objs = Label.Map.add l (Constant cb) map.objs } | SFBmind mib -> { map with objs = add_mib_nameobjects mp l mib map.objs } | SFBmodule mb -> { map with mods = Label.Map.add l (Module mb) map.mods } | SFBmodtype mtb -> { map with mods = Label.Map.add l (Modtype mtb) map.mods } in CList.fold_right add_one list empty_labmap let check_conv_error error why cst poly f env a1 a2 = try let cst' = f env (Environ.universes env) a1 a2 in if poly then if Constraint.is_empty cst' then cst else error (IncompatiblePolymorphism (env, a1, a2)) else Constraint.union cst cst' with NotConvertible -> error why | Univ.UniverseInconsistency e -> error (IncompatibleUniverses e) let check_universes error env u1 u2 = match u1, u2 with | Monomorphic _, Monomorphic _ -> env | Polymorphic auctx1, Polymorphic auctx2 -> let lbound = Environ.universes_lbound env in if not (UGraph.check_subtype ~lbound (Environ.universes env) auctx2 auctx1) then error (IncompatibleConstraints { got = auctx1; expect = auctx2; } ) else Environ.push_context ~strict:false (Univ.AUContext.repr auctx2) env | Monomorphic _, Polymorphic _ -> error (PolymorphicStatusExpected true) | Polymorphic _, Monomorphic _ -> error (PolymorphicStatusExpected false) let check_variance error v1 v2 = match v1, v2 with | None, None -> () | Some v1, Some v2 -> if not (Array.for_all2 Variance.check_subtype v2 v1) then error IncompatibleVariance | None, Some _ -> error (CumulativeStatusExpected true) | Some _, None -> error (CumulativeStatusExpected false) (* for now we do not allow reorderings *) let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2= let kn1 = KerName.make mp1 l in let kn2 = KerName.make mp2 l in let error why = error_signature_mismatch l spec2 why in let check_conv why cst poly f = check_conv_error error why cst poly f in let mib1 = match info1 with | IndType ((_,0), mib) -> Declareops.subst_mind_body subst1 mib | _ -> error (InductiveFieldExpected mib2) in let env = check_universes error env mib1.mind_universes mib2.mind_universes in let () = check_variance error mib1.mind_variance mib2.mind_variance in let inst = make_abstract_instance (Declareops.inductive_polymorphic_context mib1) in let mib2 = Declareops.subst_mind_body subst2 mib2 in let check_inductive_type cst name t1 t2 = check_conv (NotConvertibleInductiveField name) cst (inductive_is_polymorphic mib1) (infer_conv_leq ?l2r:None ?evars:None ?ts:None) env t1 t2 in let check_packet cst p1 p2 = let check f test why = if not (test (f p1) (f p2)) then error why in check (fun p -> p.mind_consnames) (Array.equal Id.equal) NotSameConstructorNamesField; check (fun p -> p.mind_typename) Id.equal NotSameInductiveNameInBlockField; (* nf_lc later *) (* nf_arity later *) (* user_lc ignored *) (* user_arity ignored *) check (fun p -> p.mind_nrealargs) Int.equal (NotConvertibleInductiveField p2.mind_typename); (* How can it fail since the type of inductive are checked below? [HH] *) (* kelim ignored *) (* listrec ignored *) (* finite done *) (* nparams done *) (* params_ctxt done because part of the inductive types *) (* Don't check the sort of the type if polymorphic *) let ty1 = type_of_inductive env ((mib1, p1), inst) in let ty2 = type_of_inductive env ((mib2, p2), inst) in let cst = check_inductive_type cst p2.mind_typename ty1 ty2 in cst in let mind = MutInd.make1 kn1 in let check_cons_types _i cst p1 p2 = Array.fold_left3 (fun cst id t1 t2 -> check_conv (NotConvertibleConstructorField id) cst (inductive_is_polymorphic mib1) (infer_conv ?l2r:None ?evars:None ?ts:None) env t1 t2) cst p2.mind_consnames (arities_of_specif (mind, inst) (mib1, p1)) (arities_of_specif (mind, inst) (mib2, p2)) in let check f test why = if not (test (f mib1) (f mib2)) then error (why (f mib2)) in check (fun mib -> mib.mind_finite<>CoFinite) (==) (fun x -> FiniteInductiveFieldExpected x); check (fun mib -> mib.mind_ntypes) Int.equal (fun x -> InductiveNumbersFieldExpected x); assert (List.is_empty mib1.mind_hyps && List.is_empty mib2.mind_hyps); assert (Array.length mib1.mind_packets >= 1 && Array.length mib2.mind_packets >= 1); (* Check that the expected numbers of uniform parameters are the same *) (* No need to check the contexts of parameters: it is checked *) (* at the time of checking the inductive arities in check_packet. *) (* Notice that we don't expect the local definitions to match: only *) (* the inductive types and constructors types have to be convertible *) check (fun mib -> mib.mind_nparams) Int.equal (fun x -> InductiveParamsNumberField x); begin let kn2' = kn_of_delta reso2 kn2 in if KerName.equal kn2 kn2' || MutInd.equal (mind_of_delta_kn reso1 kn1) (subst_mind subst2 (MutInd.make kn2 kn2')) then () else error NotEqualInductiveAliases end; (* we check that records and their field names are preserved. *) (** FIXME: this check looks nonsense *) check (fun mib -> mib.mind_record <> NotRecord) (==) (fun x -> RecordFieldExpected x); if mib1.mind_record <> NotRecord then begin let rec names_prod_letin t = match kind t with | Prod(n,_,t) -> n.binder_name::(names_prod_letin t) | LetIn(n,_,_,t) -> n.binder_name::(names_prod_letin t) | Cast(t,_,_) -> names_prod_letin t | _ -> [] in assert (Int.equal (Array.length mib1.mind_packets) 1); assert (Int.equal (Array.length mib2.mind_packets) 1); assert (Int.equal (Array.length mib1.mind_packets.(0).mind_user_lc) 1); assert (Int.equal (Array.length mib2.mind_packets.(0).mind_user_lc) 1); check (fun mib -> let nparamdecls = List.length mib.mind_params_ctxt in let names = names_prod_letin (mib.mind_packets.(0).mind_user_lc.(0)) in snd (List.chop nparamdecls names)) (List.equal Name.equal) (fun x -> RecordProjectionsExpected x); end; (* we first check simple things *) let cst = Array.fold_left2 check_packet cst mib1.mind_packets mib2.mind_packets in (* and constructor types in the end *) let cst = Array.fold_left2_i check_cons_types cst mib1.mind_packets mib2.mind_packets in cst let check_constant cst env l info1 cb2 spec2 subst1 subst2 = let error why = error_signature_mismatch l spec2 why in let check_conv cst poly f = check_conv_error error cst poly f in let check_type poly cst env t1 t2 = let err = NotConvertibleTypeField (env, t1, t2) in check_conv err cst poly (infer_conv_leq ?l2r:None ?evars:None ?ts:None) env t1 t2 in match info1 with | Constant cb1 -> let () = assert (List.is_empty cb1.const_hyps && List.is_empty cb2.const_hyps) in let cb1 = Declareops.subst_const_body subst1 cb1 in let cb2 = Declareops.subst_const_body subst2 cb2 in (* Start by checking universes *) let env = check_universes error env cb1.const_universes cb2.const_universes in let poly = Declareops.constant_is_polymorphic cb1 in (* Now check types *) let typ1 = cb1.const_type in let typ2 = cb2.const_type in let cst = check_type poly cst env typ1 typ2 in (* Now we check the bodies: - A transparent constant can only be implemented by a compatible transparent constant. - In the signature, an opaque is handled just as a parameter: anything of the right type can implement it, even if bodies differ. *) (match cb2.const_body with | Primitive _ | Undef _ | OpaqueDef _ -> cst | Def lc2 -> (match cb1.const_body with | Primitive _ | Undef _ | OpaqueDef _ -> error NotConvertibleBodyField | Def lc1 -> (* NB: cb1 might have been strengthened and appear as transparent. Anyway [check_conv] will handle that afterwards. *) let c1 = Mod_subst.force_constr lc1 in let c2 = Mod_subst.force_constr lc2 in check_conv NotConvertibleBodyField cst poly (infer_conv ?l2r:None ?evars:None ?ts:None) env c1 c2)) | IndType ((_kn,_i),_mind1) -> CErrors.user_err Pp.(str @@ "The kernel does not recognize yet that a parameter can be " ^ "instantiated by an inductive type. Hint: you can rename the " ^ "inductive type and give a definition to map the old name to the new " ^ "name.") | IndConstr (((_kn,_i),_j),_mind1) -> CErrors.user_err Pp.(str @@ "The kernel does not recognize yet that a parameter can be " ^ "instantiated by a constructor. Hint: you can rename the " ^ "constructor and give a definition to map the old name to the new " ^ "name.") let rec check_modules cst env msb1 msb2 subst1 subst2 = let mty1 = module_type_of_module msb1 in let mty2 = module_type_of_module msb2 in check_modtypes cst env mty1 mty2 subst1 subst2 false and check_signatures cst env mp1 sig1 mp2 sig2 subst1 subst2 reso1 reso2= let map1 = make_labmap mp1 sig1 in let check_one_body cst (l,spec2) = match spec2 with | SFBconst cb2 -> check_constant cst env l (get_obj mp1 map1 l) cb2 spec2 subst1 subst2 | SFBmind mib2 -> check_inductive cst env mp1 l (get_obj mp1 map1 l) mp2 mib2 spec2 subst1 subst2 reso1 reso2 | SFBmodule msb2 -> begin match get_mod mp1 map1 l with | Module msb -> check_modules cst env msb msb2 subst1 subst2 | _ -> error_signature_mismatch l spec2 ModuleFieldExpected end | SFBmodtype mtb2 -> let mtb1 = match get_mod mp1 map1 l with | Modtype mtb -> mtb | _ -> error_signature_mismatch l spec2 ModuleTypeFieldExpected in let env = add_module_type mtb2.mod_mp mtb2 (add_module_type mtb1.mod_mp mtb1 env) in check_modtypes cst env mtb1 mtb2 subst1 subst2 true in List.fold_left check_one_body cst sig2 and check_modtypes cst env mtb1 mtb2 subst1 subst2 equiv = if mtb1==mtb2 || mtb1.mod_type == mtb2.mod_type then cst else let rec check_structure cst env str1 str2 equiv subst1 subst2 = match str1,str2 with |NoFunctor list1, NoFunctor list2 -> if equiv then let subst2 = add_mp mtb2.mod_mp mtb1.mod_mp mtb1.mod_delta subst2 in let cst1 = check_signatures cst env mtb1.mod_mp list1 mtb2.mod_mp list2 subst1 subst2 mtb1.mod_delta mtb2.mod_delta in let cst2 = check_signatures cst env mtb2.mod_mp list2 mtb1.mod_mp list1 subst2 subst1 mtb2.mod_delta mtb1.mod_delta in Univ.Constraint.union cst1 cst2 else check_signatures cst env mtb1.mod_mp list1 mtb2.mod_mp list2 subst1 subst2 mtb1.mod_delta mtb2.mod_delta |MoreFunctor (arg_id1,arg_t1,body_t1), MoreFunctor (arg_id2,arg_t2,body_t2) -> let mp2 = MPbound arg_id2 in let subst1 = join (map_mbid arg_id1 mp2 arg_t2.mod_delta) subst1 in let cst = check_modtypes cst env arg_t2 arg_t1 subst2 subst1 equiv in (* contravariant *) let env = add_module_type mp2 arg_t2 env in let env = if Modops.is_functor body_t1 then env else add_module {mod_mp = mtb1.mod_mp; mod_expr = Abstract; mod_type = subst_signature subst1 body_t1; mod_type_alg = None; mod_constraints = mtb1.mod_constraints; mod_retroknowledge = ModBodyRK []; mod_delta = mtb1.mod_delta} env in check_structure cst env body_t1 body_t2 equiv subst1 subst2 | _ , _ -> error_incompatible_modtypes mtb1 mtb2 in check_structure cst env mtb1.mod_type mtb2.mod_type equiv subst1 subst2 let check_subtypes env sup super = let env = add_module_type sup.mod_mp sup env in let env = Environ.push_context_set ~strict:true super.mod_constraints env in check_modtypes Univ.Constraint.empty env (strengthen sup sup.mod_mp) super empty_subst (map_mp super.mod_mp sup.mod_mp sup.mod_delta) false coq-8.11.0/kernel/vars.ml0000644000175000017500000002414013612664533015054 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* if m>n then raise LocalOccur | _ -> Constr.iter_with_binders succ closed_rec n c in try closed_rec n c; true with LocalOccur -> false (* [closed0 M] is true iff [M] is a (de Bruijn) closed term *) let closed0 c = closedn 0 c (* (noccurn n M) returns true iff (Rel n) does NOT occur in term M *) let noccurn n term = let rec occur_rec n c = match Constr.kind c with | Constr.Rel m -> if Int.equal m n then raise LocalOccur | _ -> Constr.iter_with_binders succ occur_rec n c in try occur_rec n term; true with LocalOccur -> false (* (noccur_between n m M) returns true iff (Rel p) does NOT occur in term M for n <= p < n+m *) let noccur_between n m term = let rec occur_rec n c = match Constr.kind c with | Constr.Rel p -> if n<=p && p Constr.iter_with_binders succ occur_rec n c in try occur_rec n term; true with LocalOccur -> false (* Checking function for terms containing existential variables. The function [noccur_with_meta] considers the fact that each existential variable (as well as each isevar) in the term appears applied to its local context, which may contain the CoFix variables. These occurrences of CoFix variables are not considered *) let isMeta c = match Constr.kind c with | Constr.Meta _ -> true | _ -> false let noccur_with_meta n m term = let rec occur_rec n c = match Constr.kind c with | Constr.Rel p -> if n<=p && p (match Constr.kind f with | Constr.Cast (c,_,_) when isMeta c -> () | Constr.Meta _ -> () | _ -> Constr.iter_with_binders succ occur_rec n c) | Constr.Evar (_, _) -> () | _ -> Constr.iter_with_binders succ occur_rec n c in try (occur_rec n term; true) with LocalOccur -> false (*********************) (* Lifting *) (*********************) let exliftn = Constr.exliftn let liftn = Constr.liftn let lift = Constr.lift (*********************) (* Substituting *) (*********************) (* (subst1 M c) substitutes M for Rel(1) in c we generalise it to (substl [M1,...,Mn] c) which substitutes in parallel M1,...,Mn for respectively Rel(1),...,Rel(n) in c *) (* 1st : general case *) type info = Closed | Open | Unknown type 'a substituend = { mutable sinfo: info; sit: 'a } let lift_substituend depth s = match s.sinfo with | Closed -> s.sit | Open -> lift depth s.sit | Unknown -> let sit = s.sit in if closed0 sit then let () = s.sinfo <- Closed in sit else let () = s.sinfo <- Open in lift depth sit let make_substituend c = { sinfo=Unknown; sit=c } let substn_many lamv n c = let lv = Array.length lamv in if Int.equal lv 0 then c else let rec substrec depth c = match Constr.kind c with | Constr.Rel k -> if k<=depth then c else if k-depth <= lv then lift_substituend depth (Array.unsafe_get lamv (k-depth-1)) else Constr.mkRel (k-lv) | _ -> Constr.map_with_binders succ substrec depth c in substrec n c (* let substkey = CProfile.declare_profile "substn_many";; let substn_many lamv n c = CProfile.profile3 substkey substn_many lamv n c;; *) let make_subst = function | [] -> [||] | hd :: tl -> let len = List.length tl in let subst = Array.make (1 + len) (make_substituend hd) in let s = ref tl in for i = 1 to len do match !s with | [] -> assert false | x :: tl -> Array.unsafe_set subst i (make_substituend x); s := tl done; subst (* The type of substitutions, with term substituting most recent binder at the head *) type substl = Constr.t list let substnl laml n c = substn_many (make_subst laml) n c let substl laml c = substn_many (make_subst laml) 0 c let subst1 lam c = substn_many [|make_substituend lam|] 0 c let substnl_decl laml k r = RelDecl.map_constr (fun c -> substnl laml k c) r let substl_decl laml r = RelDecl.map_constr (fun c -> substnl laml 0 c) r let subst1_decl lam r = RelDecl.map_constr (fun c -> subst1 lam c) r (* Build a substitution from an instance, inserting missing let-ins *) let subst_of_rel_context_instance sign l = let rec aux subst sign l = let open RelDecl in match sign, l with | LocalAssum _ :: sign', a::args' -> aux (a::subst) sign' args' | LocalDef (_,c,_)::sign', args' -> aux (substl subst c :: subst) sign' args' | [], [] -> subst | _ -> CErrors.anomaly (Pp.str "Instance and signature do not match.") in aux [] (List.rev sign) l let adjust_subst_to_rel_context sign l = List.rev (subst_of_rel_context_instance sign l) let adjust_rel_to_rel_context sign n = let rec aux sign = let open RelDecl in match sign with | LocalAssum _ :: sign' -> let (n',p) = aux sign' in (n'+1,p) | LocalDef (_,_c,_)::sign' -> let (n',p) = aux sign' in (n'+1,if n' (0,n) in snd (aux sign) (* (thin_val sigma) removes identity substitutions from sigma *) let rec thin_val = function | [] -> [] | (id, c) :: tl -> match Constr.kind c with | Constr.Var v -> if Id.equal id v then thin_val tl else (id, make_substituend c) :: (thin_val tl) | _ -> (id, make_substituend c) :: (thin_val tl) let rec find_var id = function | [] -> raise Not_found | (idc, c) :: subst -> if Id.equal id idc then c else find_var id subst (* (replace_vars sigma M) applies substitution sigma to term M *) let replace_vars var_alist x = let var_alist = thin_val var_alist in match var_alist with | [] -> x | _ -> let rec substrec n c = match Constr.kind c with | Constr.Var x -> (try lift_substituend n (find_var x var_alist) with Not_found -> c) | _ -> Constr.map_with_binders succ substrec n c in substrec 0 x (* (subst_var str t) substitute (Var str) by (Rel 1) in t *) let subst_var str t = replace_vars [(str, Constr.mkRel 1)] t (* (subst_vars [id1;...;idn] t) substitute (Var idj) by (Rel j) in t *) let substn_vars p vars c = let _,subst = List.fold_left (fun (n,l) var -> ((n+1),(var,Constr.mkRel n)::l)) (p,[]) vars in replace_vars (List.rev subst) c let subst_vars subst c = substn_vars 1 subst c (** Universe substitutions *) open Constr let subst_univs_level_constr subst c = if Univ.is_empty_level_subst subst then c else let f = Univ.Instance.subst_fn (Univ.subst_univs_level_level subst) in let changed = ref false in let rec aux t = match kind t with | Const (c, u) -> if Univ.Instance.is_empty u then t else let u' = f u in if u' == u then t else (changed := true; mkConstU (c, u')) | Ind (i, u) -> if Univ.Instance.is_empty u then t else let u' = f u in if u' == u then t else (changed := true; mkIndU (i, u')) | Construct (c, u) -> if Univ.Instance.is_empty u then t else let u' = f u in if u' == u then t else (changed := true; mkConstructU (c, u')) | Sort (Sorts.Type u) -> let u' = Univ.subst_univs_level_universe subst u in if u' == u then t else (changed := true; mkSort (Sorts.sort_of_univ u')) | _ -> Constr.map aux t in let c' = aux c in if !changed then c' else c let subst_univs_level_context s = Context.Rel.map (subst_univs_level_constr s) let subst_instance_constr subst c = if Univ.Instance.is_empty subst then c else let f u = Univ.subst_instance_instance subst u in let rec aux t = match kind t with | Const (c, u) -> if Univ.Instance.is_empty u then t else let u' = f u in if u' == u then t else (mkConstU (c, u')) | Ind (i, u) -> if Univ.Instance.is_empty u then t else let u' = f u in if u' == u then t else (mkIndU (i, u')) | Construct (c, u) -> if Univ.Instance.is_empty u then t else let u' = f u in if u' == u then t else (mkConstructU (c, u')) | Sort (Sorts.Type u) -> let u' = Univ.subst_instance_universe subst u in if u' == u then t else (mkSort (Sorts.sort_of_univ u')) | _ -> Constr.map aux t in aux c let univ_instantiate_constr u c = let open Univ in assert (Int.equal (Instance.length u) (AUContext.size c.univ_abstracted_binder)); subst_instance_constr u c.univ_abstracted_value (* let substkey = CProfile.declare_profile "subst_instance_constr";; *) (* let subst_instance_constr inst c = CProfile.profile2 substkey subst_instance_constr inst c;; *) let subst_instance_context s ctx = if Univ.Instance.is_empty s then ctx else Context.Rel.map (fun x -> subst_instance_constr s x) ctx let universes_of_constr c = let open Univ in let rec aux s c = match kind c with | Const (_c, u) -> LSet.fold LSet.add (Instance.levels u) s | Ind ((_mind,_), u) | Construct (((_mind,_),_), u) -> LSet.fold LSet.add (Instance.levels u) s | Sort u when not (Sorts.is_small u) -> let u = Sorts.univ_of_sort u in LSet.fold LSet.add (Universe.levels u) s | _ -> Constr.fold aux s c in aux LSet.empty c coq-8.11.0/kernel/dune0000644000175000017500000000126013612664533014423 0ustar treinentreinen(library (name kernel) (synopsis "The Coq Kernel") (public_name coq.kernel) (wrapped false) (modules (:standard \ genOpcodeFiles uint63_31 uint63_63)) (libraries lib byterun dynlink)) (executable (name genOpcodeFiles) (modules genOpcodeFiles)) (rule (targets copcodes.ml) (action (with-stdout-to %{targets} (run ./genOpcodeFiles.exe copml)))) (rule (targets uint63.ml) (deps (:gen-file uint63_%{ocaml-config:int_size}.ml)) (action (copy# %{gen-file} %{targets}))) (documentation (package coq)) ; In dev profile, we check the kernel against a more strict set of ; warnings. (env (dev (flags :standard -w +a-4-44-50))) ; (ocaml408 (flags :standard -w +a-3-4-44-50))) coq-8.11.0/kernel/cClosure.mli0000644000175000017500000001742113612664533016035 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a (** {6 ... } *) (** Delta implies all consts (both global (= by [kernel_name]) and local (= by [Rel] or [Var])), all evars, and letin's. Rem: reduction of a Rel/Var bound to a term is Delta, but reduction of a LetIn expression is Letin reduction *) (** Sets of reduction kinds. *) module type RedFlagsSig = sig type reds type red_kind (** {7 The different kinds of reduction } *) val fBETA : red_kind val fDELTA : red_kind val fETA : red_kind (** The fETA flag is never used by the kernel reduction but pretyping does *) val fMATCH : red_kind val fFIX : red_kind val fCOFIX : red_kind val fZETA : red_kind val fCONST : Constant.t -> red_kind val fVAR : Id.t -> red_kind (** No reduction at all *) val no_red : reds (** Adds a reduction kind to a set *) val red_add : reds -> red_kind -> reds (** Removes a reduction kind to a set *) val red_sub : reds -> red_kind -> reds (** Adds a reduction kind to a set *) val red_add_transparent : reds -> TransparentState.t -> reds (** Retrieve the transparent state of the reduction flags *) val red_transparent : reds -> TransparentState.t (** Build a reduction set from scratch = iter [red_add] on [no_red] *) val mkflags : red_kind list -> reds (** Tests if a reduction kind is set *) val red_set : reds -> red_kind -> bool (** This tests if the projection is in unfolded state already or is unfodable due to delta. *) val red_projection : reds -> Projection.t -> bool end module RedFlags : RedFlagsSig open RedFlags (* These flags do not contain eta *) val all : reds val allnolet : reds val beta : reds val betadeltazeta : reds val betaiota : reds val betaiotazeta : reds val betazeta : reds val delta : reds val zeta : reds val nored : reds (***********************************************************************) type table_key = Constant.t Univ.puniverses tableKey module KeyTable : Hashtbl.S with type key = table_key (*********************************************************************** s Lazy reduction. *) (** [fconstr] is the type of frozen constr *) type fconstr (** [fconstr] can be accessed by using the function [fterm_of] and by matching on type [fterm] *) type fterm = | FRel of int | FAtom of constr (** Metas and Sorts *) | FFlex of table_key | FInd of inductive Univ.puniverses | FConstruct of constructor Univ.puniverses | FApp of fconstr * fconstr array | FProj of Projection.t * fconstr | FFix of fixpoint * fconstr subs | FCoFix of cofixpoint * fconstr subs | FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *) | FLambda of int * (Name.t Context.binder_annot * constr) list * constr * fconstr subs | FProd of Name.t Context.binder_annot * fconstr * constr * fconstr subs | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * fconstr subs | FEvar of existential * fconstr subs | FInt of Uint63.t | FFloat of Float64.t | FLIFT of int * fconstr | FCLOS of constr * fconstr subs | FLOCKED (*********************************************************************** s A [stack] is a context of arguments, arguments are pushed by [append_stack] one array at a time *) type 'a next_native_args = (CPrimitives.arg_kind * 'a) list type stack_member = | Zapp of fconstr array | ZcaseT of case_info * constr * constr array * fconstr subs | Zproj of Projection.Repr.t | Zfix of fconstr * stack | Zprimitive of CPrimitives.t * pconstant * fconstr list * fconstr next_native_args (* operator, constr def, reduced arguments rev, next arguments *) | Zshift of int | Zupdate of fconstr and stack = stack_member list val empty_stack : stack val append_stack : fconstr array -> stack -> stack val check_native_args : CPrimitives.t -> stack -> bool val get_native_args1 : CPrimitives.t -> pconstant -> stack -> fconstr list * fconstr * fconstr next_native_args * stack val stack_args_size : stack -> int val eta_expand_stack : stack -> stack (** To lazy reduce a constr, create a [clos_infos] with [create_clos_infos], inject the term to reduce with [inject]; then use a reduction function *) val inject : constr -> fconstr (** mk_atom: prevents a term from being evaluated *) val mk_atom : constr -> fconstr (** mk_red: makes a reducible term (used in newring) *) val mk_red : fterm -> fconstr val fterm_of : fconstr -> fterm val term_of_fconstr : fconstr -> constr val destFLambda : (fconstr subs -> constr -> fconstr) -> fconstr -> Name.t Context.binder_annot * fconstr * fconstr type optrel = Unknown | KnownR | KnownI val relevance_of : fconstr -> optrel val set_relevance : Sorts.relevance -> fconstr -> unit (** Global and local constant cache *) type clos_infos type clos_tab val create_clos_infos : ?evars:(existential->constr option) -> reds -> env -> clos_infos val oracle_of_infos : clos_infos -> Conv_oracle.oracle val create_tab : unit -> clos_tab val info_env : clos_infos -> env val info_flags: clos_infos -> reds val unfold_projection : clos_infos -> Projection.t -> stack_member option val infos_with_reds : clos_infos -> reds -> clos_infos (** Reduction function *) (** [norm_val] is for strong normalization *) val norm_val : clos_infos -> clos_tab -> fconstr -> constr (** [whd_val] is for weak head normalization *) val whd_val : clos_infos -> clos_tab -> fconstr -> constr (** [whd_stack] performs weak head normalization in a given stack. It stops whenever a reduction is blocked. *) val whd_stack : clos_infos -> clos_tab -> fconstr -> stack -> fconstr * stack (** [eta_expand_ind_stack env ind c s t] computes stacks corresponding to the conversion of the eta expansion of t, considered as an inhabitant of ind, and the Constructor c of this inductive type applied to arguments s. @assumes [t] is a rigid term, and not a constructor. [ind] is the inductive of the constructor term [c] @raise Not_found if the inductive is not a primitive record, or if the constructor is partially applied. *) val eta_expand_ind_stack : env -> inductive -> fconstr -> stack -> (fconstr * stack) -> stack * stack (** Conversion auxiliary functions to do step by step normalisation *) (** [unfold_reference] unfolds references in a [fconstr] *) val unfold_reference : clos_infos -> clos_tab -> table_key -> (fconstr, Util.Empty.t) constant_def (*********************************************************************** i This is for lazy debug *) val lift_fconstr : int -> fconstr -> fconstr val lift_fconstr_vect : int -> fconstr array -> fconstr array val mk_clos : fconstr subs -> constr -> fconstr val mk_clos_vect : fconstr subs -> constr array -> fconstr array val kni: clos_infos -> clos_tab -> fconstr -> stack -> fconstr * stack val knr: clos_infos -> clos_tab -> fconstr -> stack -> fconstr * stack val kl : clos_infos -> clos_tab -> fconstr -> constr val to_constr : lift -> fconstr -> constr (** End of cbn debug section i*) coq-8.11.0/kernel/csymtable.ml0000644000175000017500000001500013612664533016057 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* atom array -> vm_global -> values array -> values = "coq_eval_tcode" type global_data = { mutable glob_len : int; mutable glob_val : values array } (*******************) (* Linkage du code *) (*******************) (* Table des globaux *) (* [global_data] contient les valeurs des constantes globales (axiomes,definitions), les annotations des switch et les structured constant *) let global_data = { glob_len = 0; glob_val = Array.make 4096 crazy_val; } let get_global_data () = Vmvalues.vm_global global_data.glob_val let realloc_global_data n = let n = min (2 * n + 0x100) Sys.max_array_length in let ans = Array.make n crazy_val in let src = global_data.glob_val in let () = Array.blit src 0 ans 0 (Array.length src) in global_data.glob_val <- ans let check_global_data n = if n >= Array.length global_data.glob_val then realloc_global_data n let set_global v = let n = global_data.glob_len in check_global_data n; global_data.glob_val.(n) <- v; global_data.glob_len <- global_data.glob_len + 1; n (* table pour les structured_constant et les annotations des switchs *) module SConstTable = Hashtbl.Make (struct type t = structured_constant let equal = eq_structured_constant let hash = hash_structured_constant end) module AnnotTable = Hashtbl.Make (struct type t = annot_switch let equal = eq_annot_switch let hash = hash_annot_switch end) module ProjNameTable = Hashtbl.Make (Projection.Repr) let str_cst_tbl : int SConstTable.t = SConstTable.create 31 let annot_tbl : int AnnotTable.t = AnnotTable.create 31 (* (annot_switch * int) Hashtbl.t *) let proj_name_tbl : int ProjNameTable.t = ProjNameTable.create 31 (*************************************************************) (*** Mise a jour des valeurs des variables et des constantes *) (*************************************************************) exception NotEvaluated let key rk = match !rk with | None -> raise NotEvaluated | Some k -> try CEphemeron.get k with CEphemeron.InvalidKey -> raise NotEvaluated (************************) (* traduction des patch *) (* slot_for_*, calcul la valeur de l'objet, la place dans la table global, rend sa position dans la table *) let slot_for_str_cst key = try SConstTable.find str_cst_tbl key with Not_found -> let n = set_global (val_of_str_const key) in SConstTable.add str_cst_tbl key n; n let slot_for_annot key = try AnnotTable.find annot_tbl key with Not_found -> let n = set_global (val_of_annot_switch key) in AnnotTable.add annot_tbl key n; n let slot_for_proj_name key = try ProjNameTable.find proj_name_tbl key with Not_found -> let n = set_global (val_of_proj_name key) in ProjNameTable.add proj_name_tbl key n; n let rec slot_for_getglobal env kn = let (cb,(_,rk)) = lookup_constant_key kn env in try key rk with NotEvaluated -> (* Pp.msgnl(str"not yet evaluated");*) let pos = match cb.const_body_code with | None -> set_global (val_of_constant kn) | Some code -> match Cemitcodes.force code with | BCdefined(code,pl,fv) -> let v = eval_to_patch env (code,pl,fv) in set_global v | BCalias kn' -> slot_for_getglobal env kn' | BCconstant -> set_global (val_of_constant kn) in (*Pp.msgnl(str"value stored at: "++int pos);*) rk := Some (CEphemeron.create pos); pos and slot_for_fv env fv = let fill_fv_cache cache id v_of_id env_of_id b = let v,d = match b with | None -> v_of_id id, Id.Set.empty | Some c -> val_of_constr (env_of_id id env) c, Environ.global_vars_set env c in build_lazy_val cache (v, d); v in let val_of_rel i = val_of_rel (nb_rel env - i) in let idfun _ x = x in match fv with | FVnamed id -> let nv = lookup_named_val id env in begin match force_lazy_val nv with | None -> env |> lookup_named id |> NamedDecl.get_value |> fill_fv_cache nv id val_of_named idfun | Some (v, _) -> v end | FVrel i -> let rv = lookup_rel_val i env in begin match force_lazy_val rv with | None -> env |> lookup_rel i |> RelDecl.get_value |> fill_fv_cache rv i val_of_rel env_of_rel | Some (v, _) -> v end | FVevar evk -> val_of_evar evk | FVuniv_var _idu -> assert false and eval_to_patch env (buff,pl,fv) = let slots = function | Reloc_annot a -> slot_for_annot a | Reloc_const sc -> slot_for_str_cst sc | Reloc_getglobal kn -> slot_for_getglobal env kn | Reloc_proj_name p -> slot_for_proj_name p in let tc = patch buff pl slots in let vm_env = (* Beware, this may look like a call to [Array.map], but it's not. Calling [Array.map f] when the first argument returned by [f] is a float would lead to [vm_env] being an unboxed Double_array (Tag_val = Double_array_tag) whereas eval_tcode expects a regular array (Tag_val = 0). See test-suite/primitive/float/coq_env_double_array.v for an actual instance. *) let a = Array.make (Array.length fv) crazy_val in Array.iteri (fun i v -> a.(i) <- slot_for_fv env v) fv; a in eval_tcode tc (get_atom_rel ()) (vm_global global_data.glob_val) vm_env and val_of_constr env c = match compile ~fail_on_error:true env c with | Some v -> eval_to_patch env (to_memory v) | None -> assert false let set_transparent_const _kn = () (* !?! *) let set_opaque_const _kn = () (* !?! *) coq-8.11.0/kernel/inductive.mli0000644000175000017500000001260113612664533016243 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* types -> pinductive * constr list val find_inductive : env -> types -> pinductive * constr list val find_coinductive : env -> types -> pinductive * constr list type mind_specif = mutual_inductive_body * one_inductive_body (** {6 ... } *) (** Fetching information in the environment about an inductive type. Raises [Not_found] if the inductive type is not found. *) val lookup_mind_specif : env -> inductive -> mind_specif (** {6 Functions to build standard types related to inductive } *) val ind_subst : MutInd.t -> mutual_inductive_body -> Instance.t -> constr list val inductive_paramdecls : mutual_inductive_body puniverses -> Constr.rel_context val instantiate_inductive_constraints : mutual_inductive_body -> Instance.t -> Constraint.t val constrained_type_of_inductive : env -> mind_specif puniverses -> types constrained val constrained_type_of_inductive_knowing_parameters : env -> mind_specif puniverses -> types Lazy.t array -> types constrained val relevance_of_inductive : env -> inductive -> Sorts.relevance val type_of_inductive : env -> mind_specif puniverses -> types val type_of_inductive_knowing_parameters : env -> ?polyprop:bool -> mind_specif puniverses -> types Lazy.t array -> types val elim_sort : mind_specif -> Sorts.family val is_private : mind_specif -> bool val is_primitive_record : mind_specif -> bool (** Return type as quoted by the user *) val constrained_type_of_constructor : pconstructor -> mind_specif -> types constrained val type_of_constructor : pconstructor -> mind_specif -> types (** Return constructor types in normal form *) val arities_of_constructors : pinductive -> mind_specif -> types array (** Return constructor types in user form *) val type_of_constructors : pinductive -> mind_specif -> types array (** Transforms inductive specification into types (in nf) *) val arities_of_specif : MutInd.t puniverses -> mind_specif -> types array val inductive_params : mind_specif -> int (** [type_case_branches env (I,args) (p:A) c] computes useful types about the following Cases expression:

Cases (c :: (I args)) of b1..bn end It computes the type of every branch (pattern variables are introduced by products), the type for the whole expression, and the universe constraints generated. *) val type_case_branches : env -> pinductive * constr list -> unsafe_judgment -> constr -> types array * types val build_branches_type : pinductive -> mutual_inductive_body * one_inductive_body -> constr list -> constr -> types array (** Return the arity of an inductive type *) val mind_arity : one_inductive_body -> Constr.rel_context * Sorts.family val inductive_sort_family : one_inductive_body -> Sorts.family (** Check a [case_info] actually correspond to a Case expression on the given inductive type. *) val check_case_info : env -> pinductive -> Sorts.relevance -> case_info -> unit (** {6 Guard conditions for fix and cofix-points. } *) (** When [chk] is false, the guard condition is not actually checked. *) val check_fix : env -> fixpoint -> unit val check_cofix : env -> cofixpoint -> unit (** {6 Support for sort-polymorphic inductive types } *) (** The "polyprop" optional argument below controls the "Prop-polymorphism". By default, it is allowed. But when "polyprop=false", the following exception is raised when a polymorphic singleton inductive type becomes Prop due to parameter instantiation. This is used by the Ocaml extraction, which cannot handle (yet?) Prop-polymorphism. *) exception SingletonInductiveBecomesProp of Id.t val max_inductive_sort : Sorts.t array -> Universe.t val instantiate_universes : env -> Constr.rel_context -> template_arity -> constr Lazy.t array -> Constr.rel_context * Sorts.t (** {6 Debug} *) type size = Large | Strict type subterm_spec = Subterm of (size * wf_paths) | Dead_code | Not_subterm type guard_env = { env : env; (** dB of last fixpoint *) rel_min : int; (** dB of variables denoting subterms *) genv : subterm_spec Lazy.t list; } type stack_element = |SClosure of guard_env*constr |SArg of subterm_spec Lazy.t val subterm_specif : guard_env -> stack_element list -> constr -> subterm_spec val lambda_implicit_lift : int -> constr -> constr val abstract_mind_lc : int -> Int.t -> (rel_context * constr) array -> constr array coq-8.11.0/kernel/cbytegen.ml0000644000175000017500000010057413612664533015707 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* *) (* type = [Ct1 | .... | Ctn] *) (* Ci is the code pointer of the i-th body *) (* At runtime, a fixpoint environment (which is the same as the fixpoint *) (* itself) is a pointer to the field holding its code pointer. *) (* In each fixpoint body, de Bruijn [nbr] represents the first fixpoint *) (* and de Bruijn [1] the last one. *) (* Access to these variables is performed by the [Koffsetclosure n] *) (* instruction that shifts the environment pointer of [n] fields. *) (* This allows representing mutual fixpoints in just one block. *) (* [Ct1 | ... | Ctn] is an array holding code pointers of the fixpoint *) (* types. They are used in conversion tests (which requires that *) (* fixpoint types must be convertible). Their environment is the one of *) (* the last fixpoint : *) (* [t1|C1| ... |tc|Cc| ... |t(nbr)|C(nbr)| fv1 | fv2 | .... | fvn | type] *) (* ^ *) (* Representation of mutual cofix : *) (* a1 = [A_t | accumulate | [Cfx_t | fcofix1 ] ] *) (* ... *) (* anbr = [A_t | accumulate | [Cfx_t | fcofixnbr ] ] *) (* *) (* fcofix1 = [clos_t | code1 | a1 |...| anbr | fv1 |...| fvn | type] *) (* ^ *) (* ... *) (* fcofixnbr = [clos_t | codenbr | a1 |...| anbr | fv1 |...| fvn | type] *) (* ^ *) (* The [ai] blocks are functions that accumulate their arguments: *) (* ai arg1 argp ---> *) (* ai' = [A_t | accumulate | [Cfx_t | fcofixi] | arg1 | ... | argp ] *) (* If such a block is matched against, we have to force evaluation, *) (* function [fcofixi] is then applied to [ai'] [arg1] ... [argp] *) (* (note that [ai'] is a pointer to the closure, passed as argument) *) (* Once evaluation is completed [ai'] is updated with the result: *) (* ai' <-- *) (* [A_t | accumulate | [Cfxe_t |fcofixi|result] | arg1 | ... | argp ] *) (* This representation is nice because the application of the cofix is *) (* evaluated only once (it simulates a lazy evaluation) *) (* Moreover, when cofix don't have arguments, it is possible to create *) (* a cycle, e.g.: *) (* cofix one := cons 1 one *) (* a1 = [A_t | accumulate | [Cfx_t|fcofix1] ] *) (* fcofix1 = [clos_t | code | a1] *) (* The result of evaluating [a1] is [cons_t | 1 | a1]. *) (* When [a1] is updated : *) (* a1 = [A_t | accumulate | [Cfxe_t | fcofix1 | [cons_t | 1 | a1]] ] *) (* The cycle is created ... *) (* *) (* In Cfxe_t accumulators, we need to store [fcofixi] for testing *) (* conversion of cofixpoints (which is intentional). *) module Fv_elem = struct type t = fv_elem let compare e1 e2 = match e1, e2 with | FVnamed id1, FVnamed id2 -> Id.compare id1 id2 | FVnamed _, (FVrel _ | FVuniv_var _ | FVevar _) -> -1 | FVrel _, FVnamed _ -> 1 | FVrel r1, FVrel r2 -> Int.compare r1 r2 | FVrel _, (FVuniv_var _ | FVevar _) -> -1 | FVuniv_var i1, FVuniv_var i2 -> Int.compare i1 i2 | FVuniv_var _, (FVnamed _ | FVrel _) -> 1 | FVuniv_var _, FVevar _ -> -1 | FVevar _, (FVnamed _ | FVrel _ | FVuniv_var _) -> 1 | FVevar e1, FVevar e2 -> Evar.compare e1 e2 end module FvMap = Map.Make(Fv_elem) (*spiwack: both type have been moved from Cbytegen because I needed then for the retroknowledge *) type vm_env = { size : int; (* longueur de la liste [n] *) fv_rev : fv_elem list; (* [fvn; ... ;fv1] *) fv_fwd : int FvMap.t; (* reverse mapping *) } type comp_env = { arity : int; (* arity of the current function, 0 if none *) nb_uni_stack : int ; (* number of universes on the stack, *) (* universes are always at the bottom. *) nb_stack : int; (* number of variables on the stack *) in_stack : int list; (* position in the stack *) nb_rec : int; (* number of mutually recursive functions *) pos_rec : instruction list; (* instruction d'acces pour les variables *) (* de point fix ou de cofix *) offset : int; in_env : vm_env ref (* The free variables of the expression *) } module Config = struct let stack_threshold = 256 (* see byterun/coq_memory.h *) let stack_safety_margin = 15 end type argument = ArgLambda of lambda | ArgUniv of Univ.Level.t let empty_fv = { size= 0; fv_rev = []; fv_fwd = FvMap.empty } let push_fv d e = { size = e.size + 1; fv_rev = d :: e.fv_rev; fv_fwd = FvMap.add d e.size e.fv_fwd; } let fv r = !(r.in_env) let empty_comp_env ()= { arity = 0; nb_uni_stack = 0; nb_stack = 0; in_stack = []; nb_rec = 0; pos_rec = []; offset = 0; in_env = ref empty_fv } (* Maximal stack size reached during the current function body. Used to reallocate the stack if we lack space. *) let max_stack_size = ref 0 let set_max_stack_size stack_size = if stack_size > !max_stack_size then max_stack_size := stack_size let ensure_stack_capacity f x = let old = !max_stack_size in max_stack_size := 0; let code = f x in let used_safe = !max_stack_size + Config.stack_safety_margin in max_stack_size := old; if used_safe > Config.stack_threshold then Kensurestackcapacity used_safe :: code else code (*i Creation functions for comp_env *) let rec add_param n sz l = if Int.equal n 0 then l else add_param (n - 1) sz (n+sz::l) let comp_env_fun ?(univs=0) arity = { arity; nb_uni_stack = univs ; nb_stack = arity; in_stack = add_param arity 0 []; nb_rec = 0; pos_rec = []; offset = 1; in_env = ref empty_fv } let comp_env_fix_type rfv = { arity = 0; nb_uni_stack = 0; nb_stack = 0; in_stack = []; nb_rec = 0; pos_rec = []; offset = 1; in_env = rfv } let comp_env_fix ndef curr_pos arity rfv = let prec = ref [] in for i = ndef downto 1 do prec := Koffsetclosure (2 * (ndef - curr_pos - i)) :: !prec done; { arity; nb_uni_stack = 0; nb_stack = arity; in_stack = add_param arity 0 []; nb_rec = ndef; pos_rec = !prec; offset = 2 * (ndef - curr_pos - 1)+1; in_env = rfv } let comp_env_cofix_type ndef rfv = { arity = 0; nb_uni_stack = 0; nb_stack = 0; in_stack = []; nb_rec = 0; pos_rec = []; offset = 1+ndef; in_env = rfv } let comp_env_cofix ndef arity rfv = let prec = ref [] in for i = 1 to ndef do prec := Kenvacc i :: !prec done; { arity; nb_uni_stack = 0; nb_stack = arity; in_stack = add_param arity 0 []; nb_rec = ndef; pos_rec = !prec; offset = ndef+1; in_env = rfv } (* [push_param ] add function parameters on the stack *) let push_param n sz r = { r with nb_stack = r.nb_stack + n; in_stack = add_param n sz r.in_stack } (* [push_local sz r] add a new variable on the stack at position [sz] *) let push_local sz r = { r with nb_stack = r.nb_stack + 1; in_stack = (sz + 1) :: r.in_stack } (*i Compilation of variables *) let find_at fv env = FvMap.find fv env.fv_fwd let pos_named id r = let env = !(r.in_env) in let cid = FVnamed id in try Kenvacc(r.offset + find_at cid env) with Not_found -> let pos = env.size in r.in_env := push_fv cid env; Kenvacc (r.offset + pos) let pos_rel i r sz = if i <= r.nb_stack then Kacc(sz - (List.nth r.in_stack (i-1))) else let i = i - r.nb_stack in if i <= r.nb_rec then try List.nth r.pos_rec (i-1) with (Failure _|Invalid_argument _) -> assert false else let i = i - r.nb_rec in let db = FVrel(i) in let env = !(r.in_env) in try Kenvacc(r.offset + find_at db env) with Not_found -> let pos = env.size in r.in_env := push_fv db env; Kenvacc(r.offset + pos) let pos_universe_var i r sz = (* Compilation of a universe variable can happen either at toplevel (the current closure correspond to a constant and has local universes) or in a local closure (which has no local universes). *) if r.nb_uni_stack != 0 then (* Universe variables are represented by De Bruijn levels (not indices), starting at 0. The shape of the stack will be [v1|..|vn|u1..up|arg1..argq] with size = n + p + q, and q = r.arity. So Kacc (sz - r.arity - 1) will access the last universe. *) Kacc (sz - r.arity - (r.nb_uni_stack - i)) else let env = !(r.in_env) in let db = FVuniv_var i in try Kenvacc (r.offset + find_at db env) with Not_found -> let pos = env.size in r.in_env := push_fv db env; Kenvacc(r.offset + pos) let pos_evar evk r = let env = !(r.in_env) in let cid = FVevar evk in try Kenvacc(r.offset + find_at cid env) with Not_found -> let pos = env.size in r.in_env := push_fv cid env; Kenvacc (r.offset + pos) (*i Examination of the continuation *) (* Discard all instructions up to the next label. *) (* This function is to be applied to the continuation before adding a *) (* non-terminating instruction (branch, raise, return, appterm) *) (* in front of it. *) let discard_dead_code cont = cont (*function [] -> [] | (Klabel _ | Krestart ) :: _ as cont -> cont | _ :: cont -> discard_dead_code cont *) (* Return a label to the beginning of the given continuation. *) (* If the sequence starts with a branch, use the target of that branch *) (* as the label, thus avoiding a jump to a jump. *) let label_code = function | Klabel lbl :: _ as cont -> (lbl, cont) | Kbranch lbl :: _ as cont -> (lbl, cont) | cont -> let lbl = Label.create() in (lbl, Klabel lbl :: cont) (* Return a branch to the continuation. That is, an instruction that, when executed, branches to the continuation or performs what the continuation performs. We avoid generating branches to returns. *) (* spiwack: make_branch was only used once. Changed it back to the ZAM one to match the appropriate semantics (old one avoided the introduction of an unconditional branch operation, which seemed appropriate for the 31-bit integers' code). As a memory, I leave the former version in this comment. let make_branch cont = match cont with | (Kreturn _ as return) :: cont' -> return, cont' | Klabel lbl as b :: _ -> b, cont | _ -> let b = Klabel(Label.create()) in b,b::cont *) let rec make_branch_2 lbl n cont = function Kreturn m :: _ -> (Kreturn (n + m), cont) | Klabel _ :: c -> make_branch_2 lbl n cont c | Kpop m :: c -> make_branch_2 lbl (n + m) cont c | _ -> match lbl with Some lbl -> (Kbranch lbl, cont) | None -> let lbl = Label.create() in (Kbranch lbl, Klabel lbl :: cont) let make_branch cont = match cont with (Kbranch _ as branch) :: _ -> (branch, cont) | (Kreturn _ as return) :: _ -> (return, cont) | Klabel lbl :: _ -> make_branch_2 (Some lbl) 0 cont cont | _ -> make_branch_2 (None) 0 cont cont (* Check if we're in tailcall position *) let rec is_tailcall = function | Kreturn k :: _ -> Some k | Klabel _ :: c -> is_tailcall c | _ -> None (* Extension of the continuation *) (* Add a Kpop n instruction in front of a continuation *) let rec add_pop n = function | Kpop m :: cont -> add_pop (n+m) cont | Kreturn m:: cont -> Kreturn (n+m) ::cont | cont -> if Int.equal n 0 then cont else Kpop n :: cont let add_grab arity lbl cont = if Int.equal arity 1 then Klabel lbl :: cont else Krestart :: Klabel lbl :: Kgrab (arity - 1) :: cont let add_grabrec rec_arg arity lbl cont = if Int.equal arity 1 && rec_arg < arity then Klabel lbl :: Kgrabrec 0 :: Krestart :: cont else Krestart :: Klabel lbl :: Kgrabrec rec_arg :: Krestart :: Kgrab (arity - 1) :: cont (* continuation of a cofix *) let cont_cofix arity = (* accu = res *) (* stk = ai::args::ra::... *) (* ai = [At|accumulate|[Cfx_t|fcofix]|args] *) [ Kpush; Kpush; (* stk = res::res::ai::args::ra::... *) Kacc 2; Kfield 1; Kfield 0; Kmakeblock(2, cofix_evaluated_tag); Kpush; (* stk = [Cfxe_t|fcofix|res]::res::ai::args::ra::...*) Kacc 2; Ksetfield 1; (* ai = [At|accumulate|[Cfxe_t|fcofix|res]|args] *) (* stk = res::ai::args::ra::... *) Kacc 0; (* accu = res *) Kreturn (arity+2) ] (* Code of closures *) let fun_code = ref [] let init_fun_code () = fun_code := [] (* Compilation of constructors and inductive types *) (* If [tag] hits the OCaml limitation for non constant constructors, we switch to another representation for the remaining constructors: [last_variant_tag|tag - Obj.last_non_constant_constructor_tag|args] We subtract Obj.last_non_constant_constructor_tag for efficiency of match interpretation. *) let nest_block tag arity cont = Kconst (Const_b0 (tag - Obj.last_non_constant_constructor_tag)) :: Kmakeblock(arity+1, Obj.last_non_constant_constructor_tag) :: cont let code_makeblock ~stack_size ~arity ~tag cont = if tag < Obj.last_non_constant_constructor_tag then Kmakeblock(arity, tag) :: cont else begin set_max_stack_size (stack_size + 1); Kpush :: nest_block tag arity cont end let compile_structured_constant _cenv sc sz cont = set_max_stack_size sz; Kconst sc :: cont (* compiling application *) let comp_args comp_expr cenv args sz cont = let nargs_m_1 = Array.length args - 1 in let c = ref (comp_expr cenv args.(0) (sz + nargs_m_1) cont) in for i = 1 to nargs_m_1 do c := comp_expr cenv args.(i) (sz + nargs_m_1 - i) (Kpush :: !c) done; !c let comp_app comp_fun comp_arg cenv f args sz cont = let nargs = Array.length args in if Int.equal nargs 0 then comp_fun cenv f sz cont else match is_tailcall cont with | Some k -> comp_args comp_arg cenv args sz (Kpush :: comp_fun cenv f (sz + nargs) (Kappterm(nargs, k + nargs) :: (discard_dead_code cont))) | None -> if nargs <= 4 then comp_args comp_arg cenv args sz (Kpush :: (comp_fun cenv f (sz+nargs) (Kapply nargs :: cont))) else let lbl,cont1 = label_code cont in Kpush_retaddr lbl :: (comp_args comp_arg cenv args (sz + 3) (Kpush :: (comp_fun cenv f (sz+3+nargs) (Kapply nargs :: cont1)))) (* Compiling free variables *) let compile_fv_elem cenv fv sz cont = match fv with | FVrel i -> pos_rel i cenv sz :: cont | FVnamed id -> pos_named id cenv :: cont | FVuniv_var i -> pos_universe_var i cenv sz :: cont | FVevar evk -> pos_evar evk cenv :: cont let rec compile_fv cenv l sz cont = match l with | [] -> cont | [fvn] -> set_max_stack_size (sz + 1); compile_fv_elem cenv fvn sz cont | fvn :: tl -> compile_fv_elem cenv fvn sz (Kpush :: compile_fv cenv tl (sz + 1) cont) (* Compiling constants *) let rec get_alias env kn = let cb = lookup_constant kn env in let tps = cb.const_body_code in match tps with | None -> kn | Some tps -> (match Cemitcodes.force tps with | BCalias kn' -> get_alias env kn' | _ -> kn) (* sz is the size of the local stack *) let rec compile_lam env cenv lam sz cont = set_max_stack_size sz; match lam with | Lrel(_, i) -> pos_rel i cenv sz :: cont | Lint i -> compile_structured_constant cenv (Const_b0 i) sz cont | Lval v -> compile_structured_constant cenv (Const_val v) sz cont | Luint i -> compile_structured_constant cenv (Const_uint i) sz cont | Lfloat f -> compile_structured_constant cenv (Const_float f) sz cont | Lproj (p,arg) -> compile_lam env cenv arg sz (Kproj p :: cont) | Lvar id -> pos_named id cenv :: cont | Levar (evk, args) -> if Array.is_empty args then compile_fv_elem cenv (FVevar evk) sz cont else (** Arguments are reversed in evar instances *) let args = Array.copy args in let () = Array.rev args in comp_app compile_fv_elem (compile_lam env) cenv (FVevar evk) args sz cont | Lconst (kn,u) -> compile_constant env cenv kn u [||] sz cont | Lind (ind,u) -> if Univ.Instance.is_empty u then compile_structured_constant cenv (Const_ind ind) sz cont else comp_app compile_structured_constant compile_universe cenv (Const_ind ind) (Univ.Instance.to_array u) sz cont | Lsort (Sorts.SProp | Sorts.Prop | Sorts.Set as s) -> compile_structured_constant cenv (Const_sort s) sz cont | Lsort (Sorts.Type u) -> (* We represent universes as a global constant with local universes "compacted", i.e. as [u arg0 ... argn] where we will substitute (after evaluation) [Var 0,...,Var n] with values of [arg0,...,argn] *) let u,s = Univ.compact_univ u in let compile_get_univ cenv idx sz cont = set_max_stack_size sz; compile_fv_elem cenv (FVuniv_var idx) sz cont in if List.is_empty s then compile_structured_constant cenv (Const_sort (Sorts.sort_of_univ u)) sz cont else comp_app compile_structured_constant compile_get_univ cenv (Const_sort (Sorts.sort_of_univ u)) (Array.of_list s) sz cont | Llet (_id,def,body) -> compile_lam env cenv def sz (Kpush :: compile_lam env (push_local sz cenv) body (sz+1) (add_pop 1 cont)) | Lprod (dom,codom) -> let cont1 = Kpush :: compile_lam env cenv dom (sz+1) (Kmakeprod :: cont) in compile_lam env cenv codom sz cont1 | Llam (ids,body) -> let arity = Array.length ids in let r_fun = comp_env_fun arity in let lbl_fun = Label.create() in let cont_fun = ensure_stack_capacity (compile_lam env r_fun body arity) [Kreturn arity] in fun_code := [Ksequence(add_grab arity lbl_fun cont_fun,!fun_code)]; let fv = fv r_fun in compile_fv cenv fv.fv_rev sz (Kclosure(lbl_fun,fv.size) :: cont) | Lapp (f, args) -> begin match f with | Lconst (kn,u) -> compile_constant env cenv kn u args sz cont | _ -> comp_app (compile_lam env) (compile_lam env) cenv f args sz cont end | Lfix ((rec_args, init), (_decl, types, bodies)) -> let ndef = Array.length types in let rfv = ref empty_fv in let lbl_types = Array.make ndef Label.no in let lbl_bodies = Array.make ndef Label.no in (* Compiling types *) let env_type = comp_env_fix_type rfv in for i = 0 to ndef - 1 do let fcode = ensure_stack_capacity (compile_lam env env_type types.(i) 0) [Kstop] in let lbl,fcode = label_code fcode in lbl_types.(i) <- lbl; fun_code := [Ksequence(fcode,!fun_code)] done; (* Compiling bodies *) for i = 0 to ndef - 1 do let params,body = decompose_Llam bodies.(i) in let arity = Array.length params in let env_body = comp_env_fix ndef i arity rfv in let cont1 = ensure_stack_capacity (compile_lam env env_body body arity) [Kreturn arity] in let lbl = Label.create () in lbl_bodies.(i) <- lbl; let fcode = add_grabrec rec_args.(i) arity lbl cont1 in fun_code := [Ksequence(fcode,!fun_code)] done; let fv = !rfv in compile_fv cenv fv.fv_rev sz (Kclosurerec(fv.size,init,lbl_types,lbl_bodies) :: cont) | Lcofix(init, (_decl,types,bodies)) -> let ndef = Array.length types in let lbl_types = Array.make ndef Label.no in let lbl_bodies = Array.make ndef Label.no in (* Compiling types *) let rfv = ref empty_fv in let env_type = comp_env_cofix_type ndef rfv in for i = 0 to ndef - 1 do let fcode = ensure_stack_capacity (compile_lam env env_type types.(i) 0) [Kstop] in let lbl,fcode = label_code fcode in lbl_types.(i) <- lbl; fun_code := [Ksequence(fcode,!fun_code)] done; (* Compiling bodies *) for i = 0 to ndef - 1 do let params,body = decompose_Llam bodies.(i) in let arity = Array.length params in let env_body = comp_env_cofix ndef arity rfv in let lbl = Label.create () in let comp arity = (* 4 stack slots are needed to update the cofix when forced *) set_max_stack_size (arity + 4); compile_lam env env_body body (arity+1) (cont_cofix arity) in let cont = ensure_stack_capacity comp arity in lbl_bodies.(i) <- lbl; fun_code := [Ksequence(add_grab (arity+1) lbl cont,!fun_code)]; done; let fv = !rfv in set_max_stack_size (sz + fv.size + ndef + 2); compile_fv cenv fv.fv_rev sz (Kclosurecofix(fv.size, init, lbl_types, lbl_bodies) :: cont) | Lif(t, bt, bf) -> let branch, cont = make_branch cont in let lbl_true = Label.create() in let lbl_false = Label.create() in compile_lam env cenv t sz (Kswitch([|lbl_true;lbl_false|],[||]) :: Klabel lbl_false :: compile_lam env cenv bf sz (branch :: Klabel lbl_true :: compile_lam env cenv bt sz cont)) | Lcase(ci,rtbl,t,a,branches) -> let ind = ci.ci_ind in let mib = lookup_mind (fst ind) env in let oib = mib.mind_packets.(snd ind) in let lbl_consts = Array.make oib.mind_nb_constant Label.no in let nallblock = oib.mind_nb_args + 1 in (* +1 : accumulate *) let nconst = Array.length branches.constant_branches in let nblock = min nallblock (Obj.last_non_constant_constructor_tag + 1) in let lbl_blocks = Array.make nblock Label.no in let neblock = max 0 (nallblock - Obj.last_non_constant_constructor_tag) in let lbl_eblocks = Array.make neblock Label.no in let branch1, cont = make_branch cont in (* Compilation of the return type *) let fcode = ensure_stack_capacity (compile_lam env cenv t sz) [Kpop sz; Kstop] in let lbl_typ,fcode = label_code fcode in fun_code := [Ksequence(fcode,!fun_code)]; (* Compilation of the branches *) let lbl_sw = Label.create () in let sz_b,branch,is_tailcall = match branch1 with | Kreturn k -> assert (Int.equal k sz) ; sz, branch1, true | Kbranch _ -> sz+3, Kjump, false | _ -> assert false in let c = ref cont in (* Perform the extra match if needed (too many block constructors) *) if neblock <> 0 then begin let lbl_b, code_b = label_code ( Kpush :: Kfield 0 :: Kswitch(lbl_eblocks, [||]) :: !c) in lbl_blocks.(Obj.last_non_constant_constructor_tag) <- lbl_b; c := code_b end; (* Compilation of constant branches *) for i = nconst - 1 downto 0 do let aux = compile_lam env cenv branches.constant_branches.(i) sz_b (branch::!c) in let lbl_b,code_b = label_code aux in lbl_consts.(i) <- lbl_b; c := code_b done; (* -1 for accu branch *) for i = nallblock - 2 downto 0 do let tag = i + 1 in let (ids, body) = branches.nonconstant_branches.(i) in let arity = Array.length ids in let code_b = compile_lam env (push_param arity sz_b cenv) body (sz_b+arity) (add_pop arity (branch::!c)) in let code_b = if tag < Obj.last_non_constant_constructor_tag then begin set_max_stack_size (sz_b + arity); Kpushfields arity :: code_b end else begin set_max_stack_size (sz_b + arity + 1); Kacc 0::Kpop 1::Kpushfields(arity+1)::Kpop 1::code_b end in let lbl_b, code_b = label_code code_b in if tag < Obj.last_non_constant_constructor_tag then lbl_blocks.(tag) <- lbl_b else lbl_eblocks.(tag - Obj.last_non_constant_constructor_tag) <- lbl_b; c := code_b done; let annot = {ci = ci; rtbl = rtbl; tailcall = is_tailcall; max_stack_size = !max_stack_size - sz} in (* Compiling branch for accumulators *) let lbl_accu, code_accu = set_max_stack_size (sz+3); label_code(Kmakeswitchblock(lbl_typ,lbl_sw,annot,sz) :: branch :: !c) in lbl_blocks.(0) <- lbl_accu; c := Klabel lbl_sw :: Kswitch(lbl_consts,lbl_blocks) :: code_accu; let code_sw = match branch1 with (* spiwack : branch1 can't be a lbl anymore it's a Branch instead | Klabel lbl -> Kpush_retaddr lbl :: !c *) | Kbranch lbl -> Kpush_retaddr lbl :: !c | _ -> !c in compile_lam env cenv a sz code_sw | Lmakeblock (tag,args) -> let arity = Array.length args in let cont = code_makeblock ~stack_size:(sz+arity-1) ~arity ~tag cont in comp_args (compile_lam env) cenv args sz cont | Lprim (kn, op, args) -> comp_args (compile_lam env) cenv args sz (Kprim(op, kn)::cont) and compile_get_global cenv (kn,u) sz cont = set_max_stack_size sz; if Univ.Instance.is_empty u then Kgetglobal kn :: cont else comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont) compile_universe cenv () (Univ.Instance.to_array u) sz cont and compile_universe cenv uni sz cont = set_max_stack_size sz; match Univ.Level.var_index uni with | None -> compile_structured_constant cenv (Const_univ_level uni) sz cont | Some idx -> pos_universe_var idx cenv sz :: cont and compile_constant env cenv kn u args sz cont = set_max_stack_size sz; if Univ.Instance.is_empty u then (* normal compilation *) comp_app (fun _ _ sz cont -> compile_get_global cenv (kn,u) sz cont) (compile_lam env) cenv () args sz cont else let compile_arg cenv constr_or_uni sz cont = match constr_or_uni with | ArgLambda t -> compile_lam env cenv t sz cont | ArgUniv uni -> compile_universe cenv uni sz cont in let u = Univ.Instance.to_array u in let lu = Array.length u in let all = Array.init (lu + Array.length args) (fun i -> if i < lu then ArgUniv u.(i) else ArgLambda args.(i-lu)) in comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont) compile_arg cenv () all sz cont let is_univ_copy max u = let u = Univ.Instance.to_array u in if Array.length u = max then Array.fold_left_i (fun i acc u -> if acc then match Univ.Level.var_index u with | None -> false | Some l -> l = i else false) true u else false let dump_bytecode = ref false let dump_bytecodes init code fvs = let open Pp in (str "code =" ++ fnl () ++ pp_bytecodes init ++ fnl () ++ pp_bytecodes code ++ fnl () ++ str "fv = " ++ prlist_with_sep (fun () -> str "; ") pp_fv_elem fvs ++ fnl ()) let compile ~fail_on_error ?universes:(universes=0) env c = init_fun_code (); Label.reset_label_counter (); let cont = [Kstop] in try let cenv, init_code = if Int.equal universes 0 then let lam = lambda_of_constr ~optimize:true env c in let cenv = empty_comp_env () in cenv, ensure_stack_capacity (compile_lam env cenv lam 0) cont else (* We are going to generate a lambda, but merge the universe closure * with the function closure if it exists. *) let lam = lambda_of_constr ~optimize:true env c in let params, body = decompose_Llam lam in let arity = Array.length params in let cenv = empty_comp_env () in let full_arity = arity + universes in let r_fun = comp_env_fun ~univs:universes arity in let lbl_fun = Label.create () in let cont_fun = ensure_stack_capacity (compile_lam env r_fun body full_arity) [Kreturn full_arity] in fun_code := [Ksequence(add_grab full_arity lbl_fun cont_fun,!fun_code)]; let fv = fv r_fun in let init_code = ensure_stack_capacity (compile_fv cenv fv.fv_rev 0) (Kclosure(lbl_fun,fv.size) :: cont) in cenv, init_code in let fv = List.rev (!(cenv.in_env).fv_rev) in (if !dump_bytecode then Feedback.msg_debug (dump_bytecodes init_code !fun_code fv)) ; Some (init_code,!fun_code, Array.of_list fv) with TooLargeInductive msg -> let fn = if fail_on_error then CErrors.user_err ?loc:None ~hdr:"compile" else (fun x -> Feedback.msg_warning x) in fn msg; None let compile_constant_body ~fail_on_error env univs = function | Undef _ | OpaqueDef _ | Primitive _ -> Some BCconstant | Def sb -> let body = Mod_subst.force_constr sb in let instance_size = Univ.AUContext.size (Declareops.universes_context univs) in match kind body with | Const (kn',u) when is_univ_copy instance_size u -> (* we use the canonical name of the constant*) let con= Constant.make1 (Constant.canonical kn') in Some (BCalias (get_alias env con)) | _ -> let res = compile ~fail_on_error ~universes:instance_size env body in Option.map (fun x -> BCdefined (to_memory x)) res (* Shortcut of the previous function used during module strengthening *) let compile_alias kn = BCalias (Constant.make1 (Constant.canonical kn)) coq-8.11.0/kernel/evar.ml0000644000175000017500000000155513612664533015043 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* global -> unit val mk_open : string -> global (* Precomputed values for a compilation unit *) val clear_symbols : unit -> unit val get_value : symbols -> int -> Nativevalues.t val get_sort : symbols -> int -> Sorts.t val get_name : symbols -> int -> Name.t val get_const : symbols -> int -> Constant.t val get_match : symbols -> int -> Nativevalues.annot_sw val get_ind : symbols -> int -> inductive val get_meta : symbols -> int -> metavariable val get_evar : symbols -> int -> Evar.t val get_level : symbols -> int -> Univ.Level.t val get_proj : symbols -> int -> inductive * int val get_symbols : unit -> symbols type code_location_update type code_location_updates type linkable_code = global list * code_location_updates val clear_global_tbl : unit -> unit val empty_updates : code_location_updates val register_native_file : string -> unit val compile_constant_field : env -> string -> Constant.t -> global list -> 'a constant_body -> global list val compile_mind_field : ModPath.t -> Label.t -> global list -> mutual_inductive_body -> global list val mk_conv_code : env -> evars -> string -> constr -> constr -> linkable_code val mk_norm_code : env -> evars -> string -> constr -> linkable_code val mk_library_header : DirPath.t -> global list val mod_uid_of_dirpath : DirPath.t -> string val link_info_of_dirpath : DirPath.t -> link_info val update_locations : code_location_updates -> unit val add_header_comment : global list -> string -> global list coq-8.11.0/kernel/opaqueproof.mli0000644000175000017500000000520413612664533016612 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* proofterm -> opaquetab -> opaque * opaquetab type work_list = (Univ.Instance.t * Id.t array) Cmap.t * (Univ.Instance.t * Id.t array) Mindmap.t type cooking_info = { modlist : work_list; abstract : Constr.named_context * Univ.Instance.t * Univ.AUContext.t } type opaque_proofterm = (Constr.t * unit delayed_universes) option type indirect_accessor = { access_proof : DirPath.t -> int -> opaque_proofterm; access_discharge : cooking_info list -> (Constr.t * unit delayed_universes) -> (Constr.t * unit delayed_universes); } (** Opaque terms are indexed by their library dirpath and an integer index. The two functions above activate this indirect storage, by telling how to retrieve terms. *) (** From a [opaque] back to a [constr]. This might use the indirect opaque accessor given as an argument. *) val force_proof : indirect_accessor -> opaquetab -> opaque -> constr * unit delayed_universes val force_constraints : indirect_accessor -> opaquetab -> opaque -> Univ.ContextSet.t val subst_opaque : substitution -> opaque -> opaque val discharge_opaque : cooking_info -> opaque -> opaque val join_opaque : ?except:Future.UUIDSet.t -> opaquetab -> opaque -> unit val dump : ?except:Future.UUIDSet.t -> opaquetab -> opaque_proofterm array * int Future.UUIDMap.t coq-8.11.0/kernel/entries.ml0000644000175000017500000001046413612664533015556 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* l_sl x Int64.one); Callback.register "uint63 lsr" l_sr; Callback.register "uint63 lsr1" (fun x -> l_sr x Int64.one); Callback.register "uint63 lt" lt; Callback.register "uint63 lxor" l_xor; Callback.register "uint63 mod" rem; Callback.register "uint63 mul" mul; Callback.register "uint63 mulc_ml" mulc; Callback.register "uint63 one" one; Callback.register "uint63 sub" sub; Callback.register "uint63 subcarry" subcarry; Callback.register "uint63 tail0" tail0; Callback.register "uint63 of_float" of_float; Callback.register "uint63 to_float" to_float; Callback.register "uint63 of_int" of_int; Callback.register "uint63 to_int_min" to_int_min coq-8.11.0/kernel/modops.mli0000644000175000017500000001220613612664533015553 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* bool val destr_functor : ('ty,'a) functorize -> MBId.t * 'ty * ('ty,'a) functorize val destr_nofunctor : ('ty,'a) functorize -> 'a (** Conversions between [module_body] and [module_type_body] *) val module_type_of_module : module_body -> module_type_body val module_body_of_type : ModPath.t -> module_type_body -> module_body val check_modpath_equiv : env -> ModPath.t -> ModPath.t -> unit val implem_smartmap : (module_signature -> module_signature) -> (module_expression -> module_expression) -> (module_implementation -> module_implementation) (** {6 Substitutions } *) val subst_signature : substitution -> module_signature -> module_signature val subst_structure : substitution -> structure_body -> structure_body (** {6 Adding to an environment } *) val add_structure : ModPath.t -> structure_body -> delta_resolver -> env -> env (** adds a module and its components, but not the constraints *) val add_module : module_body -> env -> env (** same as add_module, but for a module whose native code has been linked by the native compiler. The linking information is updated. *) val add_linked_module : module_body -> link_info -> env -> env (** same, for a module type *) val add_module_type : ModPath.t -> module_type_body -> env -> env val add_retroknowledge : module_implementation module_retroknowledge -> env -> env (** {6 Strengthening } *) val strengthen : module_type_body -> ModPath.t -> module_type_body val inline_delta_resolver : env -> inline -> ModPath.t -> MBId.t -> module_type_body -> delta_resolver -> delta_resolver val strengthen_and_subst_mb : module_body -> ModPath.t -> bool -> module_body val subst_modtype_and_resolver : module_type_body -> ModPath.t -> module_type_body (** {6 Cleaning a module expression from bounded parts } For instance: functor(X:T)->struct module M:=X end) becomes: functor(X:T)->struct module M:= end) *) val clean_bounded_mod_expr : module_signature -> module_signature (** {6 Stm machinery } *) val join_structure : Future.UUIDSet.t -> Opaqueproof.opaquetab -> structure_body -> unit (** {6 Errors } *) type signature_mismatch_error = | InductiveFieldExpected of mutual_inductive_body | DefinitionFieldExpected | ModuleFieldExpected | ModuleTypeFieldExpected | NotConvertibleInductiveField of Id.t | NotConvertibleConstructorField of Id.t | NotConvertibleBodyField | NotConvertibleTypeField of env * types * types | CumulativeStatusExpected of bool | PolymorphicStatusExpected of bool | NotSameConstructorNamesField | NotSameInductiveNameInBlockField | FiniteInductiveFieldExpected of bool | InductiveNumbersFieldExpected of int | InductiveParamsNumberField of int | RecordFieldExpected of bool | RecordProjectionsExpected of Name.t list | NotEqualInductiveAliases | IncompatibleUniverses of Univ.univ_inconsistency | IncompatiblePolymorphism of env * types * types | IncompatibleConstraints of { got : Univ.AUContext.t; expect : Univ.AUContext.t } | IncompatibleVariance type module_typing_error = | SignatureMismatch of Label.t * structure_field_body * signature_mismatch_error | LabelAlreadyDeclared of Label.t | ApplicationToNotPath of module_struct_entry | NotAFunctor | IsAFunctor | IncompatibleModuleTypes of module_type_body * module_type_body | NotEqualModulePaths of ModPath.t * ModPath.t | NoSuchLabel of Label.t | IncompatibleLabels of Label.t * Label.t | NotAModule of string | NotAModuleType of string | NotAConstant of Label.t | IncorrectWithConstraint of Label.t | GenerativeModuleExpected of Label.t | LabelMissing of Label.t * string | IncludeRestrictedFunctor of ModPath.t exception ModuleTypingError of module_typing_error val error_existing_label : Label.t -> 'a val error_incompatible_modtypes : module_type_body -> module_type_body -> 'a val error_signature_mismatch : Label.t -> structure_field_body -> signature_mismatch_error -> 'a val error_incompatible_labels : Label.t -> Label.t -> 'a val error_no_such_label : Label.t -> 'a val error_not_a_module : string -> 'a val error_not_a_constant : Label.t -> 'a val error_incorrect_with_constraint : Label.t -> 'a val error_generative_module_expected : Label.t -> 'a val error_no_such_label_sub : Label.t->string->'a val error_include_restricted_functor : ModPath.t -> 'a coq-8.11.0/kernel/nativelib.mli0000644000175000017500000000357313612664533016236 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* string list) ref val load_obj : (string -> unit) ref val get_ml_filename : unit -> string * string (** [compile file code ~profile] will compile native [code] to [file], and return the name of the object file; this name depends on whether are in byte mode or not; file is expected to be .ml file *) val compile : string -> global list -> profile:bool -> string (** [compile_library lib code file] is similar to [compile file code] but will perform some extra tweaks to handle [code] as a Coq lib. *) val compile_library : Names.DirPath.t -> global list -> string -> unit val call_linker : ?fatal:bool -> Environ.env -> prefix:string -> string -> code_location_updates option -> unit val link_library : Environ.env -> prefix:string -> dirname:string -> basename:string -> unit val rt1 : Nativevalues.t ref val rt2 : Nativevalues.t ref val get_library_native_symbols : Names.DirPath.t -> Nativevalues.symbols (** Strictly for usage by code produced by native compute. *) coq-8.11.0/kernel/type_errors.ml0000644000175000017500000002007313612664533016457 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* NonInformativeToInformative | InType, InSet -> StrongEliminationOnNonSmallType (* if Set impredicative *) | _ -> WrongArity let error_unsatisfied_constraints env c = raise (TypeError (env, UnsatisfiedConstraints c)) let error_undeclared_universe env l = raise (TypeError (env, UndeclaredUniverse l)) let error_disallowed_sprop env = raise (TypeError (env, DisallowedSProp)) let error_bad_relevance env = raise (TypeError (env, BadRelevance)) let map_pguard_error f = function | NotEnoughAbstractionInFixBody -> NotEnoughAbstractionInFixBody | RecursionNotOnInductiveType c -> RecursionNotOnInductiveType (f c) | RecursionOnIllegalTerm (n, (env, c), l1, l2) -> RecursionOnIllegalTerm (n, (env, f c), l1, l2) | NotEnoughArgumentsForFixCall n -> NotEnoughArgumentsForFixCall n | CodomainNotInductiveType c -> CodomainNotInductiveType (f c) | NestedRecursiveOccurrences -> NestedRecursiveOccurrences | UnguardedRecursiveCall c -> UnguardedRecursiveCall (f c) | RecCallInTypeOfAbstraction c -> RecCallInTypeOfAbstraction (f c) | RecCallInNonRecArgOfConstructor c -> RecCallInNonRecArgOfConstructor (f c) | RecCallInTypeOfDef c -> RecCallInTypeOfDef (f c) | RecCallInCaseFun c -> RecCallInCaseFun (f c) | RecCallInCaseArg c -> RecCallInCaseArg (f c) | RecCallInCasePred c -> RecCallInCasePred (f c) | NotGuardedForm c -> NotGuardedForm (f c) | ReturnPredicateNotCoInductive c -> ReturnPredicateNotCoInductive (f c) | FixpointOnIrrelevantInductive -> FixpointOnIrrelevantInductive let map_ptype_error f = function | UnboundRel n -> UnboundRel n | UnboundVar id -> UnboundVar id | NotAType j -> NotAType (on_judgment f j) | BadAssumption j -> BadAssumption (on_judgment f j) | ReferenceVariables (id, c) -> ReferenceVariables (id, f c) | ElimArity (pi, c, j, ar) -> ElimArity (pi, f c, on_judgment f j, ar) | CaseNotInductive j -> CaseNotInductive (on_judgment f j) | WrongCaseInfo (pi, ci) -> WrongCaseInfo (pi, ci) | NumberBranches (j, n) -> NumberBranches (on_judgment f j, n) | IllFormedBranch (c, pc, t1, t2) -> IllFormedBranch (f c, pc, f t1, f t2) | Generalization ((na, t), j) -> Generalization ((na, f t), on_judgment f j) | ActualType (j, t) -> ActualType (on_judgment f j, f t) | IncorrectPrimitive (p, t) -> IncorrectPrimitive ({p with uj_type=f p.uj_type}, f t) | CantApplyBadType ((n, c1, c2), j, vj) -> CantApplyBadType ((n, f c1, f c2), on_judgment f j, Array.map (on_judgment f) vj) | CantApplyNonFunctional (j, jv) -> CantApplyNonFunctional (on_judgment f j, Array.map (on_judgment f) jv) | IllFormedRecBody (ge, na, n, env, jv) -> IllFormedRecBody (map_pguard_error f ge, na, n, env, Array.map (on_judgment f) jv) | IllTypedRecBody (n, na, jv, t) -> IllTypedRecBody (n, na, Array.map (on_judgment f) jv, Array.map f t) | UnsatisfiedConstraints g -> UnsatisfiedConstraints g | UndeclaredUniverse l -> UndeclaredUniverse l | DisallowedSProp -> DisallowedSProp | BadRelevance -> BadRelevance coq-8.11.0/kernel/primred.ml0000644000175000017500000002601613612664533015547 0ustar treinentreinen(* Reduction of native operators *) open Names open CPrimitives open Retroknowledge open Environ open CErrors let add_retroknowledge env action = match action with | Register_type(PT_int63,c) -> let retro = env.retroknowledge in let retro = match retro.retro_int63 with | None -> { retro with retro_int63 = Some c } | Some c' -> assert (Constant.equal c c'); retro in set_retroknowledge env retro | Register_type(PT_float64,c) -> let retro = env.retroknowledge in let retro = match retro.retro_float64 with | None -> { retro with retro_float64 = Some c } | Some c' -> assert (Constant.equal c c'); retro in set_retroknowledge env retro | Register_ind(pit,ind) -> let retro = env.retroknowledge in let retro = match pit with | PIT_bool -> let r = match retro.retro_bool with | None -> ((ind,1), (ind,2)) | Some (((ind',_),_) as t) -> assert (eq_ind ind ind'); t in { retro with retro_bool = Some r } | PIT_carry -> let r = match retro.retro_carry with | None -> ((ind,1), (ind,2)) | Some (((ind',_),_) as t) -> assert (eq_ind ind ind'); t in { retro with retro_carry = Some r } | PIT_pair -> let r = match retro.retro_pair with | None -> (ind,1) | Some ((ind',_) as t) -> assert (eq_ind ind ind'); t in { retro with retro_pair = Some r } | PIT_cmp -> let r = match retro.retro_cmp with | None -> ((ind,1), (ind,2), (ind,3)) | Some (((ind',_),_,_) as t) -> assert (eq_ind ind ind'); t in { retro with retro_cmp = Some r } | PIT_f_cmp -> let r = match retro.retro_f_cmp with | None -> ((ind,1), (ind,2), (ind,3), (ind,4)) | Some (((ind',_),_,_,_) as t) -> assert (eq_ind ind ind'); t in { retro with retro_f_cmp = Some r } | PIT_f_class -> let r = match retro.retro_f_class with | None -> ((ind,1), (ind,2), (ind,3), (ind,4), (ind,5), (ind,6), (ind,7), (ind,8), (ind,9)) | Some (((ind',_),_,_,_,_,_,_,_,_) as t) -> assert (eq_ind ind ind'); t in { retro with retro_f_class = Some r } in set_retroknowledge env retro let get_int_type env = match env.retroknowledge.retro_int63 with | Some c -> c | None -> anomaly Pp.(str"Reduction of primitive: int63 not registered") let get_float_type env = match env.retroknowledge.retro_float64 with | Some c -> c | None -> anomaly Pp.(str"Reduction of primitive: float64 not registered") let get_cmp_type env = match env.retroknowledge.retro_cmp with | Some (((mindcmp,_),_),_,_) -> Constant.make (MutInd.user mindcmp) (MutInd.canonical mindcmp) | None -> anomaly Pp.(str"Reduction of primitive: comparison not registered") let get_bool_constructors env = match env.retroknowledge.retro_bool with | Some r -> r | None -> anomaly Pp.(str"Reduction of primitive: bool not registered") let get_carry_constructors env = match env.retroknowledge.retro_carry with | Some r -> r | None -> anomaly Pp.(str"Reduction of primitive: carry not registered") let get_pair_constructor env = match env.retroknowledge.retro_pair with | Some c -> c | None -> anomaly Pp.(str"Reduction of primitive: pair not registered") let get_cmp_constructors env = match env.retroknowledge.retro_cmp with | Some r -> r | None -> anomaly Pp.(str"Reduction of primitive: cmp not registered") let get_f_cmp_constructors env = match env.retroknowledge.retro_f_cmp with | Some r -> r | None -> anomaly Pp.(str"Reduction of primitive: fcmp not registered") let get_f_class_constructors env = match env.retroknowledge.retro_f_class with | Some r -> r | None -> anomaly Pp.(str"Reduction of primitive: fclass not registered") exception NativeDestKO module type RedNativeEntries = sig type elem type args type evd (* will be unit in kernel, evar_map outside *) val get : args -> int -> elem val get_int : evd -> elem -> Uint63.t val get_float : evd -> elem -> Float64.t val mkInt : env -> Uint63.t -> elem val mkFloat : env -> Float64.t -> elem val mkBool : env -> bool -> elem val mkCarry : env -> bool -> elem -> elem (* true if carry *) val mkIntPair : env -> elem -> elem -> elem val mkFloatIntPair : env -> elem -> elem -> elem val mkLt : env -> elem val mkEq : env -> elem val mkGt : env -> elem val mkFLt : env -> elem val mkFEq : env -> elem val mkFGt : env -> elem val mkFNotComparable : env -> elem val mkPNormal : env -> elem val mkNNormal : env -> elem val mkPSubn : env -> elem val mkNSubn : env -> elem val mkPZero : env -> elem val mkNZero : env -> elem val mkPInf : env -> elem val mkNInf : env -> elem val mkNaN : env -> elem end module type RedNative = sig type elem type args type evd val red_prim : env -> evd -> CPrimitives.t -> args -> elem option end module RedNative (E:RedNativeEntries) : RedNative with type elem = E.elem with type args = E.args with type evd = E.evd = struct type elem = E.elem type args = E.args type evd = E.evd let get_int evd args i = E.get_int evd (E.get args i) let get_int1 evd args = get_int evd args 0 let get_int2 evd args = get_int evd args 0, get_int evd args 1 let get_int3 evd args = get_int evd args 0, get_int evd args 1, get_int evd args 2 let get_float evd args i = E.get_float evd (E.get args i) let get_float1 evd args = get_float evd args 0 let get_float2 evd args = get_float evd args 0, get_float evd args 1 let red_prim_aux env evd op args = let open CPrimitives in match op with | Int63head0 -> let i = get_int1 evd args in E.mkInt env (Uint63.head0 i) | Int63tail0 -> let i = get_int1 evd args in E.mkInt env (Uint63.tail0 i) | Int63add -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.add i1 i2) | Int63sub -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.sub i1 i2) | Int63mul -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.mul i1 i2) | Int63div -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.div i1 i2) | Int63mod -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.rem i1 i2) | Int63lsr -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_sr i1 i2) | Int63lsl -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_sl i1 i2) | Int63land -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_and i1 i2) | Int63lor -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_or i1 i2) | Int63lxor -> let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_xor i1 i2) | Int63addc -> let i1, i2 = get_int2 evd args in let s = Uint63.add i1 i2 in E.mkCarry env (Uint63.lt s i1) (E.mkInt env s) | Int63subc -> let i1, i2 = get_int2 evd args in let s = Uint63.sub i1 i2 in E.mkCarry env (Uint63.lt i1 i2) (E.mkInt env s) | Int63addCarryC -> let i1, i2 = get_int2 evd args in let s = Uint63.add (Uint63.add i1 i2) (Uint63.of_int 1) in E.mkCarry env (Uint63.le s i1) (E.mkInt env s) | Int63subCarryC -> let i1, i2 = get_int2 evd args in let s = Uint63.sub (Uint63.sub i1 i2) (Uint63.of_int 1) in E.mkCarry env (Uint63.le i1 i2) (E.mkInt env s) | Int63mulc -> let i1, i2 = get_int2 evd args in let (h, l) = Uint63.mulc i1 i2 in E.mkIntPair env (E.mkInt env h) (E.mkInt env l) | Int63diveucl -> let i1, i2 = get_int2 evd args in let q,r = Uint63.div i1 i2, Uint63.rem i1 i2 in E.mkIntPair env (E.mkInt env q) (E.mkInt env r) | Int63div21 -> let i1, i2, i3 = get_int3 evd args in let q,r = Uint63.div21 i1 i2 i3 in E.mkIntPair env (E.mkInt env q) (E.mkInt env r) | Int63addMulDiv -> let p, i, j = get_int3 evd args in E.mkInt env (Uint63.l_or (Uint63.l_sl i p) (Uint63.l_sr j (Uint63.sub (Uint63.of_int Uint63.uint_size) p))) | Int63eq -> let i1, i2 = get_int2 evd args in E.mkBool env (Uint63.equal i1 i2) | Int63lt -> let i1, i2 = get_int2 evd args in E.mkBool env (Uint63.lt i1 i2) | Int63le -> let i1, i2 = get_int2 evd args in E.mkBool env (Uint63.le i1 i2) | Int63compare -> let i1, i2 = get_int2 evd args in begin match Uint63.compare i1 i2 with | x when x < 0 -> E.mkLt env | 0 -> E.mkEq env | _ -> E.mkGt env end | Float64opp -> let f = get_float1 evd args in E.mkFloat env (Float64.opp f) | Float64abs -> let f = get_float1 evd args in E.mkFloat env (Float64.abs f) | Float64eq -> let i1, i2 = get_float2 evd args in E.mkBool env (Float64.eq i1 i2) | Float64lt -> let i1, i2 = get_float2 evd args in E.mkBool env (Float64.lt i1 i2) | Float64le -> let i1, i2 = get_float2 evd args in E.mkBool env (Float64.le i1 i2) | Float64compare -> let f1, f2 = get_float2 evd args in (match Float64.compare f1 f2 with | Float64.FEq -> E.mkFEq env | Float64.FLt -> E.mkFLt env | Float64.FGt -> E.mkFGt env | Float64.FNotComparable -> E.mkFNotComparable env) | Float64classify -> let f = get_float1 evd args in (match Float64.classify f with | Float64.PNormal -> E.mkPNormal env | Float64.NNormal -> E.mkNNormal env | Float64.PSubn -> E.mkPSubn env | Float64.NSubn -> E.mkNSubn env | Float64.PZero -> E.mkPZero env | Float64.NZero -> E.mkNZero env | Float64.PInf -> E.mkPInf env | Float64.NInf -> E.mkNInf env | Float64.NaN -> E.mkNaN env) | Float64add -> let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.add f1 f2) | Float64sub -> let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.sub f1 f2) | Float64mul -> let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.mul f1 f2) | Float64div -> let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.div f1 f2) | Float64sqrt -> let f = get_float1 evd args in E.mkFloat env (Float64.sqrt f) | Float64ofInt63 -> let i = get_int1 evd args in E.mkFloat env (Float64.of_int63 i) | Float64normfr_mantissa -> let f = get_float1 evd args in E.mkInt env (Float64.normfr_mantissa f) | Float64frshiftexp -> let f = get_float1 evd args in let (m,e) = Float64.frshiftexp f in E.mkFloatIntPair env (E.mkFloat env m) (E.mkInt env e) | Float64ldshiftexp -> let f = get_float evd args 0 in let e = get_int evd args 1 in E.mkFloat env (Float64.ldshiftexp f e) | Float64next_up -> let f = get_float1 evd args in E.mkFloat env (Float64.next_up f) | Float64next_down -> let f = get_float1 evd args in E.mkFloat env (Float64.next_down f) let red_prim env evd p args = try let r = red_prim_aux env evd p args in Some r with NativeDestKO -> None end coq-8.11.0/kernel/vconv.mli0000644000175000017500000000201013612664533015375 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* types kernel_conversion_function (** A conversion function parametrized by a universe comparator. Used outside of the kernel. *) val vm_conv_gen : conv_pb -> (types, 'a) generic_conversion_function coq-8.11.0/kernel/clambda.ml0000644000175000017500000006276613612664533015504 0ustar treinentreinenopen Util open Names open Esubst open Term open Constr open Declarations open Vmvalues open Environ open Pp let pr_con sp = str(Names.Label.to_string (Constant.label sp)) type lambda = | Lrel of Name.t * int | Lvar of Id.t | Levar of Evar.t * lambda array | Lprod of lambda * lambda | Llam of Name.t Context.binder_annot array * lambda | Llet of Name.t Context.binder_annot * lambda * lambda | Lapp of lambda * lambda array | Lconst of pconstant | Lprim of pconstant option * CPrimitives.t * lambda array (* No check if None *) | Lcase of case_info * reloc_table * lambda * lambda * lam_branches | Lif of lambda * lambda * lambda | Lfix of (int array * int) * fix_decl | Lcofix of int * fix_decl | Lint of int | Lmakeblock of int * lambda array | Luint of Uint63.t | Lfloat of Float64.t | Lval of structured_values | Lsort of Sorts.t | Lind of pinductive | Lproj of Projection.Repr.t * lambda (* We separate branches for constant and non-constant constructors. If the OCaml limitation on non-constant constructors is reached, remaining branches are stored in [extra_branches]. *) and lam_branches = { constant_branches : lambda array; nonconstant_branches : (Name.t Context.binder_annot array * lambda) array } (* extra_branches : (name array * lambda) array } *) and fix_decl = Name.t Context.binder_annot array * lambda array * lambda array (** Printing **) let pr_annot x = Name.print x.Context.binder_name let pp_names ids = prlist_with_sep (fun _ -> brk(1,1)) pr_annot (Array.to_list ids) let pp_rel name n = Name.print name ++ str "##" ++ int n let pp_sort s = match Sorts.family s with | InSet -> str "Set" | InProp -> str "Prop" | InSProp -> str "SProp" | InType -> str "Type" let rec pp_lam lam = match lam with | Lrel (id,n) -> pp_rel id n | Lvar id -> Id.print id | Levar (evk, args) -> hov 1 (str "evar(" ++ Evar.print evk ++ str "," ++ spc () ++ prlist_with_sep spc pp_lam (Array.to_list args) ++ str ")") | Lprod(dom,codom) -> hov 1 (str "forall(" ++ pp_lam dom ++ str "," ++ spc() ++ pp_lam codom ++ str ")") | Llam(ids,body) -> hov 1 (str "(fun " ++ pp_names ids ++ str " =>" ++ spc() ++ pp_lam body ++ str ")") | Llet(id,def,body) -> hov 0 (str "let " ++ pr_annot id ++ str ":=" ++ pp_lam def ++ str " in" ++ spc() ++ pp_lam body) | Lapp(f, args) -> hov 1 (str "(" ++ pp_lam f ++ spc() ++ prlist_with_sep spc pp_lam (Array.to_list args) ++ str")") | Lconst (kn,_) -> pr_con kn | Lcase(_ci, _rtbl, t, a, branches) -> let ic = ref (-1) in let ib = ref 0 in v 0 (str"<" ++ pp_lam t ++ str">" ++ cut() ++ str "Case" ++ spc () ++ pp_lam a ++ spc() ++ str "of" ++ cut() ++ v 0 ((prlist_with_sep (fun _ -> str "") (fun c -> cut () ++ str "| " ++ int (incr ic; !ic) ++ str " => " ++ pp_lam c) (Array.to_list branches.constant_branches)) ++ (prlist_with_sep (fun _ -> str "") (fun (ids,c) -> cut () ++ str "| " ++ int (incr ib; !ib) ++ str " " ++ pp_names ids ++ str " => " ++ pp_lam c) (Array.to_list branches.nonconstant_branches))) ++ cut() ++ str "end") | Lif (t, bt, bf) -> v 0 (str "(if " ++ pp_lam t ++ cut () ++ str "then " ++ pp_lam bt ++ cut() ++ str "else " ++ pp_lam bf ++ str ")") | Lfix((t,i),(lna,tl,bl)) -> let fixl = Array.mapi (fun i id -> (id,t.(i),tl.(i),bl.(i))) lna in hov 1 (str"fix " ++ int i ++ spc() ++ str"{" ++ v 0 (prlist_with_sep spc (fun (na,i,ty,bd) -> pr_annot na ++ str"/" ++ int i ++ str":" ++ pp_lam ty ++ cut() ++ str":=" ++ pp_lam bd) (Array.to_list fixl)) ++ str"}") | Lcofix (i,(lna,tl,bl)) -> let fixl = Array.mapi (fun i na -> (na,tl.(i),bl.(i))) lna in hov 1 (str"cofix " ++ int i ++ spc() ++ str"{" ++ v 0 (prlist_with_sep spc (fun (na,ty,bd) -> pr_annot na ++ str":" ++ pp_lam ty ++ cut() ++ str":=" ++ pp_lam bd) (Array.to_list fixl)) ++ str"}") | Lmakeblock(tag, args) -> hov 1 (str "(makeblock " ++ int tag ++ spc() ++ prlist_with_sep spc pp_lam (Array.to_list args) ++ str")") | Luint i -> str (Uint63.to_string i) | Lfloat f -> str (Float64.to_string f) | Lval _ -> str "values" | Lsort s -> pp_sort s | Lind ((mind,i), _) -> MutInd.print mind ++ str"#" ++ int i | Lprim(Some (kn,_u),_op,args) -> hov 1 (str "(PRIM " ++ pr_con kn ++ spc() ++ prlist_with_sep spc pp_lam (Array.to_list args) ++ str")") | Lprim(None,op,args) -> hov 1 (str "(PRIM_NC " ++ str (CPrimitives.to_string op) ++ spc() ++ prlist_with_sep spc pp_lam (Array.to_list args) ++ str")") | Lproj(p,arg) -> hov 1 (str "(proj " ++ Projection.Repr.print p ++ str "(" ++ pp_lam arg ++ str ")") | Lint i -> Pp.(str "(int:" ++ int i ++ str ")") (*s Constructors *) let mkLapp f args = if Array.length args = 0 then f else match f with | Lapp(f', args') -> Lapp (f', Array.append args' args) | _ -> Lapp(f, args) let mkLlam ids body = if Array.length ids = 0 then body else match body with | Llam(ids', body) -> Llam(Array.append ids ids', body) | _ -> Llam(ids, body) let decompose_Llam lam = match lam with | Llam(ids,body) -> ids, body | _ -> [||], lam (*s Operators on substitution *) let subst_id = subs_id 0 let lift = subs_lift let liftn = subs_liftn let cons v subst = subs_cons([|v|], subst) let shift subst = subs_shft (1, subst) (* A generic map function *) let map_lam_with_binders g f n lam = match lam with | Lrel _ | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ | Lfloat _ -> lam | Levar (evk, args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Levar (evk, args') | Lprod(dom,codom) -> let dom' = f n dom in let codom' = f n codom in if dom == dom' && codom == codom' then lam else Lprod(dom',codom') | Llam(ids,body) -> let body' = f (g (Array.length ids) n) body in if body == body' then lam else mkLlam ids body' | Llet(id,def,body) -> let def' = f n def in let body' = f (g 1 n) body in if body == body' && def == def' then lam else Llet(id,def',body') | Lapp(fct,args) -> let fct' = f n fct in let args' = Array.Smart.map (f n) args in if fct == fct' && args == args' then lam else mkLapp fct' args' | Lcase(ci,rtbl,t,a,branches) -> let const = branches.constant_branches in let nonconst = branches.nonconstant_branches in let t' = f n t in let a' = f n a in let const' = Array.Smart.map (f n) const in let on_b b = let (ids,body) = b in let body' = f (g (Array.length ids) n) body in if body == body' then b else (ids,body') in let nonconst' = Array.Smart.map on_b nonconst in let branches' = if const == const' && nonconst == nonconst' then branches else { constant_branches = const'; nonconstant_branches = nonconst' } in if t == t' && a == a' && branches == branches' then lam else Lcase(ci,rtbl,t',a',branches') | Lif(t,bt,bf) -> let t' = f n t in let bt' = f n bt in let bf' = f n bf in if t == t' && bt == bt' && bf == bf' then lam else Lif(t',bt',bf') | Lfix(init,(ids,ltypes,lbodies)) -> let ltypes' = Array.Smart.map (f n) ltypes in let lbodies' = Array.Smart.map (f (g (Array.length ids) n)) lbodies in if ltypes == ltypes' && lbodies == lbodies' then lam else Lfix(init,(ids,ltypes',lbodies')) | Lcofix(init,(ids,ltypes,lbodies)) -> let ltypes' = Array.Smart.map (f n) ltypes in let lbodies' = Array.Smart.map (f (g (Array.length ids) n)) lbodies in if ltypes == ltypes' && lbodies == lbodies' then lam else Lcofix(init,(ids,ltypes',lbodies')) | Lmakeblock(tag,args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Lmakeblock(tag,args') | Lprim(kn,op,args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Lprim(kn,op,args') | Lproj(p,arg) -> let arg' = f n arg in if arg == arg' then lam else Lproj(p,arg') (*s Lift and substitution *) let rec lam_exlift el lam = match lam with | Lrel(id,i) -> let i' = reloc_rel i el in if i == i' then lam else Lrel(id,i') | _ -> map_lam_with_binders el_liftn lam_exlift el lam let lam_lift k lam = if k = 0 then lam else lam_exlift (el_shft k el_id) lam let lam_subst_rel lam id n subst = match expand_rel n subst with | Inl(k,v) -> lam_lift k v | Inr(n',_) -> if n == n' then lam else Lrel(id, n') let rec lam_exsubst subst lam = match lam with | Lrel(id,i) -> lam_subst_rel lam id i subst | _ -> map_lam_with_binders liftn lam_exsubst subst lam let lam_subst_args subst args = if is_subs_id subst then args else Array.Smart.map (lam_exsubst subst) args (** Simplification of lambda expression *) (* [simplify subst lam] simplify the expression [lam_subst subst lam] *) (* that is : *) (* - Reduce [let] is the definition can be substituted i.e: *) (* - a variable (rel or identifier) *) (* - a constant *) (* - a structured constant *) (* - a function *) (* - Transform beta redex into [let] expression *) (* - Move arguments under [let] *) (* Invariant : Terms in [subst] are already simplified and can be *) (* substituted *) let can_subst lam = match lam with | Lrel _ | Lvar _ | Lconst _ | Luint _ | Lval _ | Lsort _ | Lind _ -> true | _ -> false let can_merge_if bt bf = match bt, bf with | Llam(_idst,_), Llam(_idsf,_) -> true | _ -> false let merge_if t bt bf = let (idst,bodyt) = decompose_Llam bt in let (idsf,bodyf) = decompose_Llam bf in let nt = Array.length idst in let nf = Array.length idsf in let common,idst,idsf = if nt = nf then idst, [||], [||] else if nt < nf then idst,[||], Array.sub idsf nt (nf - nt) else idsf, Array.sub idst nf (nt - nf), [||] in Llam(common, Lif(lam_lift (Array.length common) t, mkLlam idst bodyt, mkLlam idsf bodyf)) let rec simplify subst lam = match lam with | Lrel(id,i) -> lam_subst_rel lam id i subst | Llet(id,def,body) -> let def' = simplify subst def in if can_subst def' then simplify (cons def' subst) body else let body' = simplify (lift subst) body in if def == def' && body == body' then lam else Llet(id,def',body') | Lapp(f,args) -> begin match simplify_app subst f subst args with | Lapp(f',args') when f == f' && args == args' -> lam | lam' -> lam' end | Lif(t,bt,bf) -> let t' = simplify subst t in let bt' = simplify subst bt in let bf' = simplify subst bf in if can_merge_if bt' bf' then merge_if t' bt' bf' else if t == t' && bt == bt' && bf == bf' then lam else Lif(t',bt',bf') | _ -> map_lam_with_binders liftn simplify subst lam and simplify_app substf f substa args = match f with | Lrel(id, i) -> begin match lam_subst_rel f id i substf with | Llam(ids, body) -> reduce_lapp subst_id (Array.to_list ids) body substa (Array.to_list args) | f' -> mkLapp f' (simplify_args substa args) end | Llam(ids, body) -> reduce_lapp substf (Array.to_list ids) body substa (Array.to_list args) | Llet(id, def, body) -> let def' = simplify substf def in if can_subst def' then simplify_app (cons def' substf) body substa args else Llet(id, def', simplify_app (lift substf) body (shift substa) args) | Lapp(f, args') -> let args = Array.append (lam_subst_args substf args') (lam_subst_args substa args) in simplify_app substf f subst_id args | _ -> mkLapp (simplify substf f) (simplify_args substa args) and simplify_args subst args = Array.Smart.map (simplify subst) args and reduce_lapp substf lids body substa largs = match lids, largs with | id::lids, a::largs -> let a = simplify substa a in if can_subst a then reduce_lapp (cons a substf) lids body substa largs else let body = reduce_lapp (lift substf) lids body (shift substa) largs in Llet(id, a, body) | [], [] -> simplify substf body | _::_, _ -> Llam(Array.of_list lids, simplify (liftn (List.length lids) substf) body) | [], _ -> simplify_app substf body substa (Array.of_list largs) (* [occurrence kind k lam]: If [kind] is [true] return [true] if the variable [k] does not appear in [lam], return [false] if the variable appear one time and not under a lambda, a fixpoint, a cofixpoint; else raise Not_found. If [kind] is [false] return [false] if the variable does not appear in [lam] else raise [Not_found] *) let rec occurrence k kind lam = match lam with | Lrel (_,n) -> if n = k then if kind then false else raise Not_found else kind | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ | Lfloat _ -> kind | Levar (_, args) -> occurrence_args k kind args | Lprod(dom, codom) -> occurrence k (occurrence k kind dom) codom | Llam(ids,body) -> let _ = occurrence (k+Array.length ids) false body in kind | Llet(_,def,body) -> occurrence (k+1) (occurrence k kind def) body | Lapp(f, args) -> occurrence_args k (occurrence k kind f) args | Lprim(_,_,args) | Lmakeblock(_,args) -> occurrence_args k kind args | Lcase(_ci,_rtbl,t,a,branches) -> let kind = occurrence k (occurrence k kind t) a in let r = ref kind in Array.iter (fun c -> r := occurrence k kind c && !r) branches.constant_branches; let on_b (ids,c) = r := occurrence (k+Array.length ids) kind c && !r in Array.iter on_b branches.nonconstant_branches; !r | Lif (t, bt, bf) -> let kind = occurrence k kind t in kind && occurrence k kind bt && occurrence k kind bf | Lfix(_,(ids,ltypes,lbodies)) | Lcofix(_,(ids,ltypes,lbodies)) -> let kind = occurrence_args k kind ltypes in let _ = occurrence_args (k+Array.length ids) false lbodies in kind | Lproj(_,arg) -> occurrence k kind arg and occurrence_args k kind args = Array.fold_left (occurrence k) kind args let occur_once lam = try let _ = occurrence 1 true lam in true with Not_found -> false (* [remove_let lam] remove let expression in [lam] if the variable is *) (* used at most once time in the body, and does not appear under *) (* a lambda or a fix or a cofix *) let rec remove_let subst lam = match lam with | Lrel(id,i) -> lam_subst_rel lam id i subst | Llet(id,def,body) -> let def' = remove_let subst def in if occur_once body then remove_let (cons def' subst) body else let body' = remove_let (lift subst) body in if def == def' && body == body' then lam else Llet(id,def',body') | _ -> map_lam_with_binders liftn remove_let subst lam (*s Translation from [constr] to [lambda] *) (* Translation of constructor *) (* Limitation due to OCaml's representation of non-constant constructors: limited to 245 + 1 (0 tag) cases. *) exception TooLargeInductive of Pp.t let max_nb_const = 0x1000000 let max_nb_block = 0x1000000 + Obj.last_non_constant_constructor_tag - 1 let str_max_constructors = Format.sprintf " which has more than %i constant constructors or more than %i non-constant constructors" max_nb_const max_nb_block let check_compilable ib = if not (ib.mind_nb_args <= max_nb_block && ib.mind_nb_constant <= max_nb_const) then let msg = Pp.(str "Cannot compile code for virtual machine as it uses inductive " ++ Id.print ib.mind_typename ++ str str_max_constructors) in raise (TooLargeInductive msg) let is_value lc = match lc with | Lval _ | Lint _ | Luint _ -> true | _ -> false let get_value lc = match lc with | Luint i -> val_of_uint i | Lval v -> v | Lint i -> val_of_int i | _ -> raise Not_found let make_args start _end = Array.init (start - _end + 1) (fun i -> Lrel (Anonymous, start - i)) (* Translation of constructors *) let expand_constructor tag nparams arity = let anon = Context.make_annot Anonymous Sorts.Relevant in (* TODO relevance *) let ids = Array.make (nparams + arity) anon in if arity = 0 then mkLlam ids (Lint tag) else let args = make_args arity 1 in Llam(ids, Lmakeblock (tag, args)) let makeblock tag nparams arity args = let nargs = Array.length args in if nparams > 0 || nargs < arity then mkLapp (expand_constructor tag nparams arity) args else (* The constructor is fully applied *) if arity = 0 then Lint tag else if Array.for_all is_value args then if tag < Obj.last_non_constant_constructor_tag then Lval(val_of_block tag (Array.map get_value args)) else let args = Array.map get_value args in let args = Array.append [| val_of_int (tag - Obj.last_non_constant_constructor_tag) |] args in Lval(val_of_block Obj.last_non_constant_constructor_tag args) else Lmakeblock(tag, args) (* Compiling constants *) let rec get_alias env kn = let cb = lookup_constant kn env in let tps = cb.const_body_code in match tps with | None -> kn | Some tps -> (match Cemitcodes.force tps with | Cemitcodes.BCalias kn' -> get_alias env kn' | _ -> kn) (* Compilation of primitive *) let prim kn p args = Lprim(Some kn, p, args) let expand_prim kn op arity = (* primitives are always Relevant *) let ids = Array.make arity Context.anonR in let args = make_args arity 1 in Llam(ids, prim kn op args) let lambda_of_prim kn op args = let arity = CPrimitives.arity op in if Array.length args >= arity then prim kn op args else mkLapp (expand_prim kn op arity) args (*i Global environment *) let get_names decl = let decl = Array.of_list decl in Array.map fst decl (* Rel Environment *) module Vect = struct type 'a t = { mutable elems : 'a array; mutable size : int; } let make n a = { elems = Array.make n a; size = 0; } let extend (v : 'a t) = if v.size = Array.length v.elems then let new_size = min (2*v.size) Sys.max_array_length in if new_size <= v.size then raise (Invalid_argument "Vect.extend"); let new_elems = Array.make new_size v.elems.(0) in Array.blit v.elems 0 new_elems 0 (v.size); v.elems <- new_elems let push v a = extend v; v.elems.(v.size) <- a; v.size <- v.size + 1 let popn (v : 'a t) n = v.size <- max 0 (v.size - n) let pop v = popn v 1 let get_last (v : 'a t) n = if v.size <= n then raise (Invalid_argument "Vect.get:index out of bounds"); v.elems.(v.size - n - 1) end let dummy_lambda = Lrel(Anonymous, 0) let empty_args = [||] module Renv = struct type constructor_info = tag * int * int (* nparam nrealargs *) type t = { global_env : env; name_rel : Name.t Vect.t; construct_tbl : (constructor, constructor_info) Hashtbl.t; } let make env = { global_env = env; name_rel = Vect.make 16 Anonymous; construct_tbl = Hashtbl.create 111 } let push_rel env id = Vect.push env.name_rel id.Context.binder_name let push_rels env ids = Array.iter (push_rel env) ids let pop env = Vect.pop env.name_rel let popn env n = for _i = 1 to n do pop env done let get env n = Lrel (Vect.get_last env.name_rel (n-1), n) let get_construct_info env c = try Hashtbl.find env.construct_tbl c with Not_found -> let ((mind,j), i) = c in let oib = lookup_mind mind env.global_env in let oip = oib.mind_packets.(j) in check_compilable oip; let tag,arity = oip.mind_reloc_tbl.(i-1) in let nparams = oib.mind_nparams in let r = (tag, nparams, arity) in Hashtbl.add env.construct_tbl c r; r end open Renv let rec lambda_of_constr env c = match Constr.kind c with | Meta _ -> raise (Invalid_argument "Cbytegen.lambda_of_constr: Meta") | Evar (evk, args) -> let args = lambda_of_args env 0 args in Levar (evk, args) | Cast (c, _, _) -> lambda_of_constr env c | Rel i -> Renv.get env i | Var id -> Lvar id | Sort s -> Lsort s | Ind ind -> Lind ind | Prod(id, dom, codom) -> let ld = lambda_of_constr env dom in Renv.push_rel env id; let lc = lambda_of_constr env codom in Renv.pop env; Lprod(ld, Llam([|id|], lc)) | Lambda _ -> let params, body = decompose_lam c in let ids = get_names (List.rev params) in Renv.push_rels env ids; let lb = lambda_of_constr env body in Renv.popn env (Array.length ids); mkLlam ids lb | LetIn(id, def, _, body) -> let ld = lambda_of_constr env def in Renv.push_rel env id; let lb = lambda_of_constr env body in Renv.pop env; Llet(id, ld, lb) | App(f, args) -> lambda_of_app env f args | Const _ -> lambda_of_app env c empty_args | Construct _ -> lambda_of_app env c empty_args | Case(ci,t,a,branches) -> let ind = ci.ci_ind in let mib = lookup_mind (fst ind) env.global_env in let oib = mib.mind_packets.(snd ind) in let () = check_compilable oib in let rtbl = oib.mind_reloc_tbl in (* translation of the argument *) let la = lambda_of_constr env a in (* translation of the type *) let lt = lambda_of_constr env t in (* translation of branches *) let consts = Array.make oib.mind_nb_constant dummy_lambda in let blocks = Array.make oib.mind_nb_args ([||],dummy_lambda) in for i = 0 to Array.length rtbl - 1 do let tag, arity = rtbl.(i) in let b = lambda_of_constr env branches.(i) in if arity = 0 then consts.(tag) <- b else let b = match b with | Llam(ids, body) when Array.length ids = arity -> (ids, body) | _ -> let anon = Context.make_annot Anonymous Sorts.Relevant in (* TODO relevance *) let ids = Array.make arity anon in let args = make_args arity 1 in let ll = lam_lift arity b in (ids, mkLapp ll args) in blocks.(tag-1) <- b done; let branches = { constant_branches = consts; nonconstant_branches = blocks } in Lcase(ci, rtbl, lt, la, branches) | Fix(rec_init,(names,type_bodies,rec_bodies)) -> let ltypes = lambda_of_args env 0 type_bodies in Renv.push_rels env names; let lbodies = lambda_of_args env 0 rec_bodies in Renv.popn env (Array.length names); Lfix(rec_init, (names, ltypes, lbodies)) | CoFix(init,(names,type_bodies,rec_bodies)) -> let rec_bodies = Array.map2 (Reduction.eta_expand env.global_env) rec_bodies type_bodies in let ltypes = lambda_of_args env 0 type_bodies in Renv.push_rels env names; let lbodies = lambda_of_args env 0 rec_bodies in Renv.popn env (Array.length names); Lcofix(init, (names, ltypes, lbodies)) | Proj (p,c) -> let lc = lambda_of_constr env c in Lproj (Projection.repr p,lc) | Int i -> Luint i | Float f -> Lfloat f and lambda_of_app env f args = match Constr.kind f with | Const (kn,u as c) -> let kn = get_alias env.global_env kn in let cb = lookup_constant kn env.global_env in begin match cb.const_body with | Primitive op -> lambda_of_prim (kn,u) op (lambda_of_args env 0 args) | Def csubst when cb.const_inline_code -> lambda_of_app env (Mod_subst.force_constr csubst) args | Def _ | OpaqueDef _ | Undef _ -> mkLapp (Lconst c) (lambda_of_args env 0 args) end | Construct (c,_) -> let tag, nparams, arity = Renv.get_construct_info env c in let nargs = Array.length args in if nparams < nargs then (* got all parameters *) let args = lambda_of_args env nparams args in makeblock tag 0 arity args else makeblock tag (nparams - nargs) arity empty_args | _ -> let f = lambda_of_constr env f in let args = lambda_of_args env 0 args in mkLapp f args and lambda_of_args env start args = let nargs = Array.length args in if start < nargs then Array.init (nargs - start) (fun i -> lambda_of_constr env args.(start + i)) else empty_args (*********************************) let dump_lambda = ref false let optimize_lambda lam = let lam = simplify subst_id lam in remove_let subst_id lam let lambda_of_constr ~optimize genv c = let env = Renv.make genv in let ids = List.rev_map Context.Rel.Declaration.get_annot (rel_context genv) in Renv.push_rels env (Array.of_list ids); let lam = lambda_of_constr env c in let lam = if optimize then optimize_lambda lam else lam in if !dump_lambda then Feedback.msg_debug (pp_lam lam); lam coq-8.11.0/kernel/uGraph.ml0000644000175000017500000001654213612664533015336 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a -> 'a -> bool let g_map f g = let g' = f g.graph in if g.graph == g' then g else {g with graph=g'} let make_sprop_cumulative g = {g with sprop_cumulative=true} let check_smaller_expr g (u,n) (v,m) = let diff = n - m in match diff with | 0 -> G.check_leq g.graph u v | 1 -> G.check_lt g.graph u v | x when x < 0 -> G.check_leq g.graph u v | _ -> false let exists_bigger g ul l = Universe.exists (fun ul' -> check_smaller_expr g ul ul') l let real_check_leq g u v = Universe.for_all (fun ul -> exists_bigger g ul v) u let check_leq g u v = Universe.equal u v || (g.sprop_cumulative && Universe.is_sprop u) || (not (Universe.is_sprop u) && not (Universe.is_sprop v) && (is_type0m_univ u || real_check_leq g u v)) let check_eq g u v = Universe.equal u v || (not (Universe.is_sprop u || Universe.is_sprop v) && (real_check_leq g u v && real_check_leq g v u)) let check_eq_level g u v = u == v || (not (Level.is_sprop u || Level.is_sprop v) && G.check_eq g.graph u v) let empty_universes = {graph=G.empty; sprop_cumulative=false} let initial_universes = let big_rank = 1000000 in let g = G.empty in let g = G.add ~rank:big_rank Level.prop g in let g = G.add ~rank:big_rank Level.set g in {graph=G.enforce_lt Level.prop Level.set g; sprop_cumulative=false} let enforce_constraint (u,d,v) g = match d with | Le -> G.enforce_leq u v g | Lt -> G.enforce_lt u v g | Eq -> G.enforce_eq u v g let enforce_constraint (u,d,v as cst) g = match Level.is_sprop u, d, Level.is_sprop v with | false, _, false -> g_map (enforce_constraint cst) g | true, (Eq|Le), true -> g | true, Le, false when g.sprop_cumulative -> g | _ -> raise (UniverseInconsistency (d,Universe.make u, Universe.make v, None)) let merge_constraints csts g = Constraint.fold enforce_constraint csts g let check_constraint g (u,d,v) = match d with | Le -> G.check_leq g u v | Lt -> G.check_lt g u v | Eq -> G.check_eq g u v let check_constraint g (u,d,v as cst) = match Level.is_sprop u, d, Level.is_sprop v with | false, _, false -> check_constraint g.graph cst | true, (Eq|Le), true -> true | true, Le, false -> g.sprop_cumulative | _ -> false let check_constraints csts g = Constraint.for_all (check_constraint g) csts let leq_expr (u,m) (v,n) = let d = match m - n with | 1 -> Lt | diff -> assert (diff <= 0); Le in (u,d,v) let enforce_leq_alg u v g = let open Util in let enforce_one (u,v) = function | Inr _ as orig -> orig | Inl (cstrs,g) as orig -> if check_smaller_expr g u v then orig else (let c = leq_expr u v in match enforce_constraint c g with | g -> Inl (Constraint.add c cstrs,g) | exception (UniverseInconsistency _ as e) -> Inr e) in (* max(us) <= max(vs) <-> forall u in us, exists v in vs, u <= v *) let c = Universe.map (fun u -> Universe.map (fun v -> (u,v)) v) u in let c = List.cartesians enforce_one (Inl (Constraint.empty,g)) c in (* We pick a best constraint: smallest number of constraints, not an error if possible. *) let order x y = match x, y with | Inr _, Inr _ -> 0 | Inl _, Inr _ -> -1 | Inr _, Inl _ -> 1 | Inl (c,_), Inl (c',_) -> Int.compare (Constraint.cardinal c) (Constraint.cardinal c') in match List.min order c with | Inl x -> x | Inr e -> raise e (* sanity check wrapper *) let enforce_leq_alg u v g = let _,g as cg = enforce_leq_alg u v g in assert (check_leq g u v); cg exception AlreadyDeclared = G.AlreadyDeclared let add_universe u ~lbound ~strict g = let graph = G.add u g.graph in let d = if strict then Lt else Le in enforce_constraint (lbound,d,u) {g with graph} let add_universe_unconstrained u g = {g with graph=G.add u g.graph} exception UndeclaredLevel = G.Undeclared let check_declared_universes g l = G.check_declared g.graph (LSet.remove Level.sprop l) let constraints_of_universes g = G.constraints_of g.graph let constraints_for ~kept g = G.constraints_for ~kept:(LSet.remove Level.sprop kept) g.graph (** Subtyping of polymorphic contexts *) let check_subtype ~lbound univs ctxT ctx = if AUContext.size ctxT == AUContext.size ctx then let (inst, cst) = UContext.dest (AUContext.repr ctx) in let cstT = UContext.constraints (AUContext.repr ctxT) in let push accu v = add_universe v ~lbound ~strict:false accu in let univs = Array.fold_left push univs (Instance.to_array inst) in let univs = merge_constraints cstT univs in check_constraints cst univs else false (** Instances *) let check_eq_instances g t1 t2 = let t1 = Instance.to_array t1 in let t2 = Instance.to_array t2 in t1 == t2 || (Int.equal (Array.length t1) (Array.length t2) && let rec aux i = (Int.equal i (Array.length t1)) || (check_eq_level g t1.(i) t2.(i) && aux (i + 1)) in aux 0) let domain g = LSet.add Level.sprop (G.domain g.graph) let choose p g u = if Level.is_sprop u then if p u then Some u else None else G.choose p g.graph u let dump_universes f g = G.dump f g.graph let check_universes_invariants g = G.check_invariants ~required_canonical:Level.is_small g.graph let pr_universes prl g = G.pr prl g.graph let dummy_mp = Names.DirPath.make [Names.Id.of_string "Type"] let make_dummy i = Level.(make (UGlobal.make dummy_mp i)) let sort_universes g = g_map (G.sort make_dummy [Level.prop;Level.set]) g (** Profiling *) let merge_constraints = if Flags.profile then let key = CProfile.declare_profile "merge_constraints" in CProfile.profile2 key merge_constraints else merge_constraints let check_constraints = if Flags.profile then let key = CProfile.declare_profile "check_constraints" in CProfile.profile2 key check_constraints else check_constraints let check_eq = if Flags.profile then let check_eq_key = CProfile.declare_profile "check_eq" in CProfile.profile3 check_eq_key check_eq else check_eq let check_leq = if Flags.profile then let check_leq_key = CProfile.declare_profile "check_leq" in CProfile.profile3 check_leq_key check_leq else check_leq coq-8.11.0/kernel/subtyping.mli0000644000175000017500000000143713612664533016302 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* module_type_body -> module_type_body -> Constraint.t coq-8.11.0/kernel/nativecode.ml0000644000175000017500000023051313612664533016225 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* String.equal s1 s2 && eq_ind ind1 ind2 | Gconstant (s1, c1), Gconstant (s2, c2) -> String.equal s1 s2 && Constant.equal c1 c2 | Gproj (s1, ind1, i1), Gproj (s2, ind2, i2) -> String.equal s1 s2 && eq_ind ind1 ind2 && Int.equal i1 i2 | Gcase (None, i1), Gcase (None, i2) -> Int.equal i1 i2 | Gcase (Some l1, i1), Gcase (Some l2, i2) -> Int.equal i1 i2 && Label.equal l1 l2 | Gpred (None, i1), Gpred (None, i2) -> Int.equal i1 i2 | Gpred (Some l1, i1), Gpred (Some l2, i2) -> Int.equal i1 i2 && Label.equal l1 l2 | Gfixtype (None, i1), Gfixtype (None, i2) -> Int.equal i1 i2 | Gfixtype (Some l1, i1), Gfixtype (Some l2, i2) -> Int.equal i1 i2 && Label.equal l1 l2 | Gnorm (None, i1), Gnorm (None, i2) -> Int.equal i1 i2 | Gnorm (Some l1, i1), Gnorm (Some l2, i2) -> Int.equal i1 i2 && Label.equal l1 l2 | Gnormtbl (None, i1), Gnormtbl (None, i2) -> Int.equal i1 i2 | Gnormtbl (Some l1, i1), Gnormtbl (Some l2, i2) -> Int.equal i1 i2 && Label.equal l1 l2 | Ginternal s1, Ginternal s2 -> String.equal s1 s2 | Grel i1, Grel i2 -> Int.equal i1 i2 | Gnamed id1, Gnamed id2 -> Id.equal id1 id2 | (Gind _| Gconstant _ | Gproj _ | Gcase _ | Gpred _ | Gfixtype _ | Gnorm _ | Gnormtbl _ | Ginternal _ | Grel _ | Gnamed _), _ -> false let dummy_gname = Grel 0 open Hashset.Combine let gname_hash gn = match gn with | Gind (s, ind) -> combinesmall 1 (combine (String.hash s) (ind_hash ind)) | Gconstant (s, c) -> combinesmall 2 (combine (String.hash s) (Constant.hash c)) | Gcase (l, i) -> combinesmall 3 (combine (Option.hash Label.hash l) (Int.hash i)) | Gpred (l, i) -> combinesmall 4 (combine (Option.hash Label.hash l) (Int.hash i)) | Gfixtype (l, i) -> combinesmall 5 (combine (Option.hash Label.hash l) (Int.hash i)) | Gnorm (l, i) -> combinesmall 6 (combine (Option.hash Label.hash l) (Int.hash i)) | Gnormtbl (l, i) -> combinesmall 7 (combine (Option.hash Label.hash l) (Int.hash i)) | Ginternal s -> combinesmall 8 (String.hash s) | Grel i -> combinesmall 9 (Int.hash i) | Gnamed id -> combinesmall 10 (Id.hash id) | Gproj (s, p, i) -> combinesmall 11 (combine (String.hash s) (combine (ind_hash p) i)) let case_ctr = ref (-1) let fresh_gcase l = incr case_ctr; Gcase (l,!case_ctr) let pred_ctr = ref (-1) let fresh_gpred l = incr pred_ctr; Gpred (l,!pred_ctr) let fixtype_ctr = ref (-1) let fresh_gfixtype l = incr fixtype_ctr; Gfixtype (l,!fixtype_ctr) let norm_ctr = ref (-1) let fresh_gnorm l = incr norm_ctr; Gnorm (l,!norm_ctr) let normtbl_ctr = ref (-1) let fresh_gnormtbl l = incr normtbl_ctr; Gnormtbl (l,!normtbl_ctr) (** Symbols (pre-computed values) **) let dummy_symb = SymbValue (dummy_value ()) let eq_symbol sy1 sy2 = match sy1, sy2 with | SymbValue v1, SymbValue v2 -> (=) v1 v2 (** FIXME: how is this even valid? *) | SymbSort s1, SymbSort s2 -> Sorts.equal s1 s2 | SymbName n1, SymbName n2 -> Name.equal n1 n2 | SymbConst kn1, SymbConst kn2 -> Constant.equal kn1 kn2 | SymbMatch sw1, SymbMatch sw2 -> eq_annot_sw sw1 sw2 | SymbInd ind1, SymbInd ind2 -> eq_ind ind1 ind2 | SymbMeta m1, SymbMeta m2 -> Int.equal m1 m2 | SymbEvar evk1, SymbEvar evk2 -> Evar.equal evk1 evk2 | SymbLevel l1, SymbLevel l2 -> Univ.Level.equal l1 l2 | SymbProj (i1, k1), SymbProj (i2, k2) -> eq_ind i1 i2 && Int.equal k1 k2 | _, _ -> false let hash_symbol symb = match symb with | SymbValue v -> combinesmall 1 (Hashtbl.hash v) (** FIXME *) | SymbSort s -> combinesmall 2 (Sorts.hash s) | SymbName name -> combinesmall 3 (Name.hash name) | SymbConst c -> combinesmall 4 (Constant.hash c) | SymbMatch sw -> combinesmall 5 (hash_annot_sw sw) | SymbInd ind -> combinesmall 6 (ind_hash ind) | SymbMeta m -> combinesmall 7 m | SymbEvar evk -> combinesmall 8 (Evar.hash evk) | SymbLevel l -> combinesmall 9 (Univ.Level.hash l) | SymbProj (i, k) -> combinesmall 10 (combine (ind_hash i) k) module HashedTypeSymbol = struct type t = symbol let equal = eq_symbol let hash = hash_symbol end module HashtblSymbol = Hashtbl.Make(HashedTypeSymbol) let symb_tbl = HashtblSymbol.create 211 let clear_symbols () = HashtblSymbol.clear symb_tbl let get_value tbl i = match tbl.(i) with | SymbValue v -> v | _ -> anomaly (Pp.str "get_value failed.") let get_sort tbl i = match tbl.(i) with | SymbSort s -> s | _ -> anomaly (Pp.str "get_sort failed.") let get_name tbl i = match tbl.(i) with | SymbName id -> id | _ -> anomaly (Pp.str "get_name failed.") let get_const tbl i = match tbl.(i) with | SymbConst kn -> kn | _ -> anomaly (Pp.str "get_const failed.") let get_match tbl i = match tbl.(i) with | SymbMatch case_info -> case_info | _ -> anomaly (Pp.str "get_match failed.") let get_ind tbl i = match tbl.(i) with | SymbInd ind -> ind | _ -> anomaly (Pp.str "get_ind failed.") let get_meta tbl i = match tbl.(i) with | SymbMeta m -> m | _ -> anomaly (Pp.str "get_meta failed.") let get_evar tbl i = match tbl.(i) with | SymbEvar ev -> ev | _ -> anomaly (Pp.str "get_evar failed.") let get_level tbl i = match tbl.(i) with | SymbLevel u -> u | _ -> anomaly (Pp.str "get_level failed.") let get_proj tbl i = match tbl.(i) with | SymbProj p -> p | _ -> anomaly (Pp.str "get_proj failed.") let push_symbol x = try HashtblSymbol.find symb_tbl x with Not_found -> let i = HashtblSymbol.length symb_tbl in HashtblSymbol.add symb_tbl x i; i let symbols_tbl_name = Ginternal "symbols_tbl" let get_symbols () = let tbl = Array.make (HashtblSymbol.length symb_tbl) dummy_symb in HashtblSymbol.iter (fun x i -> tbl.(i) <- x) symb_tbl; tbl (** Lambda to Mllambda **) type primitive = | Mk_prod | Mk_sort | Mk_ind | Mk_const | Mk_sw | Mk_fix of rec_pos * int | Mk_cofix of int | Mk_rel of int | Mk_var of Id.t | Mk_proj | Is_int | Is_float | Cast_accu | Upd_cofix | Force_cofix | Mk_uint | Mk_float | Mk_int | Mk_bool | Val_to_int | Mk_meta | Mk_evar | MLand | MLnot | MLle | MLlt | MLinteq | MLlsl | MLlsr | MLland | MLlor | MLlxor | MLadd | MLsub | MLmul | MLmagic | MLarrayget | Mk_empty_instance | Coq_primitive of CPrimitives.t * (prefix * pconstant) option let eq_primitive p1 p2 = match p1, p2 with | Mk_prod, Mk_prod -> true | Mk_sort, Mk_sort -> true | Mk_ind, Mk_ind -> true | Mk_const, Mk_const -> true | Mk_sw, Mk_sw -> true | Mk_fix (rp1, i1), Mk_fix (rp2, i2) -> Int.equal i1 i2 && eq_rec_pos rp1 rp2 | Mk_cofix i1, Mk_cofix i2 -> Int.equal i1 i2 | Mk_rel i1, Mk_rel i2 -> Int.equal i1 i2 | Mk_var id1, Mk_var id2 -> Id.equal id1 id2 | Cast_accu, Cast_accu -> true | Upd_cofix, Upd_cofix -> true | Force_cofix, Force_cofix -> true | Mk_meta, Mk_meta -> true | Mk_evar, Mk_evar -> true | Mk_proj, Mk_proj -> true | MLarrayget, MLarrayget -> true | _ -> false let primitive_hash = function | Mk_prod -> 1 | Mk_sort -> 2 | Mk_ind -> 3 | Mk_const -> 4 | Mk_sw -> 5 | Mk_fix (r, i) -> let h = Array.fold_left (fun h i -> combine h (Int.hash i)) 0 r in combinesmall 6 (combine h (Int.hash i)) | Mk_cofix i -> combinesmall 7 (Int.hash i) | Mk_rel i -> combinesmall 8 (Int.hash i) | Mk_var id -> combinesmall 9 (Id.hash id) | Is_int -> 11 | Cast_accu -> 12 | Upd_cofix -> 13 | Force_cofix -> 14 | Mk_uint -> 15 | Mk_int -> 16 | Mk_bool -> 17 | Val_to_int -> 18 | Mk_meta -> 19 | Mk_evar -> 20 | MLand -> 21 | MLle -> 22 | MLlt -> 23 | MLinteq -> 24 | MLlsl -> 25 | MLlsr -> 26 | MLland -> 27 | MLlor -> 28 | MLlxor -> 29 | MLadd -> 30 | MLsub -> 31 | MLmul -> 32 | MLmagic -> 33 | Coq_primitive (prim, None) -> combinesmall 34 (CPrimitives.hash prim) | Coq_primitive (prim, Some (prefix,(kn,_))) -> combinesmall 35 (combine3 (String.hash prefix) (Constant.hash kn) (CPrimitives.hash prim)) | Mk_proj -> 36 | MLarrayget -> 37 | Mk_empty_instance -> 38 | Mk_float -> 39 | Is_float -> 40 | MLnot -> 41 type mllambda = | MLlocal of lname | MLglobal of gname | MLprimitive of primitive | MLlam of lname array * mllambda | MLletrec of (lname * lname array * mllambda) array * mllambda | MLlet of lname * mllambda * mllambda | MLapp of mllambda * mllambda array | MLif of mllambda * mllambda * mllambda | MLmatch of annot_sw * mllambda * mllambda * mllam_branches (* argument, prefix, accu branch, branches *) | MLconstruct of string * inductive * int * mllambda array (* prefix, inductive name, tag, arguments *) | MLint of int | MLuint of Uint63.t | MLfloat of Float64.t | MLsetref of string * mllambda | MLsequence of mllambda * mllambda | MLarray of mllambda array | MLisaccu of string * inductive * mllambda and 'a mllam_pattern = | ConstPattern of int | NonConstPattern of tag * 'a array and mllam_branches = (lname option mllam_pattern list * mllambda) array let push_lnames n env lns = snd (Array.fold_left (fun (i,r) x -> (i+1, LNmap.add x i r)) (n,env) lns) let opush_lnames n env lns = let oadd x i r = match x with Some ln -> LNmap.add ln i r | None -> r in snd (Array.fold_left (fun (i,r) x -> (i+1, oadd x i r)) (n,env) lns) (* Alpha-equivalence on mllambda *) (* eq_mllambda gn1 gn2 n env1 env2 t1 t2 tests if t1 = t2 modulo gn1 = gn2 *) let rec eq_mllambda gn1 gn2 n env1 env2 t1 t2 = match t1, t2 with | MLlocal ln1, MLlocal ln2 -> (try Int.equal (LNmap.find ln1 env1) (LNmap.find ln2 env2) with Not_found -> eq_lname ln1 ln2) | MLglobal gn1', MLglobal gn2' -> eq_gname gn1' gn2' || (eq_gname gn1 gn1' && eq_gname gn2 gn2') || (eq_gname gn1 gn2' && eq_gname gn2 gn1') | MLprimitive prim1, MLprimitive prim2 -> eq_primitive prim1 prim2 | MLlam (lns1, ml1), MLlam (lns2, ml2) -> Int.equal (Array.length lns1) (Array.length lns2) && let env1 = push_lnames n env1 lns1 in let env2 = push_lnames n env2 lns2 in eq_mllambda gn1 gn2 (n+Array.length lns1) env1 env2 ml1 ml2 | MLletrec (defs1, body1), MLletrec (defs2, body2) -> Int.equal (Array.length defs1) (Array.length defs2) && let lns1 = Array.map (fun (x,_,_) -> x) defs1 in let lns2 = Array.map (fun (x,_,_) -> x) defs2 in let env1 = push_lnames n env1 lns1 in let env2 = push_lnames n env2 lns2 in let n = n + Array.length defs1 in eq_letrec gn1 gn2 n env1 env2 defs1 defs2 && eq_mllambda gn1 gn2 n env1 env2 body1 body2 | MLlet (ln1, def1, body1), MLlet (ln2, def2, body2) -> eq_mllambda gn1 gn2 n env1 env2 def1 def2 && let env1 = LNmap.add ln1 n env1 in let env2 = LNmap.add ln2 n env2 in eq_mllambda gn1 gn2 (n+1) env1 env2 body1 body2 | MLapp (ml1, args1), MLapp (ml2, args2) -> eq_mllambda gn1 gn2 n env1 env2 ml1 ml2 && Array.equal (eq_mllambda gn1 gn2 n env1 env2) args1 args2 | MLif (cond1,br1,br'1), MLif (cond2,br2,br'2) -> eq_mllambda gn1 gn2 n env1 env2 cond1 cond2 && eq_mllambda gn1 gn2 n env1 env2 br1 br2 && eq_mllambda gn1 gn2 n env1 env2 br'1 br'2 | MLmatch (annot1, c1, accu1, br1), MLmatch (annot2, c2, accu2, br2) -> eq_annot_sw annot1 annot2 && eq_mllambda gn1 gn2 n env1 env2 c1 c2 && eq_mllambda gn1 gn2 n env1 env2 accu1 accu2 && eq_mllam_branches gn1 gn2 n env1 env2 br1 br2 | MLconstruct (pf1, ind1, tag1, args1), MLconstruct (pf2, ind2, tag2, args2) -> String.equal pf1 pf2 && eq_ind ind1 ind2 && Int.equal tag1 tag2 && Array.equal (eq_mllambda gn1 gn2 n env1 env2) args1 args2 | MLint i1, MLint i2 -> Int.equal i1 i2 | MLuint i1, MLuint i2 -> Uint63.equal i1 i2 | MLfloat f1, MLfloat f2 -> Float64.equal f1 f2 | MLsetref (id1, ml1), MLsetref (id2, ml2) -> String.equal id1 id2 && eq_mllambda gn1 gn2 n env1 env2 ml1 ml2 | MLsequence (ml1, ml'1), MLsequence (ml2, ml'2) -> eq_mllambda gn1 gn2 n env1 env2 ml1 ml2 && eq_mllambda gn1 gn2 n env1 env2 ml'1 ml'2 | MLarray arr1, MLarray arr2 -> Array.equal (eq_mllambda gn1 gn2 n env1 env2) arr1 arr2 | MLisaccu (s1, ind1, ml1), MLisaccu (s2, ind2, ml2) -> String.equal s1 s2 && eq_ind ind1 ind2 && eq_mllambda gn1 gn2 n env1 env2 ml1 ml2 | (MLlocal _ | MLglobal _ | MLprimitive _ | MLlam _ | MLletrec _ | MLlet _ | MLapp _ | MLif _ | MLmatch _ | MLconstruct _ | MLint _ | MLuint _ | MLfloat _ | MLsetref _ | MLsequence _ | MLarray _ | MLisaccu _), _ -> false and eq_letrec gn1 gn2 n env1 env2 defs1 defs2 = let eq_def (_,args1,ml1) (_,args2,ml2) = Int.equal (Array.length args1) (Array.length args2) && let env1 = push_lnames n env1 args1 in let env2 = push_lnames n env2 args2 in eq_mllambda gn1 gn2 (n + Array.length args1) env1 env2 ml1 ml2 in Array.equal eq_def defs1 defs2 (* we require here that patterns have the same order, which may be too strong *) and eq_mllam_branches gn1 gn2 n env1 env2 br1 br2 = let eq_cargs args1 args2 body1 body2 = Int.equal (Array.length args1) (Array.length args2) && let env1 = opush_lnames n env1 args1 in let env2 = opush_lnames n env2 args2 in eq_mllambda gn1 gn2 (n + Array.length args1) env1 env2 body1 body2 in let eq_pattern pat1 pat2 body1 body2 = match pat1, pat2 with | ConstPattern tag1, ConstPattern tag2 -> Int.equal tag1 tag2 && eq_mllambda gn1 gn2 n env1 env2 body1 body2 | NonConstPattern (tag1,args1), NonConstPattern (tag2,args2) -> Int.equal tag1 tag2 && eq_cargs args1 args2 body1 body2 | (ConstPattern _ | NonConstPattern _), _ -> false in let eq_branch (patl1,body1) (patl2,body2) = List.equal (fun pt1 pt2 -> eq_pattern pt1 pt2 body1 body2) patl1 patl2 in Array.equal eq_branch br1 br2 (* hash_mllambda gn n env t computes the hash for t ignoring occurrences of gn *) let rec hash_mllambda gn n env t = match t with | MLlocal ln -> combinesmall 1 (LNmap.find ln env) | MLglobal gn' -> combinesmall 2 (if eq_gname gn gn' then 0 else gname_hash gn') | MLprimitive prim -> combinesmall 3 (primitive_hash prim) | MLlam (lns, ml) -> let env = push_lnames n env lns in combinesmall 4 (combine (Array.length lns) (hash_mllambda gn (n+1) env ml)) | MLletrec (defs, body) -> let lns = Array.map (fun (x,_,_) -> x) defs in let env = push_lnames n env lns in let n = n + Array.length defs in let h = combine (hash_mllambda gn n env body) (Array.length defs) in combinesmall 5 (hash_mllambda_letrec gn n env h defs) | MLlet (ln, def, body) -> let hdef = hash_mllambda gn n env def in let env = LNmap.add ln n env in combinesmall 6 (combine hdef (hash_mllambda gn (n+1) env body)) | MLapp (ml, args) -> let h = hash_mllambda gn n env ml in combinesmall 7 (hash_mllambda_array gn n env h args) | MLif (cond,br,br') -> let hcond = hash_mllambda gn n env cond in let hbr = hash_mllambda gn n env br in let hbr' = hash_mllambda gn n env br' in combinesmall 8 (combine3 hcond hbr hbr') | MLmatch (annot, c, accu, br) -> let hannot = hash_annot_sw annot in let hc = hash_mllambda gn n env c in let haccu = hash_mllambda gn n env accu in combinesmall 9 (hash_mllam_branches gn n env (combine3 hannot hc haccu) br) | MLconstruct (pf, ind, tag, args) -> let hpf = String.hash pf in let hcs = ind_hash ind in let htag = Int.hash tag in combinesmall 10 (hash_mllambda_array gn n env (combine3 hpf hcs htag) args) | MLint i -> combinesmall 11 i | MLuint i -> combinesmall 12 (Uint63.hash i) | MLsetref (id, ml) -> let hid = String.hash id in let hml = hash_mllambda gn n env ml in combinesmall 13 (combine hid hml) | MLsequence (ml, ml') -> let hml = hash_mllambda gn n env ml in let hml' = hash_mllambda gn n env ml' in combinesmall 14 (combine hml hml') | MLarray arr -> combinesmall 15 (hash_mllambda_array gn n env 1 arr) | MLisaccu (s, ind, c) -> combinesmall 16 (combine (String.hash s) (combine (ind_hash ind) (hash_mllambda gn n env c))) | MLfloat f -> combinesmall 17 (Float64.hash f) and hash_mllambda_letrec gn n env init defs = let hash_def (_,args,ml) = let env = push_lnames n env args in let nargs = Array.length args in combine nargs (hash_mllambda gn (n + nargs) env ml) in Array.fold_left (fun acc t -> combine (hash_def t) acc) init defs and hash_mllambda_array gn n env init arr = Array.fold_left (fun acc t -> combine (hash_mllambda gn n env t) acc) init arr and hash_mllam_branches gn n env init br = let hash_cargs args body = let nargs = Array.length args in let env = opush_lnames n env args in let hbody = hash_mllambda gn (n + nargs) env body in combine nargs hbody in let hash_pattern pat body = match pat with | ConstPattern i -> combinesmall 1 (Int.hash i) | NonConstPattern (tag,args) -> combinesmall 2 (combine (Int.hash tag) (hash_cargs args body)) in let hash_branch acc (ptl,body) = List.fold_left (fun acc t -> combine (hash_pattern t body) acc) acc ptl in Array.fold_left hash_branch init br let fv_lam l = let rec aux l bind fv = match l with | MLlocal l -> if LNset.mem l bind then fv else LNset.add l fv | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ | MLfloat _ -> fv | MLlam (ln,body) -> let bind = Array.fold_right LNset.add ln bind in aux body bind fv | MLletrec(bodies,def) -> let bind = Array.fold_right (fun (id,_,_) b -> LNset.add id b) bodies bind in let fv_body (_,ln,body) fv = let bind = Array.fold_right LNset.add ln bind in aux body bind fv in Array.fold_right fv_body bodies (aux def bind fv) | MLlet(l,def,body) -> aux body (LNset.add l bind) (aux def bind fv) | MLapp(f,args) -> let fv_arg arg fv = aux arg bind fv in Array.fold_right fv_arg args (aux f bind fv) | MLif(t,b1,b2) -> aux t bind (aux b1 bind (aux b2 bind fv)) | MLmatch(_,a,p,bs) -> let fv = aux a bind (aux p bind fv) in let fv_bs (cargs, body) fv = let bind = List.fold_right (fun pat bind -> match pat with | ConstPattern _ -> bind | NonConstPattern(_,args) -> Array.fold_right (fun o bind -> match o with | Some l -> LNset.add l bind | _ -> bind) args bind) cargs bind in aux body bind fv in Array.fold_right fv_bs bs fv (* argument, accu branch, branches *) | MLconstruct (_,_,_,p) -> Array.fold_right (fun a fv -> aux a bind fv) p fv | MLsetref(_,l) -> aux l bind fv | MLsequence(l1,l2) -> aux l1 bind (aux l2 bind fv) | MLarray arr -> Array.fold_right (fun a fv -> aux a bind fv) arr fv | MLisaccu (_, _, body) -> aux body bind fv in aux l LNset.empty LNset.empty let mkMLlam params body = if Array.is_empty params then body else match body with | MLlam (params', body) -> MLlam(Array.append params params', body) | _ -> MLlam(params,body) let mkMLapp f args = if Array.is_empty args then f else match f with | MLapp(f,args') -> MLapp(f,Array.append args' args) | _ -> MLapp(f,args) let mkForceCofix prefix ind arg = let name = fresh_lname Anonymous in MLlet (name, arg, MLif ( MLisaccu (prefix, ind, MLlocal name), MLapp (MLprimitive Force_cofix, [|MLlocal name|]), MLlocal name)) let empty_params = [||] let decompose_MLlam c = match c with | MLlam(ids,c) -> ids,c | _ -> empty_params,c (*s Global declaration *) type global = (* | Gtblname of gname * Id.t array *) | Gtblnorm of gname * lname array * mllambda array | Gtblfixtype of gname * lname array * mllambda array | Glet of gname * mllambda | Gletcase of gname * lname array * annot_sw * mllambda * mllambda * mllam_branches | Gopen of string | Gtype of inductive * (tag * int) array (* ind name, tag and arities of constructors *) | Gcomment of string (* Alpha-equivalence on globals *) let eq_global g1 g2 = match g1, g2 with | Gtblnorm (gn1,lns1,mls1), Gtblnorm (gn2,lns2,mls2) | Gtblfixtype (gn1,lns1,mls1), Gtblfixtype (gn2,lns2,mls2) -> Int.equal (Array.length lns1) (Array.length lns2) && Int.equal (Array.length mls1) (Array.length mls2) && let env1 = push_lnames 0 LNmap.empty lns1 in let env2 = push_lnames 0 LNmap.empty lns2 in Array.for_all2 (eq_mllambda gn1 gn2 (Array.length lns1) env1 env2) mls1 mls2 | Glet (gn1, def1), Glet (gn2, def2) -> eq_mllambda gn1 gn2 0 LNmap.empty LNmap.empty def1 def2 | Gletcase (gn1,lns1,annot1,c1,accu1,br1), Gletcase (gn2,lns2,annot2,c2,accu2,br2) -> Int.equal (Array.length lns1) (Array.length lns2) && let env1 = push_lnames 0 LNmap.empty lns1 in let env2 = push_lnames 0 LNmap.empty lns2 in let t1 = MLmatch (annot1,c1,accu1,br1) in let t2 = MLmatch (annot2,c2,accu2,br2) in eq_mllambda gn1 gn2 (Array.length lns1) env1 env2 t1 t2 | Gopen s1, Gopen s2 -> String.equal s1 s2 | Gtype (ind1, arr1), Gtype (ind2, arr2) -> eq_ind ind1 ind2 && Array.equal (fun (tag1,ar1) (tag2,ar2) -> Int.equal tag1 tag2 && Int.equal ar1 ar2) arr1 arr2 | Gcomment s1, Gcomment s2 -> String.equal s1 s2 | _, _ -> false let hash_global g = match g with | Gtblnorm (gn,lns,mls) -> let nlns = Array.length lns in let nmls = Array.length mls in let env = push_lnames 0 LNmap.empty lns in let hmls = hash_mllambda_array gn nlns env (combine nlns nmls) mls in combinesmall 1 hmls | Gtblfixtype (gn,lns,mls) -> let nlns = Array.length lns in let nmls = Array.length mls in let env = push_lnames 0 LNmap.empty lns in let hmls = hash_mllambda_array gn nlns env (combine nlns nmls) mls in combinesmall 2 hmls | Glet (gn, def) -> combinesmall 3 (hash_mllambda gn 0 LNmap.empty def) | Gletcase (gn,lns,annot,c,accu,br) -> let nlns = Array.length lns in let env = push_lnames 0 LNmap.empty lns in let t = MLmatch (annot,c,accu,br) in combinesmall 4 (combine nlns (hash_mllambda gn nlns env t)) | Gopen s -> combinesmall 5 (String.hash s) | Gtype (ind, arr) -> let hash_aux acc (tag,ar) = combine3 acc (Int.hash tag) (Int.hash ar) in combinesmall 6 (combine (ind_hash ind) (Array.fold_left hash_aux 0 arr)) | Gcomment s -> combinesmall 7 (String.hash s) let global_stack = ref ([] : global list) module HashedTypeGlobal = struct type t = global let equal = eq_global let hash = hash_global end module HashtblGlobal = Hashtbl.Make(HashedTypeGlobal) let global_tbl = HashtblGlobal.create 19991 let clear_global_tbl () = HashtblGlobal.clear global_tbl let push_global gn t = try HashtblGlobal.find global_tbl t with Not_found -> (global_stack := t :: !global_stack; HashtblGlobal.add global_tbl t gn; gn) let push_global_let gn body = push_global gn (Glet (gn,body)) let push_global_fixtype gn params body = push_global gn (Gtblfixtype (gn,params,body)) let push_global_norm gn params body = push_global gn (Gtblnorm (gn, params, body)) let push_global_case gn params annot a accu bs = push_global gn (Gletcase (gn, params, annot, a, accu, bs)) (* Compares [t1] and [t2] up to alpha-equivalence. [t1] and [t2] may contain free variables. *) let eq_mllambda t1 t2 = eq_mllambda dummy_gname dummy_gname 0 LNmap.empty LNmap.empty t1 t2 (*s Compilation environment *) type env = { env_rel : mllambda list; (* (MLlocal lname) list *) env_bound : int; (* length of env_rel *) (* free variables *) env_urel : (int * mllambda) list ref; (* list of unbound rel *) env_named : (Id.t * mllambda) list ref; env_univ : lname option} let empty_env univ = { env_rel = []; env_bound = 0; env_urel = ref []; env_named = ref []; env_univ = univ } let push_rel env id = let local = fresh_lname id.binder_name in local, { env with env_rel = MLlocal local :: env.env_rel; env_bound = env.env_bound + 1 } let push_rels env ids = let lnames, env_rel = Array.fold_left (fun (names,env_rel) id -> let local = fresh_lname id.binder_name in (local::names, MLlocal local::env_rel)) ([],env.env_rel) ids in Array.of_list (List.rev lnames), { env with env_rel = env_rel; env_bound = env.env_bound + Array.length ids } let get_rel env id i = if i <= env.env_bound then List.nth env.env_rel (i-1) else let i = i - env.env_bound in try Int.List.assoc i !(env.env_urel) with Not_found -> let local = MLlocal (fresh_lname id) in env.env_urel := (i,local) :: !(env.env_urel); local let get_var env id = try Id.List.assoc id !(env.env_named) with Not_found -> let local = MLlocal (fresh_lname (Name id)) in env.env_named := (id, local)::!(env.env_named); local let fresh_univ () = fresh_lname (Name (Id.of_string "univ")) (*s Traduction of lambda to mllambda *) let get_prod_name codom = match codom with | MLlam(ids,_) -> ids.(0).lname | _ -> assert false let get_lname (_,l) = match l with | MLlocal id -> id | _ -> invalid_arg "Nativecode.get_lname" (* Collects free variables from env in an array of local names *) let fv_params env = let fvn, fvr = !(env.env_named), !(env.env_urel) in let size = List.length fvn + List.length fvr in let start,params = match env.env_univ with | None -> 0, Array.make size dummy_lname | Some u -> 1, let t = Array.make (size + 1) dummy_lname in t.(0) <- u; t in if Array.is_empty params then empty_params else begin let fvn = ref fvn in let i = ref start in while not (List.is_empty !fvn) do params.(!i) <- get_lname (List.hd !fvn); fvn := List.tl !fvn; incr i done; let fvr = ref fvr in while not (List.is_empty !fvr) do params.(!i) <- get_lname (List.hd !fvr); fvr := List.tl !fvr; incr i done; params end let generalize_fv env body = mkMLlam (fv_params env) body let empty_args = [||] let fv_args env fvn fvr = let size = List.length fvn + List.length fvr in let start,args = match env.env_univ with | None -> 0, Array.make size (MLint 0) | Some u -> 1, let t = Array.make (size + 1) (MLint 0) in t.(0) <- MLlocal u; t in if Array.is_empty args then empty_args else begin let fvn = ref fvn in let i = ref start in while not (List.is_empty !fvn) do args.(!i) <- get_var env (fst (List.hd !fvn)); fvn := List.tl !fvn; incr i done; let fvr = ref fvr in while not (List.is_empty !fvr) do let (k,_ as kml) = List.hd !fvr in let n = get_lname kml in args.(!i) <- get_rel env n.lname k; fvr := List.tl !fvr; incr i done; args end let get_value_code i = MLapp (MLglobal (Ginternal "get_value"), [|MLglobal symbols_tbl_name; MLint i|]) let get_sort_code i = MLapp (MLglobal (Ginternal "get_sort"), [|MLglobal symbols_tbl_name; MLint i|]) let get_name_code i = MLapp (MLglobal (Ginternal "get_name"), [|MLglobal symbols_tbl_name; MLint i|]) let get_const_code i = MLapp (MLglobal (Ginternal "get_const"), [|MLglobal symbols_tbl_name; MLint i|]) let get_match_code i = MLapp (MLglobal (Ginternal "get_match"), [|MLglobal symbols_tbl_name; MLint i|]) let get_ind_code i = MLapp (MLglobal (Ginternal "get_ind"), [|MLglobal symbols_tbl_name; MLint i|]) let get_meta_code i = MLapp (MLglobal (Ginternal "get_meta"), [|MLglobal symbols_tbl_name; MLint i|]) let get_evar_code i = MLapp (MLglobal (Ginternal "get_evar"), [|MLglobal symbols_tbl_name; MLint i|]) let get_level_code i = MLapp (MLglobal (Ginternal "get_level"), [|MLglobal symbols_tbl_name; MLint i|]) let get_proj_code i = MLapp (MLglobal (Ginternal "get_proj"), [|MLglobal symbols_tbl_name; MLint i|]) type rlist = | Rnil | Rcons of lname option mllam_pattern list ref * LNset.t * mllambda * rlist' and rlist' = rlist ref let rm_params fv params = Array.map (fun l -> if LNset.mem l fv then Some l else None) params let rec insert pat body rl = match !rl with | Rnil -> let fv = fv_lam body in begin match pat with | ConstPattern _ as p -> rl:= Rcons(ref [p], fv, body, ref Rnil) | NonConstPattern (tag,args) -> let args = rm_params fv args in rl:= Rcons(ref [NonConstPattern (tag,args)], fv, body, ref Rnil) end | Rcons(l,fv,body',rl) -> if eq_mllambda body body' then match pat with | ConstPattern _ as p -> l := p::!l | NonConstPattern (tag,args) -> let args = rm_params fv args in l := NonConstPattern (tag,args)::!l else insert pat body rl let rec to_list rl = match !rl with | Rnil -> [] | Rcons(l,_,body,tl) -> (!l,body)::to_list tl let merge_branches t = let newt = ref Rnil in Array.iter (fun (pat,body) -> insert pat body newt) t; Array.of_list (to_list newt) let app_prim p args = MLapp(MLprimitive p, args) type prim_aux = | PAprim of string * pconstant * CPrimitives.t * prim_aux array | PAml of mllambda let add_check cond targs args = let aux cond t a = match a with | PAml(MLint _) -> cond | PAml ml -> (* FIXME: use explicit equality function *) if List.mem (t, ml) cond then cond else (t, ml)::cond | _ -> cond in Array.fold_left2 aux cond targs args let extract_prim ml_of l = let decl = ref [] in let cond = ref [] in let type_args p = let rec aux = function [] | [_] -> [] | h :: t -> h :: aux t in Array.of_list (aux (CPrimitives.types p)) in let rec aux l = match l with | Lprim(prefix,kn,p,args) -> let targs = type_args p in let args = Array.map aux args in cond := add_check !cond targs args; PAprim(prefix,kn,p,args) | Lrel _ | Lvar _ | Luint _ | Lval _ | Lconst _ -> PAml (ml_of l) | _ -> let x = fresh_lname Anonymous in decl := (x,ml_of l)::!decl; PAml (MLlocal x) in let res = aux l in (!decl, !cond, res) let cast_to_int v = match v with | MLint _ -> v | _ -> MLapp(MLprimitive Val_to_int, [|v|]) let compile_prim decl cond paux = let rec opt_prim_aux paux = match paux with | PAprim(_prefix, _kn, op, args) -> let args = Array.map opt_prim_aux args in app_prim (Coq_primitive(op,None)) args | PAml ml -> ml and naive_prim_aux paux = match paux with | PAprim(prefix, kn, op, args) -> app_prim (Coq_primitive(op, Some (prefix,kn))) (Array.map naive_prim_aux args) | PAml ml -> ml in let compile_cond cond paux = match cond with | [] -> opt_prim_aux paux | [CPrimitives.(PITT_type PT_int63), c1] -> MLif(app_prim Is_int [|c1|], opt_prim_aux paux, naive_prim_aux paux) | _ -> let ci, cf = let is_int = function CPrimitives.(PITT_type PT_int63), _ -> true | _ -> false in List.partition is_int cond in let condi = let cond = List.fold_left (fun ml (_, c) -> app_prim MLland [| ml; cast_to_int c|]) (MLint 0) ci in app_prim MLmagic [|cond|] in let condf = match cf with | [] -> MLint 0 | [_, c1] -> app_prim Is_float [|c1|] | (_, c1) :: condf -> List.fold_left (fun ml (_, c) -> app_prim MLand [| ml; app_prim Is_float [|c|]|]) (app_prim Is_float [|c1|]) condf in match ci, cf with | [], [] -> opt_prim_aux paux | _ :: _, [] -> MLif(condi, naive_prim_aux paux, opt_prim_aux paux) | [], _ :: _ -> MLif(condf, opt_prim_aux paux, naive_prim_aux paux) | _ :: _, _ :: _ -> let cond = app_prim MLand [|condf; app_prim MLnot [|condi|]|] in MLif(cond, opt_prim_aux paux, naive_prim_aux paux) in let add_decl decl body = List.fold_left (fun body (x,d) -> MLlet(x,d,body)) body decl in (* The optimizations done for checking if integer values are closed are valid only on 64-bit architectures. So on 32-bit architectures, we fall back to less optimized checks. *) if max_int = 1073741823 (* 32-bits *) then add_decl decl (naive_prim_aux paux) else add_decl decl (compile_cond cond paux) let ml_of_instance instance u = let ml_of_level l = match Univ.Level.var_index l with | Some i -> let univ = MLapp(MLprimitive MLmagic, [|MLlocal (Option.get instance)|]) in mkMLapp (MLprimitive MLarrayget) [|univ; MLint i|] | None -> let i = push_symbol (SymbLevel l) in get_level_code i in let u = Univ.Instance.to_array u in if Array.is_empty u then [||] else let u = Array.map ml_of_level u in [|MLapp (MLprimitive MLmagic, [|MLarray u|])|] let rec ml_of_lam env l t = match t with | Lrel(id ,i) -> get_rel env id i | Lvar id -> get_var env id | Lmeta(mv,_ty) -> let tyn = fresh_lname Anonymous in let i = push_symbol (SymbMeta mv) in MLapp(MLprimitive Mk_meta, [|get_meta_code i; MLlocal tyn|]) | Levar(evk, args) -> let i = push_symbol (SymbEvar evk) in (** Arguments are *not* reversed in evar instances in native compilation *) let args = MLarray(Array.map (ml_of_lam env l) args) in MLapp(MLprimitive Mk_evar, [|get_evar_code i; args|]) | Lprod(dom,codom) -> let dom = ml_of_lam env l dom in let codom = ml_of_lam env l codom in let n = get_prod_name codom in let i = push_symbol (SymbName n) in MLapp(MLprimitive Mk_prod, [|get_name_code i;dom;codom|]) | Llam(ids,body) -> let lnames,env = push_rels env ids in MLlam(lnames, ml_of_lam env l body) | Lrec(id,body) -> let ids,body = decompose_Llam body in let lname, env = push_rel env id in let lnames, env = push_rels env ids in MLletrec([|lname, lnames, ml_of_lam env l body|], MLlocal lname) | Llet(id,def,body) -> let def = ml_of_lam env l def in let lname, env = push_rel env id in let body = ml_of_lam env l body in MLlet(lname,def,body) | Lapp(f,args) -> MLapp(ml_of_lam env l f, Array.map (ml_of_lam env l) args) | Lconst (prefix, (c, u)) -> let args = ml_of_instance env.env_univ u in mkMLapp (MLglobal(Gconstant (prefix, c))) args | Lproj (prefix, ind, i) -> MLglobal(Gproj (prefix, ind, i)) | Lprim _ -> let decl,cond,paux = extract_prim (ml_of_lam env l) t in compile_prim decl cond paux | Lcase (annot,p,a,bs) -> (* let predicate_uid fv_pred = compilation of p let rec case_uid fv a_uid = match a_uid with | Accu _ => mk_sw (predicate_uid fv_pred) (case_uid fv) a_uid | Ci argsi => compilation of branches compile case = case_uid fv (compilation of a) *) (* Compilation of the predicate *) (* Remark: if we do not want to compile the predicate we should a least compute the fv, then store the lambda representation of the predicate (not the mllambda) *) let env_p = empty_env env.env_univ in let pn = fresh_gpred l in let mlp = ml_of_lam env_p l p in let mlp = generalize_fv env_p mlp in let (pfvn,pfvr) = !(env_p.env_named), !(env_p.env_urel) in let pn = push_global_let pn mlp in (* Compilation of the case *) let env_c = empty_env env.env_univ in let a_uid = fresh_lname Anonymous in let la_uid = MLlocal a_uid in (* compilation of branches *) let nbconst = Array.length bs.constant_branches in let nbtotal = nbconst + Array.length bs.nonconstant_branches in let br = Array.init nbtotal (fun i -> if i < Array.length bs.constant_branches then (ConstPattern i, ml_of_lam env_c l bs.constant_branches.(i)) else let (params, body) = bs.nonconstant_branches.(i-nbconst) in let lnames, env_c = push_rels env_c params in (NonConstPattern (i-nbconst+1,lnames), ml_of_lam env_c l body) ) in let cn = fresh_gcase l in (* Compilation of accu branch *) let pred = MLapp(MLglobal pn, fv_args env_c pfvn pfvr) in let (fvn, fvr) = !(env_c.env_named), !(env_c.env_urel) in let cn_fv = mkMLapp (MLglobal cn) (fv_args env_c fvn fvr) in (* remark : the call to fv_args does not add free variables in env_c *) let i = push_symbol (SymbMatch annot) in let accu = MLapp(MLprimitive Mk_sw, [| get_match_code i; MLapp (MLprimitive Cast_accu, [|la_uid|]); pred; cn_fv |]) in (* let body = MLlam([|a_uid|], MLmatch(annot, la_uid, accu, bs)) in let case = generalize_fv env_c body in *) let cn = push_global_case cn (Array.append (fv_params env_c) [|a_uid|]) annot la_uid accu (merge_branches br) in (* Final result *) let arg = ml_of_lam env l a in let force = if annot.asw_finite then arg else mkForceCofix annot.asw_prefix annot.asw_ind arg in mkMLapp (MLapp (MLglobal cn, fv_args env fvn fvr)) [|force|] | Lif(t,bt,bf) -> MLif(ml_of_lam env l t, ml_of_lam env l bt, ml_of_lam env l bf) | Lfix ((rec_pos, inds, start), (ids, tt, tb)) -> (* let type_f fvt = [| type fix |] let norm_f1 fv f1 .. fn params1 = body1 .. let norm_fn fv f1 .. fn paramsn = bodyn let norm fv f1 .. fn = [|norm_f1 fv f1 .. fn; ..; norm_fn fv f1 .. fn|] compile fix = let rec f1 params1 = if is_accu rec_pos.(1) then mk_fix (type_f fvt) (norm fv) params1 else norm_f1 fv f1 .. fn params1 and .. and fn paramsn = if is_accu rec_pos.(n) then mk_fix (type_f fvt) (norm fv) paramsn else norm_fn fv f1 .. fv paramsn in start *) (* Compilation of type *) let env_t = empty_env env.env_univ in let ml_t = Array.map (ml_of_lam env_t l) tt in let params_t = fv_params env_t in let args_t = fv_args env !(env_t.env_named) !(env_t.env_urel) in let gft = fresh_gfixtype l in let gft = push_global_fixtype gft params_t ml_t in let mk_type = MLapp(MLglobal gft, args_t) in (* Compilation of norm_i *) let ndef = Array.length ids in let lf,env_n = push_rels (empty_env env.env_univ) ids in let t_params = Array.make ndef [||] in let t_norm_f = Array.make ndef (Gnorm (l,-1)) in let mk_let _envi (id,def) t = MLlet (id,def,t) in let mk_lam_or_let (params,lets,env) (id,def) = let ln,env' = push_rel env id in match def with | None -> (ln::params,lets,env') | Some lam -> (params, (ln,ml_of_lam env l lam)::lets,env') in let ml_of_fix i body = let varsi, bodyi = decompose_Llam_Llet body in let paramsi,letsi,envi = Array.fold_left mk_lam_or_let ([],[],env_n) varsi in let paramsi,letsi = Array.of_list (List.rev paramsi), Array.of_list (List.rev letsi) in t_norm_f.(i) <- fresh_gnorm l; let bodyi = ml_of_lam envi l bodyi in t_params.(i) <- paramsi; let bodyi = Array.fold_right (mk_let envi) letsi bodyi in mkMLlam paramsi bodyi in let tnorm = Array.mapi ml_of_fix tb in let fvn,fvr = !(env_n.env_named), !(env_n.env_urel) in let fv_params = fv_params env_n in let fv_args' = Array.map (fun id -> MLlocal id) fv_params in let norm_params = Array.append fv_params lf in let t_norm_f = Array.mapi (fun i body -> push_global_let (t_norm_f.(i)) (mkMLlam norm_params body)) tnorm in let norm = fresh_gnormtbl l in let norm = push_global_norm norm fv_params (Array.map (fun g -> mkMLapp (MLglobal g) fv_args') t_norm_f) in (* Compilation of fix *) let fv_args = fv_args env fvn fvr in let lf, _env = push_rels env ids in let lf_args = Array.map (fun id -> MLlocal id) lf in let mk_norm = MLapp(MLglobal norm, fv_args) in let mkrec i lname = let paramsi = t_params.(i) in let reci = MLlocal (paramsi.(rec_pos.(i))) in let pargsi = Array.map (fun id -> MLlocal id) paramsi in let (prefix, ind) = inds.(i) in let body = MLif(MLisaccu (prefix, ind, reci), mkMLapp (MLapp(MLprimitive (Mk_fix(rec_pos,i)), [|mk_type; mk_norm|])) pargsi, MLapp(MLglobal t_norm_f.(i), Array.concat [fv_args;lf_args;pargsi])) in (lname, paramsi, body) in MLletrec(Array.mapi mkrec lf, lf_args.(start)) | Lcofix (start, (ids, tt, tb)) -> (* Compilation of type *) let env_t = empty_env env.env_univ in let ml_t = Array.map (ml_of_lam env_t l) tt in let params_t = fv_params env_t in let args_t = fv_args env !(env_t.env_named) !(env_t.env_urel) in let gft = fresh_gfixtype l in let gft = push_global_fixtype gft params_t ml_t in let mk_type = MLapp(MLglobal gft, args_t) in (* Compilation of norm_i *) let ndef = Array.length ids in let lf,env_n = push_rels (empty_env env.env_univ) ids in let t_params = Array.make ndef [||] in let t_norm_f = Array.make ndef (Gnorm (l,-1)) in let ml_of_fix i body = let idsi,bodyi = decompose_Llam body in let paramsi, envi = push_rels env_n idsi in t_norm_f.(i) <- fresh_gnorm l; let bodyi = ml_of_lam envi l bodyi in t_params.(i) <- paramsi; mkMLlam paramsi bodyi in let tnorm = Array.mapi ml_of_fix tb in let fvn,fvr = !(env_n.env_named), !(env_n.env_urel) in let fv_params = fv_params env_n in let fv_args' = Array.map (fun id -> MLlocal id) fv_params in let norm_params = Array.append fv_params lf in let t_norm_f = Array.mapi (fun i body -> push_global_let (t_norm_f.(i)) (mkMLlam norm_params body)) tnorm in let norm = fresh_gnormtbl l in let norm = push_global_norm norm fv_params (Array.map (fun g -> mkMLapp (MLglobal g) fv_args') t_norm_f) in (* Compilation of fix *) let fv_args = fv_args env fvn fvr in let mk_norm = MLapp(MLglobal norm, fv_args) in let lnorm = fresh_lname Anonymous in let ltype = fresh_lname Anonymous in let lf, _env = push_rels env ids in let lf_args = Array.map (fun id -> MLlocal id) lf in let upd i _lname cont = let paramsi = t_params.(i) in let pargsi = Array.map (fun id -> MLlocal id) paramsi in let uniti = fresh_lname Anonymous in let body = MLlam(Array.append paramsi [|uniti|], MLapp(MLglobal t_norm_f.(i), Array.concat [fv_args;lf_args;pargsi])) in MLsequence(MLapp(MLprimitive Upd_cofix, [|lf_args.(i);body|]), cont) in let upd = Array.fold_right_i upd lf lf_args.(start) in let mk_let i lname cont = MLlet(lname, MLapp(MLprimitive(Mk_cofix i),[| MLlocal ltype; MLlocal lnorm|]), cont) in let init = Array.fold_right_i mk_let lf upd in MLlet(lnorm, mk_norm, MLlet(ltype, mk_type, init)) (* let mkrec i lname = let paramsi = t_params.(i) in let pargsi = Array.map (fun id -> MLlocal id) paramsi in let uniti = fresh_lname Anonymous in let body = MLapp( MLprimitive(Mk_cofix i), [|mk_type;mk_norm; MLlam([|uniti|], MLapp(MLglobal t_norm_f.(i), Array.concat [fv_args;lf_args;pargsi]))|]) in (lname, paramsi, body) in MLletrec(Array.mapi mkrec lf, lf_args.(start)) *) | Lint tag -> MLapp(MLprimitive Mk_int, [|MLint tag|]) | Lmakeblock (prefix,cn,tag,args) -> let args = Array.map (ml_of_lam env l) args in MLconstruct(prefix,cn,tag,args) | Luint i -> MLapp(MLprimitive Mk_uint, [|MLuint i|]) | Lfloat f -> MLapp(MLprimitive Mk_float, [|MLfloat f|]) | Lval v -> let i = push_symbol (SymbValue v) in get_value_code i | Lsort s -> let i = push_symbol (SymbSort s) in let uarg = match env.env_univ with | None -> MLarray [||] | Some u -> MLlocal u in let uarg = MLapp(MLprimitive MLmagic, [|uarg|]) in MLapp(MLprimitive Mk_sort, [|get_sort_code i; uarg|]) | Lind (prefix, (ind, u)) -> let uargs = ml_of_instance env.env_univ u in mkMLapp (MLglobal (Gind (prefix, ind))) uargs | Llazy -> MLglobal (Ginternal "lazy") | Lforce -> MLglobal (Ginternal "Lazy.force") let mllambda_of_lambda univ auxdefs l t = let env = empty_env univ in global_stack := auxdefs; let ml = ml_of_lam env l t in let fv_rel = !(env.env_urel) in let fv_named = !(env.env_named) in (* build the free variables *) let get_lname (_,t) = match t with | MLlocal x -> x | _ -> assert false in let params = List.append (List.map get_lname fv_rel) (List.map get_lname fv_named) in if List.is_empty params then (!global_stack, ([],[]), ml) (* final result : global list, fv, ml *) else (!global_stack, (fv_named, fv_rel), mkMLlam (Array.of_list params) ml) (** Code optimization **) (** Optimization of match and fix *) let can_subst l = match l with | MLlocal _ | MLint _ | MLuint _ | MLglobal _ -> true | _ -> false let subst s l = if LNmap.is_empty s then l else let rec aux l = match l with | MLlocal id -> (try LNmap.find id s with Not_found -> l) | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ | MLfloat _ -> l | MLlam(params,body) -> MLlam(params, aux body) | MLletrec(defs,body) -> let arec (f,params,body) = (f,params,aux body) in MLletrec(Array.map arec defs, aux body) | MLlet(id,def,body) -> MLlet(id,aux def, aux body) | MLapp(f,args) -> MLapp(aux f, Array.map aux args) | MLif(t,b1,b2) -> MLif(aux t, aux b1, aux b2) | MLmatch(annot,a,accu,bs) -> let auxb (cargs,body) = (cargs,aux body) in MLmatch(annot,a,aux accu, Array.map auxb bs) | MLconstruct(prefix,c,tag,args) -> MLconstruct(prefix,c,tag,Array.map aux args) | MLsetref(s,l1) -> MLsetref(s,aux l1) | MLsequence(l1,l2) -> MLsequence(aux l1, aux l2) | MLarray arr -> MLarray (Array.map aux arr) | MLisaccu (s, ind, l) -> MLisaccu (s, ind, aux l) in aux l let add_subst id v s = match v with | MLlocal id' when Int.equal id.luid id'.luid -> s | _ -> LNmap.add id v s let subst_norm params args s = let len = Array.length params in assert (Int.equal (Array.length args) len && Array.for_all can_subst args); let s = ref s in for i = 0 to len - 1 do s := add_subst params.(i) args.(i) !s done; !s let subst_case params args s = let len = Array.length params in assert (len > 0 && Int.equal (Array.length args) len && let r = ref true and i = ref 0 in (* we test all arguments excepted the last *) while !i < len - 1 && !r do r := can_subst args.(!i); incr i done; !r); let s = ref s in for i = 0 to len - 2 do s := add_subst params.(i) args.(i) !s done; !s, params.(len-1), args.(len-1) let empty_gdef = Int.Map.empty, Int.Map.empty let get_norm (gnorm, _) i = Int.Map.find i gnorm let get_case (_, gcase) i = Int.Map.find i gcase let all_lam n bs = let f (_, l) = match l with | MLlam(params, _) -> Int.equal (Array.length params) n | _ -> false in Array.for_all f bs let commutative_cut annot a accu bs args = let mkb (c,b) = match b with | MLlam(params, body) -> (c, Array.fold_left2 (fun body x v -> MLlet(x,v,body)) body params args) | _ -> assert false in MLmatch(annot, a, mkMLapp accu args, Array.map mkb bs) let optimize gdef l = let rec optimize s l = match l with | MLlocal id -> (try LNmap.find id s with Not_found -> l) | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ | MLfloat _ -> l | MLlam(params,body) -> MLlam(params, optimize s body) | MLletrec(decls,body) -> let opt_rec (f,params,body) = (f,params,optimize s body ) in MLletrec(Array.map opt_rec decls, optimize s body) | MLlet(id,def,body) -> let def = optimize s def in if can_subst def then optimize (add_subst id def s) body else MLlet(id,def,optimize s body) | MLapp(f, args) -> let args = Array.map (optimize s) args in begin match f with | MLglobal (Gnorm (_,i)) -> (try let params,body = get_norm gdef i in let s = subst_norm params args s in optimize s body with Not_found -> MLapp(optimize s f, args)) | MLglobal (Gcase (_,i)) -> (try let params,body = get_case gdef i in let s, id, arg = subst_case params args s in if can_subst arg then optimize (add_subst id arg s) body else MLlet(id, arg, optimize s body) with Not_found -> MLapp(optimize s f, args)) | _ -> let f = optimize s f in match f with | MLmatch (annot,a,accu,bs) -> if all_lam (Array.length args) bs then commutative_cut annot a accu bs args else MLapp(f, args) | _ -> MLapp(f, args) end | MLif(t,b1,b2) -> (* This optimization is critical: it applies to all fixpoints that start by matching on their recursive argument *) let t = optimize s t in let b1 = optimize s b1 in let b2 = optimize s b2 in begin match t, b2 with | MLisaccu (_, _, l1), MLmatch(annot, l2, _, bs) when eq_mllambda l1 l2 -> MLmatch(annot, l1, b1, bs) | _, _ -> MLif(t, b1, b2) end | MLmatch(annot,a,accu,bs) -> let opt_b (cargs,body) = (cargs,optimize s body) in MLmatch(annot, optimize s a, subst s accu, Array.map opt_b bs) | MLconstruct(prefix,c,tag,args) -> MLconstruct(prefix,c,tag,Array.map (optimize s) args) | MLsetref(r,l) -> MLsetref(r, optimize s l) | MLsequence(l1,l2) -> MLsequence(optimize s l1, optimize s l2) | MLarray arr -> MLarray (Array.map (optimize s) arr) | MLisaccu (pf, ind, l) -> MLisaccu (pf, ind, optimize s l) in optimize LNmap.empty l let optimize_stk stk = let add_global gdef g = match g with | Glet (Gnorm (_,i), body) -> let (gnorm, gcase) = gdef in (Int.Map.add i (decompose_MLlam body) gnorm, gcase) | Gletcase(Gcase (_,i), params, annot,a,accu,bs) -> let (gnorm,gcase) = gdef in (gnorm, Int.Map.add i (params,MLmatch(annot,a,accu,bs)) gcase) | Gletcase _ -> assert false | _ -> gdef in let gdef = List.fold_left add_global empty_gdef stk in let optimize_global g = match g with | Glet(Gconstant (prefix, c), body) -> Glet(Gconstant (prefix, c), optimize gdef body) | _ -> g in List.map optimize_global stk (** Printing to ocaml **) (* Redefine a bunch of functions in module Names to generate names acceptable to OCaml. *) let string_of_id s = Unicode.ascii_of_ident (Id.to_string s) let string_of_label l = string_of_id (Label.to_id l) let string_of_dirpath = function | [] -> "_" | sl -> String.concat "_" (List.rev_map string_of_id sl) (* The first letter of the file name has to be a capital to be accepted by *) (* OCaml as a module identifier. *) let string_of_dirpath s = "N"^string_of_dirpath s let mod_uid_of_dirpath dir = string_of_dirpath (DirPath.repr dir) let link_info_of_dirpath dir = Linked (mod_uid_of_dirpath dir ^ ".") let string_of_name x = match x with | Anonymous -> "anonymous" (* assert false *) | Name id -> string_of_id id let string_of_label_def l = match l with | None -> "" | Some l -> string_of_label l (* Relativization of module paths *) let rec list_of_mp acc = function | MPdot (mp,l) -> list_of_mp (string_of_label l::acc) mp | MPfile dp -> let dp = DirPath.repr dp in string_of_dirpath dp :: acc | MPbound mbid -> ("X"^string_of_id (MBId.to_id mbid))::acc let list_of_mp mp = list_of_mp [] mp let string_of_kn kn = let (mp,l) = KerName.repr kn in let mp = list_of_mp mp in String.concat "_" mp ^ "_" ^ string_of_label l let string_of_con c = string_of_kn (Constant.user c) let string_of_mind mind = string_of_kn (MutInd.user mind) let string_of_ind (mind,i) = string_of_kn (MutInd.user mind) ^ "_" ^ string_of_int i let string_of_gname g = match g with | Gind (prefix, (mind, i)) -> Format.sprintf "%sindaccu_%s_%i" prefix (string_of_mind mind) i | Gconstant (prefix, c) -> Format.sprintf "%sconst_%s" prefix (string_of_con c) | Gproj (prefix, (mind, n), i) -> Format.sprintf "%sproj_%s_%i_%i" prefix (string_of_mind mind) n i | Gcase (l,i) -> Format.sprintf "case_%s_%i" (string_of_label_def l) i | Gpred (l,i) -> Format.sprintf "pred_%s_%i" (string_of_label_def l) i | Gfixtype (l,i) -> Format.sprintf "fixtype_%s_%i" (string_of_label_def l) i | Gnorm (l,i) -> Format.sprintf "norm_%s_%i" (string_of_label_def l) i | Ginternal s -> Format.sprintf "%s" s | Gnormtbl (l,i) -> Format.sprintf "normtbl_%s_%i" (string_of_label_def l) i | Grel i -> Format.sprintf "rel_%i" i | Gnamed id -> Format.sprintf "named_%s" (string_of_id id) let pp_gname fmt g = Format.fprintf fmt "%s" (string_of_gname g) let pp_lname fmt ln = Format.fprintf fmt "x_%s_%i" (string_of_name ln.lname) ln.luid let pp_ldecls fmt ids = let len = Array.length ids in for i = 0 to len - 1 do Format.fprintf fmt " (%a : Nativevalues.t)" pp_lname ids.(i) done let string_of_construct prefix ~constant ind tag = let base = if constant then "Int" else "Construct" in Format.sprintf "%s%s_%s_%i" prefix base (string_of_ind ind) tag let string_of_accu_construct prefix ind = Format.sprintf "%sAccu_%s" prefix (string_of_ind ind) let pp_int fmt i = if i < 0 then Format.fprintf fmt "(%i)" i else Format.fprintf fmt "%i" i let pp_mllam fmt l = let rec pp_mllam fmt l = match l with | MLlocal ln -> Format.fprintf fmt "@[%a@]" pp_lname ln | MLglobal g -> Format.fprintf fmt "@[%a@]" pp_gname g | MLprimitive p -> Format.fprintf fmt "@[%a@]" pp_primitive p | MLlam(ids,body) -> Format.fprintf fmt "@[(fun%a@ ->@\n %a)@]" pp_ldecls ids pp_mllam body | MLletrec(defs, body) -> Format.fprintf fmt "@[%a@ in@\n%a@]" pp_letrec defs pp_mllam body | MLlet(id,def,body) -> Format.fprintf fmt "@[(let@ %a@ =@\n %a@ in@\n%a)@]" pp_lname id pp_mllam def pp_mllam body | MLapp(f, args) -> Format.fprintf fmt "@[%a@ %a@]" pp_mllam f (pp_args true) args | MLif(t,l1,l2) -> Format.fprintf fmt "@[(if %a then@\n %a@\nelse@\n %a)@]" pp_mllam t pp_mllam l1 pp_mllam l2 | MLmatch (annot, c, accu_br, br) -> let ind = annot.asw_ind in let prefix = annot.asw_prefix in let accu = string_of_accu_construct prefix ind in Format.fprintf fmt "@[begin match Obj.magic (%a) with@\n| %s _ ->@\n %a@\n%aend@]" pp_mllam c accu pp_mllam accu_br (pp_branches prefix ind) br | MLconstruct(prefix,ind,tag,args) -> Format.fprintf fmt "@[(Obj.magic (%s%a) : Nativevalues.t)@]" (string_of_construct prefix ~constant:false ind tag) pp_cargs args | MLint i -> pp_int fmt i | MLuint i -> Format.fprintf fmt "(%s)" (Uint63.compile i) | MLfloat f -> Format.fprintf fmt "(%s)" (Float64.compile f) | MLsetref (s, body) -> Format.fprintf fmt "@[%s@ :=@\n %a@]" s pp_mllam body | MLsequence(l1,l2) -> Format.fprintf fmt "@[%a;@\n%a@]" pp_mllam l1 pp_mllam l2 | MLarray arr -> let len = Array.length arr in Format.fprintf fmt "@[[|"; if 0 < len then begin for i = 0 to len - 2 do Format.fprintf fmt "%a;" pp_mllam arr.(i) done; pp_mllam fmt arr.(len-1) end; Format.fprintf fmt "|]@]" | MLisaccu (prefix, ind, c) -> let accu = string_of_accu_construct prefix ind in Format.fprintf fmt "@[begin match Obj.magic (%a) with@\n| %s _ ->@\n true@\n| _ ->@\n false@\nend@]" pp_mllam c accu and pp_letrec fmt defs = let len = Array.length defs in let pp_one_rec (fn, argsn, body) = Format.fprintf fmt "%a%a =@\n %a" pp_lname fn pp_ldecls argsn pp_mllam body in Format.fprintf fmt "@[let rec "; pp_one_rec defs.(0); for i = 1 to len - 1 do Format.fprintf fmt "@\nand "; pp_one_rec defs.(i) done; and pp_blam fmt l = match l with | MLprimitive (Mk_prod | Mk_sort) (* FIXME: why this special case? *) | MLlam _ | MLletrec _ | MLlet _ | MLapp _ | MLif _ -> Format.fprintf fmt "(%a)" pp_mllam l | MLconstruct(_,_,_,args) when Array.length args > 0 -> Format.fprintf fmt "(%a)" pp_mllam l | _ -> pp_mllam fmt l and pp_args sep fmt args = let sep = if sep then " " else "," in let len = Array.length args in if len > 0 then begin Format.fprintf fmt "%a" pp_blam args.(0); for i = 1 to len - 1 do Format.fprintf fmt "%s%a" sep pp_blam args.(i) done end and pp_cargs fmt args = let len = Array.length args in match len with | 0 -> () | 1 -> Format.fprintf fmt " %a" pp_blam args.(0) | _ -> Format.fprintf fmt "(%a)" (pp_args false) args and pp_cparam fmt param = match param with | Some l -> pp_mllam fmt (MLlocal l) | None -> Format.fprintf fmt "_" and pp_cparams fmt params = let len = Array.length params in match len with | 0 -> () | 1 -> Format.fprintf fmt " %a" pp_cparam params.(0) | _ -> let aux fmt params = Format.fprintf fmt "%a" pp_cparam params.(0); for i = 1 to len - 1 do Format.fprintf fmt ",%a" pp_cparam params.(i) done in Format.fprintf fmt "(%a)" aux params and pp_branches prefix ind fmt bs = let pp_branch (cargs,body) = let pp_pat fmt = function | ConstPattern i -> Format.fprintf fmt "| %s " (string_of_construct prefix ~constant:true ind i) | NonConstPattern (tag,args) -> Format.fprintf fmt "| %s%a " (string_of_construct prefix ~constant:false ind tag) pp_cparams args in let rec pp_pats fmt pats = match pats with | [] -> () | pat::pats -> Format.fprintf fmt "%a%a" pp_pat pat pp_pats pats in Format.fprintf fmt "%a ->@\n %a@\n" pp_pats cargs pp_mllam body in Array.iter pp_branch bs and pp_primitive fmt = function | Mk_prod -> Format.fprintf fmt "mk_prod_accu" | Mk_sort -> Format.fprintf fmt "mk_sort_accu" | Mk_ind -> Format.fprintf fmt "mk_ind_accu" | Mk_const -> Format.fprintf fmt "mk_constant_accu" | Mk_sw -> Format.fprintf fmt "mk_sw_accu" | Mk_fix(rec_pos,start) -> let pp_rec_pos fmt rec_pos = Format.fprintf fmt "@[[| %i" rec_pos.(0); for i = 1 to Array.length rec_pos - 1 do Format.fprintf fmt "; %i" rec_pos.(i) done; Format.fprintf fmt " |]@]" in Format.fprintf fmt "mk_fix_accu %a %i" pp_rec_pos rec_pos start | Mk_cofix(start) -> Format.fprintf fmt "mk_cofix_accu %i" start | Mk_rel i -> Format.fprintf fmt "mk_rel_accu %i" i | Mk_var id -> Format.fprintf fmt "mk_var_accu (Names.Id.of_string \"%s\")" (string_of_id id) | Mk_proj -> Format.fprintf fmt "mk_proj_accu" | Is_int -> Format.fprintf fmt "is_int" | Is_float -> Format.fprintf fmt "is_float" | Cast_accu -> Format.fprintf fmt "cast_accu" | Upd_cofix -> Format.fprintf fmt "upd_cofix" | Force_cofix -> Format.fprintf fmt "force_cofix" | Mk_uint -> Format.fprintf fmt "mk_uint" | Mk_float -> Format.fprintf fmt "mk_float" | Mk_int -> Format.fprintf fmt "mk_int" | Mk_bool -> Format.fprintf fmt "mk_bool" | Val_to_int -> Format.fprintf fmt "val_to_int" | Mk_meta -> Format.fprintf fmt "mk_meta_accu" | Mk_evar -> Format.fprintf fmt "mk_evar_accu" | MLand -> Format.fprintf fmt "(&&)" | MLnot -> Format.fprintf fmt "not" | MLle -> Format.fprintf fmt "(<=)" | MLlt -> Format.fprintf fmt "(<)" | MLinteq -> Format.fprintf fmt "(==)" | MLlsl -> Format.fprintf fmt "(lsl)" | MLlsr -> Format.fprintf fmt "(lsr)" | MLland -> Format.fprintf fmt "(land)" | MLlor -> Format.fprintf fmt "(lor)" | MLlxor -> Format.fprintf fmt "(lxor)" | MLadd -> Format.fprintf fmt "(+)" | MLsub -> Format.fprintf fmt "(-)" | MLmul -> Format.fprintf fmt "( * )" | MLmagic -> Format.fprintf fmt "Obj.magic" | MLarrayget -> Format.fprintf fmt "Array.get" | Mk_empty_instance -> Format.fprintf fmt "Univ.Instance.empty" | Coq_primitive (op,None) -> Format.fprintf fmt "no_check_%s" (CPrimitives.to_string op) | Coq_primitive (op, Some (prefix,(c,_))) -> Format.fprintf fmt "%s %a" (CPrimitives.to_string op) pp_mllam (MLglobal (Gconstant (prefix,c))) in Format.fprintf fmt "@[%a@]" pp_mllam l let pp_array fmt t = let len = Array.length t in Format.fprintf fmt "@[[|"; for i = 0 to len - 2 do Format.fprintf fmt "%a; " pp_mllam t.(i) done; if len > 0 then Format.fprintf fmt "%a" pp_mllam t.(len - 1); Format.fprintf fmt "|]@]" let pp_global fmt g = match g with | Glet (gn, c) -> let ids, c = decompose_MLlam c in Format.fprintf fmt "@[let %a%a =@\n %a@]@\n@." pp_gname gn pp_ldecls ids pp_mllam c | Gopen s -> Format.fprintf fmt "@[open %s@]@." s | Gtype (ind, lar) -> let rec aux s arity = if Int.equal arity 0 then s else aux (s^" * Nativevalues.t") (arity-1) in let pp_const_sig fmt (tag,arity) = if arity > 0 then let sig_str = aux "of Nativevalues.t" (arity-1) in let cstr = string_of_construct "" ~constant:false ind tag in Format.fprintf fmt " | %s %s@\n" cstr sig_str else let sig_str = if arity > 0 then aux "of Nativevalues.t" (arity-1) else "" in let cstr = string_of_construct "" ~constant:true ind tag in Format.fprintf fmt " | %s %s@\n" cstr sig_str in let pp_const_sigs fmt lar = Format.fprintf fmt " | %s of Nativevalues.t@\n" (string_of_accu_construct "" ind); Array.iter (pp_const_sig fmt) lar in Format.fprintf fmt "@[type ind_%s =@\n%a@]@\n@." (string_of_ind ind) pp_const_sigs lar | Gtblfixtype (g, params, t) -> Format.fprintf fmt "@[let %a %a =@\n %a@]@\n@." pp_gname g pp_ldecls params pp_array t | Gtblnorm (g, params, t) -> Format.fprintf fmt "@[let %a %a =@\n %a@]@\n@." pp_gname g pp_ldecls params pp_array t | Gletcase(gn,params,annot,a,accu,bs) -> Format.fprintf fmt "@[(* Hash = %i *)@\nlet rec %a %a =@\n %a@]@\n@." (hash_global g) pp_gname gn pp_ldecls params pp_mllam (MLmatch(annot,a,accu,bs)) | Gcomment s -> Format.fprintf fmt "@[(* %s *)@]@." s (** Compilation of elements in environment **) let rec compile_with_fv env sigma univ auxdefs l t = let (auxdefs,(fv_named,fv_rel),ml) = mllambda_of_lambda univ auxdefs l t in if List.is_empty fv_named && List.is_empty fv_rel then (auxdefs,ml) else apply_fv env sigma univ (fv_named,fv_rel) auxdefs ml and apply_fv env sigma univ (fv_named,fv_rel) auxdefs ml = let get_rel_val (n,_) auxdefs = (* match !(lookup_rel_native_val n env) with | NVKnone -> *) compile_rel env sigma univ auxdefs n (* | NVKvalue (v,d) -> assert false *) in let get_named_val (id,_) auxdefs = (* match !(lookup_named_native_val id env) with | NVKnone -> *) compile_named env sigma univ auxdefs id (* | NVKvalue (v,d) -> assert false *) in let auxdefs = List.fold_right get_rel_val fv_rel auxdefs in let auxdefs = List.fold_right get_named_val fv_named auxdefs in let lvl = Context.Rel.length (rel_context env) in let fv_rel = List.map (fun (n,_) -> MLglobal (Grel (lvl-n))) fv_rel in let fv_named = List.map (fun (id,_) -> MLglobal (Gnamed id)) fv_named in let aux_name = fresh_lname Anonymous in auxdefs, MLlet(aux_name, ml, mkMLapp (MLlocal aux_name) (Array.of_list (fv_rel@fv_named))) and compile_rel env sigma univ auxdefs n = let open Context.Rel.Declaration in let decl = lookup_rel n env in let n = List.length (rel_context env) - n in match decl with | LocalDef (_,t,_) -> let code = lambda_of_constr env sigma t in let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in Glet(Grel n, code)::auxdefs | LocalAssum _ -> Glet(Grel n, MLprimitive (Mk_rel n))::auxdefs and compile_named env sigma univ auxdefs id = let open Context.Named.Declaration in match lookup_named id env with | LocalDef (_,t,_) -> let code = lambda_of_constr env sigma t in let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in Glet(Gnamed id, code)::auxdefs | LocalAssum _ -> Glet(Gnamed id, MLprimitive (Mk_var id))::auxdefs let compile_constant env sigma prefix ~interactive con cb = let no_univs = 0 = Univ.AUContext.size (Declareops.constant_polymorphic_context cb) in begin match cb.const_body with | Def t -> let t = Mod_subst.force_constr t in let code = lambda_of_constr env sigma t in if !Flags.debug then Feedback.msg_debug (Pp.str "Generated lambda code"); let is_lazy = is_lazy t in let code = if is_lazy then mk_lazy code else code in let name = if interactive then LinkedInteractive prefix else Linked prefix in let l = Constant.label con in let auxdefs,code = if no_univs then compile_with_fv env sigma None [] (Some l) code else let univ = fresh_univ () in let (auxdefs,code) = compile_with_fv env sigma (Some univ) [] (Some l) code in (auxdefs,mkMLlam [|univ|] code) in if !Flags.debug then Feedback.msg_debug (Pp.str "Generated mllambda code"); let code = optimize_stk (Glet(Gconstant ("", con),code)::auxdefs) in if !Flags.debug then Feedback.msg_debug (Pp.str "Optimized mllambda code"); code, name | _ -> let i = push_symbol (SymbConst con) in let args = if no_univs then [|get_const_code i; MLarray [||]|] else [|get_const_code i|] in (* let t = mkMLlam [|univ|] (mkMLapp (MLprimitive Mk_const) *) [Glet(Gconstant ("", con), mkMLapp (MLprimitive Mk_const) args)], if interactive then LinkedInteractive prefix else Linked prefix end module StringOrd = struct type t = string let compare = String.compare end module StringSet = Set.Make(StringOrd) let loaded_native_files = ref StringSet.empty let is_loaded_native_file s = StringSet.mem s !loaded_native_files let register_native_file s = loaded_native_files := StringSet.add s !loaded_native_files let is_code_loaded ~interactive name = match !name with | NotLinked -> false | LinkedInteractive s -> if (interactive && is_loaded_native_file s) then true else (name := NotLinked; false) | Linked s -> if is_loaded_native_file s then true else (name := NotLinked; false) let param_name = Name (Id.of_string "params") let arg_name = Name (Id.of_string "arg") let compile_mind mb mind stack = let u = Declareops.inductive_polymorphic_context mb in (** Generate data for every block *) let f i stack ob = let ind = (mind, i) in let gtype = Gtype(ind, ob.mind_reloc_tbl) in let j = push_symbol (SymbInd ind) in let name = Gind ("", ind) in let accu = let args = if Int.equal (Univ.AUContext.size u) 0 then [|get_ind_code j; MLarray [||]|] else [|get_ind_code j|] in Glet(name, MLapp (MLprimitive Mk_ind, args)) in let nparams = mb.mind_nparams in let add_proj proj_arg acc _pb = let tbl = ob.mind_reloc_tbl in (* Building info *) let ci = { ci_ind = ind; ci_npar = nparams; ci_cstr_nargs = [|0|]; ci_relevance = ob.mind_relevance; ci_cstr_ndecls = [||] (*FIXME*); ci_pp_info = { ind_tags = []; cstr_tags = [||] (*FIXME*); style = RegularStyle } } in let asw = { asw_ind = ind; asw_prefix = ""; asw_ci = ci; asw_reloc = tbl; asw_finite = true } in let c_uid = fresh_lname Anonymous in let cf_uid = fresh_lname Anonymous in let tag, arity = tbl.(0) in assert (arity > 0); let ci_uid = fresh_lname Anonymous in let cargs = Array.init arity (fun i -> if Int.equal i proj_arg then Some ci_uid else None) in let i = push_symbol (SymbProj (ind, proj_arg)) in let accu = MLapp (MLprimitive Cast_accu, [|MLlocal cf_uid|]) in let accu_br = MLapp (MLprimitive Mk_proj, [|get_proj_code i;accu|]) in let code = MLmatch(asw,MLlocal cf_uid,accu_br,[|[NonConstPattern (tag,cargs)],MLlocal ci_uid|]) in let code = MLlet(cf_uid, mkForceCofix "" ind (MLlocal c_uid), code) in let gn = Gproj ("", ind, proj_arg) in Glet (gn, mkMLlam [|c_uid|] code) :: acc in let projs = match mb.mind_record with | NotRecord | FakeRecord -> [] | PrimRecord info -> let _, _, _, pbs = info.(i) in Array.fold_left_i add_proj [] pbs in projs @ gtype :: accu :: stack in Array.fold_left_i f stack mb.mind_packets type code_location_update = link_info ref * link_info type code_location_updates = code_location_update Mindmap_env.t * code_location_update Cmap_env.t type linkable_code = global list * code_location_updates let empty_updates = Mindmap_env.empty, Cmap_env.empty let compile_mind_deps env prefix ~interactive (comp_stack, (mind_updates, const_updates) as init) mind = let mib,nameref = lookup_mind_key mind env in if is_code_loaded ~interactive nameref || Mindmap_env.mem mind mind_updates then init else let comp_stack = compile_mind mib mind comp_stack in let name = if interactive then LinkedInteractive prefix else Linked prefix in let upd = (nameref, name) in let mind_updates = Mindmap_env.add mind upd mind_updates in (comp_stack, (mind_updates, const_updates)) (* This function compiles all necessary dependencies of t, and generates code in reverse order, as well as linking information updates *) let compile_deps env sigma prefix ~interactive init t = let rec aux env lvl init t = match kind t with | Ind ((mind,_),_u) -> compile_mind_deps env prefix ~interactive init mind | Const c -> let c,_u = get_alias env c in let cb,(nameref,_) = lookup_constant_key c env in let (_, (_, const_updates)) = init in if is_code_loaded ~interactive nameref || (Cmap_env.mem c const_updates) then init else let comp_stack, (mind_updates, const_updates) = match cb.const_body with | Def t -> aux env lvl init (Mod_subst.force_constr t) | _ -> init in let code, name = compile_constant env sigma prefix ~interactive c cb in let comp_stack = code@comp_stack in let const_updates = Cmap_env.add c (nameref, name) const_updates in comp_stack, (mind_updates, const_updates) | Construct (((mind,_),_),_u) -> compile_mind_deps env prefix ~interactive init mind | Proj (p,c) -> let init = compile_mind_deps env prefix ~interactive init (Projection.mind p) in aux env lvl init c | Case (ci, _p, _c, _ac) -> let mind = fst ci.ci_ind in let init = compile_mind_deps env prefix ~interactive init mind in fold_constr_with_binders succ (aux env) lvl init t | Var id -> let open Context.Named.Declaration in begin match lookup_named id env with | LocalDef (_,t,_) -> aux env lvl init t | _ -> init end | Rel n when n > lvl -> let open Context.Rel.Declaration in let decl = lookup_rel n env in let env = env_of_rel n env in begin match decl with | LocalDef (_,t,_) -> aux env lvl init t | LocalAssum _ -> init end | _ -> fold_constr_with_binders succ (aux env) lvl init t in aux env 0 init t let compile_constant_field env prefix con acc cb = let (gl, _) = compile_constant ~interactive:false env empty_evars prefix con cb in gl@acc let compile_mind_field mp l acc mb = let mind = MutInd.make2 mp l in compile_mind mb mind acc let mk_open s = Gopen s let mk_internal_let s code = Glet(Ginternal s, code) (* ML Code for conversion function *) let mk_conv_code env sigma prefix t1 t2 = clear_symbols (); clear_global_tbl (); let gl, (mind_updates, const_updates) = let init = ([], empty_updates) in compile_deps env sigma prefix ~interactive:true init t1 in let gl, (mind_updates, const_updates) = let init = (gl, (mind_updates, const_updates)) in compile_deps env sigma prefix ~interactive:true init t2 in let code1 = lambda_of_constr env sigma t1 in let code2 = lambda_of_constr env sigma t2 in let (gl,code1) = compile_with_fv env sigma None gl None code1 in let (gl,code2) = compile_with_fv env sigma None gl None code2 in let t1 = mk_internal_let "t1" code1 in let t2 = mk_internal_let "t2" code2 in let g1 = MLglobal (Ginternal "t1") in let g2 = MLglobal (Ginternal "t2") in let setref1 = Glet(Ginternal "_", MLsetref("rt1",g1)) in let setref2 = Glet(Ginternal "_", MLsetref("rt2",g2)) in let gl = List.rev (setref2 :: setref1 :: t2 :: t1 :: gl) in let header = Glet(Ginternal "symbols_tbl", MLapp (MLglobal (Ginternal "get_symbols"), [|MLglobal (Ginternal "()")|])) in header::gl, (mind_updates, const_updates) let mk_norm_code env sigma prefix t = clear_symbols (); clear_global_tbl (); let gl, (mind_updates, const_updates) = let init = ([], empty_updates) in compile_deps env sigma prefix ~interactive:true init t in let code = lambda_of_constr env sigma t in let (gl,code) = compile_with_fv env sigma None gl None code in let t1 = mk_internal_let "t1" code in let g1 = MLglobal (Ginternal "t1") in let setref = Glet(Ginternal "_", MLsetref("rt1",g1)) in let gl = List.rev (setref :: t1 :: gl) in let header = Glet(Ginternal "symbols_tbl", MLapp (MLglobal (Ginternal "get_symbols"), [|MLglobal (Ginternal "()")|])) in header::gl, (mind_updates, const_updates) let mk_library_header dir = let libname = Format.sprintf "(str_decode \"%s\")" (str_encode dir) in [Glet(Ginternal "symbols_tbl", MLapp (MLglobal (Ginternal "get_library_native_symbols"), [|MLglobal (Ginternal libname)|]))] let update_location (r,v) = r := v let update_locations (ind_updates,const_updates) = Mindmap_env.iter (fun _ -> update_location) ind_updates; Cmap_env.iter (fun _ -> update_location) const_updates let add_header_comment mlcode s = Gcomment s :: mlcode (* vim: set filetype=ocaml foldmethod=marker: *) coq-8.11.0/kernel/term_typing.mli0000644000175000017500000000401113612664533016606 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Constr.t -> 'a -> (Constr.t * Univ.ContextSet.t * int) type typing_context val translate_local_def : env -> Id.t -> section_def_entry -> constr * Sorts.relevance * types val translate_local_assum : env -> types -> types * Sorts.relevance val translate_constant : env -> Constant.t -> constant_entry -> 'a constant_body val translate_opaque : env -> Constant.t -> 'a opaque_entry -> unit constant_body * typing_context val translate_recipe : env -> Constant.t -> Cooking.recipe -> Opaqueproof.opaque constant_body val check_delayed : 'a effect_handler -> typing_context -> 'a proof_output -> (Constr.t * Univ.ContextSet.t Opaqueproof.delayed_universes) (** Internal functions, mentioned here for debug purpose only *) val infer_declaration : env -> constant_entry -> typing_context Cooking.result val build_constant_declaration : env -> Opaqueproof.proofterm Cooking.result -> Opaqueproof.proofterm constant_body coq-8.11.0/kernel/primred.mli0000644000175000017500000000412213612664533015712 0ustar treinentreinenopen Names open Environ (** {5 Reduction of primitives} *) val add_retroknowledge : env -> Retroknowledge.action -> env val get_int_type : env -> Constant.t val get_float_type : env -> Constant.t val get_cmp_type : env -> Constant.t val get_bool_constructors : env -> constructor * constructor val get_carry_constructors : env -> constructor * constructor val get_pair_constructor : env -> constructor val get_cmp_constructors : env -> constructor * constructor * constructor val get_f_cmp_constructors : env -> constructor * constructor * constructor * constructor val get_f_class_constructors : env -> constructor * constructor * constructor * constructor * constructor * constructor * constructor * constructor * constructor exception NativeDestKO (* Should be raised by get_* functions on failure *) module type RedNativeEntries = sig type elem type args type evd (* will be unit in kernel, evar_map outside *) val get : args -> int -> elem val get_int : evd -> elem -> Uint63.t val get_float : evd -> elem -> Float64.t val mkInt : env -> Uint63.t -> elem val mkFloat : env -> Float64.t -> elem val mkBool : env -> bool -> elem val mkCarry : env -> bool -> elem -> elem (* true if carry *) val mkIntPair : env -> elem -> elem -> elem val mkFloatIntPair : env -> elem -> elem -> elem val mkLt : env -> elem val mkEq : env -> elem val mkGt : env -> elem val mkFLt : env -> elem val mkFEq : env -> elem val mkFGt : env -> elem val mkFNotComparable : env -> elem val mkPNormal : env -> elem val mkNNormal : env -> elem val mkPSubn : env -> elem val mkNSubn : env -> elem val mkPZero : env -> elem val mkNZero : env -> elem val mkPInf : env -> elem val mkNInf : env -> elem val mkNaN : env -> elem end module type RedNative = sig type elem type args type evd val red_prim : env -> evd -> CPrimitives.t -> args -> elem option end module RedNative : functor (E:RedNativeEntries) -> RedNative with type elem = E.elem with type args = E.args with type evd = E.evd coq-8.11.0/kernel/univ.mli0000644000175000017500000003444613612664533015245 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int -> t val equal : t -> t -> bool val hash : t -> int val compare : t -> t -> int end (** Qualified global universe level *) type t (** Type of universe levels. A universe level is essentially a unique name that will be associated to constraints later on. A level can be local to a definition or global. *) val set : t val prop : t val sprop : t (** The set and prop universe levels. *) val is_small : t -> bool (** Is the universe set or prop? *) val is_sprop : t -> bool val is_prop : t -> bool val is_set : t -> bool (** Is it specifically Prop or Set *) val compare : t -> t -> int (** Comparison function *) val equal : t -> t -> bool (** Equality function *) val hash : t -> int val make : UGlobal.t -> t val pr : t -> Pp.t (** Pretty-printing *) val to_string : t -> string (** Debug printing *) val var : int -> t val var_index : t -> int option val name : t -> UGlobal.t option end (** Sets of universe levels *) module LSet : sig include CSig.SetS with type elt = Level.t val pr : (Level.t -> Pp.t) -> t -> Pp.t (** Pretty-printing *) end module Universe : sig type t (** Type of universes. A universe is defined as a set of level expressions. A level expression is built from levels and successors of level expressions, i.e.: le ::= l + n, n \in N. A universe is said atomic if it consists of a single level expression with no increment, and algebraic otherwise (think the least upper bound of a set of level expressions). *) val compare : t -> t -> int (** Comparison function *) val equal : t -> t -> bool (** Equality function on formal universes *) val hash : t -> int (** Hash function *) val make : Level.t -> t (** Create a universe representing the given level. *) val pr : t -> Pp.t (** Pretty-printing *) val pr_with : (Level.t -> Pp.t) -> t -> Pp.t val is_level : t -> bool (** Test if the universe is a level or an algebraic universe. *) val is_levels : t -> bool (** Test if the universe is a lub of levels or contains +n's. *) val level : t -> Level.t option (** Try to get a level out of a universe, returns [None] if it is an algebraic universe. *) val levels : t -> LSet.t (** Get the levels inside the universe, forgetting about increments *) val super : t -> t (** The universe strictly above *) val sup : t -> t -> t (** The l.u.b. of 2 universes *) val sprop : t val type0m : t (** image of Prop in the universes hierarchy *) val type0 : t (** image of Set in the universes hierarchy *) val type1 : t (** the universe of the type of Prop/Set *) val is_sprop : t -> bool val is_type0m : t -> bool val is_type0 : t -> bool val exists : (Level.t * int -> bool) -> t -> bool val for_all : (Level.t * int -> bool) -> t -> bool val map : (Level.t * int -> 'a) -> t -> 'a list end (** Alias name. *) val pr_uni : Universe.t -> Pp.t (** The universes hierarchy: Type 0- = Prop <= Type 0 = Set <= Type 1 <= ... Typing of universes: Type 0-, Type 0 : Type 1; Type i : Type (i+1) if i>0 *) val type0m_univ : Universe.t val type0_univ : Universe.t val type1_univ : Universe.t val is_type0_univ : Universe.t -> bool val is_type0m_univ : Universe.t -> bool val is_univ_variable : Universe.t -> bool val is_small_univ : Universe.t -> bool val sup : Universe.t -> Universe.t -> Universe.t val super : Universe.t -> Universe.t val universe_level : Universe.t -> Level.t option (** [univ_level_mem l u] Is l is mentioned in u ? *) val univ_level_mem : Level.t -> Universe.t -> bool (** [univ_level_rem u v min] removes [u] from [v], resulting in [min] if [v] was exactly [u]. *) val univ_level_rem : Level.t -> Universe.t -> Universe.t -> Universe.t (** {6 Constraints. } *) type constraint_type = AcyclicGraph.constraint_type = Lt | Le | Eq type univ_constraint = Level.t * constraint_type * Level.t module Constraint : sig include Set.S with type elt = univ_constraint end val empty_constraint : Constraint.t val union_constraint : Constraint.t -> Constraint.t -> Constraint.t val eq_constraint : Constraint.t -> Constraint.t -> bool (** A value with universe Constraint.t. *) type 'a constrained = 'a * Constraint.t (** Constrained *) val constraints_of : 'a constrained -> Constraint.t (** Enforcing Constraint.t. *) type 'a constraint_function = 'a -> 'a -> Constraint.t -> Constraint.t val enforce_eq : Universe.t constraint_function val enforce_leq : Universe.t constraint_function val enforce_eq_level : Level.t constraint_function val enforce_leq_level : Level.t constraint_function (** Type explanation is used to decorate error messages to provide useful explanation why a given constraint is rejected. It is composed of a path of universes and relation kinds [(r1,u1);..;(rn,un)] means .. <(r1) u1 <(r2) ... <(rn) un (where <(ri) is the relation symbol denoted by ri, currently only < and <=). The lowest end of the chain is supposed known (see UniverseInconsistency exn). The upper end may differ from the second univ of UniverseInconsistency because all universes in the path are canonical. Note that each step does not necessarily correspond to an actual constraint, but reflect how the system stores the graph and may result from combination of several Constraint.t... *) type explanation = (constraint_type * Level.t) list type univ_inconsistency = constraint_type * Universe.t * Universe.t * explanation Lazy.t option exception UniverseInconsistency of univ_inconsistency (** {6 Support for universe polymorphism } *) (** Polymorphic maps from universe levels to 'a *) module LMap : sig include CMap.ExtS with type key = Level.t and module Set := LSet val lunion : 'a t -> 'a t -> 'a t (** [lunion x y] favors the bindings in the first map. *) val diff : 'a t -> 'a t -> 'a t (** [diff x y] removes bindings from x that appear in y (whatever the value). *) val subst_union : 'a option t -> 'a option t -> 'a option t (** [subst_union x y] favors the bindings of the first map that are [Some], otherwise takes y's bindings. *) val pr : ('a -> Pp.t) -> 'a t -> Pp.t (** Pretty-printing *) end type 'a universe_map = 'a LMap.t (** {6 Substitution} *) type universe_subst_fn = Level.t -> Universe.t type universe_level_subst_fn = Level.t -> Level.t (** A full substitution, might involve algebraic universes *) type universe_subst = Universe.t universe_map type universe_level_subst = Level.t universe_map module Variance : sig (** A universe position in the instance given to a cumulative inductive can be the following. Note there is no Contravariant case because [forall x : A, B <= forall x : A', B'] requires [A = A'] as opposed to [A' <= A]. *) type t = Irrelevant | Covariant | Invariant (** [check_subtype x y] holds if variance [y] is also an instance of [x] *) val check_subtype : t -> t -> bool val sup : t -> t -> t val pr : t -> Pp.t end (** {6 Universe instances} *) module Instance : sig type t (** A universe instance represents a vector of argument universes to a polymorphic definition (constant, inductive or constructor). *) val empty : t val is_empty : t -> bool val of_array : Level.t array -> t val to_array : t -> Level.t array val append : t -> t -> t (** To concatenate two instances, used for discharge *) val equal : t -> t -> bool (** Equality *) val length : t -> int (** Instance length *) val hcons : t -> t (** Hash-consing. *) val hash : t -> int (** Hash value *) val share : t -> t * int (** Simultaneous hash-consing and hash-value computation *) val subst_fn : universe_level_subst_fn -> t -> t (** Substitution by a level-to-level function. *) val pr : (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t (** Pretty-printing, no comments *) val levels : t -> LSet.t (** The set of levels in the instance *) end val enforce_eq_instances : Instance.t constraint_function val enforce_eq_variance_instances : Variance.t array -> Instance.t constraint_function val enforce_leq_variance_instances : Variance.t array -> Instance.t constraint_function type 'a puniverses = 'a * Instance.t val out_punivs : 'a puniverses -> 'a val in_punivs : 'a -> 'a puniverses val eq_puniverses : ('a -> 'a -> bool) -> 'a puniverses -> 'a puniverses -> bool (** A vector of universe levels with universe Constraint.t, representing local universe variables and associated Constraint.t *) module UContext : sig type t val make : Instance.t constrained -> t val empty : t val is_empty : t -> bool val instance : t -> Instance.t val constraints : t -> Constraint.t val dest : t -> Instance.t * Constraint.t (** Keeps the order of the instances *) val union : t -> t -> t (** the number of universes in the context *) val size : t -> int end module AUContext : sig type t val repr : t -> UContext.t (** [repr ctx] is [(Var(0), ... Var(n-1) |= cstr] where [n] is the length of the context and [cstr] the abstracted Constraint.t. *) val empty : t val is_empty : t -> bool val size : t -> int (** Keeps the order of the instances *) val union : t -> t -> t val instantiate : Instance.t -> t -> Constraint.t (** Generate the set of instantiated Constraint.t **) val names : t -> Names.Name.t array (** Return the names of the bound universe variables *) end type 'a univ_abstracted = { univ_abstracted_value : 'a; univ_abstracted_binder : AUContext.t; } (** A value with bound universe levels. *) val map_univ_abstracted : ('a -> 'b) -> 'a univ_abstracted -> 'b univ_abstracted (** Universe contexts (as sets) *) (** A set of universes with universe Constraint.t. We linearize the set to a list after typechecking. Beware, representation could change. *) module ContextSet : sig type t = LSet.t constrained val empty : t val is_empty : t -> bool val singleton : Level.t -> t val of_instance : Instance.t -> t val of_set : LSet.t -> t val equal : t -> t -> bool val union : t -> t -> t val append : t -> t -> t (** Variant of {!union} which is more efficient when the left argument is much smaller than the right one. *) val diff : t -> t -> t val add_universe : Level.t -> t -> t val add_constraints : Constraint.t -> t -> t val add_instance : Instance.t -> t -> t (** Arbitrary choice of linear order of the variables *) val sort_levels : Level.t array -> Level.t array val to_context : t -> UContext.t val of_context : UContext.t -> t val constraints : t -> Constraint.t val levels : t -> LSet.t (** the number of universes in the context *) val size : t -> int end (** A value in a universe context (resp. context set). *) type 'a in_universe_context = 'a * UContext.t type 'a in_universe_context_set = 'a * ContextSet.t val extend_in_context_set : 'a in_universe_context_set -> ContextSet.t -> 'a in_universe_context_set val empty_level_subst : universe_level_subst val is_empty_level_subst : universe_level_subst -> bool (** Substitution of universes. *) val subst_univs_level_level : universe_level_subst -> Level.t -> Level.t val subst_univs_level_universe : universe_level_subst -> Universe.t -> Universe.t val subst_univs_level_constraints : universe_level_subst -> Constraint.t -> Constraint.t val subst_univs_level_abstract_universe_context : universe_level_subst -> AUContext.t -> AUContext.t val subst_univs_level_instance : universe_level_subst -> Instance.t -> Instance.t (** Level to universe substitutions. *) val empty_subst : universe_subst val is_empty_subst : universe_subst -> bool val make_subst : universe_subst -> universe_subst_fn val subst_univs_universe : universe_subst_fn -> Universe.t -> Universe.t (** Only user in the kernel is template polymorphism. Ideally we get rid of this code if it goes away. *) (** Substitution of instances *) val subst_instance_instance : Instance.t -> Instance.t -> Instance.t val subst_instance_universe : Instance.t -> Universe.t -> Universe.t val make_instance_subst : Instance.t -> universe_level_subst (** Creates [u(0) ↦ 0; ...; u(n-1) ↦ n - 1] out of [u(0); ...; u(n - 1)] *) val make_inverse_instance_subst : Instance.t -> universe_level_subst val abstract_universes : Names.Name.t array -> UContext.t -> Instance.t * AUContext.t (** TODO: move universe abstraction out of the kernel *) val make_abstract_instance : AUContext.t -> Instance.t (** [compact_univ u] remaps local variables in [u] such that their indices become consecutive. It returns the new universe and the mapping. Example: compact_univ [(Var 0, i); (Prop, 0); (Var 2; j))] = [(Var 0,i); (Prop, 0); (Var 1; j)], [0; 2] *) val compact_univ : Universe.t -> Universe.t * int list (** {6 Pretty-printing of universes. } *) val pr_constraint_type : constraint_type -> Pp.t val pr_constraints : (Level.t -> Pp.t) -> Constraint.t -> Pp.t val pr_universe_context : (Level.t -> Pp.t) -> ?variance:Variance.t array -> UContext.t -> Pp.t val pr_abstract_universe_context : (Level.t -> Pp.t) -> ?variance:Variance.t array -> AUContext.t -> Pp.t val pr_universe_context_set : (Level.t -> Pp.t) -> ContextSet.t -> Pp.t val explain_universe_inconsistency : (Level.t -> Pp.t) -> univ_inconsistency -> Pp.t val pr_universe_level_subst : universe_level_subst -> Pp.t val pr_universe_subst : universe_subst -> Pp.t (** {6 Hash-consing } *) val hcons_univ : Universe.t -> Universe.t val hcons_constraints : Constraint.t -> Constraint.t val hcons_universe_set : LSet.t -> LSet.t val hcons_universe_context : UContext.t -> UContext.t val hcons_abstract_universe_context : AUContext.t -> AUContext.t val hcons_universe_context_set : ContextSet.t -> ContextSet.t coq-8.11.0/kernel/nativelib.ml0000644000175000017500000001563713612664533016071 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* anomaly (Pp.str "get_load_paths not initialized.") : unit -> string list) let open_header = ["Nativevalues"; "Nativecode"; "Nativelib"; "Nativeconv"] let open_header = List.map mk_open open_header (* Directory where compiled files are stored *) let output_dir = ".coq-native" (* Extension of generated ml files, stored for debugging purposes *) let source_ext = ".native" let ( / ) = Filename.concat (* Directory for temporary files for the conversion and normalisation (as opposed to compiling the library itself, which uses [output_dir]). *) let my_temp_dir = lazy (CUnix.mktemp_dir "Coq_native" "") let () = at_exit (fun () -> if Lazy.is_val my_temp_dir then try let d = Lazy.force my_temp_dir in Array.iter (fun f -> Sys.remove (Filename.concat d f)) (Sys.readdir d); Unix.rmdir d with e -> Feedback.msg_warning Pp.(str "Native compile: failed to cleanup: " ++ str(Printexc.to_string e) ++ fnl())) (* We have to delay evaluation of include_dirs because coqlib cannot be guessed until flags have been properly initialized. It also lets us avoid forcing [my_temp_dir] if we don't need it (eg stdlib file without native compute or native conv uses). *) let include_dirs () = let base = [Envars.coqlib () / "kernel"; Envars.coqlib () / "library"] in if Lazy.is_val my_temp_dir then (Lazy.force my_temp_dir) :: base else base (* Pointer to the function linking an ML object into coq's toplevel *) let load_obj = ref (fun _x -> () : string -> unit) let rt1 = ref (dummy_value ()) let rt2 = ref (dummy_value ()) let get_ml_filename () = let temp_dir = Lazy.force my_temp_dir in let filename = Filename.temp_file ~temp_dir "Coq_native" source_ext in let prefix = Filename.chop_extension (Filename.basename filename) ^ "." in filename, prefix let write_ml_code fn ?(header=[]) code = let header = open_header@header in let ch_out = open_out fn in let fmt = Format.formatter_of_out_channel ch_out in List.iter (pp_global fmt) (header@code); close_out ch_out let error_native_compiler_failed e = let msg = match e with | Inl (Unix.WEXITED 127) -> Pp.(strbrk "The OCaml compiler was not found. Make sure it is installed, together with findlib.") | Inl (Unix.WEXITED n) -> Pp.(strbrk "Native compiler exited with status" ++ str" " ++ int n) | Inl (Unix.WSIGNALED n) -> Pp.(strbrk "Native compiler killed by signal" ++ str" " ++ int n) | Inl (Unix.WSTOPPED n) -> Pp.(strbrk "Native compiler stopped by signal" ++ str" " ++ int n) | Inr e -> Pp.(strbrk "Native compiler failed with error: " ++ strbrk (Unix.error_message e)) in CErrors.user_err msg let call_compiler ?profile:(profile=false) ml_filename = let load_path = !get_load_paths () in let load_path = List.map (fun dn -> dn / output_dir) load_path in let include_dirs = List.flatten (List.map (fun x -> ["-I"; x]) (include_dirs () @ load_path)) in let f = Filename.chop_extension ml_filename in let link_filename = f ^ ".cmo" in let link_filename = Dynlink.adapt_filename link_filename in let remove f = if Sys.file_exists f then Sys.remove f in remove link_filename; remove (f ^ ".cmi"); let initial_args = if Dynlink.is_native then ["opt"; "-shared"] else ["ocamlc"; "-c"] in let profile_args = if profile then ["-g"] else [] in let flambda_args = if Sys.(backend_type = Native) then ["-Oclassic"] else [] in let args = initial_args @ profile_args @ flambda_args @ ("-o"::link_filename ::"-rectypes" ::"-w"::"a" ::include_dirs) @ ["-impl"; ml_filename] in if !Flags.debug then Feedback.msg_debug (Pp.str (Envars.ocamlfind () ^ " " ^ (String.concat " " args))); try let res = CUnix.sys_command (Envars.ocamlfind ()) args in match res with | Unix.WEXITED 0 -> link_filename | Unix.WEXITED _n | Unix.WSIGNALED _n | Unix.WSTOPPED _n -> error_native_compiler_failed (Inl res) with Unix.Unix_error (e,_,_) -> error_native_compiler_failed (Inr e) let compile fn code ~profile:profile = write_ml_code fn code; let r = call_compiler ~profile fn in if (not !Flags.debug) && Sys.file_exists fn then Sys.remove fn; r let compile_library dir code fn = let header = mk_library_header dir in let fn = fn ^ source_ext in let basename = Filename.basename fn in let dirname = Filename.dirname fn in let dirname = dirname / output_dir in let () = try Unix.mkdir dirname 0o755 with Unix.Unix_error (Unix.EEXIST, _, _) -> () in let fn = dirname / basename in write_ml_code fn ~header code; let _ = call_compiler fn in if (not !Flags.debug) && Sys.file_exists fn then Sys.remove fn let native_symbols = ref Names.DPmap.empty let get_library_native_symbols dir = try Names.DPmap.find dir !native_symbols with Not_found -> CErrors.user_err ~hdr:"get_library_native_symbols" Pp.((str "Linker error in the native compiler. Are you using Require inside a nested Module declaration?") ++ fnl () ++ (str "This use case is not supported, but disabling the native compiler may help.")) (* call_linker links dynamically the code for constants in environment or a *) (* conversion test. *) let call_linker ?(fatal=true) env ~prefix f upds = native_symbols := env.Environ.native_symbols; rt1 := dummy_value (); rt2 := dummy_value (); if not (Sys.file_exists f) then begin let msg = "Cannot find native compiler file " ^ f in if fatal then CErrors.user_err Pp.(str msg) else if !Flags.debug then Feedback.msg_debug (Pp.str msg) end else (try if Dynlink.is_native then Dynlink.loadfile f else !load_obj f; register_native_file prefix with Dynlink.Error _ as exn -> let exn = CErrors.push exn in if fatal then iraise exn else if !Flags.debug then Feedback.msg_debug CErrors.(iprint exn)); match upds with Some upds -> update_locations upds | _ -> () let link_library env ~prefix ~dirname ~basename = let f = dirname / output_dir / basename in call_linker env ~fatal:false ~prefix f None coq-8.11.0/kernel/opaqueproof.ml0000644000175000017500000001162513612664533016445 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* int -> opaque_proofterm; access_discharge : cooking_info list -> (Constr.t * unit delayed_universes) -> (Constr.t * unit delayed_universes); } let drop_mono = function | PrivateMonomorphic _ -> PrivateMonomorphic () | PrivatePolymorphic _ as ctx -> ctx type proofterm = (constr * Univ.ContextSet.t delayed_universes) Future.computation type opaque = | Indirect of substitution list * cooking_info list * DirPath.t * int (* subst, discharge, lib, index *) type opaquetab = { opaque_val : proofterm Int.Map.t; (** Actual proof terms *) opaque_len : int; (** Size of the above map *) opaque_dir : DirPath.t; } let empty_opaquetab = { opaque_val = Int.Map.empty; opaque_len = 0; opaque_dir = DirPath.initial; } let not_here () = CErrors.user_err Pp.(str "Cannot access opaque delayed proof") let create dp cu tab = let hcons (c, u) = let c = Constr.hcons c in let u = match u with | PrivateMonomorphic u -> PrivateMonomorphic (Univ.hcons_universe_context_set u) | PrivatePolymorphic (n, u) -> PrivatePolymorphic (n, Univ.hcons_universe_context_set u) in (c, u) in let cu = Future.chain cu hcons in let id = tab.opaque_len in let opaque_val = Int.Map.add id cu tab.opaque_val in let opaque_dir = if DirPath.equal dp tab.opaque_dir then tab.opaque_dir else if DirPath.equal tab.opaque_dir DirPath.initial then dp else CErrors.anomaly (Pp.str "Using the same opaque table for multiple dirpaths.") in let ntab = { opaque_val; opaque_dir; opaque_len = id + 1 } in Indirect ([], [], dp, id), ntab let subst_opaque sub = function | Indirect (s, ci, dp, i) -> Indirect (sub :: s, ci, dp, i) let discharge_opaque info = function | Indirect (s, ci, dp, i) -> assert (CList.is_empty s); Indirect ([], info :: ci, dp, i) let join except cu = match except with | None -> ignore (Future.join cu) | Some except -> if Future.UUIDSet.mem (Future.uuid cu) except then () else ignore (Future.join cu) let join_opaque ?except { opaque_val = prfs; opaque_dir = odp; _ } = function | Indirect (_,_,dp,i) -> if DirPath.equal dp odp then let fp = Int.Map.find i prfs in join except fp let force_proof access { opaque_val = prfs; opaque_dir = odp; _ } = function | Indirect (l,d,dp,i) -> let c, u = if DirPath.equal dp odp then let cu = Int.Map.find i prfs in let (c, u) = Future.force cu in access.access_discharge d (c, drop_mono u) else let cu = access.access_proof dp i in match cu with | None -> not_here () | Some (c, u) -> access.access_discharge d (c, u) in let c = force_constr (List.fold_right subst_substituted l (from_val c)) in (c, u) let get_mono (_, u) = match u with | PrivateMonomorphic ctx -> ctx | PrivatePolymorphic _ -> Univ.ContextSet.empty let force_constraints _access { opaque_val = prfs; opaque_dir = odp; _ } = function | Indirect (_,_,dp,i) -> if DirPath.equal dp odp then let cu = Int.Map.find i prfs in get_mono (Future.force cu) else Univ.ContextSet.empty module FMap = Future.UUIDMap let dump ?(except = Future.UUIDSet.empty) { opaque_val = otab; opaque_len = n; _ } = let opaque_table = Array.make n None in let f2t_map = ref FMap.empty in let iter n cu = let uid = Future.uuid cu in let () = f2t_map := FMap.add (Future.uuid cu) n !f2t_map in let c = if Future.is_val cu then let (c, priv) = Future.force cu in let priv = drop_mono priv in Some (c, priv) else if Future.UUIDSet.mem uid except then None else CErrors.anomaly Pp.(str"Proof object "++int n++str" is not checked nor to be checked") in opaque_table.(n) <- c in let () = Int.Map.iter iter otab in opaque_table, !f2t_map coq-8.11.0/kernel/kernel.mllib0000644000175000017500000000100213612664533016040 0ustar treinentreinenNames TransparentState Uint63 Float64 CPrimitives Univ UGraph Esubst Sorts Evar Context Constr Vars Term Mod_subst Vmvalues Cbytecodes Copcodes Cemitcodes Opaqueproof Declarations Entries Nativevalues Declareops Retroknowledge Conv_oracle Environ Primred CClosure Retypeops Reduction Clambda Nativelambda Cbytegen Nativecode Nativelib Csymtable Vm Vconv Nativeconv Type_errors Modops Inductive Typeops IndTyping Indtypes InferCumulativity Cooking Term_typing Subtyping Mod_typing Nativelibrary Section Safe_typing coq-8.11.0/kernel/cClosure.ml0000644000175000017500000014432113612664533015664 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* red_kind val fVAR : Id.t -> red_kind val no_red : reds val red_add : reds -> red_kind -> reds val red_sub : reds -> red_kind -> reds val red_add_transparent : reds -> TransparentState.t -> reds val red_transparent : reds -> TransparentState.t val mkflags : red_kind list -> reds val red_set : reds -> red_kind -> bool val red_projection : reds -> Projection.t -> bool end module RedFlags : RedFlagsSig = struct (* [r_const=(true,cl)] means all constants but those in [cl] *) (* [r_const=(false,cl)] means only those in [cl] *) (* [r_delta=true] just mean [r_const=(true,[])] *) open TransparentState type reds = { r_beta : bool; r_delta : bool; r_eta : bool; r_const : TransparentState.t; r_zeta : bool; r_match : bool; r_fix : bool; r_cofix : bool } type red_kind = BETA | DELTA | ETA | MATCH | FIX | COFIX | ZETA | CONST of Constant.t | VAR of Id.t let fBETA = BETA let fDELTA = DELTA let fETA = ETA let fMATCH = MATCH let fFIX = FIX let fCOFIX = COFIX let fZETA = ZETA let fCONST kn = CONST kn let fVAR id = VAR id let no_red = { r_beta = false; r_delta = false; r_eta = false; r_const = all_opaque; r_zeta = false; r_match = false; r_fix = false; r_cofix = false } let red_add red = function | BETA -> { red with r_beta = true } | ETA -> { red with r_eta = true } | DELTA -> { red with r_delta = true; r_const = all_transparent } | CONST kn -> let r = red.r_const in { red with r_const = { r with tr_cst = Cpred.add kn r.tr_cst } } | MATCH -> { red with r_match = true } | FIX -> { red with r_fix = true } | COFIX -> { red with r_cofix = true } | ZETA -> { red with r_zeta = true } | VAR id -> let r = red.r_const in { red with r_const = { r with tr_var = Id.Pred.add id r.tr_var } } let red_sub red = function | BETA -> { red with r_beta = false } | ETA -> { red with r_eta = false } | DELTA -> { red with r_delta = false } | CONST kn -> let r = red.r_const in { red with r_const = { r with tr_cst = Cpred.remove kn r.tr_cst } } | MATCH -> { red with r_match = false } | FIX -> { red with r_fix = false } | COFIX -> { red with r_cofix = false } | ZETA -> { red with r_zeta = false } | VAR id -> let r = red.r_const in { red with r_const = { r with tr_var = Id.Pred.remove id r.tr_var } } let red_transparent red = red.r_const let red_add_transparent red tr = { red with r_const = tr } let mkflags = List.fold_left red_add no_red let red_set red = function | BETA -> incr_cnt red.r_beta beta | ETA -> incr_cnt red.r_eta eta | CONST kn -> let c = is_transparent_constant red.r_const kn in incr_cnt c delta | VAR id -> (* En attendant d'avoir des kn pour les Var *) let c = is_transparent_variable red.r_const id in incr_cnt c delta | ZETA -> incr_cnt red.r_zeta zeta | MATCH -> incr_cnt red.r_match nb_match | FIX -> incr_cnt red.r_fix fix | COFIX -> incr_cnt red.r_cofix cofix | DELTA -> (* Used for Rel/Var defined in context *) incr_cnt red.r_delta delta let red_projection red p = if Projection.unfolded p then true else red_set red (fCONST (Projection.constant p)) end open RedFlags let all = mkflags [fBETA;fDELTA;fZETA;fMATCH;fFIX;fCOFIX] let allnolet = mkflags [fBETA;fDELTA;fMATCH;fFIX;fCOFIX] let beta = mkflags [fBETA] let betadeltazeta = mkflags [fBETA;fDELTA;fZETA] let betaiota = mkflags [fBETA;fMATCH;fFIX;fCOFIX] let betaiotazeta = mkflags [fBETA;fMATCH;fFIX;fCOFIX;fZETA] let betazeta = mkflags [fBETA;fZETA] let delta = mkflags [fDELTA] let zeta = mkflags [fZETA] let nored = no_red (* Flags of reduction and cache of constants: 'a is a type that may be * mapped to constr. 'a infos implements a cache for constants and * abstractions, storing a representation (of type 'a) of the body of * this constant or abstraction. * * i_tab is the cache table of the results * * ref_value_cache searches in the tab, otherwise uses i_repr to * compute the result and store it in the table. If the constant can't * be unfolded, returns None, but does not store this failure. * This * doesn't take the RESET into account. You mustn't keep such a table * after a Reset. * This type is not exported. Only its two * instantiations (cbv or lazy) are. *) type table_key = Constant.t Univ.puniverses tableKey let eq_pconstant_key (c,u) (c',u') = eq_constant_key c c' && Univ.Instance.equal u u' module IdKeyHash = struct open Hashset.Combine type t = table_key let equal = Names.eq_table_key eq_pconstant_key let hash = function | ConstKey (c, _) -> combinesmall 1 (Constant.UserOrd.hash c) | VarKey id -> combinesmall 2 (Id.hash id) | RelKey i -> combinesmall 3 (Int.hash i) end module KeyTable = Hashtbl.Make(IdKeyHash) open Context.Named.Declaration let assoc_defined id env = match Environ.lookup_named id env with | LocalDef (_, c, _) -> c | LocalAssum _ -> raise Not_found (**********************************************************************) (* Lazy reduction: the one used in kernel operations *) (* type of shared terms. fconstr and frterm are mutually recursive. * Clone of the constr structure, but completely mutable, and * annotated with reduction state (reducible or not). * - FLIFT is a delayed shift; allows sharing between 2 lifted copies * of a given term. * - FCLOS is a delayed substitution applied to a constr * - FLOCKED is used to erase the content of a reference that must * be updated. This is to allow the garbage collector to work * before the term is computed. *) (* Norm means the term is fully normalized and cannot create a redex when substituted Cstr means the term is in head normal form and that it can create a redex when substituted (i.e. constructor, fix, lambda) Whnf means we reached the head normal form and that it cannot create a redex when substituted Red is used for terms that might be reduced *) type red_state = Norm | Cstr | Whnf | Red let neutr = function | Whnf|Norm -> Whnf | Red|Cstr -> Red type optrel = Unknown | KnownR | KnownI let opt_of_rel = function | Sorts.Relevant -> KnownR | Sorts.Irrelevant -> KnownI module Mark : sig type t val mark : red_state -> optrel -> t val relevance : t -> optrel val red_state : t -> red_state val neutr : t -> t val set_norm : t -> t end = struct type t = int let[@inline] of_state = function | Norm -> 0b00 | Cstr -> 0b01 | Whnf -> 0b10 | Red -> 0b11 let[@inline] of_relevance = function | Unknown -> 0 | KnownR -> 0b01 | KnownI -> 0b10 let[@inline] mark state relevance = (of_state state) * 4 + (of_relevance relevance) let[@inline] relevance x = match x land 0b11 with | 0b00 -> Unknown | 0b01 -> KnownR | 0b10 -> KnownI | _ -> assert false let[@inline] red_state x = match x land 0b1100 with | 0b0000 -> Norm | 0b0100 -> Cstr | 0b1000 -> Whnf | 0b1100 -> Red | _ -> assert false let[@inline] neutr x = x lor 0b1000 (* Whnf|Norm -> Whnf | Red|Cstr -> Red *) let[@inline] set_norm x = x land 0b0011 end let mark = Mark.mark type fconstr = { mutable mark : Mark.t; mutable term: fterm; } and fterm = | FRel of int | FAtom of constr (* Metas and Sorts *) | FFlex of table_key | FInd of pinductive | FConstruct of pconstructor | FApp of fconstr * fconstr array | FProj of Projection.t * fconstr | FFix of fixpoint * fconstr subs | FCoFix of cofixpoint * fconstr subs | FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *) | FLambda of int * (Name.t Context.binder_annot * constr) list * constr * fconstr subs | FProd of Name.t Context.binder_annot * fconstr * constr * fconstr subs | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * fconstr subs | FEvar of existential * fconstr subs | FInt of Uint63.t | FFloat of Float64.t | FLIFT of int * fconstr | FCLOS of constr * fconstr subs | FLOCKED let fterm_of v = v.term let set_norm v = v.mark <- Mark.set_norm v.mark let is_val v = match Mark.red_state v.mark with Norm -> true | Cstr | Whnf | Red -> false let mk_atom c = {mark=mark Norm Unknown;term=FAtom c} let mk_red f = {mark=mark Red Unknown;term=f} (* Could issue a warning if no is still Red, pointing out that we loose sharing. *) let update ~share v1 mark t = if share then (v1.mark <- mark; v1.term <- t; v1) else {mark;term=t;} (** Reduction cache *) type infos_cache = { i_env : env; i_sigma : existential -> constr option; i_share : bool; } type clos_infos = { i_flags : reds; i_cache : infos_cache } type clos_tab = (fconstr, Empty.t) constant_def KeyTable.t let info_flags info = info.i_flags let info_env info = info.i_cache.i_env (**********************************************************************) (* The type of (machine) stacks (= lambda-bar-calculus' contexts) *) type 'a next_native_args = (CPrimitives.arg_kind * 'a) list type stack_member = | Zapp of fconstr array | ZcaseT of case_info * constr * constr array * fconstr subs | Zproj of Projection.Repr.t | Zfix of fconstr * stack | Zprimitive of CPrimitives.t * pconstant * fconstr list * fconstr next_native_args (* operator, constr def, arguments already seen (in rev order), next arguments *) | Zshift of int | Zupdate of fconstr and stack = stack_member list let empty_stack = [] let append_stack v s = if Int.equal (Array.length v) 0 then s else match s with | Zapp l :: s -> Zapp (Array.append v l) :: s | (ZcaseT _ | Zproj _ | Zfix _ | Zshift _ | Zupdate _ | Zprimitive _) :: _ | [] -> Zapp v :: s (* Collapse the shifts in the stack *) let zshift n s = match (n,s) with (0,_) -> s | (_,Zshift(k)::s) -> Zshift(n+k)::s | (_,(ZcaseT _ | Zproj _ | Zfix _ | Zapp _ | Zupdate _ | Zprimitive _) :: _) | _,[] -> Zshift(n)::s let rec stack_args_size = function | Zapp v :: s -> Array.length v + stack_args_size s | Zshift(_)::s -> stack_args_size s | Zupdate(_)::s -> stack_args_size s | (ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | [] -> 0 (* Lifting. Preserves sharing (useful only for cell with norm=Red). lft_fconstr always create a new cell, while lift_fconstr avoids it when the lift is 0. *) let rec lft_fconstr n ft = let r = Mark.relevance ft.mark in match ft.term with | (FInd _|FConstruct _|FFlex(ConstKey _|VarKey _)|FInt _|FFloat _) -> ft | FRel i -> {mark=mark Norm r;term=FRel(i+n)} | FLambda(k,tys,f,e) -> {mark=mark Cstr r; term=FLambda(k,tys,f,subs_shft(n,e))} | FFix(fx,e) -> {mark=mark Cstr r; term=FFix(fx,subs_shft(n,e))} | FCoFix(cfx,e) -> {mark=mark Cstr r; term=FCoFix(cfx,subs_shft(n,e))} | FLIFT(k,m) -> lft_fconstr (n+k) m | FLOCKED -> assert false | FFlex (RelKey _) | FAtom _ | FApp _ | FProj _ | FCaseT _ | FProd _ | FLetIn _ | FEvar _ | FCLOS _ -> {mark=ft.mark; term=FLIFT(n,ft)} let lift_fconstr k f = if Int.equal k 0 then f else lft_fconstr k f let lift_fconstr_vect k v = if Int.equal k 0 then v else Array.Fun1.map lft_fconstr k v let clos_rel e i = match expand_rel i e with | Inl(n,mt) -> lift_fconstr n mt | Inr(k,None) -> {mark=mark Norm Unknown; term= FRel k} | Inr(k,Some p) -> lift_fconstr (k-p) {mark=mark Red Unknown;term=FFlex(RelKey p)} (* since the head may be reducible, we might introduce lifts of 0 *) let compact_stack head stk = let rec strip_rec depth = function | Zshift(k)::s -> strip_rec (depth+k) s | Zupdate(m)::s -> (* Be sure to create a new cell otherwise sharing would be lost by the update operation *) let h' = lft_fconstr depth head in (** The stack contains [Zupdate] marks only if in sharing mode *) let _ = update ~share:true m h'.mark h'.term in strip_rec depth s | ((ZcaseT _ | Zproj _ | Zfix _ | Zapp _ | Zprimitive _) :: _ | []) as stk -> zshift depth stk in strip_rec 0 stk (* Put an update mark in the stack, only if needed *) let zupdate info m s = let share = info.i_cache.i_share in if share && begin match Mark.red_state m.mark with Red -> true | Norm | Whnf | Cstr -> false end then let s' = compact_stack m s in let _ = m.term <- FLOCKED in Zupdate(m)::s' else s let mk_lambda env t = let (rvars,t') = Term.decompose_lam t in FLambda(List.length rvars, List.rev rvars, t', env) let destFLambda clos_fun t = match [@ocaml.warning "-4"] t.term with FLambda(_,[(na,ty)],b,e) -> (na,clos_fun e ty,clos_fun (subs_lift e) b) | FLambda(n,(na,ty)::tys,b,e) -> (na,clos_fun e ty,{mark=t.mark;term=FLambda(n-1,tys,b,subs_lift e)}) | _ -> assert false (* t must be a FLambda and binding list cannot be empty *) (* Optimization: do not enclose variables in a closure. Makes variable access much faster *) let mk_clos e t = match kind t with | Rel i -> clos_rel e i | Var x -> {mark = mark Red Unknown; term = FFlex (VarKey x) } | Const c -> {mark = mark Red Unknown; term = FFlex (ConstKey c) } | Meta _ | Sort _ -> {mark = mark Norm KnownR; term = FAtom t } | Ind kn -> {mark = mark Norm KnownR; term = FInd kn } | Construct kn -> {mark = mark Cstr Unknown; term = FConstruct kn } | Int i -> {mark = mark Cstr Unknown; term = FInt i} | Float f -> {mark = mark Cstr Unknown; term = FFloat f} | (CoFix _|Lambda _|Fix _|Prod _|Evar _|App _|Case _|Cast _|LetIn _|Proj _) -> {mark = mark Red Unknown; term = FCLOS(t,e)} let inject c = mk_clos (subs_id 0) c (** Hand-unrolling of the map function to bypass the call to the generic array allocation *) let mk_clos_vect env v = match v with | [||] -> [||] | [|v0|] -> [|mk_clos env v0|] | [|v0; v1|] -> [|mk_clos env v0; mk_clos env v1|] | [|v0; v1; v2|] -> [|mk_clos env v0; mk_clos env v1; mk_clos env v2|] | [|v0; v1; v2; v3|] -> [|mk_clos env v0; mk_clos env v1; mk_clos env v2; mk_clos env v3|] | v -> Array.Fun1.map mk_clos env v let ref_value_cache ({ i_cache = cache; _ }) tab ref = try KeyTable.find tab ref with Not_found -> let v = try let body = match ref with | RelKey n -> let open! Context.Rel.Declaration in let i = n - 1 in let (d, _) = try Range.get cache.i_env.env_rel_context.env_rel_map i with Invalid_argument _ -> raise Not_found in begin match d with | LocalAssum _ -> raise Not_found | LocalDef (_, t, _) -> lift n t end | VarKey id -> assoc_defined id cache.i_env | ConstKey cst -> constant_value_in cache.i_env cst in Def (inject body) with | NotEvaluableConst (IsPrimitive op) (* Const *) -> Primitive op | Not_found (* List.assoc *) | NotEvaluableConst _ (* Const *) -> Undef None in KeyTable.add tab ref v; v (* The inverse of mk_clos: move back to constr *) let rec to_constr lfts v = match v.term with | FRel i -> mkRel (reloc_rel i lfts) | FFlex (RelKey p) -> mkRel (reloc_rel p lfts) | FFlex (VarKey x) -> mkVar x | FAtom c -> exliftn lfts c | FFlex (ConstKey op) -> mkConstU op | FInd op -> mkIndU op | FConstruct op -> mkConstructU op | FCaseT (ci,p,c,ve,env) -> if is_subs_id env && is_lift_id lfts then mkCase (ci, p, to_constr lfts c, ve) else let subs = comp_subs lfts env in mkCase (ci, subst_constr subs p, to_constr lfts c, Array.map (fun b -> subst_constr subs b) ve) | FFix ((op,(lna,tys,bds)) as fx, e) -> if is_subs_id e && is_lift_id lfts then mkFix fx else let n = Array.length bds in let subs_ty = comp_subs lfts e in let subs_bd = comp_subs (el_liftn n lfts) (subs_liftn n e) in let tys = Array.Fun1.map subst_constr subs_ty tys in let bds = Array.Fun1.map subst_constr subs_bd bds in mkFix (op, (lna, tys, bds)) | FCoFix ((op,(lna,tys,bds)) as cfx, e) -> if is_subs_id e && is_lift_id lfts then mkCoFix cfx else let n = Array.length bds in let subs_ty = comp_subs lfts e in let subs_bd = comp_subs (el_liftn n lfts) (subs_liftn n e) in let tys = Array.Fun1.map subst_constr subs_ty tys in let bds = Array.Fun1.map subst_constr subs_bd bds in mkCoFix (op, (lna, tys, bds)) | FApp (f,ve) -> mkApp (to_constr lfts f, Array.Fun1.map to_constr lfts ve) | FProj (p,c) -> mkProj (p,to_constr lfts c) | FLambda (len, tys, f, e) -> if is_subs_id e && is_lift_id lfts then Term.compose_lam (List.rev tys) f else let subs = comp_subs lfts e in let tys = List.mapi (fun i (na, c) -> na, subst_constr (subs_liftn i subs) c) tys in let f = subst_constr (subs_liftn len subs) f in Term.compose_lam (List.rev tys) f | FProd (n, t, c, e) -> if is_subs_id e && is_lift_id lfts then mkProd (n, to_constr lfts t, c) else let subs' = comp_subs lfts e in mkProd (n, to_constr lfts t, subst_constr (subs_lift subs') c) | FLetIn (n,b,t,f,e) -> let subs = comp_subs (el_lift lfts) (subs_lift e) in mkLetIn (n, to_constr lfts b, to_constr lfts t, subst_constr subs f) | FEvar ((ev,args),env) -> let subs = comp_subs lfts env in mkEvar(ev,Array.map (fun a -> subst_constr subs a) args) | FLIFT (k,a) -> to_constr (el_shft k lfts) a | FInt i -> Constr.mkInt i | FFloat f -> Constr.mkFloat f | FCLOS (t,env) -> if is_subs_id env && is_lift_id lfts then t else let subs = comp_subs lfts env in subst_constr subs t | FLOCKED -> assert false (*mkVar(Id.of_string"_LOCK_")*) and subst_constr subst c = match [@ocaml.warning "-4"] Constr.kind c with | Rel i -> begin match expand_rel i subst with | Inl (k, lazy v) -> Vars.lift k v | Inr (m, _) -> mkRel m end | _ -> Constr.map_with_binders Esubst.subs_lift subst_constr subst c and comp_subs el s = Esubst.lift_subst (fun el c -> lazy (to_constr el c)) el s (* This function defines the correspondence between constr and fconstr. When we find a closure whose substitution is the identity, then we directly return the constr to avoid possibly huge reallocation. *) let term_of_fconstr c = to_constr el_id c (* fstrong applies unfreeze_fun recursively on the (freeze) term and * yields a term. Assumes that the unfreeze_fun never returns a * FCLOS term. let rec fstrong unfreeze_fun lfts v = to_constr (fstrong unfreeze_fun) lfts (unfreeze_fun v) *) let rec zip m stk = match stk with | [] -> m | Zapp args :: s -> zip {mark=Mark.neutr m.mark; term=FApp(m, args)} s | ZcaseT(ci,p,br,e)::s -> let t = FCaseT(ci, p, m, br, e) in let mark = mark (neutr (Mark.red_state m.mark)) Unknown in zip {mark; term=t} s | Zproj p :: s -> let mark = mark (neutr (Mark.red_state m.mark)) Unknown in zip {mark; term=FProj(Projection.make p true,m)} s | Zfix(fx,par)::s -> zip fx (par @ append_stack [|m|] s) | Zshift(n)::s -> zip (lift_fconstr n m) s | Zupdate(rf)::s -> (** The stack contains [Zupdate] marks only if in sharing mode *) zip (update ~share:true rf m.mark m.term) s | Zprimitive(_op,c,rargs,kargs)::s -> let args = List.rev_append rargs (m::List.map snd kargs) in let f = {mark = mark Red Unknown;term = FFlex (ConstKey c)} in zip {mark=mark (neutr (Mark.red_state m.mark)) KnownR; term = FApp (f, Array.of_list args)} s let fapp_stack (m,stk) = zip m stk (*********************************************************************) (* The assertions in the functions below are granted because they are called only when m is a constructor, a cofix (strip_update_shift_app), a fix (get_nth_arg) or an abstraction (strip_update_shift, through get_arg). *) (* optimised for the case where there are no shifts... *) let strip_update_shift_app_red head stk = let rec strip_rec rstk h depth = function | Zshift(k) as e :: s -> strip_rec (e::rstk) (lift_fconstr k h) (depth+k) s | (Zapp args :: s) -> strip_rec (Zapp args :: rstk) {mark=h.mark;term=FApp(h,args)} depth s | Zupdate(m)::s -> (** The stack contains [Zupdate] marks only if in sharing mode *) strip_rec rstk (update ~share:true m h.mark h.term) depth s | ((ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | []) as stk -> (depth,List.rev rstk, stk) in strip_rec [] head 0 stk let strip_update_shift_app head stack = assert (match Mark.red_state head.mark with Red -> false | Norm | Cstr | Whnf -> true); strip_update_shift_app_red head stack let get_nth_arg head n stk = assert (match Mark.red_state head.mark with Red -> false | Norm | Cstr | Whnf -> true); let rec strip_rec rstk h n = function | Zshift(k) as e :: s -> strip_rec (e::rstk) (lift_fconstr k h) n s | Zapp args::s' -> let q = Array.length args in if n >= q then strip_rec (Zapp args::rstk) {mark=h.mark;term=FApp(h,args)} (n-q) s' else let bef = Array.sub args 0 n in let aft = Array.sub args (n+1) (q-n-1) in let stk' = List.rev (if Int.equal n 0 then rstk else (Zapp bef :: rstk)) in (Some (stk', args.(n)), append_stack aft s') | Zupdate(m)::s -> (** The stack contains [Zupdate] mark only if in sharing mode *) strip_rec rstk (update ~share:true m h.mark h.term) n s | ((ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | []) as s -> (None, List.rev rstk @ s) in strip_rec [] head n stk (* Beta reduction: look for an applied argument in the stack. Since the encountered update marks are removed, h must be a whnf *) let rec get_args n tys f e = function | Zupdate r :: s -> (** The stack contains [Zupdate] mark only if in sharing mode *) let _hd = update ~share:true r (mark Cstr (Mark.relevance r.mark)) (FLambda(n,tys,f,e)) in get_args n tys f e s | Zshift k :: s -> get_args n tys f (subs_shft (k,e)) s | Zapp l :: s -> let na = Array.length l in if n == na then (Inl (subs_cons(l,e)),s) else if n < na then (* more arguments *) let args = Array.sub l 0 n in let eargs = Array.sub l n (na-n) in (Inl (subs_cons(args,e)), Zapp eargs :: s) else (* more lambdas *) let etys = List.skipn na tys in get_args (n-na) etys f (subs_cons(l,e)) s | ((ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | []) as stk -> (Inr {mark=mark Cstr Unknown;term=FLambda(n,tys,f,e)}, stk) (* Eta expansion: add a reference to implicit surrounding lambda at end of stack *) let rec eta_expand_stack = function | (Zapp _ | Zfix _ | ZcaseT _ | Zproj _ | Zshift _ | Zupdate _ | Zprimitive _ as e) :: s -> e :: eta_expand_stack s | [] -> [Zshift 1; Zapp [|{mark=mark Norm Unknown; term= FRel 1}|]] (* Get the arguments of a native operator *) let rec skip_native_args rargs nargs = match nargs with | (kd, a) :: nargs' -> if kd = CPrimitives.Kwhnf then rargs, nargs else skip_native_args (a::rargs) nargs' | [] -> rargs, [] let get_native_args op c stk = let kargs = CPrimitives.kind op in let rec get_args rnargs kargs args = match kargs, args with | kd::kargs, a::args -> get_args ((kd,a)::rnargs) kargs args | _, _ -> rnargs, kargs, args in let rec strip_rec rnargs h depth kargs = function | Zshift k :: s -> strip_rec (List.map (fun (kd,f) -> kd,lift_fconstr k f) rnargs) (lift_fconstr k h) (depth+k) kargs s | Zapp args :: s' -> begin match get_args rnargs kargs (Array.to_list args) with | rnargs, [], [] -> (skip_native_args [] (List.rev rnargs), s') | rnargs, [], eargs -> (skip_native_args [] (List.rev rnargs), Zapp (Array.of_list eargs) :: s') | rnargs, kargs, _ -> strip_rec rnargs {mark = h.mark;term=FApp(h, args)} depth kargs s' end | Zupdate(m) :: s -> strip_rec rnargs (update ~share:true m h.mark h.term) depth kargs s | (Zprimitive _ | ZcaseT _ | Zproj _ | Zfix _) :: _ | [] -> assert false in strip_rec [] {mark = mark Red Unknown;term = FFlex(ConstKey c)} 0 kargs stk let get_native_args1 op c stk = match get_native_args op c stk with | ((rargs, (kd,a):: nargs), stk) -> assert (kd = CPrimitives.Kwhnf); (rargs, a, nargs, stk) | _ -> assert false let check_native_args op stk = let nargs = CPrimitives.arity op in let rargs = stack_args_size stk in nargs <= rargs (* Iota reduction: extract the arguments to be passed to the Case branches *) let rec reloc_rargs_rec depth = function | Zapp args :: s -> Zapp (lift_fconstr_vect depth args) :: reloc_rargs_rec depth s | Zshift(k)::s -> if Int.equal k depth then s else reloc_rargs_rec (depth-k) s | ((ZcaseT _ | Zproj _ | Zfix _ | Zupdate _ | Zprimitive _) :: _ | []) as stk -> stk let reloc_rargs depth stk = if Int.equal depth 0 then stk else reloc_rargs_rec depth stk let rec try_drop_parameters depth n = function | Zapp args::s -> let q = Array.length args in if n > q then try_drop_parameters depth (n-q) s else if Int.equal n q then reloc_rargs depth s else let aft = Array.sub args n (q-n) in reloc_rargs depth (append_stack aft s) | Zshift(k)::s -> try_drop_parameters (depth-k) n s | [] -> if Int.equal n 0 then [] else raise Not_found | (ZcaseT _ | Zproj _ | Zfix _ | Zupdate _ | Zprimitive _) :: _ -> assert false (* strip_update_shift_app only produces Zapp and Zshift items *) let drop_parameters depth n argstk = try try_drop_parameters depth n argstk with Not_found -> (* we know that n < stack_args_size(argstk) (if well-typed term) *) anomaly (Pp.str "ill-typed term: found a match on a partially applied constructor.") (** [eta_expand_ind_stack env ind c s t] computes stacks corresponding to the conversion of the eta expansion of t, considered as an inhabitant of ind, and the Constructor c of this inductive type applied to arguments s. @assumes [t] is an irreducible term, and not a constructor. [ind] is the inductive of the constructor term [c] @raise Not_found if the inductive is not a primitive record, or if the constructor is partially applied. *) let eta_expand_ind_stack env ind m s (f, s') = let open Declarations in let mib = lookup_mind (fst ind) env in (* disallow eta-exp for non-primitive records *) if not (mib.mind_finite == BiFinite) then raise Not_found; match Declareops.inductive_make_projections ind mib with | Some projs -> (* (Construct, pars1 .. parsm :: arg1...argn :: []) ~= (f, s') -> arg1..argn ~= (proj1 t...projn t) where t = zip (f,s') *) let pars = mib.Declarations.mind_nparams in let right = fapp_stack (f, s') in let (depth, args, _s) = strip_update_shift_app m s in (** Try to drop the params, might fail on partially applied constructors. *) let argss = try_drop_parameters depth pars args in let hstack = Array.map (fun p -> { mark = mark Red Unknown; (* right can't be a constructor though *) term = FProj (Projection.make p true, right) }) projs in argss, [Zapp hstack] | None -> raise Not_found (* disallow eta-exp for non-primitive records *) let rec project_nth_arg n = function | Zapp args :: s -> let q = Array.length args in if n >= q then project_nth_arg (n - q) s else (* n < q *) args.(n) | (ZcaseT _ | Zproj _ | Zfix _ | Zupdate _ | Zshift _ | Zprimitive _) :: _ | [] -> assert false (* After drop_parameters we have a purely applicative stack *) (* Iota reduction: expansion of a fixpoint. * Given a fixpoint and a substitution, returns the corresponding * fixpoint body, and the substitution in which it should be * evaluated: its first variables are the fixpoint bodies * * FCLOS(fix Fi {F0 := T0 .. Fn-1 := Tn-1}, S) * -> (S. FCLOS(F0,S) . ... . FCLOS(Fn-1,S), Ti) *) (* does not deal with FLIFT *) let contract_fix_vect fix = let (thisbody, make_body, env, nfix) = match [@ocaml.warning "-4"] fix with | FFix (((reci,i),(nas,_,bds as rdcl)),env) -> (bds.(i), (fun j -> { mark = mark Cstr (opt_of_rel nas.(j).binder_relevance); term = FFix (((reci,j),rdcl),env) }), env, Array.length bds) | FCoFix ((i,(nas,_,bds as rdcl)),env) -> (bds.(i), (fun j -> { mark = mark Cstr (opt_of_rel nas.(j).binder_relevance); term = FCoFix ((j,rdcl),env) }), env, Array.length bds) | _ -> assert false in (subs_cons(Array.init nfix make_body, env), thisbody) let unfold_projection info p = if red_projection info.i_flags p then Some (Zproj (Projection.repr p)) else None (*********************************************************************) (* A machine that inspects the head of a term until it finds an atom or a subterm that may produce a redex (abstraction, constructor, cofix, letin, constant), or a neutral term (product, inductive) *) let rec knh info m stk = match m.term with | FLIFT(k,a) -> knh info a (zshift k stk) | FCLOS(t,e) -> knht info e t (zupdate info m stk) | FLOCKED -> assert false | FApp(a,b) -> knh info a (append_stack b (zupdate info m stk)) | FCaseT(ci,p,t,br,e) -> knh info t (ZcaseT(ci,p,br,e)::zupdate info m stk) | FFix(((ri,n),_),_) -> (match get_nth_arg m ri.(n) stk with (Some(pars,arg),stk') -> knh info arg (Zfix(m,pars)::stk') | (None, stk') -> (m,stk')) | FProj (p,c) -> (match unfold_projection info p with | None -> (m, stk) | Some s -> knh info c (s :: zupdate info m stk)) (* cases where knh stops *) | (FFlex _|FLetIn _|FConstruct _|FEvar _| FCoFix _|FLambda _|FRel _|FAtom _|FInd _|FProd _|FInt _|FFloat _) -> (m, stk) (* The same for pure terms *) and knht info e t stk = match kind t with | App(a,b) -> knht info e a (append_stack (mk_clos_vect e b) stk) | Case(ci,p,t,br) -> knht info e t (ZcaseT(ci, p, br, e)::stk) | Fix fx -> knh info { mark = mark Cstr Unknown; term = FFix (fx, e) } stk | Cast(a,_,_) -> knht info e a stk | Rel n -> knh info (clos_rel e n) stk | Proj (p, c) -> knh info { mark = mark Red Unknown; term = FProj (p, mk_clos e c) } stk | (Ind _|Const _|Construct _|Var _|Meta _ | Sort _ | Int _|Float _) -> (mk_clos e t, stk) | CoFix cfx -> { mark = mark Cstr Unknown; term = FCoFix (cfx,e) }, stk | Lambda _ -> { mark = mark Cstr Unknown; term = mk_lambda e t }, stk | Prod (n, t, c) -> { mark = mark Whnf KnownR; term = FProd (n, mk_clos e t, c, e) }, stk | LetIn (n,b,t,c) -> { mark = mark Red Unknown; term = FLetIn (n, mk_clos e b, mk_clos e t, c, e) }, stk | Evar ev -> { mark = mark Red Unknown; term = FEvar (ev, e) }, stk let inject c = mk_clos (subs_id 0) c (************************************************************************) (* Reduction of Native operators *) open Primred module FNativeEntries = struct type elem = fconstr type args = fconstr array type evd = unit let get = Array.get let get_int () e = match [@ocaml.warning "-4"] e.term with | FInt i -> i | _ -> raise Primred.NativeDestKO let get_float () e = match [@ocaml.warning "-4"] e.term with | FFloat f -> f | _ -> raise Primred.NativeDestKO let dummy = {mark = mark Norm KnownR; term = FRel 0} let current_retro = ref Retroknowledge.empty let defined_int = ref false let fint = ref dummy let init_int retro = match retro.Retroknowledge.retro_int63 with | Some c -> defined_int := true; fint := { mark = mark Norm KnownR; term = FFlex (ConstKey (Univ.in_punivs c)) } | None -> defined_int := false let defined_float = ref false let ffloat = ref dummy let init_float retro = match retro.Retroknowledge.retro_float64 with | Some c -> defined_float := true; ffloat := { mark = mark Norm KnownR; term = FFlex (ConstKey (Univ.in_punivs c)) } | None -> defined_float := false let defined_bool = ref false let ftrue = ref dummy let ffalse = ref dummy let init_bool retro = match retro.Retroknowledge.retro_bool with | Some (ct,cf) -> defined_bool := true; ftrue := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs ct) }; ffalse := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cf) } | None -> defined_bool :=false let defined_carry = ref false let fC0 = ref dummy let fC1 = ref dummy let init_carry retro = match retro.Retroknowledge.retro_carry with | Some(c0,c1) -> defined_carry := true; fC0 := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs c0) }; fC1 := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs c1) } | None -> defined_carry := false let defined_pair = ref false let fPair = ref dummy let init_pair retro = match retro.Retroknowledge.retro_pair with | Some c -> defined_pair := true; fPair := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs c) } | None -> defined_pair := false let defined_cmp = ref false let fEq = ref dummy let fLt = ref dummy let fGt = ref dummy let fcmp = ref dummy let init_cmp retro = match retro.Retroknowledge.retro_cmp with | Some (cEq, cLt, cGt) -> defined_cmp := true; fEq := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cEq) }; fLt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cLt) }; fGt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cGt) }; let (icmp, _) = cEq in fcmp := { mark = mark Norm KnownR; term = FInd (Univ.in_punivs icmp) } | None -> defined_cmp := false let defined_f_cmp = ref false let fFEq = ref dummy let fFLt = ref dummy let fFGt = ref dummy let fFNotComparable = ref dummy let init_f_cmp retro = match retro.Retroknowledge.retro_f_cmp with | Some (cFEq, cFLt, cFGt, cFNotComparable) -> defined_f_cmp := true; fFEq := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFEq) }; fFLt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFLt) }; fFGt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFGt) }; fFNotComparable := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFNotComparable) }; | None -> defined_f_cmp := false let defined_f_class = ref false let fPNormal = ref dummy let fNNormal = ref dummy let fPSubn = ref dummy let fNSubn = ref dummy let fPZero = ref dummy let fNZero = ref dummy let fPInf = ref dummy let fNInf = ref dummy let fNaN = ref dummy let init_f_class retro = match retro.Retroknowledge.retro_f_class with | Some (cPNormal, cNNormal, cPSubn, cNSubn, cPZero, cNZero, cPInf, cNInf, cNaN) -> defined_f_class := true; fPNormal := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPNormal) }; fNNormal := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNNormal) }; fPSubn := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPSubn) }; fNSubn := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNSubn) }; fPZero := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPZero) }; fNZero := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNZero) }; fPInf := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPInf) }; fNInf := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNInf) }; fNaN := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNaN) }; | None -> defined_f_class := false let defined_refl = ref false let frefl = ref dummy let init_refl retro = match retro.Retroknowledge.retro_refl with | Some crefl -> defined_refl := true; frefl := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs crefl) } | None -> defined_refl := false let init env = current_retro := env.retroknowledge; init_int !current_retro; init_float !current_retro; init_bool !current_retro; init_carry !current_retro; init_pair !current_retro; init_cmp !current_retro; init_f_cmp !current_retro; init_f_class !current_retro; init_refl !current_retro let check_env env = if not (!current_retro == env.retroknowledge) then init env let check_int env = check_env env; assert (!defined_int) let check_float env = check_env env; assert (!defined_float) let check_bool env = check_env env; assert (!defined_bool) let check_carry env = check_env env; assert (!defined_carry && !defined_int) let check_pair env = check_env env; assert (!defined_pair && !defined_int) let check_cmp env = check_env env; assert (!defined_cmp) let check_f_cmp env = check_env env; assert (!defined_f_cmp) let check_f_class env = check_env env; assert (!defined_f_class) let mkInt env i = check_int env; { mark = mark Cstr KnownR; term = FInt i } let mkFloat env f = check_float env; { mark = mark Norm KnownR; term = FFloat f } let mkBool env b = check_bool env; if b then !ftrue else !ffalse let mkCarry env b e = check_carry env; {mark = mark Cstr KnownR; term = FApp ((if b then !fC1 else !fC0),[|!fint;e|])} let mkIntPair env e1 e2 = check_pair env; { mark = mark Cstr KnownR; term = FApp(!fPair, [|!fint;!fint;e1;e2|]) } let mkFloatIntPair env f i = check_pair env; check_float env; { mark = mark Cstr KnownR; term = FApp(!fPair, [|!ffloat;!fint;f;i|]) } let mkLt env = check_cmp env; !fLt let mkEq env = check_cmp env; !fEq let mkGt env = check_cmp env; !fGt let mkFLt env = check_f_cmp env; !fFLt let mkFEq env = check_f_cmp env; !fFEq let mkFGt env = check_f_cmp env; !fFGt let mkFNotComparable env = check_f_cmp env; !fFNotComparable let mkPNormal env = check_f_class env; !fPNormal let mkNNormal env = check_f_class env; !fNNormal let mkPSubn env = check_f_class env; !fPSubn let mkNSubn env = check_f_class env; !fNSubn let mkPZero env = check_f_class env; !fPZero let mkNZero env = check_f_class env; !fNZero let mkPInf env = check_f_class env; !fPInf let mkNInf env = check_f_class env; !fNInf let mkNaN env = check_f_class env; !fNaN end module FredNative = RedNative(FNativeEntries) (************************************************************************) (* Computes a weak head normal form from the result of knh. *) let rec knr info tab m stk = match m.term with | FLambda(n,tys,f,e) when red_set info.i_flags fBETA -> (match get_args n tys f e stk with Inl e', s -> knit info tab e' f s | Inr lam, s -> (lam,s)) | FFlex(ConstKey (kn,_ as c)) when red_set info.i_flags (fCONST kn) -> (match ref_value_cache info tab (ConstKey c) with | Def v -> kni info tab v stk | Primitive op when check_native_args op stk -> let rargs, a, nargs, stk = get_native_args1 op c stk in kni info tab a (Zprimitive(op,c,rargs,nargs)::stk) | Undef _ | OpaqueDef _ | Primitive _ -> (set_norm m; (m,stk))) | FFlex(VarKey id) when red_set info.i_flags (fVAR id) -> (match ref_value_cache info tab (VarKey id) with | Def v -> kni info tab v stk | Primitive _ -> assert false | OpaqueDef _ | Undef _ -> (set_norm m; (m,stk))) | FFlex(RelKey k) when red_set info.i_flags fDELTA -> (match ref_value_cache info tab (RelKey k) with | Def v -> kni info tab v stk | Primitive _ -> assert false | OpaqueDef _ | Undef _ -> (set_norm m; (m,stk))) | FConstruct((_ind,c),_u) -> let use_match = red_set info.i_flags fMATCH in let use_fix = red_set info.i_flags fFIX in if use_match || use_fix then (match [@ocaml.warning "-4"] strip_update_shift_app m stk with | (depth, args, ZcaseT(ci,_,br,e)::s) when use_match -> assert (ci.ci_npar>=0); let rargs = drop_parameters depth ci.ci_npar args in knit info tab e br.(c-1) (rargs@s) | (_, cargs, Zfix(fx,par)::s) when use_fix -> let rarg = fapp_stack(m,cargs) in let stk' = par @ append_stack [|rarg|] s in let (fxe,fxbd) = contract_fix_vect fx.term in knit info tab fxe fxbd stk' | (depth, args, Zproj p::s) when use_match -> let rargs = drop_parameters depth (Projection.Repr.npars p) args in let rarg = project_nth_arg (Projection.Repr.arg p) rargs in kni info tab rarg s | (_,args,s) -> (m,args@s)) else (m,stk) | FCoFix _ when red_set info.i_flags fCOFIX -> (match strip_update_shift_app m stk with | (_, args, (((ZcaseT _|Zproj _)::_) as stk')) -> let (fxe,fxbd) = contract_fix_vect m.term in knit info tab fxe fxbd (args@stk') | (_,args, ((Zapp _ | Zfix _ | Zshift _ | Zupdate _ | Zprimitive _) :: _ | [] as s)) -> (m,args@s)) | FLetIn (_,v,_,bd,e) when red_set info.i_flags fZETA -> knit info tab (subs_cons([|v|],e)) bd stk | FEvar(ev,env) -> (match info.i_cache.i_sigma ev with Some c -> knit info tab env c stk | None -> (m,stk)) | FInt _ | FFloat _ -> (match [@ocaml.warning "-4"] strip_update_shift_app m stk with | (_, _, Zprimitive(op,c,rargs,nargs)::s) -> let (rargs, nargs) = skip_native_args (m::rargs) nargs in begin match nargs with | [] -> let args = Array.of_list (List.rev rargs) in begin match FredNative.red_prim (info_env info) () op args with | Some m -> kni info tab m s | None -> let f = {mark = mark Whnf KnownR; term = FFlex (ConstKey c)} in let m = {mark = mark Whnf KnownR; term = FApp(f,args)} in (m,s) end | (kd,a)::nargs -> assert (kd = CPrimitives.Kwhnf); kni info tab a (Zprimitive(op,c,rargs,nargs)::s) end | (_, _, s) -> (m, s)) | FLOCKED | FRel _ | FAtom _ | FFlex (RelKey _ | ConstKey _ | VarKey _) | FInd _ | FApp _ | FProj _ | FFix _ | FCoFix _ | FCaseT _ | FLambda _ | FProd _ | FLetIn _ | FLIFT _ | FCLOS _ -> (m, stk) (* Computes the weak head normal form of a term *) and kni info tab m stk = let (hm,s) = knh info m stk in knr info tab hm s and knit info tab e t stk = let (ht,s) = knht info e t stk in knr info tab ht s let kh info tab v stk = fapp_stack(kni info tab v stk) (************************************************************************) let rec zip_term zfun m stk = match stk with | [] -> m | Zapp args :: s -> zip_term zfun (mkApp(m, Array.map zfun args)) s | ZcaseT(ci,p,br,e)::s -> let t = mkCase(ci, zfun (mk_clos e p), m, Array.map (fun b -> zfun (mk_clos e b)) br) in zip_term zfun t s | Zproj p::s -> let t = mkProj (Projection.make p true, m) in zip_term zfun t s | Zfix(fx,par)::s -> let h = mkApp(zip_term zfun (zfun fx) par,[|m|]) in zip_term zfun h s | Zshift(n)::s -> zip_term zfun (lift n m) s | Zupdate(_rf)::s -> zip_term zfun m s | Zprimitive(_,c,rargs, kargs)::s -> let kargs = List.map (fun (_,a) -> zfun a) kargs in let args = List.fold_left (fun args a -> zfun a ::args) (m::kargs) rargs in let h = mkApp (mkConstU c, Array.of_list args) in zip_term zfun h s (* Computes the strong normal form of a term. 1- Calls kni 2- tries to rebuild the term. If a closure still has to be computed, calls itself recursively. *) let rec kl info tab m = let share = info.i_cache.i_share in if is_val m then (incr prune; term_of_fconstr m) else let (nm,s) = kni info tab m [] in let () = if share then ignore (fapp_stack (nm, s)) in (* to unlock Zupdates! *) zip_term (kl info tab) (norm_head info tab nm) s (* no redex: go up for atoms and already normalized terms, go down otherwise. *) and norm_head info tab m = if is_val m then (incr prune; term_of_fconstr m) else match m.term with | FLambda(_n,tys,f,e) -> let (e',info,rvtys) = List.fold_left (fun (e,info,ctxt) (na,ty) -> (subs_lift e, info, (na,kl info tab (mk_clos e ty))::ctxt)) (e,info,[]) tys in let bd = kl info tab (mk_clos e' f) in List.fold_left (fun b (na,ty) -> mkLambda(na,ty,b)) bd rvtys | FLetIn(na,a,b,f,e) -> let c = mk_clos (subs_lift e) f in mkLetIn(na, kl info tab a, kl info tab b, kl info tab c) | FProd(na,dom,rng,e) -> mkProd(na, kl info tab dom, kl info tab (mk_clos (subs_lift e) rng)) | FCoFix((n,(na,tys,bds)),e) -> let ftys = Array.Fun1.map mk_clos e tys in let fbds = Array.Fun1.map mk_clos (subs_liftn (Array.length na) e) bds in mkCoFix(n,(na, CArray.map (kl info tab) ftys, CArray.map (kl info tab) fbds)) | FFix((n,(na,tys,bds)),e) -> let ftys = Array.Fun1.map mk_clos e tys in let fbds = Array.Fun1.map mk_clos (subs_liftn (Array.length na) e) bds in mkFix(n,(na, CArray.map (kl info tab) ftys, CArray.map (kl info tab) fbds)) | FEvar((i,args),env) -> mkEvar(i, Array.map (fun a -> kl info tab (mk_clos env a)) args) | FProj (p,c) -> mkProj (p, kl info tab c) | FLOCKED | FRel _ | FAtom _ | FFlex _ | FInd _ | FConstruct _ | FApp _ | FCaseT _ | FLIFT _ | FCLOS _ | FInt _ | FFloat _ -> term_of_fconstr m (* Initialization and then normalization *) (* weak reduction *) let whd_val info tab v = with_stats (lazy (term_of_fconstr (kh info tab v []))) (* strong reduction *) let norm_val info tab v = with_stats (lazy (kl info tab v)) let whd_stack infos tab m stk = match Mark.red_state m.mark with | Whnf | Norm -> (** No need to perform [kni] nor to unlock updates because every head subterm of [m] is [Whnf] or [Norm] *) knh infos m stk | Red | Cstr -> let k = kni infos tab m stk in let () = if infos.i_cache.i_share then ignore (fapp_stack k) in (* to unlock Zupdates! *) k let create_clos_infos ?(evars=fun _ -> None) flgs env = let share = (Environ.typing_flags env).Declarations.share_reduction in let cache = { i_env = env; i_sigma = evars; i_share = share; } in { i_flags = flgs; i_cache = cache } let create_tab () = KeyTable.create 17 let oracle_of_infos infos = Environ.oracle infos.i_cache.i_env let infos_with_reds infos reds = { infos with i_flags = reds } let unfold_reference info tab key = match key with | ConstKey (kn,_) -> if red_set info.i_flags (fCONST kn) then ref_value_cache info tab key else Undef None | VarKey i -> if red_set info.i_flags (fVAR i) then ref_value_cache info tab key else Undef None | RelKey _ -> ref_value_cache info tab key let relevance_of f = Mark.relevance f.mark let set_relevance r f = f.mark <- Mark.mark (Mark.red_state f.mark) (opt_of_rel r) coq-8.11.0/kernel/nativelambda.ml0000644000175000017500000005400413612664533016532 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* constr option; evars_metas : metavariable -> types } (*s Constructors *) let mkLapp f args = if Array.is_empty args then f else match f with | Lapp(f', args') -> Lapp (f', Array.append args' args) | _ -> Lapp(f, args) let mkLlam ids body = if Array.is_empty ids then body else match body with | Llam(ids', body) -> Llam(Array.append ids ids', body) | _ -> Llam(ids, body) let decompose_Llam lam = match lam with | Llam(ids,body) -> ids, body | _ -> [||], lam let rec decompose_Llam_Llet lam = match lam with | Llam(ids,body) -> let vars, body = decompose_Llam_Llet body in Array.fold_right (fun x l -> (x, None) :: l) ids vars, body | Llet(id,def,body) -> let vars, body = decompose_Llam_Llet body in (id,Some def) :: vars, body | _ -> [], lam let decompose_Llam_Llet lam = let vars, body = decompose_Llam_Llet lam in Array.of_list vars, body (*s Operators on substitution *) let subst_id = subs_id 0 let lift = subs_lift let liftn = subs_liftn let cons v subst = subs_cons([|v|], subst) let shift subst = subs_shft (1, subst) (* Linked code location utilities *) let get_mind_prefix env mind = let _,name = lookup_mind_key mind env in match !name with | NotLinked -> "" | Linked s -> s | LinkedInteractive s -> s let get_const_prefix env c = let _,(nameref,_) = lookup_constant_key c env in match !nameref with | NotLinked -> "" | Linked s -> s | LinkedInteractive s -> s (* A generic map function *) let map_lam_with_binders g f n lam = match lam with | Lrel _ | Lvar _ | Lconst _ | Lproj _ | Lval _ | Lsort _ | Lind _ | Luint _ | Llazy | Lforce | Lmeta _ | Lint _ | Lfloat _ -> lam | Lprod(dom,codom) -> let dom' = f n dom in let codom' = f n codom in if dom == dom' && codom == codom' then lam else Lprod(dom',codom') | Llam(ids,body) -> let body' = f (g (Array.length ids) n) body in if body == body' then lam else mkLlam ids body' | Lrec(id,body) -> let body' = f (g 1 n) body in if body == body' then lam else Lrec(id,body') | Llet(id,def,body) -> let def' = f n def in let body' = f (g 1 n) body in if body == body' && def == def' then lam else Llet(id,def',body') | Lapp(fct,args) -> let fct' = f n fct in let args' = Array.Smart.map (f n) args in if fct == fct' && args == args' then lam else mkLapp fct' args' | Lprim(prefix,kn,op,args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Lprim(prefix,kn,op,args') | Lcase(annot,t,a,branches) -> let const = branches.constant_branches in let nonconst = branches.nonconstant_branches in let t' = f n t in let a' = f n a in let const' = Array.Smart.map (f n) const in let on_b b = let (ids,body) = b in let body' = f (g (Array.length ids) n) body in if body == body' then b else (ids,body') in let nonconst' = Array.Smart.map on_b nonconst in let branches' = if const == const' && nonconst == nonconst' then branches else { constant_branches = const'; nonconstant_branches = nonconst' } in if t == t' && a == a' && branches == branches' then lam else Lcase(annot,t',a',branches') | Lif(t,bt,bf) -> let t' = f n t in let bt' = f n bt in let bf' = f n bf in if t == t' && bt == bt' && bf == bf' then lam else Lif(t',bt',bf') | Lfix(init,(ids,ltypes,lbodies)) -> let ltypes' = Array.Smart.map (f n) ltypes in let lbodies' = Array.Smart.map (f (g (Array.length ids) n)) lbodies in if ltypes == ltypes' && lbodies == lbodies' then lam else Lfix(init,(ids,ltypes',lbodies')) | Lcofix(init,(ids,ltypes,lbodies)) -> let ltypes' = Array.Smart.map (f n) ltypes in let lbodies' = Array.Smart.map (f (g (Array.length ids) n)) lbodies in if ltypes == ltypes' && lbodies == lbodies' then lam else Lcofix(init,(ids,ltypes',lbodies')) | Lmakeblock(prefix,cn,tag,args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Lmakeblock(prefix,cn,tag,args') | Levar (evk, args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Levar (evk, args') (*s Lift and substitution *) let rec lam_exlift el lam = match lam with | Lrel(id,i) -> let i' = reloc_rel i el in if i == i' then lam else Lrel(id,i') | _ -> map_lam_with_binders el_liftn lam_exlift el lam let lam_lift k lam = if Int.equal k 0 then lam else lam_exlift (el_shft k el_id) lam let lam_subst_rel lam id n subst = match expand_rel n subst with | Inl(k,v) -> lam_lift k v | Inr(n',_) -> if n == n' then lam else Lrel(id, n') let rec lam_exsubst subst lam = match lam with | Lrel(id,i) -> lam_subst_rel lam id i subst | _ -> map_lam_with_binders liftn lam_exsubst subst lam let lam_subst_args subst args = if is_subs_id subst then args else Array.Smart.map (lam_exsubst subst) args (** Simplification of lambda expression *) (* [simplify subst lam] simplify the expression [lam_subst subst lam] *) (* that is : *) (* - Reduce [let] is the definition can be substituted i.e: *) (* - a variable (rel or Id.t) *) (* - a constant *) (* - a structured constant *) (* - a function *) (* - Transform beta redex into [let] expression *) (* - Move arguments under [let] *) (* Invariant : Terms in [subst] are already simplified and can be *) (* substituted *) let can_subst lam = match lam with | Lrel _ | Lvar _ | Lconst _ | Lproj _ | Lval _ | Lsort _ | Lind _ | Llam _ | Lmeta _ | Levar _ -> true | _ -> false let can_merge_if bt bf = match bt, bf with | Llam(_idst,_), Llam(_idsf,_) -> true | _ -> false let merge_if t bt bf = let (idst,bodyt) = decompose_Llam bt in let (idsf,bodyf) = decompose_Llam bf in let nt = Array.length idst in let nf = Array.length idsf in let common,idst,idsf = if Int.equal nt nf then idst, [||], [||] else if nt < nf then idst,[||], Array.sub idsf nt (nf - nt) else idsf, Array.sub idst nf (nt - nf), [||] in Llam(common, Lif(lam_lift (Array.length common) t, mkLlam idst bodyt, mkLlam idsf bodyf)) let rec simplify subst lam = match lam with | Lrel(id,i) -> lam_subst_rel lam id i subst | Llet(id,def,body) -> let def' = simplify subst def in if can_subst def' then simplify (cons def' subst) body else let body' = simplify (lift subst) body in if def == def' && body == body' then lam else Llet(id,def',body') | Lapp(f,args) -> begin match simplify_app subst f subst args with | Lapp(f',args') when f == f' && args == args' -> lam | lam' -> lam' end | Lif(t,bt,bf) -> let t' = simplify subst t in let bt' = simplify subst bt in let bf' = simplify subst bf in if can_merge_if bt' bf' then merge_if t' bt' bf' else if t == t' && bt == bt' && bf == bf' then lam else Lif(t',bt',bf') | _ -> map_lam_with_binders liftn simplify subst lam and simplify_app substf f substa args = match f with | Lrel(id, i) -> begin match lam_subst_rel f id i substf with | Llam(ids, body) -> reduce_lapp subst_id (Array.to_list ids) body substa (Array.to_list args) | f' -> mkLapp f' (simplify_args substa args) end | Llam(ids, body) -> reduce_lapp substf (Array.to_list ids) body substa (Array.to_list args) | Llet(id, def, body) -> let def' = simplify substf def in if can_subst def' then simplify_app (cons def' substf) body substa args else Llet(id, def', simplify_app (lift substf) body (shift substa) args) | Lapp(f, args') -> let args = Array.append (lam_subst_args substf args') (lam_subst_args substa args) in simplify_app substf f subst_id args (* TODO | Lproj -> simplify if the argument is known or a known global *) | _ -> mkLapp (simplify substf f) (simplify_args substa args) and simplify_args subst args = Array.Smart.map (simplify subst) args and reduce_lapp substf lids body substa largs = match lids, largs with | id::lids, a::largs -> let a = simplify substa a in if can_subst a then reduce_lapp (cons a substf) lids body substa largs else let body = reduce_lapp (lift substf) lids body (shift substa) largs in Llet(id, a, body) | [], [] -> simplify substf body | _::_, _ -> Llam(Array.of_list lids, simplify (liftn (List.length lids) substf) body) | [], _::_ -> simplify_app substf body substa (Array.of_list largs) (*s Translation from [constr] to [lambda] *) (* Translation of constructor *) let is_value lc = match lc with | Lval _ | Lint _ | Luint _ | Lfloat _ -> true | _ -> false let get_value lc = match lc with | Lval v -> v | Lint tag -> Nativevalues.mk_int tag | Luint i -> Nativevalues.mk_uint i | Lfloat f -> Nativevalues.mk_float f | _ -> raise Not_found let make_args start _end = Array.init (start - _end + 1) (fun i -> Lrel (Anonymous, start - i)) (* Translation of constructors *) let expand_constructor prefix ind tag nparams arity = let anon = Context.make_annot Anonymous Sorts.Relevant in (* TODO relevance *) let ids = Array.make (nparams + arity) anon in if Int.equal arity 0 then mkLlam ids (Lint tag) else let args = make_args arity 1 in Llam(ids, Lmakeblock (prefix, ind, tag, args)) (* [nparams] is the number of parameters still expected *) let makeblock env ind tag nparams arity args = let nargs = Array.length args in if nparams > 0 || nargs < arity then let prefix = get_mind_prefix env (fst ind) in mkLapp (expand_constructor prefix ind tag nparams arity) args else (* The constructor is fully applied *) if Int.equal arity 0 then Lint tag else if Array.for_all is_value args then let dummy_val = Obj.magic 0 in let args = (* Don't simplify this to Array.map, cf. the related comment in function eval_to_patch, file kernel/csymtable.ml *) let a = Array.make (Array.length args) dummy_val in Array.iteri (fun i v -> a.(i) <- get_value v) args; a in Lval (Nativevalues.mk_block tag args) else let prefix = get_mind_prefix env (fst ind) in Lmakeblock(prefix, ind, tag, args) (* Translation of constants *) let rec get_alias env (kn, u as p) = let tps = (lookup_constant kn env).const_body_code in match tps with | None -> p | Some tps -> match Cemitcodes.force tps with | Cemitcodes.BCalias kn' -> get_alias env (kn', u) | _ -> p let prim env kn p args = let prefix = get_const_prefix env (fst kn) in Lprim(prefix, kn, p, args) let expand_prim env kn op arity = (* primitives are always Relevant *) let ids = Array.make arity Context.anonR in let args = make_args arity 1 in Llam(ids, prim env kn op args) let lambda_of_prim env kn op args = let arity = CPrimitives.arity op in if Array.length args >= arity then prim env kn op args else mkLapp (expand_prim env kn op arity) args (*i Global environment *) let get_names decl = let decl = Array.of_list decl in Array.map fst decl let empty_args = [||] module Cache = struct module ConstrHash = struct type t = constructor let equal = eq_constructor let hash = constructor_hash end module ConstrTable = Hashtbl.Make(ConstrHash) type constructor_info = tag * int * int (* nparam nrealargs *) let get_construct_info cache env c : constructor_info = try ConstrTable.find cache c with Not_found -> let ((mind,j), i) = c in let oib = lookup_mind mind env in let oip = oib.mind_packets.(j) in let tag,arity = oip.mind_reloc_tbl.(i-1) in let nparams = oib.mind_nparams in let r = (tag, nparams, arity) in ConstrTable.add cache c r; r end let is_lazy t = match Constr.kind t with | App _ | LetIn _ | Case _ | Proj _ -> true | _ -> false let evar_value sigma ev = sigma.evars_val ev let meta_type sigma mv = sigma.evars_metas mv let empty_evars = { evars_val = (fun _ -> None); evars_metas = (fun _ -> assert false) } (** Extract the inductive type over which a fixpoint is decreasing *) let rec get_fix_struct env i t = match kind (Reduction.whd_all env t) with | Prod (na, dom, t) -> if Int.equal i 0 then let dom = Reduction.whd_all env dom in let (dom, _) = decompose_appvect dom in match kind dom with | Ind (ind, _) -> ind | _ -> assert false else let env = Environ.push_rel (RelDecl.LocalAssum (na, dom)) env in get_fix_struct env (i - 1) t | _ -> assert false let rec lambda_of_constr cache env sigma c = match kind c with | Meta mv -> let ty = meta_type sigma mv in Lmeta (mv, lambda_of_constr cache env sigma ty) | Evar (evk,args as ev) -> (match evar_value sigma ev with | None -> let args = Array.map (lambda_of_constr cache env sigma) args in Levar(evk, args) | Some t -> lambda_of_constr cache env sigma t) | Cast (c, _, _) -> lambda_of_constr cache env sigma c | Rel i -> Lrel (RelDecl.get_name (Environ.lookup_rel i env), i) | Var id -> Lvar id | Sort s -> Lsort s | Ind (ind,_u as pind) -> let prefix = get_mind_prefix env (fst ind) in Lind (prefix, pind) | Prod(id, dom, codom) -> let ld = lambda_of_constr cache env sigma dom in let env = Environ.push_rel (RelDecl.LocalAssum (id, dom)) env in let lc = lambda_of_constr cache env sigma codom in Lprod(ld, Llam([|id|], lc)) | Lambda _ -> let params, body = Term.decompose_lam c in let fold (na, t) env = Environ.push_rel (RelDecl.LocalAssum (na, t)) env in let env = List.fold_right fold params env in let lb = lambda_of_constr cache env sigma body in let ids = get_names (List.rev params) in mkLlam ids lb | LetIn(id, def, t, body) -> let ld = lambda_of_constr cache env sigma def in let env = Environ.push_rel (RelDecl.LocalDef (id, def, t)) env in let lb = lambda_of_constr cache env sigma body in Llet(id, ld, lb) | App(f, args) -> lambda_of_app cache env sigma f args | Const _ -> lambda_of_app cache env sigma c empty_args | Construct _ -> lambda_of_app cache env sigma c empty_args | Proj (p, c) -> let ind = Projection.inductive p in let prefix = get_mind_prefix env (fst ind) in mkLapp (Lproj (prefix, ind, Projection.arg p)) [|lambda_of_constr cache env sigma c|] | Case(ci,t,a,branches) -> let (mind,i as ind) = ci.ci_ind in let mib = lookup_mind mind env in let oib = mib.mind_packets.(i) in let tbl = oib.mind_reloc_tbl in (* Building info *) let prefix = get_mind_prefix env mind in let annot_sw = { asw_ind = ind; asw_ci = ci; asw_reloc = tbl; asw_finite = mib.mind_finite <> CoFinite; asw_prefix = prefix} in (* translation of the argument *) let la = lambda_of_constr cache env sigma a in (* translation of the type *) let lt = lambda_of_constr cache env sigma t in (* translation of branches *) let dummy = Lrel(Anonymous,0) in let consts = Array.make oib.mind_nb_constant dummy in let blocks = Array.make oib.mind_nb_args ([||],dummy) in let rtbl = oib.mind_reloc_tbl in for i = 0 to Array.length rtbl - 1 do let tag, arity = rtbl.(i) in let b = lambda_of_constr cache env sigma branches.(i) in if arity = 0 then consts.(tag) <- b else let b = match b with | Llam(ids, body) when Array.length ids = arity -> (ids, body) | _ -> let anon = Context.make_annot Anonymous Sorts.Relevant in (* TODO relevance *) let ids = Array.make arity anon in let args = make_args arity 1 in let ll = lam_lift arity b in (ids, mkLapp ll args) in blocks.(tag-1) <- b done; let branches = { constant_branches = consts; nonconstant_branches = blocks } in Lcase(annot_sw, lt, la, branches) | Fix((pos, i), (names,type_bodies,rec_bodies)) -> let ltypes = lambda_of_args cache env sigma 0 type_bodies in let map i t = let ind = get_fix_struct env i t in let prefix = get_mind_prefix env (fst ind) in (prefix, ind) in let inds = Array.map2 map pos type_bodies in let env = Environ.push_rec_types (names, type_bodies, rec_bodies) env in let lbodies = lambda_of_args cache env sigma 0 rec_bodies in Lfix((pos, inds, i), (names, ltypes, lbodies)) | CoFix(init,(names,type_bodies,rec_bodies)) -> let rec_bodies = Array.map2 (Reduction.eta_expand env) rec_bodies type_bodies in let ltypes = lambda_of_args cache env sigma 0 type_bodies in let env = Environ.push_rec_types (names, type_bodies, rec_bodies) env in let lbodies = lambda_of_args cache env sigma 0 rec_bodies in Lcofix(init, (names, ltypes, lbodies)) | Int i -> Luint i | Float f -> Lfloat f and lambda_of_app cache env sigma f args = match kind f with | Const (_kn,_u as c) -> let kn,u = get_alias env c in let cb = lookup_constant kn env in begin match cb.const_body with | Primitive op -> lambda_of_prim env c op (lambda_of_args cache env sigma 0 args) | Def csubst -> (* TODO optimize if f is a proj and argument is known *) if cb.const_inline_code then lambda_of_app cache env sigma (Mod_subst.force_constr csubst) args else let prefix = get_const_prefix env kn in let t = if is_lazy (Mod_subst.force_constr csubst) then mkLapp Lforce [|Lconst (prefix, (kn,u))|] else Lconst (prefix, (kn,u)) in mkLapp t (lambda_of_args cache env sigma 0 args) | OpaqueDef _ | Undef _ -> let prefix = get_const_prefix env kn in mkLapp (Lconst (prefix, (kn,u))) (lambda_of_args cache env sigma 0 args) end | Construct ((ind,_ as c),_) -> let tag, nparams, arity = Cache.get_construct_info cache env c in let nargs = Array.length args in if nparams < nargs then (* got all parameters *) let args = lambda_of_args cache env sigma nparams args in makeblock env ind tag 0 arity args else makeblock env ind tag (nparams - nargs) arity empty_args | _ -> let f = lambda_of_constr cache env sigma f in let args = lambda_of_args cache env sigma 0 args in mkLapp f args and lambda_of_args cache env sigma start args = let nargs = Array.length args in if start < nargs then Array.init (nargs - start) (fun i -> lambda_of_constr cache env sigma args.(start + i)) else empty_args let optimize lam = let lam = simplify subst_id lam in (* if Flags.vm_draw_opt () then (msgerrnl (str "Simplify = \n" ++ pp_lam lam);flush_all()); let lam = remove_let subst_id lam in if Flags.vm_draw_opt () then (msgerrnl (str "Remove let = \n" ++ pp_lam lam);flush_all()); *) lam let lambda_of_constr env sigma c = let cache = Cache.ConstrTable.create 91 in let lam = lambda_of_constr cache env sigma c in (* if Flags.vm_draw_opt () then begin (msgerrnl (str "Constr = \n" ++ pr_constr c);flush_all()); (msgerrnl (str "Lambda = \n" ++ pp_lam lam);flush_all()); end; *) optimize lam let mk_lazy c = mkLapp Llazy [|c|] coq-8.11.0/kernel/names.ml0000644000175000017500000007426213612664533015216 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* true | Some _ -> false let of_bytes s = let s = Bytes.to_string s in check_valid s; String.hcons s let of_string s = let () = check_valid s in String.hcons s let of_string_soft s = let () = check_valid ~strict:false s in String.hcons s let to_string id = id let print id = str id module Self = struct type t = string let compare = compare end module Set = Set.Make(Self) module Map = CMap.Make(Self) module Pred = Predicate.Make(Self) module List = String.List let hcons = String.hcons end (** Representation and operations on identifiers that are allowed to be anonymous (i.e. "_" in concrete syntax). *) module Name = struct type t = Anonymous (** anonymous identifier *) | Name of Id.t (** non-anonymous identifier *) let mk_name id = Name id let is_anonymous = function | Anonymous -> true | Name _ -> false let is_name = is_anonymous %> not let compare n1 n2 = match n1, n2 with | Anonymous, Anonymous -> 0 | Name id1, Name id2 -> Id.compare id1 id2 | Anonymous, Name _ -> -1 | Name _, Anonymous -> 1 let equal n1 n2 = match n1, n2 with | Anonymous, Anonymous -> true | Name id1, Name id2 -> String.equal id1 id2 | _ -> false let hash = function | Anonymous -> 0 | Name id -> Id.hash id let print = function | Anonymous -> str "_" | Name id -> Id.print id module Self_Hashcons = struct type nonrec t = t type u = Id.t -> Id.t let hashcons hident = function | Name id -> Name (hident id) | n -> n let eq n1 n2 = n1 == n2 || match (n1,n2) with | (Name id1, Name id2) -> id1 == id2 | (Anonymous,Anonymous) -> true | _ -> false let hash = hash end module Hname = Hashcons.Make(Self_Hashcons) let hcons = Hashcons.simple_hcons Hname.generate Hname.hcons Id.hcons end (** Alias, to import constructors. *) type name = Name.t = Anonymous | Name of Id.t (** {6 Various types based on identifiers } *) type variable = Id.t type module_ident = Id.t module ModIdset = Id.Set module ModIdmap = Id.Map (** {6 Directory paths = section names paths } *) (** Dirpaths are lists of module identifiers. The actual representation is reversed to optimise sharing: Coq.A.B is ["B";"A";"Coq"] *) let default_module_name = "If you see this, it's a bug" module DirPath = struct type t = module_ident list let compare = List.compare Id.compare let equal = List.equal Id.equal let rec hash accu = function | [] -> accu | id :: dp -> let accu = Hashset.Combine.combine (Id.hash id) accu in hash accu dp let hash dp = hash 0 dp let make x = x let repr x = x let empty = [] let is_empty = List.is_empty let to_string = function | [] -> "<>" | sl -> String.concat "." (List.rev_map Id.to_string sl) let print dp = str (to_string dp) let initial = [default_module_name] module Hdir = Hashcons.Hlist(Id) let hcons = Hashcons.recursive_hcons Hdir.generate Hdir.hcons Id.hcons end (** {6 Unique names for bound modules } *) module MBId = struct type t = int * Id.t * DirPath.t let gen = let seed = ref 0 in fun () -> let ans = !seed in let () = incr seed in ans let make dir s = (gen(), s, dir) let repr mbid = mbid let to_string (_i, s, p) = DirPath.to_string p ^ "." ^ s let debug_to_string (i, s, p) = "<"^DirPath.to_string p ^"#" ^ s ^"#"^ string_of_int i^">" let compare (x : t) (y : t) = if x == y then 0 else match (x, y) with | (nl, idl, dpl), (nr, idr, dpr) -> let ans = Int.compare nl nr in if not (Int.equal ans 0) then ans else let ans = Id.compare idl idr in if not (Int.equal ans 0) then ans else DirPath.compare dpl dpr let equal x y = x == y || let (i1, id1, p1) = x in let (i2, id2, p2) = y in Int.equal i1 i2 && Id.equal id1 id2 && DirPath.equal p1 p2 let to_id (_, s, _) = s open Hashset.Combine let hash (i, id, dp) = combine3 (Int.hash i) (Id.hash id) (DirPath.hash dp) module Self_Hashcons = struct type nonrec t = t type u = (Id.t -> Id.t) * (DirPath.t -> DirPath.t) let hashcons (hid,hdir) (n,s,dir) = (n,hid s,hdir dir) let eq ((n1,s1,dir1) as x) ((n2,s2,dir2) as y) = (x == y) || (Int.equal n1 n2 && s1 == s2 && dir1 == dir2) let hash = hash end module HashMBId = Hashcons.Make(Self_Hashcons) let hcons = Hashcons.simple_hcons HashMBId.generate HashMBId.hcons (Id.hcons, DirPath.hcons) end module MBImap = CMap.Make(MBId) module MBIset = Set.Make(MBId) (** {6 Names of structure elements } *) module Label = struct include Id let make = Id.of_string let of_id id = id let to_id id = id end (** {6 The module part of the kernel name } *) module ModPath = struct type t = | MPfile of DirPath.t | MPbound of MBId.t | MPdot of t * Label.t type module_path = t let rec is_bound = function | MPbound _ -> true | MPdot(mp,_) -> is_bound mp | _ -> false let rec to_string = function | MPfile sl -> DirPath.to_string sl | MPbound uid -> MBId.to_string uid | MPdot (mp,l) -> to_string mp ^ "." ^ Label.to_string l let rec debug_to_string = function | MPfile sl -> DirPath.to_string sl | MPbound uid -> MBId.debug_to_string uid | MPdot (mp,l) -> debug_to_string mp ^ "." ^ Label.to_string l (** we compare labels first if both are MPdots *) let rec compare mp1 mp2 = if mp1 == mp2 then 0 else match mp1, mp2 with | MPfile p1, MPfile p2 -> DirPath.compare p1 p2 | MPbound id1, MPbound id2 -> MBId.compare id1 id2 | MPdot (mp1, l1), MPdot (mp2, l2) -> let c = String.compare l1 l2 in if not (Int.equal c 0) then c else compare mp1 mp2 | MPfile _, _ -> -1 | MPbound _, MPfile _ -> 1 | MPbound _, MPdot _ -> -1 | MPdot _, _ -> 1 let rec equal mp1 mp2 = mp1 == mp2 || match mp1, mp2 with | MPfile p1, MPfile p2 -> DirPath.equal p1 p2 | MPbound id1, MPbound id2 -> MBId.equal id1 id2 | MPdot (mp1, l1), MPdot (mp2, l2) -> String.equal l1 l2 && equal mp1 mp2 | (MPfile _ | MPbound _ | MPdot _), _ -> false open Hashset.Combine let rec hash = function | MPfile dp -> combinesmall 1 (DirPath.hash dp) | MPbound id -> combinesmall 2 (MBId.hash id) | MPdot (mp, lbl) -> combinesmall 3 (combine (hash mp) (Label.hash lbl)) let initial = MPfile DirPath.initial let rec dp = function | MPfile sl -> sl | MPbound (_,_,dp) -> dp | MPdot (mp,_l) -> dp mp module Self_Hashcons = struct type t = module_path type u = (DirPath.t -> DirPath.t) * (MBId.t -> MBId.t) * (string -> string) let rec hashcons (hdir,huniqid,hstr as hfuns) = function | MPfile dir -> MPfile (hdir dir) | MPbound m -> MPbound (huniqid m) | MPdot (md,l) -> MPdot (hashcons hfuns md, hstr l) let eq d1 d2 = d1 == d2 || match d1,d2 with | MPfile dir1, MPfile dir2 -> dir1 == dir2 | MPbound m1, MPbound m2 -> m1 == m2 | MPdot (mod1,l1), MPdot (mod2,l2) -> l1 == l2 && equal mod1 mod2 | _ -> false let hash = hash end module HashMP = Hashcons.Make(Self_Hashcons) let hcons = Hashcons.simple_hcons HashMP.generate HashMP.hcons (DirPath.hcons,MBId.hcons,String.hcons) end module DPset = Set.Make(DirPath) module DPmap = Map.Make(DirPath) module MPset = Set.Make(ModPath) module MPmap = CMap.Make(ModPath) (** {6 Kernel names } *) module KerName = struct type t = { modpath : ModPath.t; knlabel : Label.t; mutable refhash : int; (** Lazily computed hash. If unset, it is set to negative values. *) } type kernel_name = t let make modpath knlabel = { modpath; knlabel; refhash = -1; } let repr kn = (kn.modpath, kn.knlabel) let modpath kn = kn.modpath let label kn = kn.knlabel let to_string_gen mp_to_string kn = mp_to_string kn.modpath ^ "." ^ Label.to_string kn.knlabel let to_string kn = to_string_gen ModPath.to_string kn let debug_to_string kn = to_string_gen ModPath.debug_to_string kn let print kn = str (to_string kn) let debug_print kn = str (debug_to_string kn) let compare (kn1 : kernel_name) (kn2 : kernel_name) = if kn1 == kn2 then 0 else let c = String.compare kn1.knlabel kn2.knlabel in if not (Int.equal c 0) then c else ModPath.compare kn1.modpath kn2.modpath let equal kn1 kn2 = let h1 = kn1.refhash in let h2 = kn2.refhash in if 0 <= h1 && 0 <= h2 && not (Int.equal h1 h2) then false else Label.equal kn1.knlabel kn2.knlabel && ModPath.equal kn1.modpath kn2.modpath open Hashset.Combine let hash kn = let h = kn.refhash in if h < 0 then let { modpath = mp; knlabel = lbl; _ } = kn in let h = combine (ModPath.hash mp) (Label.hash lbl) in (* Ensure positivity on all platforms. *) let h = h land 0x3FFFFFFF in let () = kn.refhash <- h in h else h module Self_Hashcons = struct type t = kernel_name type u = (ModPath.t -> ModPath.t) * (DirPath.t -> DirPath.t) * (string -> string) let hashcons (hmod,_hdir,hstr) kn = let { modpath = mp; knlabel = l; refhash; } = kn in { modpath = hmod mp; knlabel = hstr l; refhash; } let eq kn1 kn2 = kn1.modpath == kn2.modpath && kn1.knlabel == kn2.knlabel let hash = hash end module HashKN = Hashcons.Make(Self_Hashcons) let hcons = Hashcons.simple_hcons HashKN.generate HashKN.hcons (ModPath.hcons,DirPath.hcons,String.hcons) end module KNmap = HMap.Make(KerName) module KNpred = Predicate.Make(KerName) module KNset = KNmap.Set (** {6 Kernel pairs } *) (** For constant and inductive names, we use a kernel name couple (kn1,kn2) where kn1 corresponds to the name used at toplevel (i.e. what the user see) and kn2 corresponds to the canonical kernel name i.e. in the environment we have {% kn1 \rhd_{\delta}^* kn2 \rhd_{\delta} t %} Invariants : - the user and canonical kn may differ only on their [module_path], the dirpaths and labels should be the same - when user and canonical parts differ, we cannot be in a section anymore, hence the dirpath must be empty - two pairs with the same user part should have the same canonical part in a given environment (though with backtracking, the hash-table can contains pairs with same user part but different canonical part from a previous state of the session) Note: since most of the time the canonical and user parts are equal, we handle this case with a particular constructor to spare some memory *) module KerPair = struct type t = | Same of KerName.t (** user = canonical *) | Dual of KerName.t * KerName.t (** user then canonical *) type kernel_pair = t let canonical = function | Same kn -> kn | Dual (_,kn) -> kn let user = function | Same kn -> kn | Dual (kn,_) -> kn let same kn = Same kn let make knu knc = if KerName.equal knu knc then Same knc else Dual (knu,knc) let make1 = same let make2 mp l = same (KerName.make mp l) let repr2 kp = KerName.repr (user kp) let label kp = KerName.label (user kp) let modpath kp = KerName.modpath (user kp) let change_label kp lbl = let (mp1,l1) = KerName.repr (user kp) and (mp2,l2) = KerName.repr (canonical kp) in assert (String.equal l1 l2); if String.equal lbl l1 then kp else let kn = KerName.make mp1 lbl in if mp1 == mp2 then same kn else make kn (KerName.make mp2 lbl) let to_string kp = KerName.to_string (user kp) let print kp = str (to_string kp) let debug_to_string = function | Same kn -> "(" ^ KerName.debug_to_string kn ^ ")" | Dual (knu,knc) -> "(" ^ KerName.debug_to_string knu ^ "," ^ KerName.debug_to_string knc ^ ")" let debug_print kp = str (debug_to_string kp) (** For ordering kernel pairs, both user or canonical parts may make sense, according to your needs: user for the environments, canonical for other uses (ex: non-logical things). *) module UserOrd = struct type t = kernel_pair let compare x y = KerName.compare (user x) (user y) let equal x y = x == y || KerName.equal (user x) (user y) let hash x = KerName.hash (user x) end module CanOrd = struct type t = kernel_pair let compare x y = KerName.compare (canonical x) (canonical y) let equal x y = x == y || KerName.equal (canonical x) (canonical y) let hash x = KerName.hash (canonical x) end module SyntacticOrd = struct let compare x y = match x, y with | Same knx, Same kny -> KerName.compare knx kny | Dual (knux,kncx), Dual (knuy,kncy) -> let c = KerName.compare knux knuy in if not (Int.equal c 0) then c else KerName.compare kncx kncy | Same _, _ -> -1 | Dual _, _ -> 1 let equal x y = x == y || compare x y = 0 let hash = function | Same kn -> KerName.hash kn | Dual (knu, knc) -> Hashset.Combine.combine (KerName.hash knu) (KerName.hash knc) end (** Default (logical) comparison and hash is on the canonical part *) let equal = CanOrd.equal let hash = CanOrd.hash module Self_Hashcons = struct type t = kernel_pair type u = KerName.t -> KerName.t let hashcons hkn = function | Same kn -> Same (hkn kn) | Dual (knu,knc) -> make (hkn knu) (hkn knc) let eq x y = (* physical comparison on subterms *) x == y || match x,y with | Same x, Same y -> x == y | Dual (ux,cx), Dual (uy,cy) -> ux == uy && cx == cy | (Same _ | Dual _), _ -> false (** Hash-consing (despite having the same user part implies having the same canonical part is a logical invariant of the system, it is not necessarily an invariant in memory, so we treat kernel names as they are syntactically for hash-consing) *) let hash = function | Same kn -> KerName.hash kn | Dual (knu, knc) -> Hashset.Combine.combine (KerName.hash knu) (KerName.hash knc) end module HashKP = Hashcons.Make(Self_Hashcons) end (** {6 Constant Names} *) module Constant = KerPair module Cmap = HMap.Make(Constant.CanOrd) (** A map whose keys are constants (values of the {!Constant.t} type). Keys are ordered wrt. "canonical form" of the constant. *) module Cmap_env = HMap.Make(Constant.UserOrd) (** A map whose keys are constants (values of the {!Constant.t} type). Keys are ordered wrt. "user form" of the constant. *) module Cpred = Predicate.Make(Constant.CanOrd) module Cset = Cmap.Set module Cset_env = Cmap_env.Set (** {6 Names of mutual inductive types } *) module MutInd = KerPair module Mindmap = HMap.Make(MutInd.CanOrd) module Mindset = Mindmap.Set module Mindmap_env = HMap.Make(MutInd.UserOrd) (** Designation of a (particular) inductive type. *) type inductive = MutInd.t (* the name of the inductive type *) * int (* the position of this inductive type within the block of mutually-recursive inductive types. BEWARE: indexing starts from 0. *) (** Designation of a (particular) constructor of a (particular) inductive type. *) type constructor = inductive (* designates the inductive type *) * int (* the index of the constructor BEWARE: indexing starts from 1. *) let ind_modpath (mind,_) = MutInd.modpath mind let constr_modpath (ind,_) = ind_modpath ind let ith_mutual_inductive (mind, _) i = (mind, i) let ith_constructor_of_inductive ind i = (ind, i) let inductive_of_constructor (ind, _i) = ind let index_of_constructor (_ind, i) = i let eq_ind (m1, i1) (m2, i2) = Int.equal i1 i2 && MutInd.equal m1 m2 let eq_user_ind (m1, i1) (m2, i2) = Int.equal i1 i2 && MutInd.UserOrd.equal m1 m2 let eq_syntactic_ind (m1, i1) (m2, i2) = Int.equal i1 i2 && MutInd.SyntacticOrd.equal m1 m2 let ind_ord (m1, i1) (m2, i2) = let c = Int.compare i1 i2 in if Int.equal c 0 then MutInd.CanOrd.compare m1 m2 else c let ind_user_ord (m1, i1) (m2, i2) = let c = Int.compare i1 i2 in if Int.equal c 0 then MutInd.UserOrd.compare m1 m2 else c let ind_syntactic_ord (m1, i1) (m2, i2) = let c = Int.compare i1 i2 in if Int.equal c 0 then MutInd.SyntacticOrd.compare m1 m2 else c let ind_hash (m, i) = Hashset.Combine.combine (MutInd.hash m) (Int.hash i) let ind_user_hash (m, i) = Hashset.Combine.combine (MutInd.UserOrd.hash m) (Int.hash i) let ind_syntactic_hash (m, i) = Hashset.Combine.combine (MutInd.SyntacticOrd.hash m) (Int.hash i) let eq_constructor (ind1, j1) (ind2, j2) = Int.equal j1 j2 && eq_ind ind1 ind2 let eq_user_constructor (ind1, j1) (ind2, j2) = Int.equal j1 j2 && eq_user_ind ind1 ind2 let eq_syntactic_constructor (ind1, j1) (ind2, j2) = Int.equal j1 j2 && eq_syntactic_ind ind1 ind2 let constructor_ord (ind1, j1) (ind2, j2) = let c = Int.compare j1 j2 in if Int.equal c 0 then ind_ord ind1 ind2 else c let constructor_user_ord (ind1, j1) (ind2, j2) = let c = Int.compare j1 j2 in if Int.equal c 0 then ind_user_ord ind1 ind2 else c let constructor_syntactic_ord (ind1, j1) (ind2, j2) = let c = Int.compare j1 j2 in if Int.equal c 0 then ind_syntactic_ord ind1 ind2 else c let constructor_hash (ind, i) = Hashset.Combine.combine (ind_hash ind) (Int.hash i) let constructor_user_hash (ind, i) = Hashset.Combine.combine (ind_user_hash ind) (Int.hash i) let constructor_syntactic_hash (ind, i) = Hashset.Combine.combine (ind_syntactic_hash ind) (Int.hash i) module InductiveOrdered = struct type t = inductive let compare = ind_ord end module InductiveOrdered_env = struct type t = inductive let compare = ind_user_ord end module Indset = Set.Make(InductiveOrdered) module Indset_env = Set.Make(InductiveOrdered_env) module Indmap = Map.Make(InductiveOrdered) module Indmap_env = Map.Make(InductiveOrdered_env) module ConstructorOrdered = struct type t = constructor let compare = constructor_ord end module ConstructorOrdered_env = struct type t = constructor let compare = constructor_user_ord end module Constrset = Set.Make(ConstructorOrdered) module Constrset_env = Set.Make(ConstructorOrdered_env) module Constrmap = Map.Make(ConstructorOrdered) module Constrmap_env = Map.Make(ConstructorOrdered_env) (** {6 Hash-consing of name objects } *) module Hind = Hashcons.Make( struct type t = inductive type u = MutInd.t -> MutInd.t let hashcons hmind (mind, i) = (hmind mind, i) let eq (mind1,i1) (mind2,i2) = mind1 == mind2 && Int.equal i1 i2 let hash = ind_hash end) module Hconstruct = Hashcons.Make( struct type t = constructor type u = inductive -> inductive let hashcons hind (ind, j) = (hind ind, j) let eq (ind1, j1) (ind2, j2) = ind1 == ind2 && Int.equal j1 j2 let hash = constructor_hash end) let hcons_con = Hashcons.simple_hcons Constant.HashKP.generate Constant.HashKP.hcons KerName.hcons let hcons_mind = Hashcons.simple_hcons MutInd.HashKP.generate MutInd.HashKP.hcons KerName.hcons let hcons_ind = Hashcons.simple_hcons Hind.generate Hind.hcons hcons_mind let hcons_construct = Hashcons.simple_hcons Hconstruct.generate Hconstruct.hcons hcons_ind (*****************) type 'a tableKey = | ConstKey of 'a | VarKey of Id.t | RelKey of Int.t type inv_rel_key = int (* index in the [rel_context] part of environment starting by the end, {\em inverse} of de Bruijn indice *) let eq_table_key f ik1 ik2 = if ik1 == ik2 then true else match ik1,ik2 with | ConstKey c1, ConstKey c2 -> f c1 c2 | VarKey id1, VarKey id2 -> Id.equal id1 id2 | RelKey k1, RelKey k2 -> Int.equal k1 k2 | _ -> false let eq_mind_chk = MutInd.UserOrd.equal let eq_ind_chk (kn1,i1) (kn2,i2) = Int.equal i1 i2 && eq_mind_chk kn1 kn2 (*******************************************************************) (** Compatibility layers *) let eq_constant_key = Constant.UserOrd.equal (** Compatibility layer for [ModPath] *) type module_path = ModPath.t = | MPfile of DirPath.t | MPbound of MBId.t | MPdot of module_path * Label.t (** Compatibility layer for [Constant] *) module Projection = struct module Repr = struct type t = { proj_ind : inductive; proj_npars : int; proj_arg : int; proj_name : Label.t; } let make proj_ind ~proj_npars ~proj_arg proj_name = {proj_ind;proj_npars;proj_arg;proj_name} let inductive c = c.proj_ind let mind c = fst c.proj_ind let constant c = KerPair.change_label (mind c) c.proj_name let label c = c.proj_name let npars c = c.proj_npars let arg c = c.proj_arg let equal a b = eq_ind a.proj_ind b.proj_ind && Int.equal a.proj_arg b.proj_arg let hash p = Hashset.Combine.combinesmall p.proj_arg (ind_hash p.proj_ind) module SyntacticOrd = struct let compare a b = let c = ind_syntactic_ord a.proj_ind b.proj_ind in if c == 0 then Int.compare a.proj_arg b.proj_arg else c let equal a b = a.proj_arg == b.proj_arg && eq_syntactic_ind a.proj_ind b.proj_ind let hash p = Hashset.Combine.combinesmall p.proj_arg (ind_hash p.proj_ind) end module CanOrd = struct let compare a b = let c = ind_ord a.proj_ind b.proj_ind in if c == 0 then Int.compare a.proj_arg b.proj_arg else c let equal a b = a.proj_arg == b.proj_arg && eq_ind a.proj_ind b.proj_ind let hash p = Hashset.Combine.combinesmall p.proj_arg (ind_hash p.proj_ind) end module UserOrd = struct let compare a b = let c = ind_user_ord a.proj_ind b.proj_ind in if c == 0 then Int.compare a.proj_arg b.proj_arg else c let equal a b = a.proj_arg == b.proj_arg && eq_user_ind a.proj_ind b.proj_ind let hash p = Hashset.Combine.combinesmall p.proj_arg (ind_user_hash p.proj_ind) end let compare a b = let c = ind_ord a.proj_ind b.proj_ind in if c == 0 then Int.compare a.proj_arg b.proj_arg else c module Self_Hashcons = struct type nonrec t = t type u = (inductive -> inductive) * (Id.t -> Id.t) let hashcons (hind,hid) p = { proj_ind = hind p.proj_ind; proj_npars = p.proj_npars; proj_arg = p.proj_arg; proj_name = hid p.proj_name } let eq p p' = p == p' || (p.proj_ind == p'.proj_ind && p.proj_npars == p'.proj_npars && p.proj_arg == p'.proj_arg && p.proj_name == p'.proj_name) let hash = hash end module HashRepr = Hashcons.Make(Self_Hashcons) let hcons = Hashcons.simple_hcons HashRepr.generate HashRepr.hcons (hcons_ind,Id.hcons) let map_npars f p = let ind = fst p.proj_ind in let npars = p.proj_npars in let ind', npars' = f ind npars in if ind == ind' && npars == npars' then p else {p with proj_ind = (ind',snd p.proj_ind); proj_npars = npars'} let map f p = map_npars (fun mind n -> f mind, n) p let to_string p = Constant.to_string (constant p) let print p = Constant.print (constant p) end type t = Repr.t * bool let make c b = (c, b) let mind (c,_) = Repr.mind c let inductive (c,_) = Repr.inductive c let npars (c,_) = Repr.npars c let arg (c,_) = Repr.arg c let constant (c,_) = Repr.constant c let label (c,_) = Repr.label c let repr = fst let unfolded = snd let unfold (c, b as p) = if b then p else (c, true) let equal (c, b) (c', b') = Repr.equal c c' && b == b' let repr_equal p p' = Repr.equal (repr p) (repr p') let hash (c, b) = (if b then 0 else 1) + Repr.hash c module SyntacticOrd = struct let compare (c, b) (c', b') = if b = b' then Repr.SyntacticOrd.compare c c' else -1 let equal (c, b as x) (c', b' as x') = x == x' || b = b' && Repr.SyntacticOrd.equal c c' let hash (c, b) = (if b then 0 else 1) + Repr.SyntacticOrd.hash c end module CanOrd = struct let compare (c, b) (c', b') = if b = b' then Repr.CanOrd.compare c c' else -1 let equal (c, b as x) (c', b' as x') = x == x' || b = b' && Repr.CanOrd.equal c c' let hash (c, b) = (if b then 0 else 1) + Repr.CanOrd.hash c end module Self_Hashcons = struct type nonrec t = t type u = Repr.t -> Repr.t let hashcons hc (c,b) = (hc c,b) let eq ((c,b) as x) ((c',b') as y) = x == y || (c == c' && b == b') let hash = hash end module HashProjection = Hashcons.Make(Self_Hashcons) let hcons = Hashcons.simple_hcons HashProjection.generate HashProjection.hcons Repr.hcons let compare (c, b) (c', b') = if b == b' then Repr.compare c c' else if b then 1 else -1 let map f (c, b as x) = let c' = Repr.map f c in if c' == c then x else (c', b) let map_npars f (c, b as x) = let c' = Repr.map_npars f c in if c' == c then x else (c', b) let to_string p = Constant.to_string (constant p) let print p = Constant.print (constant p) end module GlobRefInternal = struct type t = | VarRef of variable (** A reference to the section-context. *) | ConstRef of Constant.t (** A reference to the environment. *) | IndRef of inductive (** A reference to an inductive type. *) | ConstructRef of constructor (** A reference to a constructor of an inductive type. *) let equal gr1 gr2 = gr1 == gr2 || match gr1,gr2 with | ConstRef con1, ConstRef con2 -> Constant.equal con1 con2 | IndRef kn1, IndRef kn2 -> eq_ind kn1 kn2 | ConstructRef kn1, ConstructRef kn2 -> eq_constructor kn1 kn2 | VarRef v1, VarRef v2 -> Id.equal v1 v2 | (ConstRef _ | IndRef _ | ConstructRef _ | VarRef _), _ -> false let global_eq_gen eq_cst eq_ind eq_cons x y = x == y || match x, y with | ConstRef cx, ConstRef cy -> eq_cst cx cy | IndRef indx, IndRef indy -> eq_ind indx indy | ConstructRef consx, ConstructRef consy -> eq_cons consx consy | VarRef v1, VarRef v2 -> Id.equal v1 v2 | (VarRef _ | ConstRef _ | IndRef _ | ConstructRef _), _ -> false let global_ord_gen ord_cst ord_ind ord_cons x y = if x == y then 0 else match x, y with | VarRef v1, VarRef v2 -> Id.compare v1 v2 | VarRef _, _ -> -1 | _, VarRef _ -> 1 | ConstRef cx, ConstRef cy -> ord_cst cx cy | ConstRef _, _ -> -1 | _, ConstRef _ -> 1 | IndRef indx, IndRef indy -> ord_ind indx indy | IndRef _, _ -> -1 | _ , IndRef _ -> 1 | ConstructRef consx, ConstructRef consy -> ord_cons consx consy let global_hash_gen hash_cst hash_ind hash_cons gr = let open Hashset.Combine in match gr with | ConstRef c -> combinesmall 1 (hash_cst c) | IndRef i -> combinesmall 2 (hash_ind i) | ConstructRef c -> combinesmall 3 (hash_cons c) | VarRef id -> combinesmall 4 (Id.hash id) end module GlobRef = struct type t = GlobRefInternal.t = | VarRef of variable (** A reference to the section-context. *) | ConstRef of Constant.t (** A reference to the environment. *) | IndRef of inductive (** A reference to an inductive type. *) | ConstructRef of constructor (** A reference to a constructor of an inductive type. *) let equal = GlobRefInternal.equal (* By default, [global_reference] are ordered on their canonical part *) module Ordered = struct open Constant.CanOrd type t = GlobRefInternal.t let compare gr1 gr2 = GlobRefInternal.global_ord_gen compare ind_ord constructor_ord gr1 gr2 let equal gr1 gr2 = GlobRefInternal.global_eq_gen equal eq_ind eq_constructor gr1 gr2 let hash gr = GlobRefInternal.global_hash_gen hash ind_hash constructor_hash gr end module Ordered_env = struct open Constant.UserOrd type t = GlobRefInternal.t let compare gr1 gr2 = GlobRefInternal.global_ord_gen compare ind_user_ord constructor_user_ord gr1 gr2 let equal gr1 gr2 = GlobRefInternal.global_eq_gen equal eq_user_ind eq_user_constructor gr1 gr2 let hash gr = GlobRefInternal.global_hash_gen hash ind_user_hash constructor_user_hash gr end module Map = HMap.Make(Ordered) module Set = Map.Set (* Alternative sets and maps indexed by the user part of the kernel names *) module Map_env = HMap.Make(Ordered_env) module Set_env = Map_env.Set end type evaluable_global_reference = | EvalVarRef of Id.t | EvalConstRef of Constant.t (* Better to have it here that in closure, since used in grammar.cma *) let eq_egr e1 e2 = match e1, e2 with EvalConstRef con1, EvalConstRef con2 -> Constant.equal con1 con2 | EvalVarRef id1, EvalVarRef id2 -> Id.equal id1 id2 | _, _ -> false (** Located identifiers and objects with syntax. *) type lident = Id.t CAst.t type lname = Name.t CAst.t type lstring = string CAst.t coq-8.11.0/kernel/typeops.ml0000644000175000017500000006155613612664533015620 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* try conv_leq false env t1 t2 with NotConvertible -> raise (NotConvertibleVect i)) () v1 v2 let check_constraints cst env = if Environ.check_constraints cst env then () else error_unsatisfied_constraints env cst (* This should be a type (a priori without intention to be an assumption) *) let check_type env c t = match kind(whd_all env t) with | Sort s -> s | _ -> error_not_type env (make_judge c t) (* This should be a type intended to be assumed. The error message is not as useful as for [type_judgment]. *) let infer_assumption env t ty = try let s = check_type env t ty in (match s with Sorts.SProp -> Irrelevant | _ -> Relevant) with TypeError _ -> error_assumption env (make_judge t ty) let warn_bad_relevance_name = "bad-relevance" let warn_bad_relevance = CWarnings.create ~name:warn_bad_relevance_name ~category:"debug" ~default:CWarnings.Disabled Pp.(function | None -> str "Bad relevance in case annotation." | Some x -> str "Bad relevance for binder " ++ Name.print x.binder_name ++ str ".") let warn_bad_relevance_ci ?loc () = warn_bad_relevance ?loc None let warn_bad_relevance ?loc x = warn_bad_relevance ?loc (Some x) let check_assumption env x t ty = let r = x.binder_relevance in let r' = infer_assumption env t ty in let x = if Sorts.relevance_equal r r' then x else (warn_bad_relevance x; {x with binder_relevance = r'}) in x (************************************************) (* Incremental typing rules: builds a typing judgment given the *) (* judgments for the subterms. *) (*s Type of sorts *) (* Prop and Set *) let type1 = mkSort Sorts.type1 (* Type of Type(i). *) let type_of_type u = let uu = Universe.super u in mkType uu let type_of_sort = function | SProp | Prop | Set -> type1 | Type u -> type_of_type u (*s Type of a de Bruijn index. *) let type_of_relative env n = try env |> lookup_rel n |> RelDecl.get_type |> lift n with Not_found -> error_unbound_rel env n (* Type of variables *) let type_of_variable env id = try named_type id env with Not_found -> error_unbound_var env id (* Management of context of variables. *) (* Checks if a context of variables can be instantiated by the variables of the current env. Order does not have to be checked assuming that all names are distinct *) let check_hyps_inclusion env ?evars f c sign = let conv env a b = conv env ?evars a b in Context.Named.fold_outside (fun d1 () -> let open Context.Named.Declaration in let id = NamedDecl.get_id d1 in try let d2 = lookup_named id env in conv env (get_type d2) (get_type d1); (match d2,d1 with | LocalAssum _, LocalAssum _ -> () | LocalAssum _, LocalDef _ -> (* This is wrong, because we don't know if the body is needed or not for typechecking: *) () | LocalDef _, LocalAssum _ -> raise NotConvertible | LocalDef (_,b2,_), LocalDef (_,b1,_) -> conv env b2 b1); with Not_found | NotConvertible | Option.Heterogeneous -> error_reference_variables env id (f c)) sign ~init:() (* Instantiation of terms on real arguments. *) (* Make a type polymorphic if an arity *) (* Type of constants *) let type_of_constant env (kn,_u as cst) = let cb = lookup_constant kn env in let () = check_hyps_inclusion env mkConstU cst cb.const_hyps in let ty, cu = constant_type env cst in let () = check_constraints cu env in ty let type_of_constant_in env (kn,_u as cst) = let cb = lookup_constant kn env in let () = check_hyps_inclusion env mkConstU cst cb.const_hyps in constant_type_in env cst (* Type of a lambda-abstraction. *) (* [judge_of_abstraction env name var j] implements the rule env, name:typ |- j.uj_val:j.uj_type env, |- (name:typ)j.uj_type : s ----------------------------------------------------------------------- env |- [name:typ]j.uj_val : (name:typ)j.uj_type Since all products are defined in the Calculus of Inductive Constructions and no upper constraint exists on the sort $s$, we don't need to compute $s$ *) let type_of_abstraction _env name var ty = mkProd (name, var, ty) (* Type of an application. *) let make_judgev c t = Array.map2 make_judge c t let rec check_empty_stack = function | [] -> true | CClosure.Zupdate _ :: s -> check_empty_stack s | _ -> false let type_of_apply env func funt argsv argstv = let open CClosure in let len = Array.length argsv in let infos = create_clos_infos all env in let tab = create_tab () in let rec apply_rec i typ = if Int.equal i len then term_of_fconstr typ else let typ, stk = whd_stack infos tab typ [] in (** The return stack is known to be empty *) let () = assert (check_empty_stack stk) in match fterm_of typ with | FProd (_, c1, c2, e) -> let arg = argsv.(i) in let argt = argstv.(i) in let c1 = term_of_fconstr c1 in begin match conv_leq false env argt c1 with | () -> apply_rec (i+1) (mk_clos (Esubst.subs_cons ([| inject arg |], e)) c2) | exception NotConvertible -> error_cant_apply_bad_type env (i+1,c1,argt) (make_judge func funt) (make_judgev argsv argstv) end | _ -> error_cant_apply_not_functional env (make_judge func funt) (make_judgev argsv argstv) in apply_rec 0 (inject funt) (* Type of primitive constructs *) let type_of_prim_type _env = function | CPrimitives.PT_int63 -> Constr.mkSet | CPrimitives.PT_float64 -> Constr.mkSet let type_of_int env = match env.retroknowledge.Retroknowledge.retro_int63 with | Some c -> mkConst c | None -> CErrors.user_err Pp.(str"The type int must be registered before this construction can be typechecked.") let type_of_float env = match env.retroknowledge.Retroknowledge.retro_float64 with | Some c -> mkConst c | None -> raise (Invalid_argument "Typeops.type_of_float: float64 not_defined") let type_of_prim env t = let int_ty () = type_of_int env in let float_ty () = type_of_float env in let bool_ty () = match env.retroknowledge.Retroknowledge.retro_bool with | Some ((ind,_),_) -> Constr.mkInd ind | None -> CErrors.user_err Pp.(str"The type bool must be registered before this primitive.") in let compare_ty () = match env.retroknowledge.Retroknowledge.retro_cmp with | Some ((ind,_),_,_) -> Constr.mkInd ind | None -> CErrors.user_err Pp.(str"The type compare must be registered before this primitive.") in let f_compare_ty () = match env.retroknowledge.Retroknowledge.retro_f_cmp with | Some ((ind,_),_,_,_) -> Constr.mkInd ind | None -> CErrors.user_err Pp.(str"The type float_comparison must be registered before this primitive.") in let f_class_ty () = match env.retroknowledge.Retroknowledge.retro_f_class with | Some ((ind,_),_,_,_,_,_,_,_,_) -> Constr.mkInd ind | None -> CErrors.user_err Pp.(str"The type float_class must be registered before this primitive.") in let pair_ty fst_ty snd_ty = match env.retroknowledge.Retroknowledge.retro_pair with | Some (ind,_) -> Constr.mkApp(Constr.mkInd ind, [|fst_ty;snd_ty|]) | None -> CErrors.user_err Pp.(str"The type pair must be registered before this primitive.") in let carry_ty int_ty = match env.retroknowledge.Retroknowledge.retro_carry with | Some ((ind,_),_) -> Constr.mkApp(Constr.mkInd ind, [|int_ty|]) | None -> CErrors.user_err Pp.(str"The type carry must be registered before this primitive.") in let open CPrimitives in let tr_prim_type = function | PT_int63 -> int_ty () | PT_float64 -> float_ty () in let tr_ind (type t) (i : t prim_ind) (a : t) = match i, a with | PIT_bool, () -> bool_ty () | PIT_carry, t -> carry_ty (tr_prim_type t) | PIT_pair, (t1, t2) -> pair_ty (tr_prim_type t1) (tr_prim_type t2) | PIT_cmp, () -> compare_ty () | PIT_f_cmp, () -> f_compare_ty () | PIT_f_class, () -> f_class_ty () in let tr_type = function | PITT_ind (i, a) -> tr_ind i a | PITT_type t -> tr_prim_type t in let rec nary_op = function | [] -> assert false | [ret_ty] -> tr_type ret_ty | arg_ty :: r -> let arg_ty = tr_type arg_ty in Constr.mkProd(Context.nameR (Id.of_string "x"), arg_ty, nary_op r) in nary_op (types t) let type_of_prim_or_type env = let open CPrimitives in function | OT_type t -> type_of_prim_type env t | OT_op op -> type_of_prim env op let judge_of_int env i = make_judge (Constr.mkInt i) (type_of_int env) let judge_of_float env f = make_judge (Constr.mkFloat f) (type_of_float env) (* Type of product *) let sort_of_product env domsort rangsort = match (domsort, rangsort) with | (_, SProp) | (SProp, _) -> rangsort (* Product rule (s,Prop,Prop) *) | (_, Prop) -> rangsort (* Product rule (Prop/Set,Set,Set) *) | ((Prop | Set), Set) -> rangsort (* Product rule (Type,Set,?) *) | (Type u1, Set) -> if is_impredicative_set env then (* Rule is (Type,Set,Set) in the Set-impredicative calculus *) rangsort else (* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *) Sorts.sort_of_univ (Universe.sup Universe.type0 u1) (* Product rule (Prop,Type_i,Type_i) *) | (Set, Type u2) -> Sorts.sort_of_univ (Universe.sup Universe.type0 u2) (* Product rule (Prop,Type_i,Type_i) *) | (Prop, Type _) -> rangsort (* Product rule (Type_i,Type_i,Type_i) *) | (Type u1, Type u2) -> Sorts.sort_of_univ (Universe.sup u1 u2) (* [judge_of_product env name (typ1,s1) (typ2,s2)] implements the rule env |- typ1:s1 env, name:typ1 |- typ2 : s2 ------------------------------------------------------------------------- s' >= (s1,s2), env |- (name:typ)j.uj_val : s' where j.uj_type is convertible to a sort s2 *) let type_of_product env _name s1 s2 = let s = sort_of_product env s1 s2 in mkSort s (* Type of a type cast *) (* [judge_of_cast env (c,typ1) (typ2,s)] implements the rule env |- c:typ1 env |- typ2:s env |- typ1 <= typ2 --------------------------------------------------------------------- env |- c:typ2 *) let check_cast env c ct k expected_type = try match k with | VMcast -> Vconv.vm_conv CUMUL env ct expected_type | DEFAULTcast -> default_conv ~l2r:false CUMUL env ct expected_type | REVERTcast -> default_conv ~l2r:true CUMUL env ct expected_type | NATIVEcast -> let sigma = Nativelambda.empty_evars in Nativeconv.native_conv CUMUL sigma env ct expected_type with NotConvertible -> error_actual_type env (make_judge c ct) expected_type (* Inductive types. *) (* The type is parametric over the uniform parameters whose conclusion is in Type; to enforce the internal constraints between the parameters and the instances of Type occurring in the type of the constructors, we use the level variables _statically_ assigned to the conclusions of the parameters as mediators: e.g. if a parameter has conclusion Type(alpha), static constraints of the form alpha<=v exist between alpha and the Type's occurring in the constructor types; when the parameters is finally instantiated by a term of conclusion Type(u), then the constraints u<=alpha is computed in the App case of execute; from this constraints, the expected dynamic constraints of the form u<=v are enforced *) let type_of_inductive_knowing_parameters env (ind,u as indu) args = let (mib,_mip) as spec = lookup_mind_specif env ind in check_hyps_inclusion env mkIndU indu mib.mind_hyps; let t,cst = Inductive.constrained_type_of_inductive_knowing_parameters env (spec,u) args in check_constraints cst env; t let type_of_inductive env (ind,u as indu) = let (mib,mip) = lookup_mind_specif env ind in check_hyps_inclusion env mkIndU indu mib.mind_hyps; let t,cst = Inductive.constrained_type_of_inductive env ((mib,mip),u) in check_constraints cst env; t (* Constructors. *) let type_of_constructor env (c,_u as cu) = let () = let ((kn,_),_) = c in let mib = lookup_mind kn env in check_hyps_inclusion env mkConstructU cu mib.mind_hyps in let specif = lookup_mind_specif env (inductive_of_constructor c) in let t,cst = constrained_type_of_constructor cu specif in let () = check_constraints cst env in t (* Case. *) let check_branch_types env (ind,u) c ct lft explft = try conv_leq_vecti env lft explft with NotConvertibleVect i -> error_ill_formed_branch env c ((ind,i+1),u) lft.(i) explft.(i) | Invalid_argument _ -> error_number_branches env (make_judge c ct) (Array.length explft) let type_of_case env ci p pt c ct _lf lft = let (pind, _ as indspec) = try find_rectype env ct with Not_found -> error_case_not_inductive env (make_judge c ct) in let _, sp = try dest_arity env pt with NotArity -> error_elim_arity env pind c (make_judge p pt) None in let rp = Sorts.relevance_of_sort sp in let ci = if ci.ci_relevance == rp then ci else (warn_bad_relevance_ci (); {ci with ci_relevance=rp}) in let () = check_case_info env pind rp ci in let (bty,rslty) = type_case_branches env indspec (make_judge p pt) c in let () = check_branch_types env pind c ct lft bty in ci, rslty let type_of_projection env p c ct = let pty = lookup_projection p env in let (ind,u), args = try find_rectype env ct with Not_found -> error_case_not_inductive env (make_judge c ct) in assert(eq_ind (Projection.inductive p) ind); let ty = Vars.subst_instance_constr u pty in substl (c :: CList.rev args) ty (* Fixpoints. *) (* Checks the type of a general (co)fixpoint, i.e. without checking *) (* the specific guard condition. *) let check_fixpoint env lna lar vdef vdeft = let lt = Array.length vdeft in assert (Int.equal (Array.length lar) lt); try conv_leq_vecti env vdeft (Array.map (fun ty -> lift lt ty) lar) with NotConvertibleVect i -> error_ill_typed_rec_body env i lna (make_judgev vdef vdeft) lar (* Global references *) let type_of_global_in_context env r = let open Names.GlobRef in match r with | VarRef id -> Environ.named_type id env, Univ.AUContext.empty | ConstRef c -> let cb = Environ.lookup_constant c env in let univs = Declareops.constant_polymorphic_context cb in cb.Declarations.const_type, univs | IndRef ind -> let (mib,_ as specif) = Inductive.lookup_mind_specif env ind in let univs = Declareops.inductive_polymorphic_context mib in let inst = Univ.make_abstract_instance univs in let env = Environ.push_context ~strict:false (Univ.AUContext.repr univs) env in Inductive.type_of_inductive env (specif, inst), univs | ConstructRef cstr -> let (mib,_ as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in let univs = Declareops.inductive_polymorphic_context mib in let inst = Univ.make_abstract_instance univs in Inductive.type_of_constructor (cstr,inst) specif, univs (************************************************************************) (************************************************************************) let check_binder_annot s x = let r = x.binder_relevance in let r' = Sorts.relevance_of_sort s in if r' == r then x else (warn_bad_relevance x; {x with binder_relevance = r'}) (* The typing machine. *) (* ATTENTION : faudra faire le typage du contexte des Const, Ind et Constructsi un jour cela devient des constructions arbitraires et non plus des variables *) let rec execute env cstr = let open Context.Rel.Declaration in match kind cstr with (* Atomic terms *) | Sort s -> (match s with | SProp -> if not (Environ.sprop_allowed env) then error_disallowed_sprop env | _ -> ()); cstr, type_of_sort s | Rel n -> cstr, type_of_relative env n | Var id -> cstr, type_of_variable env id | Const c -> cstr, type_of_constant env c | Proj (p, c) -> let c', ct = execute env c in let cstr = if c == c' then cstr else mkProj (p,c') in cstr, type_of_projection env p c' ct (* Lambda calculus operators *) | App (f,args) -> let args', argst = execute_array env args in let f', ft = match kind f with | Ind ind when Environ.template_polymorphic_pind ind env -> let args = Array.map (fun t -> lazy t) argst in f, type_of_inductive_knowing_parameters env ind args | _ -> (* No template polymorphism *) execute env f in let cstr = if f == f' && args == args' then cstr else mkApp (f',args') in cstr, type_of_apply env f' ft args' argst | Lambda (name,c1,c2) -> let c1', s = execute_is_type env c1 in let name' = check_binder_annot s name in let env1 = push_rel (LocalAssum (name',c1')) env in let c2', c2t = execute env1 c2 in let cstr = if name == name' && c1 == c1' && c2 == c2' then cstr else mkLambda(name',c1',c2') in cstr, type_of_abstraction env name' c1 c2t | Prod (name,c1,c2) -> let c1', vars = execute_is_type env c1 in let name' = check_binder_annot vars name in let env1 = push_rel (LocalAssum (name',c1')) env in let c2', vars' = execute_is_type env1 c2 in let cstr = if name == name' && c1 == c1' && c2 == c2' then cstr else mkProd(name',c1',c2') in cstr, type_of_product env name' vars vars' | LetIn (name,c1,c2,c3) -> let c1', c1t = execute env c1 in let c2', c2s = execute_is_type env c2 in let name' = check_binder_annot c2s name in let () = check_cast env c1' c1t DEFAULTcast c2' in let env1 = push_rel (LocalDef (name',c1',c2')) env in let c3', c3t = execute env1 c3 in let cstr = if name == name' && c1 == c1' && c2 == c2' && c3 == c3' then cstr else mkLetIn(name',c1',c2',c3') in cstr, subst1 c1 c3t | Cast (c,k,t) -> let c', ct = execute env c in let t', _ts = execute_is_type env t in let () = check_cast env c' ct k t' in let cstr = if c == c' && t == t' then cstr else mkCast(c',k,t') in cstr, t' (* Inductive types *) | Ind ind -> cstr, type_of_inductive env ind | Construct c -> cstr, type_of_constructor env c | Case (ci,p,c,lf) -> let c', ct = execute env c in let p', pt = execute env p in let lf', lft = execute_array env lf in let ci', t = type_of_case env ci p' pt c' ct lf' lft in let cstr = if ci == ci' && c == c' && p == p' && lf == lf' then cstr else mkCase(ci',p',c',lf') in cstr, t | Fix ((_vn,i as vni),recdef as fix) -> let (fix_ty,recdef') = execute_recdef env recdef i in let cstr, fix = if recdef == recdef' then cstr, fix else let fix = (vni,recdef') in mkFix fix, fix in check_fix env fix; cstr, fix_ty | CoFix (i,recdef as cofix) -> let (fix_ty,recdef') = execute_recdef env recdef i in let cstr, cofix = if recdef == recdef' then cstr, cofix else let cofix = (i,recdef') in mkCoFix cofix, cofix in check_cofix env cofix; cstr, fix_ty (* Primitive types *) | Int _ -> cstr, type_of_int env | Float _ -> cstr, type_of_float env (* Partial proofs: unsupported by the kernel *) | Meta _ -> anomaly (Pp.str "the kernel does not support metavariables.") | Evar _ -> anomaly (Pp.str "the kernel does not support existential variables.") and execute_is_type env constr = let c, t = execute env constr in c, check_type env constr t and execute_recdef env (names,lar,vdef as recdef) i = let lar', lart = execute_array env lar in let names' = Array.Smart.map_i (fun i na -> check_assumption env na lar'.(i) lart.(i)) names in let env1 = push_rec_types (names',lar',vdef) env in (* vdef is ignored *) let vdef', vdeft = execute_array env1 vdef in let () = check_fixpoint env1 names' lar' vdef' vdeft in let recdef = if names == names' && lar == lar' && vdef == vdef' then recdef else (names',lar',vdef') in (lar'.(i),recdef) and execute_array env cs = let tys = Array.make (Array.length cs) mkProp in let cs = Array.Smart.map_i (fun i c -> let c, ty = execute env c in tys.(i) <- ty; c) cs in cs, tys (* Derived functions *) let check_wellformed_universes env c = let univs = universes_of_constr c in try UGraph.check_declared_universes (universes env) univs with UGraph.UndeclaredLevel u -> error_undeclared_universe env u let infer env constr = let () = check_wellformed_universes env constr in let constr, t = execute env constr in make_judge constr t let infer = if Flags.profile then let infer_key = CProfile.declare_profile "Fast_infer" in CProfile.profile2 infer_key (fun b c -> infer b c) else (fun b c -> infer b c) let assumption_of_judgment env {uj_val=c; uj_type=t} = infer_assumption env c t let type_judgment env {uj_val=c; uj_type=t} = let s = check_type env c t in {utj_val = c; utj_type = s } let infer_type env constr = let () = check_wellformed_universes env constr in let constr, t = execute env constr in let s = check_type env constr t in {utj_val = constr; utj_type = s} (* Typing of several terms. *) let check_context env rels = let open Context.Rel.Declaration in Context.Rel.fold_outside (fun d (env,rels) -> match d with | LocalAssum (x,ty) -> let jty = infer_type env ty in let x = check_binder_annot jty.utj_type x in push_rel d env, LocalAssum (x,jty.utj_val) :: rels | LocalDef (x,bd,ty) -> let j1 = infer env bd in let jty = infer_type env ty in conv_leq false env j1.uj_type ty; let x = check_binder_annot jty.utj_type x in push_rel d env, LocalDef (x,j1.uj_val,jty.utj_val) :: rels) rels ~init:(env,[]) let judge_of_prop = make_judge mkProp type1 let judge_of_set = make_judge mkSet type1 let judge_of_type u = make_judge (mkType u) (type_of_type u) let judge_of_relative env k = make_judge (mkRel k) (type_of_relative env k) let judge_of_variable env x = make_judge (mkVar x) (type_of_variable env x) let judge_of_constant env cst = make_judge (mkConstU cst) (type_of_constant env cst) let judge_of_projection env p cj = make_judge (mkProj (p,cj.uj_val)) (type_of_projection env p cj.uj_val cj.uj_type) let dest_judgev v = Array.map j_val v, Array.map j_type v let judge_of_apply env funj argjv = let args, argtys = dest_judgev argjv in make_judge (mkApp (funj.uj_val, args)) (type_of_apply env funj.uj_val funj.uj_type args argtys) (* let judge_of_abstraction env x varj bodyj = *) (* make_judge (mkLambda (x, varj.utj_val, bodyj.uj_val)) *) (* (type_of_abstraction env x varj.utj_val bodyj.uj_type) *) (* let judge_of_product env x varj outj = *) (* make_judge (mkProd (x, varj.utj_val, outj.utj_val)) *) (* (mkSort (sort_of_product env varj.utj_type outj.utj_type)) *) (* let judge_of_letin env name defj typj j = *) (* make_judge (mkLetIn (name, defj.uj_val, typj.utj_val, j.uj_val)) *) (* (subst1 defj.uj_val j.uj_type) *) let judge_of_cast env cj k tj = let () = check_cast env cj.uj_val cj.uj_type k tj.utj_val in let c = match k with | REVERTcast -> cj.uj_val | _ -> mkCast (cj.uj_val, k, tj.utj_val) in make_judge c tj.utj_val let judge_of_inductive env indu = make_judge (mkIndU indu) (type_of_inductive env indu) let judge_of_constructor env cu = make_judge (mkConstructU cu) (type_of_constructor env cu) let judge_of_case env ci pj cj lfj = let lf, lft = dest_judgev lfj in let ci, t = type_of_case env ci pj.uj_val pj.uj_type cj.uj_val cj.uj_type lf lft in make_judge (mkCase (ci, (*nf_betaiota*) pj.uj_val, cj.uj_val, lft)) t (* Building type of primitive operators and type *) let check_primitive_type env op_t t = let inft = type_of_prim_or_type env op_t in try default_conv ~l2r:false CUMUL env inft t with NotConvertible -> error_incorrect_primitive env (make_judge op_t inft) t coq-8.11.0/kernel/term.ml0000644000175000017500000003277413612664533015064 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* mkProd (na, t, c) | LocalDef (na,b,t) -> mkLetIn (na, b, t, c) let mkNamedProd_or_LetIn decl c = let open Context.Named.Declaration in match decl with | LocalAssum (id,t) -> mkNamedProd id t c | LocalDef (id,b,t) -> mkNamedLetIn id b t c (* Constructs either [(x:t)c] or [c] where [x] is replaced by [b] *) let mkProd_wo_LetIn decl c = let open Context.Rel.Declaration in match decl with | LocalAssum (na,t) -> mkProd (na, t, c) | LocalDef (_na,b,_t) -> subst1 b c let mkNamedProd_wo_LetIn decl c = let open Context.Named.Declaration in match decl with | LocalAssum (id,t) -> mkNamedProd id t c | LocalDef (id,b,_) -> subst1 b (subst_var id.binder_name c) (* non-dependent product t1 -> t2 *) let mkArrow t1 r t2 = mkProd (make_annot Anonymous r, t1, t2) let mkArrowR t1 t2 = mkArrow t1 Sorts.Relevant t2 (* Constructs either [[x:t]c] or [[x=b:t]c] *) let mkLambda_or_LetIn decl c = let open Context.Rel.Declaration in match decl with | LocalAssum (na,t) -> mkLambda (na, t, c) | LocalDef (na,b,t) -> mkLetIn (na, b, t, c) let mkNamedLambda_or_LetIn decl c = let open Context.Named.Declaration in match decl with | LocalAssum (id,t) -> mkNamedLambda id t c | LocalDef (id,b,t) -> mkNamedLetIn id b t c (* prodn n [xn:Tn;..;x1:T1;Gamma] b = (x1:T1)..(xn:Tn)b *) let prodn n env b = let rec prodrec = function | (0, _env, b) -> b | (n, ((v,t)::l), b) -> prodrec (n-1, l, mkProd (v,t,b)) | _ -> assert false in prodrec (n,env,b) (* compose_prod [xn:Tn;..;x1:T1] b = (x1:T1)..(xn:Tn)b *) let compose_prod l b = prodn (List.length l) l b (* lamn n [xn:Tn;..;x1:T1;Gamma] b = [x1:T1]..[xn:Tn]b *) let lamn n env b = let rec lamrec = function | (0, _env, b) -> b | (n, ((v,t)::l), b) -> lamrec (n-1, l, mkLambda (v,t,b)) | _ -> assert false in lamrec (n,env,b) (* compose_lam [xn:Tn;..;x1:T1] b = [x1:T1]..[xn:Tn]b *) let compose_lam l b = lamn (List.length l) l b let applist (f,l) = mkApp (f, Array.of_list l) let applistc f l = mkApp (f, Array.of_list l) let appvect = mkApp let appvectc f l = mkApp (f,l) (* to_lambda n (x1:T1)...(xn:Tn)T = * [x1:T1]...[xn:Tn]T *) let rec to_lambda n prod = if Int.equal n 0 then prod else match kind prod with | Prod (na,ty,bd) -> mkLambda (na,ty,to_lambda (n-1) bd) | Cast (c,_,_) -> to_lambda n c | _ -> user_err ~hdr:"to_lambda" (mt ()) let rec to_prod n lam = if Int.equal n 0 then lam else match kind lam with | Lambda (na,ty,bd) -> mkProd (na,ty,to_prod (n-1) bd) | Cast (c,_,_) -> to_prod n c | _ -> user_err ~hdr:"to_prod" (mt ()) let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c) let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c) (* Application with expected on-the-fly reduction *) let lambda_applist c l = let rec app subst c l = match kind c, l with | Lambda(_,_,c), arg::l -> app (arg::subst) c l | _, [] -> substl subst c | _ -> anomaly (Pp.str "Not enough lambda's.") in app [] c l let lambda_appvect c v = lambda_applist c (Array.to_list v) let lambda_applist_assum n c l = let rec app n subst t l = if Int.equal n 0 then if l == [] then substl subst t else anomaly (Pp.str "Too many arguments.") else match kind t, l with | Lambda(_,_,c), arg::l -> app (n-1) (arg::subst) c l | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l | _, [] -> anomaly (Pp.str "Not enough arguments.") | _ -> anomaly (Pp.str "Not enough lambda/let's.") in app n [] c l let lambda_appvect_assum n c v = lambda_applist_assum n c (Array.to_list v) (* prod_applist T [ a1 ; ... ; an ] -> (T a1 ... an) *) let prod_applist c l = let rec app subst c l = match kind c, l with | Prod(_,_,c), arg::l -> app (arg::subst) c l | _, [] -> substl subst c | _ -> anomaly (Pp.str "Not enough prod's.") in app [] c l (* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *) let prod_appvect c v = prod_applist c (Array.to_list v) let prod_applist_assum n c l = let rec app n subst t l = if Int.equal n 0 then if l == [] then substl subst t else anomaly (Pp.str "Too many arguments.") else match kind t, l with | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l | _, [] -> anomaly (Pp.str "Not enough arguments.") | _ -> anomaly (Pp.str "Not enough prod/let's.") in app n [] c l let prod_appvect_assum n c v = prod_applist_assum n c (Array.to_list v) (*********************************) (* Other term destructors *) (*********************************) (* Transforms a product term (x1:T1)..(xn:Tn)T into the pair ([(xn,Tn);...;(x1,T1)],T), where T is not a product *) let decompose_prod = let rec prodec_rec l c = match kind c with | Prod (x,t,c) -> prodec_rec ((x,t)::l) c | Cast (c,_,_) -> prodec_rec l c | _ -> l,c in prodec_rec [] (* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair ([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *) let decompose_lam = let rec lamdec_rec l c = match kind c with | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) c | Cast (c,_,_) -> lamdec_rec l c | _ -> l,c in lamdec_rec [] (* Given a positive integer n, transforms a product term (x1:T1)..(xn:Tn)T into the pair ([(xn,Tn);...;(x1,T1)],T) *) let decompose_prod_n n = if n < 0 then user_err (str "decompose_prod_n: integer parameter must be positive"); let rec prodec_rec l n c = if Int.equal n 0 then l,c else match kind c with | Prod (x,t,c) -> prodec_rec ((x,t)::l) (n-1) c | Cast (c,_,_) -> prodec_rec l n c | _ -> user_err (str "decompose_prod_n: not enough products") in prodec_rec [] n (* Given a positive integer n, transforms a lambda term [x1:T1]..[xn:Tn]T into the pair ([(xn,Tn);...;(x1,T1)],T) *) let decompose_lam_n n = if n < 0 then user_err (str "decompose_lam_n: integer parameter must be positive"); let rec lamdec_rec l n c = if Int.equal n 0 then l,c else match kind c with | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c | Cast (c,_,_) -> lamdec_rec l n c | _ -> user_err (str "decompose_lam_n: not enough abstractions") in lamdec_rec [] n (* Transforms a product term (x1:T1)..(xn:Tn)T into the pair ([(xn,Tn);...;(x1,T1)],T), where T is not a product *) let decompose_prod_assum = let open Context.Rel.Declaration in let rec prodec_rec l c = match kind c with | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) c | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c | Cast (c,_,_) -> prodec_rec l c | _ -> l,c in prodec_rec Context.Rel.empty (* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair ([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *) let decompose_lam_assum = let rec lamdec_rec l c = let open Context.Rel.Declaration in match kind c with | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) c | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c | Cast (c,_,_) -> lamdec_rec l c | _ -> l,c in lamdec_rec Context.Rel.empty (* Given a positive integer n, decompose a product or let-in term of the form [forall (x1:T1)..(xi:=ci:Ti)..(xn:Tn), T] into the pair of the quantifying context [(xn,None,Tn);..;(xi,Some ci,Ti);..;(x1,None,T1)] and of the inner type [T]) *) let decompose_prod_n_assum n = if n < 0 then user_err (str "decompose_prod_n_assum: integer parameter must be positive"); let rec prodec_rec l n c = if Int.equal n 0 then l,c else let open Context.Rel.Declaration in match kind c with | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c | Cast (c,_,_) -> prodec_rec l n c | _ -> user_err (str "decompose_prod_n_assum: not enough assumptions") in prodec_rec Context.Rel.empty n (* Given a positive integer n, decompose a lambda or let-in term [fun (x1:T1)..(xi:=ci:Ti)..(xn:Tn) => T] into the pair of the abstracted context [(xn,None,Tn);...;(xi,Some ci,Ti);...;(x1,None,T1)] and of the inner body [T]. Lets in between are not expanded but turn into local definitions, but n is the actual number of destructurated lambdas. *) let decompose_lam_n_assum n = if n < 0 then user_err (str "decompose_lam_n_assum: integer parameter must be positive"); let rec lamdec_rec l n c = if Int.equal n 0 then l,c else let open Context.Rel.Declaration in match kind c with | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) n c | Cast (c,_,_) -> lamdec_rec l n c | _c -> user_err (str "decompose_lam_n_assum: not enough abstractions") in lamdec_rec Context.Rel.empty n (* Same, counting let-in *) let decompose_lam_n_decls n = if n < 0 then user_err (str "decompose_lam_n_decls: integer parameter must be positive"); let rec lamdec_rec l n c = if Int.equal n 0 then l,c else let open Context.Rel.Declaration in match kind c with | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c | Cast (c,_,_) -> lamdec_rec l n c | _ -> user_err (str "decompose_lam_n_decls: not enough abstractions") in lamdec_rec Context.Rel.empty n let prod_assum t = fst (decompose_prod_assum t) let prod_n_assum n t = fst (decompose_prod_n_assum n t) let strip_prod_assum t = snd (decompose_prod_assum t) let strip_prod t = snd (decompose_prod t) let strip_prod_n n t = snd (decompose_prod_n n t) let lam_assum t = fst (decompose_lam_assum t) let lam_n_assum n t = fst (decompose_lam_n_assum n t) let strip_lam_assum t = snd (decompose_lam_assum t) let strip_lam t = snd (decompose_lam t) let strip_lam_n n t = snd (decompose_lam_n n t) (***************************) (* Arities *) (***************************) (* An "arity" is a term of the form [[x1:T1]...[xn:Tn]s] with [s] a sort. Such a term can canonically be seen as the pair of a context of types and of a sort *) type arity = Constr.rel_context * Sorts.t let destArity = let open Context.Rel.Declaration in let rec prodec_rec l c = match kind c with | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) c | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) c | Cast (c,_,_) -> prodec_rec l c | Sort s -> l,s | _ -> anomaly ~label:"destArity" (Pp.str "not an arity.") in prodec_rec [] let mkArity (sign,s) = it_mkProd_or_LetIn (mkSort s) sign let rec isArity c = match kind c with | Prod (_,_,c) -> isArity c | LetIn (_,b,_,c) -> isArity (subst1 b c) | Cast (c,_,_) -> isArity c | Sort _ -> true | _ -> false (** Kind of type *) (* Experimental, used in Presburger contrib *) type ('constr, 'types) kind_of_type = | SortType of Sorts.t | CastType of 'types * 'types | ProdType of Name.t Context.binder_annot * 'types * 'types | LetInType of Name.t Context.binder_annot * 'constr * 'types * 'types | AtomicType of 'constr * 'constr array let kind_of_type t = match kind t with | Sort s -> SortType s | Cast (c,_,t) -> CastType (c, t) | Prod (na,t,c) -> ProdType (na, t, c) | LetIn (na,b,t,c) -> LetInType (na, b, t, c) | App (c,l) -> AtomicType (c, l) | (Rel _ | Meta _ | Var _ | Evar _ | Const _ | Proj _ | Case _ | Fix _ | CoFix _ | Ind _) -> AtomicType (t,[||]) | (Lambda _ | Construct _ | Int _ | Float _) -> failwith "Not a type" coq-8.11.0/kernel/term.mli0000644000175000017500000002041513612664533015222 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* t2], an alias for [forall (_:t1), t2]. Beware [t_2] is NOT lifted. Eg: in context [A:Prop], [A->A] is built by [(mkArrow (mkRel 1) (mkRel 2))] *) val mkArrow : types -> Sorts.relevance -> types -> constr val mkArrowR : types -> types -> constr (** For an always-relevant domain *) (** Named version of the functions from [Term]. *) val mkNamedLambda : Id.t Context.binder_annot -> types -> constr -> constr val mkNamedLetIn : Id.t Context.binder_annot -> constr -> types -> constr -> constr val mkNamedProd : Id.t Context.binder_annot -> types -> types -> types (** Constructs either [(x:t)c] or [[x=b:t]c] *) val mkProd_or_LetIn : Constr.rel_declaration -> types -> types val mkProd_wo_LetIn : Constr.rel_declaration -> types -> types val mkNamedProd_or_LetIn : Constr.named_declaration -> types -> types val mkNamedProd_wo_LetIn : Constr.named_declaration -> types -> types (** Constructs either [[x:t]c] or [[x=b:t]c] *) val mkLambda_or_LetIn : Constr.rel_declaration -> constr -> constr val mkNamedLambda_or_LetIn : Constr.named_declaration -> constr -> constr (** {5 Other term constructors. } *) (** [applist (f,args)] and its variants work as [mkApp] *) val applist : constr * constr list -> constr val applistc : constr -> constr list -> constr val appvect : constr * constr array -> constr val appvectc : constr -> constr array -> constr (** [prodn n l b] = [forall (x_1:T_1)...(x_n:T_n), b] where [l] is [(x_n,T_n)...(x_1,T_1)...]. *) val prodn : int -> (Name.t Context.binder_annot * constr) list -> constr -> constr (** [compose_prod l b] @return [forall (x_1:T_1)...(x_n:T_n), b] where [l] is [(x_n,T_n)...(x_1,T_1)]. Inverse of [decompose_prod]. *) val compose_prod : (Name.t Context.binder_annot * constr) list -> constr -> constr (** [lamn n l b] @return [fun (x_1:T_1)...(x_n:T_n) => b] where [l] is [(x_n,T_n)...(x_1,T_1)...]. *) val lamn : int -> (Name.t Context.binder_annot * constr) list -> constr -> constr (** [compose_lam l b] @return [fun (x_1:T_1)...(x_n:T_n) => b] where [l] is [(x_n,T_n)...(x_1,T_1)]. Inverse of [it_destLam] *) val compose_lam : (Name.t Context.binder_annot * constr) list -> constr -> constr (** [to_lambda n l] @return [fun (x_1:T_1)...(x_n:T_n) => T] where [l] is [forall (x_1:T_1)...(x_n:T_n), T] *) val to_lambda : int -> constr -> constr (** [to_prod n l] @return [forall (x_1:T_1)...(x_n:T_n), T] where [l] is [fun (x_1:T_1)...(x_n:T_n) => T] *) val to_prod : int -> constr -> constr val it_mkLambda_or_LetIn : constr -> Constr.rel_context -> constr val it_mkProd_or_LetIn : types -> Constr.rel_context -> types (** In [lambda_applist c args], [c] is supposed to have the form [λΓ.c] with [Γ] without let-in; it returns [c] with the variables of [Γ] instantiated by [args]. *) val lambda_applist : constr -> constr list -> constr val lambda_appvect : constr -> constr array -> constr (** In [lambda_applist_assum n c args], [c] is supposed to have the form [λΓ.c] with [Γ] of length [n] and possibly with let-ins; it returns [c] with the assumptions of [Γ] instantiated by [args] and the local definitions of [Γ] expanded. *) val lambda_applist_assum : int -> constr -> constr list -> constr val lambda_appvect_assum : int -> constr -> constr array -> constr (** pseudo-reduction rule *) (** [prod_appvect] [forall (x1:B1;...;xn:Bn), B] [a1...an] @return [B[a1...an]] *) val prod_appvect : types -> constr array -> types val prod_applist : types -> constr list -> types (** In [prod_appvect_assum n c args], [c] is supposed to have the form [∀Γ.c] with [Γ] of length [n] and possibly with let-ins; it returns [c] with the assumptions of [Γ] instantiated by [args] and the local definitions of [Γ] expanded. *) val prod_appvect_assum : int -> types -> constr array -> types val prod_applist_assum : int -> types -> constr list -> types (** {5 Other term destructors. } *) (** Transforms a product term {% $ %}(x_1:T_1)..(x_n:T_n)T{% $ %} into the pair {% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}, where {% $ %}T{% $ %} is not a product. *) val decompose_prod : constr -> (Name.t Context.binder_annot * constr) list * constr (** Transforms a lambda term {% $ %}[x_1:T_1]..[x_n:T_n]T{% $ %} into the pair {% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}, where {% $ %}T{% $ %} is not a lambda. *) val decompose_lam : constr -> (Name.t Context.binder_annot * constr) list * constr (** Given a positive integer n, decompose a product term {% $ %}(x_1:T_1)..(x_n:T_n)T{% $ %} into the pair {% $ %}([(xn,Tn);...;(x1,T1)],T){% $ %}. Raise a user error if not enough products. *) val decompose_prod_n : int -> constr -> (Name.t Context.binder_annot * constr) list * constr (** Given a positive integer {% $ %}n{% $ %}, decompose a lambda term {% $ %}[x_1:T_1]..[x_n:T_n]T{% $ %} into the pair {% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}. Raise a user error if not enough lambdas. *) val decompose_lam_n : int -> constr -> (Name.t Context.binder_annot * constr) list * constr (** Extract the premisses and the conclusion of a term of the form "(xi:Ti) ... (xj:=cj:Tj) ..., T" where T is not a product nor a let *) val decompose_prod_assum : types -> Constr.rel_context * types (** Idem with lambda's and let's *) val decompose_lam_assum : constr -> Constr.rel_context * constr (** Idem but extract the first [n] premisses, counting let-ins. *) val decompose_prod_n_assum : int -> types -> Constr.rel_context * types (** Idem for lambdas, _not_ counting let-ins *) val decompose_lam_n_assum : int -> constr -> Constr.rel_context * constr (** Idem, counting let-ins *) val decompose_lam_n_decls : int -> constr -> Constr.rel_context * constr (** Return the premisses/parameters of a type/term (let-in included) *) val prod_assum : types -> Constr.rel_context val lam_assum : constr -> Constr.rel_context (** Return the first n-th premisses/parameters of a type (let included and counted) *) val prod_n_assum : int -> types -> Constr.rel_context (** Return the first n-th premisses/parameters of a term (let included but not counted) *) val lam_n_assum : int -> constr -> Constr.rel_context (** Remove the premisses/parameters of a type/term *) val strip_prod : types -> types val strip_lam : constr -> constr (** Remove the first n-th premisses/parameters of a type/term *) val strip_prod_n : int -> types -> types val strip_lam_n : int -> constr -> constr (** Remove the premisses/parameters of a type/term (including let-in) *) val strip_prod_assum : types -> types val strip_lam_assum : constr -> constr (** {5 ... } *) (** An "arity" is a term of the form [[x1:T1]...[xn:Tn]s] with [s] a sort. Such a term can canonically be seen as the pair of a context of types and of a sort *) type arity = Constr.rel_context * Sorts.t (** Build an "arity" from its canonical form *) val mkArity : arity -> types (** Destruct an "arity" into its canonical form *) val destArity : types -> arity (** Tell if a term has the form of an arity *) val isArity : types -> bool (** {5 Kind of type} *) type ('constr, 'types) kind_of_type = | SortType of Sorts.t | CastType of 'types * 'types | ProdType of Name.t Context.binder_annot * 'types * 'types | LetInType of Name.t Context.binder_annot * 'constr * 'types * 'types | AtomicType of 'constr * 'constr array val kind_of_type : types -> (constr, types) kind_of_type (* Deprecated *) type sorts_family = Sorts.family = InSProp | InProp | InSet | InType [@@ocaml.deprecated "Alias for Sorts.family"] type sorts = Sorts.t = private | SProp | Prop | Set | Type of Univ.Universe.t (** Type *) [@@ocaml.deprecated "Alias for Sorts.t"] coq-8.11.0/kernel/constr.ml0000644000175000017500000015565113612664533015425 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* = ci_cstr_nargs.(i) *) type case_info = { ci_ind : inductive; (* inductive type to which belongs the value that is being matched *) ci_npar : int; (* number of parameters of the above inductive type *) ci_cstr_ndecls : int array; (* For each constructor, the corresponding integer determines the number of values that can be bound in a match-construct. NOTE: parameters of the inductive type are therefore excluded from the count *) ci_cstr_nargs : int array; (* for each constructor, the corresponding integers determines the number of values that can be applied to the constructor, in addition to the parameters of the related inductive type NOTE: "lets" are therefore excluded from the count NOTE: parameters of the inductive type are also excluded from the count *) ci_relevance : Sorts.relevance; ci_pp_info : case_printing (* not interpreted by the kernel *) } (********************************************************************) (* Constructions as implemented *) (********************************************************************) (* [constr array] is an instance matching definitional [named_context] in the same order (i.e. last argument first) *) type 'constr pexistential = existential_key * 'constr array type ('constr, 'types) prec_declaration = Name.t binder_annot array * 'types array * 'constr array type ('constr, 'types) pfixpoint = (int array * int) * ('constr, 'types) prec_declaration type ('constr, 'types) pcofixpoint = int * ('constr, 'types) prec_declaration type 'a puniverses = 'a Univ.puniverses type pconstant = Constant.t puniverses type pinductive = inductive puniverses type pconstructor = constructor puniverses (* [Var] is used for named variables and [Rel] for variables as de Bruijn indices. *) type ('constr, 'types, 'sort, 'univs) kind_of_term = | Rel of int | Var of Id.t | Meta of metavariable | Evar of 'constr pexistential | Sort of 'sort | Cast of 'constr * cast_kind * 'types | Prod of Name.t binder_annot * 'types * 'types | Lambda of Name.t binder_annot * 'types * 'constr | LetIn of Name.t binder_annot * 'constr * 'types * 'constr | App of 'constr * 'constr array | Const of (Constant.t * 'univs) | Ind of (inductive * 'univs) | Construct of (constructor * 'univs) | Case of case_info * 'constr * 'constr * 'constr array | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint | Proj of Projection.t * 'constr | Int of Uint63.t | Float of Float64.t (* constr is the fixpoint of the previous type. Requires option -rectypes of the Caml compiler to be set *) type t = (t, t, Sorts.t, Instance.t) kind_of_term type constr = t type existential = existential_key * constr array type types = constr type rec_declaration = (constr, types) prec_declaration type fixpoint = (constr, types) pfixpoint type cofixpoint = (constr, types) pcofixpoint (*********************) (* Term constructors *) (*********************) (* Constructs a de Bruijn index with number n *) let rels = [|Rel 1;Rel 2;Rel 3;Rel 4;Rel 5;Rel 6;Rel 7; Rel 8; Rel 9;Rel 10;Rel 11;Rel 12;Rel 13;Rel 14;Rel 15; Rel 16|] let mkRel n = if 0 mkSProp | Sorts.Prop -> mkProp (* Easy sharing *) | Sorts.Set -> mkSet | Sorts.Type _ as s -> Sort s (* Constructs the term t1::t2, i.e. the term t1 casted with the type t2 *) (* (that means t2 is declared as the type of t1) *) let mkCast (t1,k2,t2) = match t1 with | Cast (c,k1, _) when (k1 == VMcast || k1 == NATIVEcast) && k1 == k2 -> Cast (c,k1,t2) | _ -> Cast (t1,k2,t2) (* Constructs the product (x:t1)t2 *) let mkProd (x,t1,t2) = Prod (x,t1,t2) (* Constructs the abstraction [x:t1]t2 *) let mkLambda (x,t1,t2) = Lambda (x,t1,t2) (* Constructs [x=c_1:t]c_2 *) let mkLetIn (x,c1,t,c2) = LetIn (x,c1,t,c2) (* If lt = [t1; ...; tn], constructs the application (t1 ... tn) *) (* We ensure applicative terms have at least one argument and the function is not itself an applicative term *) let mkApp (f, a) = if Int.equal (Array.length a) 0 then f else match f with | App (g, cl) -> App (g, Array.append cl a) | _ -> App (f, a) let map_puniverses f (x,u) = (f x, u) let in_punivs a = (a, Univ.Instance.empty) (* Constructs a constant *) let mkConst c = Const (in_punivs c) let mkConstU c = Const c (* Constructs an applied projection *) let mkProj (p,c) = Proj (p,c) (* Constructs an existential variable *) let mkEvar e = Evar e (* Constructs the ith (co)inductive type of the block named kn *) let mkInd m = Ind (in_punivs m) let mkIndU m = Ind m (* Constructs the jth constructor of the ith (co)inductive type of the block named kn. *) let mkConstruct c = Construct (in_punivs c) let mkConstructU c = Construct c let mkConstructUi ((ind,u),i) = Construct ((ind,i),u) (* Constructs the term

Case c of c1 | c2 .. | cn end *) let mkCase (ci, p, c, ac) = Case (ci, p, c, ac) (* If recindxs = [|i1,...in|] funnames = [|f1,...fn|] typarray = [|t1,...tn|] bodies = [|b1,...bn|] then mkFix ((recindxs,i),(funnames,typarray,bodies)) constructs the ith function of the block Fixpoint f1 [ctx1] : t1 := b1 with f2 [ctx2] : t2 := b2 ... with fn [ctxn] : tn := bn. where the length of the jth context is ij. *) let mkFix fix = Fix fix (* If funnames = [|f1,...fn|] typarray = [|t1,...tn|] bodies = [|b1,...bn|] then mkCoFix (i,(funnames,typsarray,bodies)) constructs the ith function of the block CoFixpoint f1 : t1 := b1 with f2 : t2 := b2 ... with fn : tn := bn. *) let mkCoFix cofix= CoFix cofix (* Constructs an existential variable named "?n" *) let mkMeta n = Meta n (* Constructs a Variable named id *) let mkVar id = Var id let mkRef (gr,u) = let open GlobRef in match gr with | ConstRef c -> mkConstU (c,u) | IndRef ind -> mkIndU (ind,u) | ConstructRef c -> mkConstructU (c,u) | VarRef x -> mkVar x (* Constructs a primitive integer *) let mkInt i = Int i (* Constructs a primitive float number *) let mkFloat f = Float f (************************************************************************) (* kind_of_term = constructions as seen by the user *) (************************************************************************) (* User view of [constr]. For [App], it is ensured there is at least one argument and the function is not itself an applicative term *) let kind c = c let rec kind_nocast_gen kind c = match kind c with | Cast (c, _, _) -> kind_nocast_gen kind c | App (h, outer) as k -> (match kind_nocast_gen kind h with | App (h, inner) -> App (h, Array.append inner outer) | _ -> k) | k -> k let kind_nocast c = kind_nocast_gen kind c (* The other way around. We treat specifically smart constructors *) let of_kind = function | App (f, a) -> mkApp (f, a) | Cast (c, knd, t) -> mkCast (c, knd, t) | k -> k (**********************************************************************) (* Non primitive term destructors *) (**********************************************************************) (* Destructor operations : partial functions Raise [DestKO] if the const has not the expected form *) exception DestKO let isMeta c = match kind c with Meta _ -> true | _ -> false (* Destructs a type *) let isSort c = match kind c with | Sort _ -> true | _ -> false let rec isprop c = match kind c with | Sort (Sorts.Prop | Sorts.Set) -> true | Cast (c,_,_) -> isprop c | _ -> false let rec is_Prop c = match kind c with | Sort Sorts.Prop -> true | Cast (c,_,_) -> is_Prop c | _ -> false let rec is_Set c = match kind c with | Sort Sorts.Set -> true | Cast (c,_,_) -> is_Set c | _ -> false let rec is_Type c = match kind c with | Sort (Sorts.Type _) -> true | Cast (c,_,_) -> is_Type c | _ -> false let is_small = Sorts.is_small let iskind c = isprop c || is_Type c (* Tests if an evar *) let isEvar c = match kind c with Evar _ -> true | _ -> false let isEvar_or_Meta c = match kind c with | Evar _ | Meta _ -> true | _ -> false let isCast c = match kind c with Cast _ -> true | _ -> false (* Tests if a de Bruijn index *) let isRel c = match kind c with Rel _ -> true | _ -> false let isRelN n c = match kind c with Rel n' -> Int.equal n n' | _ -> false (* Tests if a variable *) let isVar c = match kind c with Var _ -> true | _ -> false let isVarId id c = match kind c with Var id' -> Id.equal id id' | _ -> false (* Tests if an inductive *) let isInd c = match kind c with Ind _ -> true | _ -> false let isProd c = match kind c with | Prod _ -> true | _ -> false let isLambda c = match kind c with | Lambda _ -> true | _ -> false let isLetIn c = match kind c with LetIn _ -> true | _ -> false let isApp c = match kind c with App _ -> true | _ -> false let isConst c = match kind c with Const _ -> true | _ -> false let isConstruct c = match kind c with Construct _ -> true | _ -> false let isCase c = match kind c with Case _ -> true | _ -> false let isProj c = match kind c with Proj _ -> true | _ -> false let isFix c = match kind c with Fix _ -> true | _ -> false let isCoFix c = match kind c with CoFix _ -> true | _ -> false (* Destructs a de Bruijn index *) let destRel c = match kind c with | Rel n -> n | _ -> raise DestKO (* Destructs an existential variable *) let destMeta c = match kind c with | Meta n -> n | _ -> raise DestKO (* Destructs a variable *) let destVar c = match kind c with | Var id -> id | _ -> raise DestKO let destSort c = match kind c with | Sort s -> s | _ -> raise DestKO (* Destructs a casted term *) let destCast c = match kind c with | Cast (t1,k,t2) -> (t1,k,t2) | _ -> raise DestKO (* Destructs the product (x:t1)t2 *) let destProd c = match kind c with | Prod (x,t1,t2) -> (x,t1,t2) | _ -> raise DestKO (* Destructs the abstraction [x:t1]t2 *) let destLambda c = match kind c with | Lambda (x,t1,t2) -> (x,t1,t2) | _ -> raise DestKO (* Destructs the let [x:=b:t1]t2 *) let destLetIn c = match kind c with | LetIn (x,b,t1,t2) -> (x,b,t1,t2) | _ -> raise DestKO (* Destructs an application *) let destApp c = match kind c with | App (f,a) -> (f, a) | _ -> raise DestKO (* Destructs a constant *) let destConst c = match kind c with | Const kn -> kn | _ -> raise DestKO (* Destructs an existential variable *) let destEvar c = match kind c with | Evar (_kn, _a as r) -> r | _ -> raise DestKO (* Destructs a (co)inductive type named kn *) let destInd c = match kind c with | Ind (_kn, _a as r) -> r | _ -> raise DestKO (* Destructs a constructor *) let destConstruct c = match kind c with | Construct (_kn, _a as r) -> r | _ -> raise DestKO (* Destructs a term

Case c of lc1 | lc2 .. | lcn end *) let destCase c = match kind c with | Case (ci,p,c,v) -> (ci,p,c,v) | _ -> raise DestKO let destProj c = match kind c with | Proj (p, c) -> (p, c) | _ -> raise DestKO let destFix c = match kind c with | Fix fix -> fix | _ -> raise DestKO let destCoFix c = match kind c with | CoFix cofix -> cofix | _ -> raise DestKO let destRef c = let open GlobRef in match kind c with | Var x -> VarRef x, Univ.Instance.empty | Const (c,u) -> ConstRef c, u | Ind (ind,u) -> IndRef ind, u | Construct (c,u) -> ConstructRef c, u | _ -> raise DestKO (******************************************************************) (* Flattening and unflattening of embedded applications and casts *) (******************************************************************) let decompose_app c = match kind c with | App (f,cl) -> (f, Array.to_list cl) | _ -> (c,[]) let decompose_appvect c = match kind c with | App (f,cl) -> (f, cl) | _ -> (c,[||]) (****************************************************************************) (* Functions to recur through subterms *) (****************************************************************************) (* [fold f acc c] folds [f] on the immediate subterms of [c] starting from [acc] and proceeding from left to right according to the usual representation of the constructions; it is not recursive *) let fold f acc c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> acc | Cast (c,_,t) -> f (f acc c) t | Prod (_,t,c) -> f (f acc t) c | Lambda (_,t,c) -> f (f acc t) c | LetIn (_,b,t,c) -> f (f (f acc b) t) c | App (c,l) -> Array.fold_left f (f acc c) l | Proj (_p,c) -> f acc c | Evar (_,l) -> Array.fold_left f acc l | Case (_,p,c,bl) -> Array.fold_left f (f (f acc p) c) bl | Fix (_,(_lna,tl,bl)) -> Array.fold_left2 (fun acc t b -> f (f acc t) b) acc tl bl | CoFix (_,(_lna,tl,bl)) -> Array.fold_left2 (fun acc t b -> f (f acc t) b) acc tl bl (* [iter f c] iters [f] on the immediate subterms of [c]; it is not recursive and the order with which subterms are processed is not specified *) let iter f c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> () | Cast (c,_,t) -> f c; f t | Prod (_,t,c) -> f t; f c | Lambda (_,t,c) -> f t; f c | LetIn (_,b,t,c) -> f b; f t; f c | App (c,l) -> f c; Array.iter f l | Proj (_p,c) -> f c | Evar (_,l) -> Array.iter f l | Case (_,p,c,bl) -> f p; f c; Array.iter f bl | Fix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl | CoFix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl (* [iter_with_binders g f n c] iters [f n] on the immediate subterms of [c]; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to [n]) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified *) let iter_with_binders g f n c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> () | Cast (c,_,t) -> f n c; f n t | Prod (_,t,c) -> f n t; f (g n) c | Lambda (_,t,c) -> f n t; f (g n) c | LetIn (_,b,t,c) -> f n b; f n t; f (g n) c | App (c,l) -> f n c; Array.Fun1.iter f n l | Evar (_,l) -> Array.Fun1.iter f n l | Case (_,p,c,bl) -> f n p; f n c; Array.Fun1.iter f n bl | Proj (_p,c) -> f n c | Fix (_,(_,tl,bl)) -> Array.Fun1.iter f n tl; Array.Fun1.iter f (iterate g (Array.length tl) n) bl | CoFix (_,(_,tl,bl)) -> Array.Fun1.iter f n tl; Array.Fun1.iter f (iterate g (Array.length tl) n) bl (* [fold_constr_with_binders g f n acc c] folds [f n] on the immediate subterms of [c] starting from [acc] and proceeding from left to right according to the usual representation of the constructions as [fold_constr] but it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to [n]) at each binder traversal; it is not recursive *) let fold_constr_with_binders g f n acc c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> acc | Cast (c,_, t) -> f n (f n acc c) t | Prod (_na,t,c) -> f (g n) (f n acc t) c | Lambda (_na,t,c) -> f (g n) (f n acc t) c | LetIn (_na,b,t,c) -> f (g n) (f n (f n acc b) t) c | App (c,l) -> Array.fold_left (f n) (f n acc c) l | Proj (_p,c) -> f n acc c | Evar (_,l) -> Array.fold_left (f n) acc l | Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl | Fix (_,(_,tl,bl)) -> let n' = iterate g (Array.length tl) n in let fd = Array.map2 (fun t b -> (t,b)) tl bl in Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd | CoFix (_,(_,tl,bl)) -> let n' = iterate g (Array.length tl) n in let fd = Array.map2 (fun t b -> (t,b)) tl bl in Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd (* [map f c] maps [f] on the immediate subterms of [c]; it is not recursive and the order with which subterms are processed is not specified *) let rec map_under_context f n d = if n = 0 then f d else match kind d with | LetIn (na,b,t,c) -> let b' = f b in let t' = f t in let c' = map_under_context f (n-1) c in if b' == b && t' == t && c' == c then d else mkLetIn (na,b',t',c') | Lambda (na,t,b) -> let t' = f t in let b' = map_under_context f (n-1) b in if t' == t && b' == b then d else mkLambda (na,t',b') | _ -> CErrors.anomaly (Pp.str "Ill-formed context") let map_branches f ci bl = let nl = Array.map List.length ci.ci_pp_info.cstr_tags in let bl' = Array.map2 (map_under_context f) nl bl in if Array.for_all2 (==) bl' bl then bl else bl' let map_return_predicate f ci p = map_under_context f (List.length ci.ci_pp_info.ind_tags) p let rec map_under_context_with_binders g f l n d = if n = 0 then f l d else match kind d with | LetIn (na,b,t,c) -> let b' = f l b in let t' = f l t in let c' = map_under_context_with_binders g f (g l) (n-1) c in if b' == b && t' == t && c' == c then d else mkLetIn (na,b',t',c') | Lambda (na,t,b) -> let t' = f l t in let b' = map_under_context_with_binders g f (g l) (n-1) b in if t' == t && b' == b then d else mkLambda (na,t',b') | _ -> CErrors.anomaly (Pp.str "Ill-formed context") let map_branches_with_binders g f l ci bl = let tags = Array.map List.length ci.ci_pp_info.cstr_tags in let bl' = Array.map2 (map_under_context_with_binders g f l) tags bl in if Array.for_all2 (==) bl' bl then bl else bl' let map_return_predicate_with_binders g f l ci p = map_under_context_with_binders g f l (List.length ci.ci_pp_info.ind_tags) p let rec map_under_context_with_full_binders g f l n d = if n = 0 then f l d else match kind d with | LetIn (na,b,t,c) -> let b' = f l b in let t' = f l t in let c' = map_under_context_with_full_binders g f (g (Context.Rel.Declaration.LocalDef (na,b,t)) l) (n-1) c in if b' == b && t' == t && c' == c then d else mkLetIn (na,b',t',c') | Lambda (na,t,b) -> let t' = f l t in let b' = map_under_context_with_full_binders g f (g (Context.Rel.Declaration.LocalAssum (na,t)) l) (n-1) b in if t' == t && b' == b then d else mkLambda (na,t',b') | _ -> CErrors.anomaly (Pp.str "Ill-formed context") let map_branches_with_full_binders g f l ci bl = let tags = Array.map List.length ci.ci_pp_info.cstr_tags in let bl' = Array.map2 (map_under_context_with_full_binders g f l) tags bl in if Array.for_all2 (==) bl' bl then bl else bl' let map_return_predicate_with_full_binders g f l ci p = map_under_context_with_full_binders g f l (List.length ci.ci_pp_info.ind_tags) p let map_gen userview f c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> c | Cast (b,k,t) -> let b' = f b in let t' = f t in if b'==b && t' == t then c else mkCast (b', k, t') | Prod (na,t,b) -> let b' = f b in let t' = f t in if b'==b && t' == t then c else mkProd (na, t', b') | Lambda (na,t,b) -> let b' = f b in let t' = f t in if b'==b && t' == t then c else mkLambda (na, t', b') | LetIn (na,b,t,k) -> let b' = f b in let t' = f t in let k' = f k in if b'==b && t' == t && k'==k then c else mkLetIn (na, b', t', k') | App (b,l) -> let b' = f b in let l' = Array.Smart.map f l in if b'==b && l'==l then c else mkApp (b', l') | Proj (p,t) -> let t' = f t in if t' == t then c else mkProj (p, t') | Evar (e,l) -> let l' = Array.Smart.map f l in if l'==l then c else mkEvar (e, l') | Case (ci,p,b,bl) when userview -> let b' = f b in let p' = map_return_predicate f ci p in let bl' = map_branches f ci bl in if b'==b && p'==p && bl'==bl then c else mkCase (ci, p', b', bl') | Case (ci,p,b,bl) -> let b' = f b in let p' = f p in let bl' = Array.Smart.map f bl in if b'==b && p'==p && bl'==bl then c else mkCase (ci, p', b', bl') | Fix (ln,(lna,tl,bl)) -> let tl' = Array.Smart.map f tl in let bl' = Array.Smart.map f bl in if tl'==tl && bl'==bl then c else mkFix (ln,(lna,tl',bl')) | CoFix(ln,(lna,tl,bl)) -> let tl' = Array.Smart.map f tl in let bl' = Array.Smart.map f bl in if tl'==tl && bl'==bl then c else mkCoFix (ln,(lna,tl',bl')) let map_user_view = map_gen true let map = map_gen false (* Like {!map} but with an accumulator. *) let fold_map f accu c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> accu, c | Cast (b,k,t) -> let accu, b' = f accu b in let accu, t' = f accu t in if b'==b && t' == t then accu, c else accu, mkCast (b', k, t') | Prod (na,t,b) -> let accu, b' = f accu b in let accu, t' = f accu t in if b'==b && t' == t then accu, c else accu, mkProd (na, t', b') | Lambda (na,t,b) -> let accu, b' = f accu b in let accu, t' = f accu t in if b'==b && t' == t then accu, c else accu, mkLambda (na, t', b') | LetIn (na,b,t,k) -> let accu, b' = f accu b in let accu, t' = f accu t in let accu, k' = f accu k in if b'==b && t' == t && k'==k then accu, c else accu, mkLetIn (na, b', t', k') | App (b,l) -> let accu, b' = f accu b in let accu, l' = Array.Smart.fold_left_map f accu l in if b'==b && l'==l then accu, c else accu, mkApp (b', l') | Proj (p,t) -> let accu, t' = f accu t in if t' == t then accu, c else accu, mkProj (p, t') | Evar (e,l) -> let accu, l' = Array.Smart.fold_left_map f accu l in if l'==l then accu, c else accu, mkEvar (e, l') | Case (ci,p,b,bl) -> let accu, b' = f accu b in let accu, p' = f accu p in let accu, bl' = Array.Smart.fold_left_map f accu bl in if b'==b && p'==p && bl'==bl then accu, c else accu, mkCase (ci, p', b', bl') | Fix (ln,(lna,tl,bl)) -> let accu, tl' = Array.Smart.fold_left_map f accu tl in let accu, bl' = Array.Smart.fold_left_map f accu bl in if tl'==tl && bl'==bl then accu, c else accu, mkFix (ln,(lna,tl',bl')) | CoFix(ln,(lna,tl,bl)) -> let accu, tl' = Array.Smart.fold_left_map f accu tl in let accu, bl' = Array.Smart.fold_left_map f accu bl in if tl'==tl && bl'==bl then accu, c else accu, mkCoFix (ln,(lna,tl',bl')) (* [map_with_binders g f n c] maps [f n] on the immediate subterms of [c]; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to [n]) at each binder traversal; it is not recursive and the order with which subterms are processed is not specified *) let map_with_binders g f l c0 = match kind c0 with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> c0 | Cast (c, k, t) -> let c' = f l c in let t' = f l t in if c' == c && t' == t then c0 else mkCast (c', k, t') | Prod (na, t, c) -> let t' = f l t in let c' = f (g l) c in if t' == t && c' == c then c0 else mkProd (na, t', c') | Lambda (na, t, c) -> let t' = f l t in let c' = f (g l) c in if t' == t && c' == c then c0 else mkLambda (na, t', c') | LetIn (na, b, t, c) -> let b' = f l b in let t' = f l t in let c' = f (g l) c in if b' == b && t' == t && c' == c then c0 else mkLetIn (na, b', t', c') | App (c, al) -> let c' = f l c in let al' = Array.Fun1.Smart.map f l al in if c' == c && al' == al then c0 else mkApp (c', al') | Proj (p, t) -> let t' = f l t in if t' == t then c0 else mkProj (p, t') | Evar (e, al) -> let al' = Array.Fun1.Smart.map f l al in if al' == al then c0 else mkEvar (e, al') | Case (ci, p, c, bl) -> let p' = f l p in let c' = f l c in let bl' = Array.Fun1.Smart.map f l bl in if p' == p && c' == c && bl' == bl then c0 else mkCase (ci, p', c', bl') | Fix (ln, (lna, tl, bl)) -> let tl' = Array.Fun1.Smart.map f l tl in let l' = iterate g (Array.length tl) l in let bl' = Array.Fun1.Smart.map f l' bl in if tl' == tl && bl' == bl then c0 else mkFix (ln,(lna,tl',bl')) | CoFix(ln,(lna,tl,bl)) -> let tl' = Array.Fun1.Smart.map f l tl in let l' = iterate g (Array.length tl) l in let bl' = Array.Fun1.Smart.map f l' bl in mkCoFix (ln,(lna,tl',bl')) (*********************) (* Lifting *) (*********************) (* The generic lifting function *) let rec exliftn el c = let open Esubst in match kind c with | Rel i -> mkRel(reloc_rel i el) | _ -> map_with_binders el_lift exliftn el c (* Lifting the binding depth across k bindings *) let liftn n k c = let open Esubst in match el_liftn (pred k) (el_shft n el_id) with | ELID -> c | el -> exliftn el c let lift n = liftn n 1 let fold_with_full_binders g f n acc c = let open Context.Rel.Declaration in match kind c with | Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _ -> acc | Cast (c,_, t) -> f n (f n acc c) t | Prod (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c | App (c,l) -> Array.fold_left (f n) (f n acc c) l | Proj (_,c) -> f n acc c | Evar (_,l) -> Array.fold_left (f n) acc l | Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl | Fix (_,(lna,tl,bl)) -> let n' = CArray.fold_left2_i (fun i c n t -> g (LocalAssum (n,lift i t)) c) n lna tl in let fd = Array.map2 (fun t b -> (t,b)) tl bl in Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd | CoFix (_,(lna,tl,bl)) -> let n' = CArray.fold_left2_i (fun i c n t -> g (LocalAssum (n,lift i t)) c) n lna tl in let fd = Array.map2 (fun t b -> (t,b)) tl bl in Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd type 'univs instance_compare_fn = GlobRef.t -> int -> 'univs -> 'univs -> bool type 'constr constr_compare_fn = int -> 'constr -> 'constr -> bool (* [compare_head_gen_evar k1 k2 u s e eq leq c1 c2] compare [c1] and [c2] (using [k1] to expose the structure of [c1] and [k2] to expose the structure [c2]) using [eq] to compare the immediate subterms of [c1] of [c2] for conversion if needed, [leq] for cumulativity, [u] to compare universe instances, and [s] to compare sorts; Cast's, application associativity, binders name and Cases annotations are not taken into account. Note that as [kind1] and [kind2] are potentially different, we cannot use, in recursive case, the optimisation that physically equal arrays are equals (hence the calls to {!Array.equal_norefl}). *) let compare_head_gen_leq_with kind1 kind2 leq_universes leq_sorts eq leq nargs t1 t2 = match kind_nocast_gen kind1 t1, kind_nocast_gen kind2 t2 with | Cast _, _ | _, Cast _ -> assert false (* kind_nocast *) | Rel n1, Rel n2 -> Int.equal n1 n2 | Meta m1, Meta m2 -> Int.equal m1 m2 | Var id1, Var id2 -> Id.equal id1 id2 | Int i1, Int i2 -> Uint63.equal i1 i2 | Float f1, Float f2 -> Float64.equal f1 f2 | Sort s1, Sort s2 -> leq_sorts s1 s2 | Prod (_,t1,c1), Prod (_,t2,c2) -> eq 0 t1 t2 && leq 0 c1 c2 | Lambda (_,t1,c1), Lambda (_,t2,c2) -> eq 0 t1 t2 && eq 0 c1 c2 | LetIn (_,b1,t1,c1), LetIn (_,b2,t2,c2) -> eq 0 b1 b2 && eq 0 t1 t2 && leq nargs c1 c2 (* Why do we suddenly make a special case for Cast here? *) | App (c1, l1), App (c2, l2) -> let len = Array.length l1 in Int.equal len (Array.length l2) && leq (nargs+len) c1 c2 && Array.equal_norefl (eq 0) l1 l2 | Proj (p1,c1), Proj (p2,c2) -> Projection.equal p1 p2 && eq 0 c1 c2 | Evar (e1,l1), Evar (e2,l2) -> Evar.equal e1 e2 && Array.equal (eq 0) l1 l2 | Const (c1,u1), Const (c2,u2) -> (* The args length currently isn't used but may as well pass it. *) Constant.equal c1 c2 && leq_universes (GlobRef.ConstRef c1) nargs u1 u2 | Ind (c1,u1), Ind (c2,u2) -> eq_ind c1 c2 && leq_universes (GlobRef.IndRef c1) nargs u1 u2 | Construct (c1,u1), Construct (c2,u2) -> eq_constructor c1 c2 && leq_universes (GlobRef.ConstructRef c1) nargs u1 u2 | Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) -> eq 0 p1 p2 && eq 0 c1 c2 && Array.equal (eq 0) bl1 bl2 | Fix ((ln1, i1),(_,tl1,bl1)), Fix ((ln2, i2),(_,tl2,bl2)) -> Int.equal i1 i2 && Array.equal Int.equal ln1 ln2 && Array.equal_norefl (eq 0) tl1 tl2 && Array.equal_norefl (eq 0) bl1 bl2 | CoFix(ln1,(_,tl1,bl1)), CoFix(ln2,(_,tl2,bl2)) -> Int.equal ln1 ln2 && Array.equal_norefl (eq 0) tl1 tl2 && Array.equal_norefl (eq 0) bl1 bl2 | (Rel _ | Meta _ | Var _ | Sort _ | Prod _ | Lambda _ | LetIn _ | App _ | Proj _ | Evar _ | Const _ | Ind _ | Construct _ | Case _ | Fix _ | CoFix _ | Int _ | Float _), _ -> false (* [compare_head_gen_leq u s eq leq c1 c2] compare [c1] and [c2] using [eq] to compare the immediate subterms of [c1] of [c2] for conversion if needed, [leq] for cumulativity, [u] to compare universe instances and [s] to compare sorts; Cast's, application associativity, binders name and Cases annotations are not taken into account *) let compare_head_gen_leq leq_universes leq_sorts eq leq t1 t2 = compare_head_gen_leq_with kind kind leq_universes leq_sorts eq leq t1 t2 (* [compare_head_gen u s f c1 c2] compare [c1] and [c2] using [f] to compare the immediate subterms of [c1] of [c2] if needed, [u] to compare universe instances and [s] to compare sorts; Cast's, application associativity, binders name and Cases annotations are not taken into account. [compare_head_gen_with] is a variant taking kind-of-term functions, to expose subterms of [c1] and [c2], as arguments. *) let compare_head_gen_with kind1 kind2 eq_universes eq_sorts eq t1 t2 = compare_head_gen_leq_with kind1 kind2 eq_universes eq_sorts eq eq t1 t2 let compare_head_gen eq_universes eq_sorts eq t1 t2 = compare_head_gen_leq eq_universes eq_sorts eq eq t1 t2 let compare_head = compare_head_gen (fun _ _ -> Univ.Instance.equal) Sorts.equal (*******************************) (* alpha conversion functions *) (*******************************) (* alpha conversion : ignore print names and casts *) let rec eq_constr nargs m n = (m == n) || compare_head_gen (fun _ _ -> Instance.equal) Sorts.equal eq_constr nargs m n let equal n m = eq_constr 0 m n (* to avoid tracing a recursive fun *) let eq_constr_univs univs m n = if m == n then true else let eq_universes _ _ = UGraph.check_eq_instances univs in let eq_sorts s1 s2 = s1 == s2 || UGraph.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in let rec eq_constr' nargs m n = m == n || compare_head_gen eq_universes eq_sorts eq_constr' nargs m n in compare_head_gen eq_universes eq_sorts eq_constr' 0 m n let leq_constr_univs univs m n = if m == n then true else let eq_universes _ _ = UGraph.check_eq_instances univs in let eq_sorts s1 s2 = s1 == s2 || UGraph.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in let leq_sorts s1 s2 = s1 == s2 || UGraph.check_leq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in let rec eq_constr' nargs m n = m == n || compare_head_gen eq_universes eq_sorts eq_constr' nargs m n in let rec compare_leq nargs m n = compare_head_gen_leq eq_universes leq_sorts eq_constr' leq_constr' nargs m n and leq_constr' nargs m n = m == n || compare_leq nargs m n in compare_leq 0 m n let eq_constr_univs_infer univs m n = if m == n then true, Constraint.empty else let cstrs = ref Constraint.empty in let eq_universes _ _ = UGraph.check_eq_instances univs in let eq_sorts s1 s2 = if Sorts.equal s1 s2 then true else let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in if UGraph.check_eq univs u1 u2 then true else (cstrs := Univ.enforce_eq u1 u2 !cstrs; true) in let rec eq_constr' nargs m n = m == n || compare_head_gen eq_universes eq_sorts eq_constr' nargs m n in let res = compare_head_gen eq_universes eq_sorts eq_constr' 0 m n in res, !cstrs let leq_constr_univs_infer univs m n = if m == n then true, Constraint.empty else let cstrs = ref Constraint.empty in let eq_universes _ _ l l' = UGraph.check_eq_instances univs l l' in let eq_sorts s1 s2 = if Sorts.equal s1 s2 then true else let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in if UGraph.check_eq univs u1 u2 then true else (cstrs := Univ.enforce_eq u1 u2 !cstrs; true) in let leq_sorts s1 s2 = if Sorts.equal s1 s2 then true else let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in if UGraph.check_leq univs u1 u2 then true else (try let c, _ = UGraph.enforce_leq_alg u1 u2 univs in cstrs := Univ.Constraint.union c !cstrs; true with Univ.UniverseInconsistency _ -> false) in let rec eq_constr' nargs m n = m == n || compare_head_gen eq_universes eq_sorts eq_constr' nargs m n in let rec compare_leq nargs m n = compare_head_gen_leq eq_universes leq_sorts eq_constr' leq_constr' nargs m n and leq_constr' nargs m n = m == n || compare_leq nargs m n in let res = compare_leq 0 m n in res, !cstrs let rec eq_constr_nounivs m n = (m == n) || compare_head_gen (fun _ _ _ _ -> true) (fun _ _ -> true) (fun _ -> eq_constr_nounivs) 0 m n let constr_ord_int f t1 t2 = let (=?) f g i1 i2 j1 j2= let c = f i1 i2 in if Int.equal c 0 then g j1 j2 else c in let (==?) fg h i1 i2 j1 j2 k1 k2= let c=fg i1 i2 j1 j2 in if Int.equal c 0 then h k1 k2 else c in let fix_cmp (a1, i1) (a2, i2) = ((Array.compare Int.compare) =? Int.compare) a1 a2 i1 i2 in match kind t1, kind t2 with | Cast (c1,_,_), _ -> f c1 t2 | _, Cast (c2,_,_) -> f t1 c2 (* Why this special case? *) | App (Cast(c1,_,_),l1), _ -> f (mkApp (c1,l1)) t2 | _, App (Cast(c2, _,_),l2) -> f t1 (mkApp (c2,l2)) | Rel n1, Rel n2 -> Int.compare n1 n2 | Rel _, _ -> -1 | _, Rel _ -> 1 | Var id1, Var id2 -> Id.compare id1 id2 | Var _, _ -> -1 | _, Var _ -> 1 | Meta m1, Meta m2 -> Int.compare m1 m2 | Meta _, _ -> -1 | _, Meta _ -> 1 | Evar (e1,l1), Evar (e2,l2) -> (Evar.compare =? (Array.compare f)) e1 e2 l1 l2 | Evar _, _ -> -1 | _, Evar _ -> 1 | Sort s1, Sort s2 -> Sorts.compare s1 s2 | Sort _, _ -> -1 | _, Sort _ -> 1 | Prod (_,t1,c1), Prod (_,t2,c2) | Lambda (_,t1,c1), Lambda (_,t2,c2) -> (f =? f) t1 t2 c1 c2 | Prod _, _ -> -1 | _, Prod _ -> 1 | Lambda _, _ -> -1 | _, Lambda _ -> 1 | LetIn (_,b1,t1,c1), LetIn (_,b2,t2,c2) -> ((f =? f) ==? f) b1 b2 t1 t2 c1 c2 | LetIn _, _ -> -1 | _, LetIn _ -> 1 | App (c1,l1), App (c2,l2) -> (f =? (Array.compare f)) c1 c2 l1 l2 | App _, _ -> -1 | _, App _ -> 1 | Const (c1,_u1), Const (c2,_u2) -> Constant.CanOrd.compare c1 c2 | Const _, _ -> -1 | _, Const _ -> 1 | Ind (ind1, _u1), Ind (ind2, _u2) -> ind_ord ind1 ind2 | Ind _, _ -> -1 | _, Ind _ -> 1 | Construct (ct1,_u1), Construct (ct2,_u2) -> constructor_ord ct1 ct2 | Construct _, _ -> -1 | _, Construct _ -> 1 | Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) -> ((f =? f) ==? (Array.compare f)) p1 p2 c1 c2 bl1 bl2 | Case _, _ -> -1 | _, Case _ -> 1 | Fix (ln1,(_,tl1,bl1)), Fix (ln2,(_,tl2,bl2)) -> ((fix_cmp =? (Array.compare f)) ==? (Array.compare f)) ln1 ln2 tl1 tl2 bl1 bl2 | Fix _, _ -> -1 | _, Fix _ -> 1 | CoFix(ln1,(_,tl1,bl1)), CoFix(ln2,(_,tl2,bl2)) -> ((Int.compare =? (Array.compare f)) ==? (Array.compare f)) ln1 ln2 tl1 tl2 bl1 bl2 | CoFix _, _ -> -1 | _, CoFix _ -> 1 | Proj (p1,c1), Proj (p2,c2) -> (Projection.compare =? f) p1 p2 c1 c2 | Proj _, _ -> -1 | _, Proj _ -> 1 | Int i1, Int i2 -> Uint63.compare i1 i2 | Int _, _ -> -1 | _, Int _ -> 1 | Float f1, Float f2 -> Float64.total_compare f1 f2 let rec compare m n= constr_ord_int compare m n (*******************) (* hash-consing *) (*******************) (* Hash-consing of [constr] does not use the module [Hashcons] because [Hashcons] is not efficient on deep tree-like data structures. Indeed, [Hashcons] is based the (very efficient) generic hash function [Hashtbl.hash], which computes the hash key through a depth bounded traversal of the data structure to be hashed. As a consequence, for a deep [constr] like the natural number 1000 (S (S (... (S O)))), the same hash is assigned to all the sub [constr]s greater than the maximal depth handled by [Hashtbl.hash]. This entails a huge number of collisions in the hash table and leads to cubic hash-consing in this worst-case. In order to compute a hash key that is independent of the data structure depth while being constant-time, an incremental hashing function must be devised. A standard implementation creates a cache of the hashing function by decorating each node of the hash-consed data structure with its hash key. In that case, the hash function can deduce the hash key of a toplevel data structure by a local computation based on the cache held on its substructures. Unfortunately, this simple implementation introduces a space overhead that is damageable for the hash-consing of small [constr]s (the most common case). One can think of an heterogeneous distribution of caches on smartly chosen nodes, but this is forbidden by the use of generic equality in Coq source code. (Indeed, this forces each [constr] to have a unique canonical representation.) Given that hash-consing proceeds inductively, we can nonetheless computes the hash key incrementally during hash-consing by changing a little the signature of the hash-consing function: it now returns both the hash-consed term and its hash key. This simple solution is implemented in the following code: it does not introduce a space overhead in [constr], that's why the efficiency is unchanged for small [constr]s. Besides, it does handle deep [constr]s without introducing an unreasonable number of collisions in the hash table. Some benchmarks make us think that this implementation of hash-consing is linear in the size of the hash-consed data structure for our daily use of Coq. *) let array_eqeq t1 t2 = t1 == t2 || (Int.equal (Array.length t1) (Array.length t2) && let rec aux i = (Int.equal i (Array.length t1)) || (t1.(i) == t2.(i) && aux (i + 1)) in aux 0) let hasheq t1 t2 = match t1, t2 with | Rel n1, Rel n2 -> n1 == n2 | Meta m1, Meta m2 -> m1 == m2 | Var id1, Var id2 -> id1 == id2 | Sort s1, Sort s2 -> s1 == s2 | Cast (c1,k1,t1), Cast (c2,k2,t2) -> c1 == c2 && k1 == k2 && t1 == t2 | Prod (n1,t1,c1), Prod (n2,t2,c2) -> n1 == n2 && t1 == t2 && c1 == c2 | Lambda (n1,t1,c1), Lambda (n2,t2,c2) -> n1 == n2 && t1 == t2 && c1 == c2 | LetIn (n1,b1,t1,c1), LetIn (n2,b2,t2,c2) -> n1 == n2 && b1 == b2 && t1 == t2 && c1 == c2 | App (c1,l1), App (c2,l2) -> c1 == c2 && array_eqeq l1 l2 | Proj (p1,c1), Proj(p2,c2) -> p1 == p2 && c1 == c2 | Evar (e1,l1), Evar (e2,l2) -> e1 == e2 && array_eqeq l1 l2 | Const (c1,u1), Const (c2,u2) -> c1 == c2 && u1 == u2 | Ind (ind1,u1), Ind (ind2,u2) -> ind1 == ind2 && u1 == u2 | Construct (cstr1,u1), Construct (cstr2,u2) -> cstr1 == cstr2 && u1 == u2 | Case (ci1,p1,c1,bl1), Case (ci2,p2,c2,bl2) -> ci1 == ci2 && p1 == p2 && c1 == c2 && array_eqeq bl1 bl2 | Fix ((ln1, i1),(lna1,tl1,bl1)), Fix ((ln2, i2),(lna2,tl2,bl2)) -> Int.equal i1 i2 && Array.equal Int.equal ln1 ln2 && array_eqeq lna1 lna2 && array_eqeq tl1 tl2 && array_eqeq bl1 bl2 | CoFix(ln1,(lna1,tl1,bl1)), CoFix(ln2,(lna2,tl2,bl2)) -> Int.equal ln1 ln2 && array_eqeq lna1 lna2 && array_eqeq tl1 tl2 && array_eqeq bl1 bl2 | Int i1, Int i2 -> i1 == i2 | Float f1, Float f2 -> Float64.equal f1 f2 | (Rel _ | Meta _ | Var _ | Sort _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Proj _ | Evar _ | Const _ | Ind _ | Construct _ | Case _ | Fix _ | CoFix _ | Int _ | Float _), _ -> false (** Note that the following Make has the side effect of creating once and for all the table we'll use for hash-consing all constr *) module HashsetTerm = Hashset.Make(struct type t = constr let eq = hasheq end) module HashsetTermArray = Hashset.Make(struct type t = constr array let eq = array_eqeq end) let term_table = HashsetTerm.create 19991 (* The associative table to hashcons terms. *) let term_array_table = HashsetTermArray.create 4999 (* The associative table to hashcons term arrays. *) open Hashset.Combine let hash_cast_kind = function | VMcast -> 0 | NATIVEcast -> 1 | DEFAULTcast -> 2 | REVERTcast -> 3 let sh_instance = Univ.Instance.share (* [hashcons hash_consing_functions constr] computes an hash-consed representation for [constr] using [hash_consing_functions] on leaves. *) let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) = let rec hash_term t = match t with | Var i -> (Var (sh_id i), combinesmall 1 (Id.hash i)) | Sort s -> (Sort (sh_sort s), combinesmall 2 (Sorts.hash s)) | Cast (c, k, t) -> let c, hc = sh_rec c in let t, ht = sh_rec t in (Cast (c, k, t), combinesmall 3 (combine3 hc (hash_cast_kind k) ht)) | Prod (na,t,c) -> let t, ht = sh_rec t and c, hc = sh_rec c in (Prod (sh_na na, t, c), combinesmall 4 (combine3 (hash_annot Name.hash na) ht hc)) | Lambda (na,t,c) -> let t, ht = sh_rec t and c, hc = sh_rec c in (Lambda (sh_na na, t, c), combinesmall 5 (combine3 (hash_annot Name.hash na) ht hc)) | LetIn (na,b,t,c) -> let b, hb = sh_rec b in let t, ht = sh_rec t in let c, hc = sh_rec c in (LetIn (sh_na na, b, t, c), combinesmall 6 (combine4 (hash_annot Name.hash na) hb ht hc)) | App (c,l) -> let c, hc = sh_rec c in let l, hl = hash_term_array l in (App (c,l), combinesmall 7 (combine hl hc)) | Evar (e,l) -> let l, hl = hash_term_array l in (Evar (e,l), combinesmall 8 (combine (Evar.hash e) hl)) | Const (c,u) -> let c' = sh_con c in let u', hu = sh_instance u in (Const (c', u'), combinesmall 9 (combine (Constant.SyntacticOrd.hash c) hu)) | Ind (ind,u) -> let u', hu = sh_instance u in (Ind (sh_ind ind, u'), combinesmall 10 (combine (ind_syntactic_hash ind) hu)) | Construct (c,u) -> let u', hu = sh_instance u in (Construct (sh_construct c, u'), combinesmall 11 (combine (constructor_syntactic_hash c) hu)) | Case (ci,p,c,bl) -> let p, hp = sh_rec p and c, hc = sh_rec c in let bl,hbl = hash_term_array bl in let hbl = combine (combine hc hp) hbl in (Case (sh_ci ci, p, c, bl), combinesmall 12 hbl) | Fix (ln,(lna,tl,bl)) -> let bl,hbl = hash_term_array bl in let tl,htl = hash_term_array tl in let () = Array.iteri (fun i x -> Array.unsafe_set lna i (sh_na x)) lna in let fold accu na = combine (hash_annot Name.hash na) accu in let hna = Array.fold_left fold 0 lna in let h = combine3 hna hbl htl in (Fix (ln,(lna,tl,bl)), combinesmall 13 h) | CoFix(ln,(lna,tl,bl)) -> let bl,hbl = hash_term_array bl in let tl,htl = hash_term_array tl in let () = Array.iteri (fun i x -> Array.unsafe_set lna i (sh_na x)) lna in let fold accu na = combine (hash_annot Name.hash na) accu in let hna = Array.fold_left fold 0 lna in let h = combine3 hna hbl htl in (CoFix (ln,(lna,tl,bl)), combinesmall 14 h) | Meta n -> (t, combinesmall 15 n) | Rel n -> (t, combinesmall 16 n) | Proj (p,c) -> let c, hc = sh_rec c in let p' = Projection.hcons p in (Proj (p', c), combinesmall 17 (combine (Projection.SyntacticOrd.hash p') hc)) | Int i -> let (h,l) = Uint63.to_int2 i in (t, combinesmall 18 (combine h l)) | Float f -> (t, combinesmall 19 (Float64.hash f)) and sh_rec t = let (y, h) = hash_term t in (* [h] must be positive. *) let h = h land 0x3FFFFFFF in (HashsetTerm.repr h y term_table, h) (* Note : During hash-cons of arrays, we modify them *in place* *) and hash_term_array t = let accu = ref 0 in for i = 0 to Array.length t - 1 do let x, h = sh_rec (Array.unsafe_get t i) in accu := combine !accu h; Array.unsafe_set t i x done; (* [h] must be positive. *) let h = !accu land 0x3FFFFFFF in (HashsetTermArray.repr h t term_array_table, h) in (* Make sure our statically allocated Rels (1 to 16) are considered as canonical, and hence hash-consed to themselves *) ignore (hash_term_array rels); fun t -> fst (sh_rec t) (* Exported hashing fonction on constr, used mainly in plugins. Appears to have slight differences from [snd (hash_term t)] above ? *) let rec hash t = match kind t with | Var i -> combinesmall 1 (Id.hash i) | Sort s -> combinesmall 2 (Sorts.hash s) | Cast (c, k, t) -> let hc = hash c in let ht = hash t in combinesmall 3 (combine3 hc (hash_cast_kind k) ht) | Prod (_, t, c) -> combinesmall 4 (combine (hash t) (hash c)) | Lambda (_, t, c) -> combinesmall 5 (combine (hash t) (hash c)) | LetIn (_, b, t, c) -> combinesmall 6 (combine3 (hash b) (hash t) (hash c)) | App (Cast(c, _, _),l) -> hash (mkApp (c,l)) | App (c,l) -> combinesmall 7 (combine (hash_term_array l) (hash c)) | Evar (e,l) -> combinesmall 8 (combine (Evar.hash e) (hash_term_array l)) | Const (c,u) -> combinesmall 9 (combine (Constant.hash c) (Instance.hash u)) | Ind (ind,u) -> combinesmall 10 (combine (ind_hash ind) (Instance.hash u)) | Construct (c,u) -> combinesmall 11 (combine (constructor_hash c) (Instance.hash u)) | Case (_ , p, c, bl) -> combinesmall 12 (combine3 (hash c) (hash p) (hash_term_array bl)) | Fix (_ln ,(_, tl, bl)) -> combinesmall 13 (combine (hash_term_array bl) (hash_term_array tl)) | CoFix(_ln, (_, tl, bl)) -> combinesmall 14 (combine (hash_term_array bl) (hash_term_array tl)) | Meta n -> combinesmall 15 n | Rel n -> combinesmall 16 n | Proj (p,c) -> combinesmall 17 (combine (Projection.hash p) (hash c)) | Int i -> combinesmall 18 (Uint63.hash i) | Float f -> combinesmall 19 (Float64.hash f) and hash_term_array t = Array.fold_left (fun acc t -> combine (hash t) acc) 0 t module CaseinfoHash = struct type t = case_info type u = inductive -> inductive let hashcons hind ci = { ci with ci_ind = hind ci.ci_ind } let pp_info_equal info1 info2 = List.equal (==) info1.ind_tags info2.ind_tags && Array.equal (List.equal (==)) info1.cstr_tags info2.cstr_tags && info1.style == info2.style let eq ci ci' = ci.ci_ind == ci'.ci_ind && ci.ci_relevance == ci'.ci_relevance && Int.equal ci.ci_npar ci'.ci_npar && Array.equal Int.equal ci.ci_cstr_ndecls ci'.ci_cstr_ndecls && (* we use [Array.equal] on purpose *) Array.equal Int.equal ci.ci_cstr_nargs ci'.ci_cstr_nargs && (* we use [Array.equal] on purpose *) pp_info_equal ci.ci_pp_info ci'.ci_pp_info (* we use (=) on purpose *) open Hashset.Combine let hash_bool b = if b then 0 else 1 let hash_bool_list = List.fold_left (fun n b -> combine n (hash_bool b)) let hash_pp_info info = let h1 = match info.style with | LetStyle -> 0 | IfStyle -> 1 | LetPatternStyle -> 2 | MatchStyle -> 3 | RegularStyle -> 4 in let h2 = hash_bool_list 0 info.ind_tags in let h3 = Array.fold_left hash_bool_list 0 info.cstr_tags in combine3 h1 h2 h3 let hash ci = let h1 = ind_hash ci.ci_ind in let h2 = Int.hash ci.ci_npar in let h3 = Array.fold_left combine 0 ci.ci_cstr_ndecls in let h4 = Array.fold_left combine 0 ci.ci_cstr_nargs in let h5 = hash_pp_info ci.ci_pp_info in combinesmall (Sorts.relevance_hash ci.ci_relevance) (combine5 h1 h2 h3 h4 h5) end module Hcaseinfo = Hashcons.Make(CaseinfoHash) let case_info_hash = CaseinfoHash.hash let hcons_caseinfo = Hashcons.simple_hcons Hcaseinfo.generate Hcaseinfo.hcons hcons_ind module Hannotinfo = struct type t = Name.t binder_annot type u = Name.t -> Name.t let hash = hash_annot Name.hash let eq = eq_annot (fun na1 na2 -> na1 == na2) let hashcons h {binder_name=na;binder_relevance} = {binder_name=h na;binder_relevance} end module Hannot = Hashcons.Make(Hannotinfo) let hcons_annot = Hashcons.simple_hcons Hannot.generate Hannot.hcons Name.hcons let hcons = hashcons (Sorts.hcons, hcons_caseinfo, hcons_construct, hcons_ind, hcons_con, hcons_annot, Id.hcons) (* let hcons_types = hcons_constr *) type rel_declaration = (constr, types) Context.Rel.Declaration.pt type named_declaration = (constr, types) Context.Named.Declaration.pt type compacted_declaration = (constr, types) Context.Compacted.Declaration.pt type rel_context = rel_declaration list type named_context = named_declaration list type compacted_context = compacted_declaration list (** Minimalistic constr printer, typically for debugging *) let debug_print_fix pr_constr ((t,i),(lna,tl,bl)) = let open Pp in let fixl = Array.mapi (fun i na -> (na.binder_name,t.(i),tl.(i),bl.(i))) lna in hov 1 (str"fix " ++ int i ++ spc() ++ str"{" ++ v 0 (prlist_with_sep spc (fun (na,i,ty,bd) -> Name.print na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++ cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++ str"}") let pr_puniverses p u = if Univ.Instance.is_empty u then p else Pp.(p ++ str"(*" ++ Univ.Instance.pr Univ.Level.pr u ++ str"*)") let rec debug_print c = let open Pp in match kind c with | Rel n -> str "#"++int n | Meta n -> str "Meta(" ++ int n ++ str ")" | Var id -> Id.print id | Sort s -> Sorts.debug_print s | Cast (c,_, t) -> hov 1 (str"(" ++ debug_print c ++ cut() ++ str":" ++ debug_print t ++ str")") | Prod ({binder_name=Name id;_},t,c) -> hov 1 (str"forall " ++ Id.print id ++ str":" ++ debug_print t ++ str"," ++ spc() ++ debug_print c) | Prod ({binder_name=Anonymous;_},t,c) -> hov 0 (str"(" ++ debug_print t ++ str " ->" ++ spc() ++ debug_print c ++ str")") | Lambda (na,t,c) -> hov 1 (str"fun " ++ Name.print na.binder_name ++ str":" ++ debug_print t ++ str" =>" ++ spc() ++ debug_print c) | LetIn (na,b,t,c) -> hov 0 (str"let " ++ Name.print na.binder_name ++ str":=" ++ debug_print b ++ str":" ++ brk(1,2) ++ debug_print t ++ cut() ++ debug_print c) | App (c,l) -> hov 1 (str"(" ++ debug_print c ++ spc() ++ prlist_with_sep spc debug_print (Array.to_list l) ++ str")") | Evar (e,l) -> hov 1 (str"Evar#" ++ int (Evar.repr e) ++ str"{" ++ prlist_with_sep spc debug_print (Array.to_list l) ++str"}") | Const (c,u) -> str"Cst(" ++ pr_puniverses (Constant.debug_print c) u ++ str")" | Ind ((sp,i),u) -> str"Ind(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i) u ++ str")" | Construct (((sp,i),j),u) -> str"Constr(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i ++ str"," ++ int j) u ++ str")" | Proj (p,c) -> str"Proj(" ++ Constant.debug_print (Projection.constant p) ++ str"," ++ bool (Projection.unfolded p) ++ debug_print c ++ str")" | Case (_ci,p,c,bl) -> v 0 (hv 0 (str"<"++debug_print p++str">"++ cut() ++ str"Case " ++ debug_print c ++ str"of") ++ cut() ++ prlist_with_sep (fun _ -> brk(1,2)) debug_print (Array.to_list bl) ++ cut() ++ str"end") | Fix f -> debug_print_fix debug_print f | CoFix(i,(lna,tl,bl)) -> let fixl = Array.mapi (fun i na -> (na,tl.(i),bl.(i))) lna in hov 1 (str"cofix " ++ int i ++ spc() ++ str"{" ++ v 0 (prlist_with_sep spc (fun (na,ty,bd) -> Name.print na.binder_name ++ str":" ++ debug_print ty ++ cut() ++ str":=" ++ debug_print bd) (Array.to_list fixl)) ++ str"}") | Int i -> str"Int("++str (Uint63.to_string i) ++ str")" | Float i -> str"Float("++str (Float64.to_string i) ++ str")" coq-8.11.0/kernel/vmvalues.mli0000644000175000017500000001163513612664533016121 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Pp.t type reloc_table = (tag * int) array type annot_switch = {ci : case_info; rtbl : reloc_table; tailcall : bool; max_stack_size : int} val eq_structured_constant : structured_constant -> structured_constant -> bool val hash_structured_constant : structured_constant -> int val eq_annot_switch : annot_switch -> annot_switch -> bool val hash_annot_switch : annot_switch -> int val fun_val : vfun -> values val fix_val : vfix -> values val cofix_upd_val : to_update -> values val fun_env : vfun -> vm_env val fix_env : vfix -> vm_env val cofix_env : vcofix -> vm_env val cofix_upd_env : to_update -> vm_env val vm_global : values array -> vm_global (** Cast a value known to be a function, unsafe in general *) val fun_of_val : values -> vfun val crazy_val : values (** Machine code *) type tcode type vswitch = { sw_type_code : tcode; sw_code : tcode; sw_annot : annot_switch; sw_stk : vstack; sw_env : vm_env } external mkAccuCode : int -> tcode = "coq_makeaccu" val fun_code : vfun -> tcode val fix_code : vfix -> tcode val cofix_upd_code : to_update -> tcode type id_key = | ConstKey of Constant.t | VarKey of Id.t | RelKey of Int.t | EvarKey of Evar.t val eq_id_key : id_key -> id_key -> bool type atom = | Aid of id_key | Aind of inductive | Asort of Sorts.t val get_atom_rel : unit -> atom array (** Global table of rels *) (** Zippers *) type zipper = | Zapp of arguments | Zfix of vfix * arguments (** might be empty *) | Zswitch of vswitch | Zproj of Projection.Repr.t (* name of the projection *) type stack = zipper list type whd = | Vprod of vprod | Vfun of vfun | Vfix of vfix * arguments option | Vcofix of vcofix * to_update * arguments option | Vconstr_const of int | Vconstr_block of vblock | Vint64 of int64 | Vfloat64 of float | Vatom_stk of atom * stack | Vuniv_level of Univ.Level.t (** For debugging purposes only *) val pr_atom : atom -> Pp.t val pr_whd : whd -> Pp.t val pr_stack : stack -> Pp.t (** Constructors *) val val_of_str_const : structured_constant -> values val val_of_rel : int -> values val val_of_named : Id.t -> values val val_of_constant : Constant.t -> values val val_of_evar : Evar.t -> values val val_of_proj : Projection.Repr.t -> values -> values val val_of_atom : atom -> values val val_of_int : int -> structured_values val val_of_block : tag -> structured_values array -> structured_values val val_of_uint : Uint63.t -> structured_values external val_of_annot_switch : annot_switch -> values = "%identity" external val_of_proj_name : Projection.Repr.t -> values = "%identity" (** Destructors *) val whd_val : values -> whd val uni_lvl_val : values -> Univ.Level.t (** Arguments *) val nargs : arguments -> int val arg : arguments -> int -> values (** Product *) val dom : vprod -> values val codom : vprod -> vfun (** Fun *) external closure_arity : vfun -> int = "coq_closure_arity" (** Fix *) val current_fix : vfix -> int val check_fix : vfix -> vfix -> bool val rec_args : vfix -> int array val first_fix : vfix -> vfix val fix_types : vfix -> tcode array val cofix_types : vcofix -> tcode array external offset_closure_fix : vfix -> int -> vm_env = "coq_offset_closure" val mk_fix_body : int -> int -> vfix -> vfun array (** CoFix *) val current_cofix : vcofix -> int val check_cofix : vcofix -> vcofix -> bool val mk_cofix_body : (vfun -> vstack -> values) -> int -> int -> vcofix -> values array (** Block *) val btag : vblock -> int val bsize : vblock -> int val bfield : vblock -> int -> values (** Switch *) val check_switch : vswitch -> vswitch -> bool val branch_arg : int -> tag * int -> values coq-8.11.0/kernel/retypeops.ml0000644000175000017500000001132713612664533016136 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* relevance_of_constant env c | VarKey x -> relevance_of_var env x | RelKey p -> relevance_of_rel env p let rec relevance_of_fterm env extra lft f = let open CClosure in match CClosure.relevance_of f with | KnownR -> Sorts.Relevant | KnownI -> Sorts.Irrelevant | Unknown -> let r = match fterm_of f with | FRel n -> Range.get extra (Esubst.reloc_rel n lft - 1) | FAtom c -> relevance_of_term_extra env extra lft (Esubst.subs_id 0) c | FFlex key -> relevance_of_flex env key | FInt _ | FFloat _ -> Sorts.Relevant | FInd _ | FProd _ -> Sorts.Relevant (* types are always relevant *) | FConstruct (c,_) -> relevance_of_constructor env c | FApp (f, _) -> relevance_of_fterm env extra lft f | FProj (p, _) -> relevance_of_projection env p | FFix (((_,i),(lna,_,_)), _) -> (lna.(i)).binder_relevance | FCoFix ((i,(lna,_,_)), _) -> (lna.(i)).binder_relevance | FCaseT (ci, _, _, _, _) -> ci.ci_relevance | FLambda (len, tys, bdy, e) -> let extra = List.fold_left (fun accu (x, _) -> Range.cons (binder_relevance x) accu) extra tys in let lft = Esubst.el_liftn len lft in let e = Esubst.subs_liftn len e in relevance_of_term_extra env extra lft e bdy | FLetIn (x, _, _, bdy, e) -> relevance_of_term_extra env (Range.cons x.binder_relevance extra) (Esubst.el_lift lft) (Esubst.subs_lift e) bdy | FLIFT (k, f) -> relevance_of_fterm env extra (Esubst.el_shft k lft) f | FCLOS (c, e) -> relevance_of_term_extra env extra lft e c | FEvar (_, _) -> Sorts.Relevant (* let's assume evars are relevant for now *) | FLOCKED -> assert false in CClosure.set_relevance r f; r and relevance_of_term_extra env extra lft subs c = match kind c with | Rel n -> begin match Esubst.expand_rel n subs with | Inl (k, f) -> relevance_of_fterm env extra (Esubst.el_liftn k lft) f | Inr (n, None) -> Range.get extra (Esubst.reloc_rel n lft - 1) | Inr (_, Some p) -> relevance_of_rel env p end | Var x -> relevance_of_var env x | Sort _ | Ind _ | Prod _ -> Sorts.Relevant (* types are always relevant *) | Cast (c, _, _) -> relevance_of_term_extra env extra lft subs c | Lambda ({binder_relevance=r;_}, _, bdy) -> relevance_of_term_extra env (Range.cons r extra) (Esubst.el_lift lft) (Esubst.subs_lift subs) bdy | LetIn ({binder_relevance=r;_}, _, _, bdy) -> relevance_of_term_extra env (Range.cons r extra) (Esubst.el_lift lft) (Esubst.subs_lift subs) bdy | App (c, _) -> relevance_of_term_extra env extra lft subs c | Const (c,_) -> relevance_of_constant env c | Construct (c,_) -> relevance_of_constructor env c | Case (ci, _, _, _) -> ci.ci_relevance | Fix ((_,i),(lna,_,_)) -> (lna.(i)).binder_relevance | CoFix (i,(lna,_,_)) -> (lna.(i)).binder_relevance | Proj (p, _) -> relevance_of_projection env p | Int _ | Float _ -> Sorts.Relevant | Meta _ | Evar _ -> Sorts.Relevant (* let's assume metas and evars are relevant for now *) let relevance_of_fterm env extra lft c = if Environ.sprop_allowed env then relevance_of_fterm env extra lft c else Sorts.Relevant let relevance_of_term env c = if Environ.sprop_allowed env then relevance_of_term_extra env Range.empty Esubst.el_id (Esubst.subs_id 0) c else Sorts.Relevant coq-8.11.0/CREDITS0000644000175000017500000001757413612664533013324 0ustar treinentreinenThe "Coq proof assistant" was jointly developed by - INRIA Formel, Coq, LogiCal, ProVal, TypiCal, Marelle, pi.r2, Ascola, Galinette projects (starting 1985), - Laboratoire de l'Informatique du Parallelisme (LIP) associated to CNRS and ENS Lyon (Sep. 1989 to Aug. 1997), - Laboratoire de Recherche en Informatique (LRI) associated to CNRS and university Paris Sud (since Sep. 1997), - Laboratoire d'Informatique de l'Ecole Polytechnique (LIX) associated to CNRS and Ecole Polytechnique (since Jan. 2003). - Laboratoire PPS associated to CNRS and University Paris Diderot (Jan. 2009 - Dec. 2015 when it was merged into IRIF). - Institut de Recherche en Informatique Fondamentale (IRIF), associated to CNRS and University Paris Diderot (since Jan. 2016). - And many contributors from various institutions. All files but the material of the reference manual are distributed under the term of the GNU Lesser General Public License Version 2.1. The material of the reference manual is distributed under the terms of the Open Publication License v1.0 or above, as indicated in file doc/LICENCE. The following directories contain independent contributions supported by the Coq development team. All of them are released under the terms of the GNU Lesser General Public License Version 2.1. plugins/cc developed by Pierre Corbineau (ENS Cachan, 2001, LRI, 2001-2005, Radboud University at Nijmegen, 2005-2008, Grenoble 1, 2010-2014) plugins/extraction developed by Pierre Letouzey (LRI, 2000-2004, PPS, 2005-now) plugins/firstorder developed by Pierre Corbineau (LRI, 2003-2008) plugins/funind developed by Pierre Courtieu (INRIA-Lemme, 2003-2004, CNAM, 2006-now), Julien Forest (INRIA-Everest, 2006, CNAM, 2007-2008, ENSIIE, 2008-now) and Yves Bertot (INRIA-Marelle, 2005-2006) plugins/micromega developed by Frédéric Besson (IRISA/INRIA, 2006-now), with some extensions by Evgeny Makarov (INRIA, 2007); sum-of-squares solver and interface to the csdp solver uses code from John Harrison (University of Cambridge, 1998) plugins/nsatz developed by Loïc Pottier (INRIA-Marelle, 2009-2011) plugins/omega developed by Pierre Crégut (France Telecom R&D, 1996) plugins/rtauto developed by Pierre Corbineau (LRI, 2005) plugins/setoid_ring developed by Benjamin Grégoire (INRIA-Everest, 2005-2006), Assia Mahboubi, Laurent Théry (INRIA-Marelle, 2006) and Bruno Barras (INRIA LogiCal, 2005-2006), plugins/ssr developed by Georges Gonthier (Microsoft Research - Inria Joint Centre, 2007-2013, Inria, 2013-now), Assia Mahboubi and Enrico Tassi (Inria, 2011-now). plugins/ssrmatching developed by Georges Gonthier (Microsoft Research - Inria Joint Centre, 2007-2011, Inria, 2013-now), and Enrico Tassi (Inria-Marelle, 2011-now) theories/ZArith started by Pierre Crégut (France Telecom R&D, 1996) theories/Strings developed by Laurent Théry (INRIA-Lemme, 2003) theories/Numbers/Cyclic developed by Benjamin Grégoire (INRIA-Everest, 2007), Laurent Théry (INRIA-Marelle, 2007-2008), Arnaud Spiwack (INRIA-LogiCal, 2007) and Pierre Letouzey (PPS, 2008) ide/utils some files come from Maxence Guesdon's Cameleon tool The development of Coq significantly benefited from feedback, suggestions or short contributions from the following non exhaustive list of persons and groups: C. Alvarado, C. Auger, F. Blanqui, P. Castéran, C. Cohen, J. Courant, J. Duprat, F. Garillot, G. Gonthier, J. Goubault, J.-P. Jouannaud, S. Lescuyer, A. Miquel, J.-F. Monin, P.-Y. Strub the Foundations Group (Radboud University, Nijmegen, The Netherlands), Laboratoire J.-A. Dieudonné (University of Nice-Sophia Antipolis), L. Lee (https://orcid.org/0000-0002-7128-9257, 2018), INRIA-Gallium project, the CS dept at Yale, the CIS dept at U. Penn, the CSE dept at Harvard, the CS dept at Princeton, the CS dept at MIT as well as a lot of users on coq-club, coqdev, coq-bugs The following people have contributed to the development of different versions of the Coq Proof assistant during the indicated time: Bruno Barras (INRIA, 1995-now) Yves Bertot (INRIA, 2000-now) Pierre Boutillier (INRIA-PPS, 2010-2015) Xavier Clerc (INRIA, 2012-2014) Tej Chajed (MIT, 2016-now) Jacek Chrzaszcz (LRI, 1998-2003) Thierry Coquand (INRIA, 1985-1989) Pierre Corbineau (LRI, 2003-2005, Nijmegen, 2005-2008, Grenoble 1, 2008-2011) Cristina Cornes (INRIA, 1993-1996) Yann Coscoy (INRIA Sophia-Antipolis, 1995-1996) Pierre Courtieu (CNAM, 2006-now) David Delahaye (INRIA, 1997-2002) Maxime Dénès (INRIA, 2013-now) Daniel de Rauglaudre (INRIA, 1996-1998, 2012, 2016) Olivier Desmettre (INRIA, 2001-2003) Gilles Dowek (INRIA, 1991-1994) Jim Fehrle (2018-now) Amy Felty (INRIA, 1993) Jean-Christophe Filliâtre (ENS Lyon, 1994-1997, LRI, 1997-2008) Emilio Jesús Gallego Arias (MINES ParisTech 2015-now) Gaetan Gilbert (INRIA-Galinette, 2016-now) Eduardo Giménez (ENS Lyon, 1993-1996, INRIA, 1997-1998) Stéphane Glondu (INRIA-PPS, 2007-2013) Benjamin Grégoire (INRIA, 2003-2011) Jason Gross (MIT 2013-now) Hugo Herbelin (INRIA, 1996-now) Sébastien Hinderer (INRIA, 2014) Gérard Huet (INRIA, 1985-1997) Konstantinos Kallas (U. Penn, 2019) Matej Košík (INRIA, 2015-2017) Leonidas Lampropoulos (University of Pennsylvania, 2018) Pierre Letouzey (LRI, 2000-2004, PPS, 2005-2008, INRIA-PPS then IRIF, 2009-2018) Yao Li (ORCID: https://orcid.org/0000-0001-8720-883X, University of Pennsylvania, 2018) Yishuai Li (ORCID: https://orcid.org/0000-0002-5728-5903 U. Penn, 2018-2019) Patrick Loiseleur (Paris Sud, 1997-1999) Andreas Lynge (Aarhus University, 2019) Evgeny Makarov (INRIA, 2007) Gregory Malecha (Harvard University 2013-2015, University of California, San Diego 2016) Cyprien Mangin (INRIA-PPS then IRIF, 2015-now) Pascal Manoury (INRIA, 1993) Claude Marché (INRIA, 2003-2004 & LRI, 2004) Micaela Mayero (INRIA, 1997-2002) Guillaume Melquiond (INRIA, 2009-now) Benjamin Monate (LRI, 2003) César Muñoz (INRIA, 1994-1995) Chetan Murthy (INRIA, 1992-1994) Julien Narboux (INRIA, 2005-2006, Strasbourg, 2007-2011) Jean-Marc Notin (CNRS, 2006-now) Catherine Parent-Vigouroux (ENS Lyon, 1992-1995) Christine Paulin-Mohring (INRIA, 1985-1989, ENS Lyon, 1989-1997, LRI, 1997-2006) Pierre-Marie Pédrot (INRIA-PPS, 2011-2015, INRIA-Ascola, 2015-2016, University of Ljubljana, 2016-2017, MPI-SWS, 2017-2018, INRIA 2018-now) Clément Pit-Claudel (MIT, 2015-now) Matthias Puech (INRIA-Bologna, 2008-2011) Yann Régis-Gianas (INRIA-PPS then IRIF, 2009-2016) Clément Renard (INRIA, 2001-2004) Talia Ringer (University of Washington, 2019) Claudio Sacerdoti Coen (INRIA, 2004-2005) Amokrane Saïbi (INRIA, 1993-1998) Vincent Semeria (2018-now) Vincent Siles (INRIA, 2007) Élie Soubiran (INRIA, 2007-2010) Matthieu Sozeau (INRIA, 2005-now) Arnaud Spiwack (INRIA-LIX-Chalmers University, 2006-2010, INRIA, 2011-2014, MINES ParisTech 2014-2015, Tweag/IO 2015-now) Paul Steckler (MIT 2016-2018) Enrico Tassi (INRIA, 2011-now) Amin Timany (Katholieke Universiteit Leuven, 2017) Benjamin Werner (INRIA, 1989-1994) Nickolai Zeldovich (MIT 2014-2016) Théo Zimmermann (ORCID: https://orcid.org/0000-0002-3580-8806, INRIA-PPS then IRIF, 2015-now) *************************************************************************** INRIA refers to: Institut National de la Recherche en Informatique et Automatique CNRS refers to: Centre National de la Recherche Scientifique LRI refers to: Laboratoire de Recherche en Informatique, UMR 8623 CNRS and Université Paris-Sud ENS Lyon refers to: Ecole Normale Supérieure de Lyon PPS refers to: Laboratoire Preuve, Programmation, Système, UMR 7126, CNRS and Université Paris 7 **************************************************************************** coq-8.11.0/default.nix0000644000175000017500000000745313612664533014443 0ustar treinentreinen# How to use? # If you have Nix installed, you can get in an environment with everything # needed to compile Coq and CoqIDE by running: # $ nix-shell # at the root of the Coq repository. # How to tweak default arguments? # nix-shell supports the --arg option (see Nix doc) that allows you for # instance to do this: # $ nix-shell --arg ocamlPackages "(import {}).ocaml-ng.ocamlPackages_4_05" --arg buildIde false # You can also compile Coq and "install" it by running: # $ make clean # (only needed if you have left-over compilation files) # $ nix-build # at the root of the Coq repository. # nix-build also supports the --arg option, so you will be able to do: # $ nix-build --arg doInstallCheck false # if you want to speed up things by not running the test-suite. # Once the build is finished, you will find, in the current directory, # a symlink to where Coq was installed. { pkgs ? import ./dev/nixpkgs.nix {} , ocamlPackages ? pkgs.ocamlPackages , buildIde ? true , buildDoc ? true , doInstallCheck ? true , shell ? false # We don't use lib.inNixShell because that would also apply # when in a nix-shell of some package depending on this one. , coq-version ? "8.11-git" }: with pkgs; with stdenv.lib; stdenv.mkDerivation rec { name = "coq"; buildInputs = [ hostname python2 # update-compat.py python3 time # coq-makefile timing tools dune ] ++ (with ocamlPackages; [ ocaml findlib num ]) ++ optionals buildIde [ ocamlPackages.lablgtk3-sourceview3 glib gnome3.defaultIconTheme wrapGAppsHook ] ++ optionals buildDoc [ # Sphinx doc dependencies pkgconfig (python3.withPackages (ps: [ ps.sphinx ps.sphinx_rtd_theme ps.pexpect ps.beautifulsoup4 ps.antlr4-python3-runtime ps.sphinxcontrib-bibtex ])) antlr4 ocamlPackages.odoc ] ++ optionals doInstallCheck ( # Test-suite dependencies # ncurses is required to build an OCaml REPL optional (!versionAtLeast ocaml.version "4.07") ncurses ++ [ ocamlPackages.ounit rsync which ] ) ++ optionals shell ( [ jq curl gitFull gnupg ] # Dependencies of the merging script ++ (with ocamlPackages; [ merlin ocp-indent ocp-index utop ]) # Dev tools ++ [ graphviz ] # Useful for STM debugging ); src = if shell then null else with builtins; filterSource (path: _: !elem (baseNameOf path) [".git" "result" "bin" "_build" "_build_ci" "nix"]) ./.; preConfigure = '' patchShebangs dev/tools/ ''; prefixKey = "-prefix "; enableParallelBuilding = true; buildFlags = [ "world" "byte" ] ++ optional buildDoc "doc-html"; installTargets = [ "install" "install-byte" ] ++ optional buildDoc "install-doc-html"; createFindlibDestdir = !shell; postInstall = "ln -s $out/lib/coq $OCAMLFIND_DESTDIR/coq"; inherit doInstallCheck; preInstallCheck = '' patchShebangs tools/ patchShebangs test-suite/ export OCAMLPATH=$OCAMLFIND_DESTDIR:$OCAMLPATH ''; installCheckTarget = [ "check" ]; passthru = { inherit coq-version ocamlPackages; dontFilter = true; # Useful to use mkCoqPackages from }; setupHook = writeText "setupHook.sh" " addCoqPath () { if test -d \"$1/lib/coq/${coq-version}/user-contrib\"; then export COQPATH=\"$COQPATH\${COQPATH:+:}$1/lib/coq/${coq-version}/user-contrib/\" fi } addEnvHooks \"$targetOffset\" addCoqPath "; meta = { description = "Coq proof assistant"; longDescription = '' Coq is a formal proof management system. It provides a formal language to write mathematical definitions, executable algorithms and theorems together with an environment for semi-interactive development of machine-checked proofs. ''; homepage = http://coq.inria.fr; license = licenses.lgpl21; platforms = platforms.unix; }; } coq-8.11.0/ide/0000755000175000017500000000000013612664533013027 5ustar treinentreinencoq-8.11.0/ide/configwin_types.ml0000644000175000017500000001456013612664533016576 0ustar treinentreinen(*********************************************************************************) (* Cameleon *) (* *) (* Copyright (C) 2005 Institut National de Recherche en Informatique et *) (* en Automatique. All rights reserved. *) (* *) (* This program is free software; you can redistribute it and/or modify *) (* it under the terms of the GNU Library General Public License as *) (* published by the Free Software Foundation; either version 2 of the *) (* License, or any later version. *) (* *) (* This program is distributed in the hope that it will be useful, *) (* but WITHOUT ANY WARRANTY; without even the implied warranty of *) (* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) (* GNU Library General Public License for more details. *) (* *) (* You should have received a copy of the GNU Library General Public *) (* License along with this program; if not, write to the Free Software *) (* Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA *) (* 02111-1307 USA *) (* *) (* Contact: Maxence.Guesdon@inria.fr *) (* *) (*********************************************************************************) (** This module contains the types used in Configwin. *) (** This type represents a string or filename parameter, or any other type, depending on the given conversion functions. *) type 'a string_param = { string_label : string; (** the label of the parameter *) mutable string_value : 'a; (** the current value of the parameter *) string_editable : bool ; (** indicates if the value can be changed *) string_f_apply : ('a -> unit) ; (** the function to call to apply the new value of the parameter *) string_help : string option ; (** optional help string *) string_expand : bool ; (** expand or not *) string_to_string : 'a -> string ; string_of_string : string -> 'a ; } ;; (** This type represents a boolean parameter. *) type bool_param = { bool_label : string; (** the label of the parameter *) mutable bool_value : bool; (** the current value of the parameter *) bool_editable : bool ; (** indicates if the value can be changed *) bool_f_apply : (bool -> unit) ; (** the function to call to apply the new value of the parameter *) bool_help : string option ; (** optional help string *) } ;; (** This type represents a parameter whose value is a list of ['a]. *) type 'a list_param = { list_label : string; (** the label of the parameter *) mutable list_value : 'a list; (** the current value of the parameter *) list_titles : string list option; (** the titles of columns, if they must be displayed *) list_f_edit : ('a -> 'a) option; (** optional edition function *) list_eq : ('a -> 'a -> bool) ; (** the comparison function used to get list without doubles *) list_strings : ('a -> string list); (** the function to get a string list from a ['a]. *) list_color : ('a -> string option) ; (** a function to get the optional color of an element *) list_editable : bool ; (** indicates if the value can be changed *) list_f_add : unit -> 'a list ; (** the function to call to add list *) list_f_apply : ('a list -> unit) ; (** the function to call to apply the new value of the parameter *) list_help : string option ; (** optional help string *) } ;; type combo_param = { combo_label : string ; mutable combo_value : string ; combo_choices : string list ; combo_editable : bool ; combo_blank_allowed : bool ; combo_new_allowed : bool ; combo_f_apply : (string -> unit); combo_help : string option ; (** optional help string *) combo_expand : bool ; (** expand the entry widget or not *) } ;; type custom_param = { custom_box : GPack.box ; custom_f_apply : (unit -> unit) ; custom_expand : bool ; custom_framed : string option ; (** optional label for an optional frame *) } ;; type modifiers_param = { md_label : string ; (** the label of the parameter *) mutable md_value : Gdk.Tags.modifier list ; (** The value, as a list of modifiers and a key code *) md_editable : bool ; (** indicates if the value can be changed *) md_f_apply : Gdk.Tags.modifier list -> unit ; (** the function to call to apply the new value of the parameter *) md_help : string option ; (** optional help string *) md_expand : bool ; (** expand or not *) md_allow : Gdk.Tags.modifier list } (** This type represents the different kinds of parameters. *) type parameter_kind = String_param of string string_param | List_param of (unit -> ) | Bool_param of bool_param | Text_param of string string_param | Combo_param of combo_param | Custom_param of custom_param | Modifiers_param of modifiers_param ;; (** This type represents the structure of the configuration window. *) type configuration_structure = | Section of string * GtkStock.id option * parameter_kind list (** label of the section, icon, parameters *) | Section_list of string * GtkStock.id option * configuration_structure list (** label of the section, list of the sub sections *) ;; (** To indicate what button was pushed by the user when the window is closed. *) type return_button = Return_apply (** The user clicked on Apply at least once before closing the window with Cancel or the window manager. *) | Return_ok (** The user closed the window with the ok button. *) | Return_cancel (** The user closed the window with the cancel button or the window manager but never clicked on the apply button.*) coq-8.11.0/ide/coqide.ml0000644000175000017500000014263113612664533014634 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* None in match term with | None -> () | Some t -> ignore (f t) let cb_on_current_term f _ = on_current_term f (** Nota: using && here has the advantage of working both under win32 and unix. If someday we want the main command to be tried even if the "cd" has failed, then we should use " ; " under unix but " & " under win32 (cf. #2363). *) let local_cd file = "cd " ^ Filename.quote (Filename.dirname file) ^ " && " let pr_exit_status = function | Unix.WEXITED 0 -> " succeeded" | _ -> " failed" let make_coqtop_args fname = let open CoqProject_file in let base_args = match read_project#get with | Ignore_args -> !sup_args | Append_args -> !sup_args | Subst_args -> [] in let proj, args = if read_project#get = Ignore_args then "", base_args else match !custom_project_file, fname with | Some (d,proj), _ -> d, coqtop_args_from_project proj @ base_args | None, None -> "", base_args | None, Some the_file -> match CoqProject_file.find_project_file ~from:(Filename.dirname the_file) ~projfile_name:project_file_name#get with | None -> "", base_args | Some proj -> let warning_fn x = Feedback.msg_warning Pp.(str x) in proj, coqtop_args_from_project (read_project_file ~warning_fn proj) @ base_args in let args = match fname with | None -> args | Some fname -> if List.exists (String.equal "-top") args then args else (* We basically copy the code of Names.check_valid since it is not exported *) (* to coqide. This is to prevent a possible failure of parsing "-topfile" *) (* at initialization of coqtop (see #10286) *) match Unicode.ident_refutation (Filename.chop_extension (Filename.basename fname)) with | Some (_,x) -> output_string stderr (x^"\n"); exit 1 | None -> "-topfile"::fname::args in proj, args (** Setting drag & drop on widgets *) let load_file_cb : (string -> unit) ref = ref ignore let drop_received context ~x ~y data ~info ~time = if data#format = 8 then begin let files = Str.split (Str.regexp "\r?\n") data#data in let path = Str.regexp "^file://\\(.*\\)$" in List.iter (fun f -> if Str.string_match path f 0 then !load_file_cb (Str.matched_group 1 f) ) files; context#finish ~success:true ~del:false ~time end else context#finish ~success:false ~del:false ~time let drop_targets = [ { Gtk.target = "text/uri-list"; Gtk.flags = []; Gtk.info = 0} ] let set_drag (w : GObj.drag_ops) = w#dest_set drop_targets ~actions:[`COPY;`MOVE]; w#connect#data_received ~callback:drop_received (** Session management *) let create_session f = let project_file, args = make_coqtop_args f in if project_file <> "" then flash_info (Printf.sprintf "Reading options from %s" project_file); let ans = Session.create f args in let _ = set_drag ans.script#drag in ans (** Auxiliary functions for the File operations *) module FileAux = struct let load_file ?(maycreate=false) f = let f = CUnix.correct_path f (Sys.getcwd ()) in try Minilib.log "Loading file starts"; let is_f = CUnix.same_file f in let rec search_f i = function | [] -> false | sn :: sessions -> match sn.fileops#filename with | Some fn when is_f fn -> notebook#goto_page i; true | _ -> search_f (i+1) sessions in if not (search_f 0 notebook#pages) then begin Minilib.log "Loading: get raw content"; let b = Buffer.create 1024 in if Sys.file_exists f then Ideutils.read_file f b else if not maycreate then flash_info ("Load failed: no such file"); Minilib.log "Loading: convert content"; let s = do_convert (Buffer.contents b) in Minilib.log "Loading: create view"; let session = create_session (Some f) in let index = notebook#append_term session in notebook#goto_page index; Minilib.log "Loading: stats"; session.fileops#update_stats; let input_buffer = session.buffer in Minilib.log "Loading: fill buffer"; input_buffer#set_text s; input_buffer#set_modified false; input_buffer#place_cursor ~where:input_buffer#start_iter; Sentence.tag_all input_buffer; session.script#clear_undo (); Minilib.log "Loading: success"; end with e -> flash_info ("Load failed: "^(Printexc.to_string e)) let confirm_save ok = if ok then flash_info "Saved" else warning "Save Failed" let select_and_save ?parent ~saveas ?filename sn = let do_save = if saveas then sn.fileops#saveas ?parent else sn.fileops#save in let title = if saveas then "Save file as" else "Save file" in match select_file_for_save ~title ?parent ?filename () with |None -> false |Some f -> let ok = do_save f in confirm_save ok; if ok then sn.tab_label#set_text (Filename.basename f); ok let check_save ?parent ~saveas sn = try match sn.fileops#filename with |None -> select_and_save ?parent ~saveas sn |Some f -> let ok = sn.fileops#save f in confirm_save ok; ok with _ -> warning "Save Failed"; false exception DontQuit let check_quit ?parent saveall = (try save_pref () with e -> flash_info ("Cannot save preferences (" ^ Printexc.to_string e ^ ")")); let is_modified sn = sn.buffer#modified in if List.exists is_modified notebook#pages then begin let answ = Configwin_ihm.question_box ~title:"Quit" ~buttons:["Save Named Buffers and Quit"; "Quit without Saving"; "Don't Quit"] ~default:0 ~icon:(warn_image ())#coerce ?parent "There are unsaved buffers" in match answ with | 1 -> saveall () | 2 -> () | _ -> raise DontQuit end; List.iter (fun sn -> Coq.close_coqtop sn.coqtop) notebook#pages (* For MacOS, just to be sure, we close all coqtops (again?) *) let close_and_quit () = List.iter (fun sn -> Coq.close_coqtop sn.coqtop) notebook#pages; exit 0 let crash_save exitcode = Minilib.log "Starting emergency save of buffers in .crashcoqide files"; let idx = let r = ref 0 in fun () -> incr r; string_of_int !r in let save_session sn = let filename = match sn.fileops#filename with | None -> "Unnamed_coqscript_" ^ idx () ^ ".crashcoqide" | Some f -> f^".crashcoqide" in try if try_export filename (sn.buffer#get_text ()) then Minilib.log ("Saved "^filename) else Minilib.log ("Could not save "^filename) with _ -> Minilib.log ("Could not save "^filename) in List.iter save_session notebook#pages; Minilib.log "End emergency save"; exit exitcode end let () = load_file_cb := (fun s -> FileAux.load_file s) (** Callbacks for the File menu *) module File = struct let newfile _ = let session = create_session None in let index = notebook#append_term session in notebook#goto_page index let load ?parent _ = let filename = try notebook#current_term.fileops#filename with Invalid_argument _ -> None in match select_file_for_open ~title:"Load file" ?parent ?filename () with | None -> () | Some f -> FileAux.load_file f let save ?parent _ = on_current_term (FileAux.check_save ?parent ~saveas:false) let saveas ?parent sn = try let filename = sn.fileops#filename in ignore (FileAux.select_and_save ?parent ~saveas:true ?filename sn) with _ -> warning "Save Failed" let saveas ?parent = cb_on_current_term (saveas ?parent) let saveall _ = List.iter (fun sn -> match sn.fileops#filename with | None -> () | Some f -> ignore (sn.fileops#save f)) notebook#pages let () = Coq.save_all := saveall let revert_all ?parent _ = List.iter (fun sn -> if sn.fileops#changed_on_disk then sn.fileops#revert ?parent ()) notebook#pages let quit ?parent _ = try FileAux.check_quit ?parent saveall; exit 0 with FileAux.DontQuit -> () let close_buffer ?parent sn = let do_remove () = notebook#remove_page notebook#current_page in if not sn.buffer#modified then do_remove () else let answ = Configwin_ihm.question_box ~title:"Close" ~buttons:["Save Buffer and Close"; "Close without Saving"; "Don't Close"] ~default:0 ~icon:(warn_image ())#coerce ?parent "This buffer has unsaved modifications" in match answ with | 1 when FileAux.check_save ?parent ~saveas:true sn -> do_remove () | 2 -> do_remove () | _ -> () let close_buffer ?parent = cb_on_current_term (close_buffer ?parent) let export kind sn = match sn.fileops#filename with |None -> flash_info "Cannot print: this buffer has no name" |Some f -> let basef = Filename.basename f in let output = let basef_we = try Filename.chop_extension basef with _ -> basef in match kind with | "latex" -> basef_we ^ ".tex" | "dvi" | "ps" | "pdf" | "html" -> basef_we ^ "." ^ kind | _ -> assert false in let cmd = local_cd f ^ cmd_coqdoc#get ^ " --" ^ kind ^ " -o " ^ (Filename.quote output) ^ " " ^ (Filename.quote basef) ^ " 2>&1" in sn.messages#default_route#set (Pp.str ("Running: "^cmd)); let finally st = flash_info (cmd ^ pr_exit_status st) in run_command (fun msg -> sn.messages#default_route#add_string msg) finally cmd let export kind = cb_on_current_term (export kind) let print sn = match sn.fileops#filename with |None -> flash_info "Cannot print: this buffer has no name" |Some f_name -> let cmd = local_cd f_name ^ cmd_coqdoc#get ^ " -ps " ^ Filename.quote (Filename.basename f_name) ^ " | " ^ cmd_print#get in let w = GWindow.window ~title:"Print" ~modal:true ~position:`CENTER ~wmclass:("CoqIDE","CoqIDE") () in let v = GPack.vbox ~spacing:10 ~border_width:10 ~packing:w#add () in let _ = GMisc.label ~text:"Print using the following command:" ~justify:`LEFT ~packing:v#add () in let e = GEdit.entry ~text:cmd ~editable:true ~width_chars:80 ~packing:v#add () in let h = GPack.hbox ~spacing:10 ~packing:v#add () in let ko = GButton.button ~stock:`CANCEL ~label:"Cancel" ~packing:h#add () in let ok = GButton.button ~stock:`PRINT ~label:"Print" ~packing:h#add () in let callback_print () = w#destroy (); let cmd = e#text in let finally st = flash_info (cmd ^ pr_exit_status st) in run_command ignore finally cmd in let _ = ko#connect#clicked ~callback:w#destroy in let _ = ok#connect#clicked ~callback:callback_print in w#misc#show () let print = cb_on_current_term print let highlight sn = Sentence.tag_all sn.buffer; sn.script#recenter_insert let highlight = cb_on_current_term highlight end (** Timers *) let reset_revert_timer () = FileOps.revert_timer.kill (); if global_auto_revert#get then FileOps.revert_timer.run ~ms:global_auto_revert_delay#get ~callback:(fun () -> File.revert_all (); true) let reset_autosave_timer () = let autosave sn = try sn.fileops#auto_save with _ -> () in let autosave_all () = List.iter autosave notebook#pages; true in FileOps.autosave_timer.kill (); if auto_save#get then FileOps.autosave_timer.run ~ms:auto_save_delay#get ~callback:autosave_all (** Export of functions used in [coqide_main] : *) let forbid_quit () = try FileAux.check_quit File.saveall; false with FileAux.DontQuit -> true let close_and_quit = FileAux.close_and_quit let crash_save = FileAux.crash_save let do_load f = FileAux.load_file f (** Callbacks for external commands *) module External = struct let coq_makefile sn = match sn.fileops#filename with |None -> flash_info "Cannot make makefile: this buffer has no name" |Some f -> let cmd = local_cd f ^ cmd_coqmakefile#get in let finally st = flash_info (cmd_coqmakefile#get ^ pr_exit_status st) in run_command ignore finally cmd let coq_makefile = cb_on_current_term coq_makefile let editor ?parent sn = match sn.fileops#filename with |None -> warning "Call to external editor available only on named files" |Some f -> File.save (); let f = Filename.quote f in let cmd = Util.subst_command_placeholder cmd_editor#get f in run_command ignore (fun _ -> sn.fileops#revert ?parent ()) cmd let editor ?parent = cb_on_current_term (editor ?parent) let compile sn = File.save (); match sn.fileops#filename with |None -> flash_info "Active buffer has no name" |Some f -> let args = Coq.get_arguments sn.coqtop in let cmd = cmd_coqc#get ^ " " ^ String.concat " " args ^ " " ^ (Filename.quote f) ^ " 2>&1" in let buf = Buffer.create 1024 in sn.messages#default_route#set (Pp.str ("Running: "^cmd)); let display s = sn.messages#default_route#add_string s; Buffer.add_string buf s in let finally st = if st = Unix.WEXITED 0 then flash_info (f ^ " successfully compiled") else begin flash_info (f ^ " failed to compile"); sn.messages#default_route#set (Pp.str "Compilation output:\n"); sn.messages#default_route#add (Pp.str (Buffer.contents buf)); end in run_command display finally cmd let compile = cb_on_current_term compile (** [last_make_buf] contains the output of the last make compilation. [last_make] is the same, but as a string, refreshed only when searching the next error. *) let last_make_buf = Buffer.create 1024 let last_make = ref "" let last_make_index = ref 0 let last_make_dir = ref "" let make sn = match sn.fileops#filename with |None -> flash_info "Cannot make: this buffer has no name" |Some f -> File.saveall (); let cmd = local_cd f ^ cmd_make#get ^ " 2>&1" in sn.messages#default_route#set (Pp.str "Compilation output:\n"); Buffer.reset last_make_buf; last_make := ""; last_make_index := 0; last_make_dir := Filename.dirname f; let display s = sn.messages#default_route#add_string s; Buffer.add_string last_make_buf s in let finally st = flash_info (cmd_make#get ^ pr_exit_status st) in run_command display finally cmd let make = cb_on_current_term make let search_compile_error_regexp = Str.regexp "File \"\\([^\"]+\\)\", line \\([0-9]+\\), characters \\([0-9]+\\)-\\([0-9]+\\)" let search_next_error () = if String.length !last_make <> Buffer.length last_make_buf then last_make := Buffer.contents last_make_buf; let _ = Str.search_forward search_compile_error_regexp !last_make !last_make_index in let f = Str.matched_group 1 !last_make and l = int_of_string (Str.matched_group 2 !last_make) and b = int_of_string (Str.matched_group 3 !last_make) and e = int_of_string (Str.matched_group 4 !last_make) and msg_index = Str.match_beginning () in last_make_index := Str.group_end 4; (Filename.concat !last_make_dir f, l, b, e, String.sub !last_make msg_index (String.length !last_make - msg_index)) let next_error sn = try let file,line,start,stop,error_msg = search_next_error () in FileAux.load_file file; let b = sn.buffer in let starti = b#get_iter_at_byte ~line:(line-1) start in let stopi = b#get_iter_at_byte ~line:(line-1) stop in b#apply_tag Tags.Script.error ~start:starti ~stop:stopi; b#place_cursor ~where:starti; sn.messages#default_route#set (Pp.str error_msg); sn.script#misc#grab_focus () with Not_found -> last_make_index := 0; sn.messages#default_route#set (Pp.str "No more errors.\n") let next_error = cb_on_current_term next_error end (** Callbacks for the Navigation menu *) let update_status sn = let display msg = pop_info (); push_info msg in let next = function | Interface.Fail x -> sn.coqops#handle_failure x | Interface.Good status -> let path = match status.Interface.status_path with | [] | _ :: [] -> "" (* Drop the topmost level, usually "Top" *) | _ :: l -> " in " ^ String.concat "." l in let name = match status.Interface.status_proofname with | None -> "" | Some n -> ", proving " ^ n in display ("Ready"^ (if microPG#get then ", [μPG]" else "") ^ path ^ name); Coq.return () in Coq.bind (Coq.status false) next let find_next_occurrence ~backward sn = (* go to the next occurrence of the current word, forward or backward *) let b = sn.buffer in let start = find_word_start (b#get_iter_at_mark `INSERT) in let stop = find_word_end start in let text = b#get_text ~start ~stop () in let search = if backward then start#backward_search else stop#forward_search in match search text with |None -> () |Some(where, _) -> b#place_cursor ~where; sn.script#recenter_insert let send_to_coq_aux f sn = let info () = Minilib.log ("Coq busy, discarding query") in let f = Coq.seq (f sn) (update_status sn) in Coq.try_grab sn.coqtop f info let send_to_coq f = on_current_term (send_to_coq_aux f) module Nav = struct let forward_one _ = send_to_coq (fun sn -> sn.coqops#process_next_phrase) let backward_one _ = send_to_coq (fun sn -> sn.coqops#backtrack_last_phrase) let goto _ = send_to_coq (fun sn -> sn.coqops#go_to_insert) let goto_end _ = send_to_coq (fun sn -> sn.coqops#process_until_end_or_error) let previous_occ = cb_on_current_term (find_next_occurrence ~backward:true) let next_occ = cb_on_current_term (find_next_occurrence ~backward:false) let restart sn = Minilib.log "Reset Initial"; Coq.reset_coqtop sn.coqtop let restart _ = on_current_term restart let interrupt sn = Minilib.log "User break received"; Coq.break_coqtop sn.coqtop CString.(Set.elements (Map.domain sn.jobpage#data)) let interrupt = cb_on_current_term interrupt let join_document _ = send_to_coq (fun sn -> sn.coqops#join_document) end let printopts_callback opts v = let b = v#get_active in let () = List.iter (fun o -> Coq.PrintOpt.set o b) opts in send_to_coq (fun sn -> sn.coqops#show_goals) (** Templates menu *) let get_current_word term = (* First look to find if autocompleting *) match term.script#proposal with | Some p -> p | None -> (* Then look at the current selected word *) let buf1 = term.script#buffer in let buf2 = term.proof#buffer in if buf1#has_selection then let (start, stop) = buf1#selection_bounds in buf1#get_text ~slice:true ~start ~stop () else if buf2#has_selection then let (start, stop) = buf2#selection_bounds in buf2#get_text ~slice:true ~start ~stop () else if term.messages#has_selection then term.messages#get_selected_text (* Otherwise try to find the word around the cursor *) else let it = term.script#buffer#get_iter_at_mark `INSERT in let start = find_word_start it in let stop = find_word_end start in term.script#buffer#get_text ~slice:true ~start ~stop () let print_branch c l = Format.fprintf c " | @[%a@]=> _@\n" (Minilib.print_list (fun c s -> Format.fprintf c "%s@ " s)) l let print_branches c cases = Format.fprintf c "@[match var with@\n%aend@]@." (Minilib.print_list print_branch) cases let display_match sn = function |Interface.Fail _ -> flash_info "Not an inductive type"; Coq.return () |Interface.Good cases -> let text = let buf = Buffer.create 1024 in let () = print_branches (Format.formatter_of_buffer buf) cases in Buffer.contents buf in Minilib.log ("match template :\n" ^ text); let b = sn.buffer in let _ = b#delete_selection () in let m = b#create_mark (b#get_iter_at_mark `INSERT) in if b#insert_interactive text then begin let i = b#get_iter (`MARK m) in let _ = i#nocopy#forward_chars 9 in let _ = b#place_cursor ~where:i in b#move_mark ~where:(i#backward_chars 3) `SEL_BOUND end; b#delete_mark (`MARK m); Coq.return () let match_callback sn = let w = get_current_word sn in let coqtop = sn.coqtop in let query = Coq.bind (Coq.mkcases w) (display_match sn) in Coq.try_grab coqtop query ignore let match_callback = cb_on_current_term match_callback (** Queries *) module Query = struct let doquery query sn = sn.messages#default_route#clear; Coq.try_grab sn.coqtop (sn.coqops#raw_coq_query query ~route_id:0 ~next:(function | Interface.Fail (_, _, err) -> let err = Ideutils.validate err in sn.messages#default_route#add err; Coq.return () | Interface.Good () -> Coq.return ())) ignore let queryif command sn = Option.iter (fun query -> doquery (query ^ ".") sn) begin try let i = CString.string_index_from command 0 "..." in let word = get_current_word sn in if word = "" then None else Some (CString.sub command 0 i ^ " " ^ word) with Not_found -> Some command end let query command _ = cb_on_current_term (queryif command) () end (** Misc *) module MiscMenu = struct let detach_view sn = sn.control#detach () let detach_view = cb_on_current_term detach_view let log_file_message () = if !Minilib.debug then let file = match !logfile with None -> "stderr" | Some f -> f in "\nDebug mode is on, log file is "^file else "" let initial_about () = let initial_string = "Welcome to CoqIDE, an Integrated Development Environment for Coq" in let coq_version = Coq.short_version () in let version_info = if Glib.Utf8.validate coq_version then "\nYou are running " ^ coq_version else "" in let msg = initial_string ^ version_info ^ log_file_message () in on_current_term (fun term -> term.messages#default_route#add_string msg) let coq_icon () = (* May raise Nof_found *) let name = "coq.png" in let chk d = Sys.file_exists (Filename.concat d name) in let dir = List.find chk (Minilib.coqide_data_dirs ()) in Filename.concat dir name let about _ = let dialog = GWindow.about_dialog () in let _ = dialog#connect#response ~callback:(fun _ -> dialog#destroy ()) in let _ = try dialog#set_logo (GdkPixbuf.from_file (coq_icon ())) with _ -> () in let copyright = "Coq is developed by the Coq Development Team\n\ (INRIA - CNRS - LIX - LRI - PPS)" in let authors = [ "Benjamin Monate"; "Jean-Christophe Filliâtre"; "Pierre Letouzey"; "Claude Marché"; "Bruno Barras"; "Pierre Corbineau"; "Julien Narboux"; "Hugo Herbelin"; "Enrico Tassi"; ] in dialog#set_name "CoqIDE"; dialog#set_comments "The Coq Integrated Development Environment"; dialog#set_website Coq_config.wwwcoq; dialog#set_version Coq_config.version; dialog#set_copyright copyright; dialog#set_authors authors; dialog#show () let apply_unicode_binding = cb_on_current_term (fun t -> t.script#apply_unicode_binding()) let comment = cb_on_current_term (fun t -> t.script#comment ()) let uncomment = cb_on_current_term (fun t -> t.script#uncomment ()) let coqtop_arguments sn = let dialog = GWindow.dialog ~title:"Coqtop arguments" () in let coqtop = sn.coqtop in (* Text entry *) let args = Coq.get_arguments coqtop in let text = String.concat " " args in let entry = GEdit.entry ~text ~packing:dialog#vbox#add () in (* Buttons *) let box = dialog#action_area in let ok = GButton.button ~stock:`OK ~packing:box#add () in let ok_cb () = let nargs = String.split_on_char ' ' entry#text in if nargs <> args then let failed = Coq.filter_coq_opts nargs in match failed with | [] -> let () = Coq.set_arguments coqtop nargs in dialog#destroy () | args -> let args = String.concat " " args in let msg = Printf.sprintf "Invalid arguments: %s" args in let () = sn.messages#default_route#clear in sn.messages#default_route#push Feedback.Error (Pp.str msg) else dialog#destroy () in let _ = entry#connect#activate ~callback:ok_cb in let _ = ok#connect#clicked ~callback:ok_cb in let cancel = GButton.button ~stock:`CANCEL ~packing:box#add () in let cancel_cb () = dialog#destroy () in let _ = cancel#connect#clicked ~callback:cancel_cb in dialog#show () let coqtop_arguments = cb_on_current_term coqtop_arguments let show_hide_query_pane sn = let ccw = sn.command in if ccw#visible then ccw#hide else ccw#show let zoom_fit sn = let script = sn.script in let space = script#misc#allocation.Gtk.width in let cols = script#right_margin_position in let pango_ctx = script#misc#pango_context in let layout = pango_ctx#create_layout#as_layout in let fsize = Pango.Font.get_size (Pango.Font.from_string text_font#get) in Pango.Layout.set_text layout (String.make cols 'X'); let tlen = fst (Pango.Layout.get_pixel_size layout) in Pango.Font.set_size (Pango.Font.from_string text_font#get) (fsize * space / tlen / Pango.scale * Pango.scale); save_pref () end (** Refresh functions *) let refresh_notebook_pos () = let pos = match vertical_tabs#get, opposite_tabs#get with | false, false -> `TOP | false, true -> `BOTTOM | true , false -> `LEFT | true , true -> `RIGHT in notebook#set_tab_pos pos (** Wrappers around GAction functions for creating menus *) let menu = GAction.add_actions let item = GAction.add_action let radio = GAction.add_radio_action (** Toggle items in menus for printing options *) let toggle_item = GAction.add_toggle_action (** Search the first '_' in a label string and return the following character as shortcut, plus the string without the '_' *) let get_shortcut s = try let n = String.length s in let i = String.index s '_' in let k = String.make 1 s.[i+1] in let s' = (String.sub s 0 i) ^ (String.sub s (i+1) (n-i-1)) in Some k, s' with _ -> None,s module Opt = Coq.PrintOpt let toggle_items menu_name l = let f d = let label = d.Opt.label in let k, name = get_shortcut label in let accel = Option.map ((^) modifier_for_display#get) k in toggle_item name ~label ?accel ~active:d.Opt.init ~callback:(printopts_callback d.Opt.opts) menu_name in List.iter f l let no_under = Util.String.map (fun x -> if x = '_' then '-' else x) (** Create alphabetical menu items with elements in sub-items. [l] is a list of lists, one per initial letter *) let alpha_items menu_name item_name l = let mk_item text = let text' = let len = String.length text in let buf = Buffer.create (len + 1) in let escaped = ref false in String.iter (fun c -> if !escaped then let () = Buffer.add_char buf c in escaped := false else if c = '_' then escaped := true else Buffer.add_char buf c ) text; if text.[len - 1] = '.' then Buffer.add_char buf '\n' else Buffer.add_char buf ' '; Buffer.contents buf in let callback _ = on_current_term (fun sn -> sn.buffer#insert_interactive text') in item (item_name^" "^(no_under text)) ~label:text ~callback menu_name in let mk_items = function | [] -> () | [s] -> mk_item s | s::_ as ll -> let name = Printf.sprintf "%s %c" item_name s.[0] in let label = Printf.sprintf "_%c..." s.[0] in item name ~label menu_name; List.iter mk_item ll in List.iter mk_items l (** Create a menu item that will insert a given text, and select the zone given by (offset,len). The first word in the text is used as item keyword. Caveat: the offset is now from the start of the text. *) let template_item (text, offset, len, key) = let modifier = modifier_for_templates#get in let idx = String.index text ' ' in let name = String.sub text 0 idx in let label = "_"^name^" __" in let negoffset = String.length text - offset - len in let callback sn = let b = sn.buffer in if b#insert_interactive text then begin let iter = b#get_iter_at_mark `INSERT in ignore (iter#nocopy#backward_chars negoffset); b#move_mark `INSERT ~where:iter; ignore (iter#nocopy#backward_chars len); b#move_mark `SEL_BOUND ~where:iter; end in item name ~label ~callback:(cb_on_current_term callback) ~accel:(modifier^key) (** Create menu items for pairs (query, shortcut key). *) let user_queries_items menu_name item_name l = let mk_item (query, key) = let callback = Query.query query in let accel = if not (CString.is_empty key) then Some (modifier_for_queries#get^key) else None in item (item_name^" "^(no_under query)) ~label:query ?accel ~callback menu_name in List.iter mk_item l let emit_to_focus window sgn = let focussed_widget = GtkWindow.Window.get_focus window#as_window in let obj = Gobject.unsafe_cast focussed_widget in try GtkSignal.emit_unit obj ~sgn with _ -> () (** {2 Creation of the main coqide window } *) let build_ui () = let w = GWindow.window ~wmclass:("CoqIde","CoqIde") ~width:window_width#get ~height:window_height#get ~title:"CoqIde" () in let () = try w#set_icon (Some (GdkPixbuf.from_file (MiscMenu.coq_icon ()))) with _ -> () in let _ = w#event#connect#delete ~callback:(fun _ -> File.quit ~parent:w (); true) in let _ = w#misc#connect#size_allocate ~callback:(fun rect -> window_size := (rect.Gtk.width, rect.Gtk.height)) in let _ = set_drag w#drag in let vbox = GPack.vbox ~homogeneous:false ~packing:w#add () in let file_menu = GAction.action_group ~name:"File" () in let edit_menu = GAction.action_group ~name:"Edit" () in let view_menu = GAction.action_group ~name:"View" () in let export_menu = GAction.action_group ~name:"Export" () in let navigation_menu = GAction.action_group ~name:"Navigation" () in let tactics_menu = GAction.action_group ~name:"Tactics" () in let templates_menu = GAction.action_group ~name:"Templates" () in let tools_menu = GAction.action_group ~name:"Tools" () in let queries_menu = GAction.action_group ~name:"Queries" () in let compile_menu = GAction.action_group ~name:"Compile" () in let windows_menu = GAction.action_group ~name:"Windows" () in let help_menu = GAction.action_group ~name:"Help" () in let all_menus = [ file_menu; edit_menu; view_menu; export_menu; navigation_menu; tactics_menu; templates_menu; tools_menu; queries_menu; compile_menu; windows_menu; help_menu; ] in menu file_menu [ item "File" ~label:"_File"; item "New" ~callback:File.newfile ~stock:`NEW; item "Open" ~callback:(File.load ~parent:w) ~stock:`OPEN; item "Save" ~callback:(File.save ~parent:w) ~stock:`SAVE ~tooltip:"Save current buffer"; item "Save as" ~label:"S_ave as" ~stock:`SAVE_AS ~callback:(File.saveas ~parent:w); item "Save all" ~label:"Sa_ve all" ~callback:File.saveall; item "Revert all buffers" ~label:"_Revert all buffers" ~callback:(File.revert_all ~parent:w) ~stock:`REVERT_TO_SAVED; item "Close buffer" ~label:"_Close buffer" ~stock:`CLOSE ~callback:(File.close_buffer ~parent:w) ~tooltip:"Close current buffer"; item "Print..." ~label:"_Print..." ~callback:File.print ~stock:`PRINT ~accel:"p"; item "Rehighlight" ~label:"Reh_ighlight" ~accel:"l" ~callback:File.highlight ~stock:`REFRESH; item "Quit" ~stock:`QUIT ~callback:(File.quit ~parent:w); ]; menu export_menu [ item "Export to" ~label:"E_xport to"; item "Html" ~label:"_Html" ~callback:(File.export "html"); item "Latex" ~label:"_LaTeX" ~callback:(File.export "latex"); item "Dvi" ~label:"_Dvi" ~callback:(File.export "dvi"); item "Pdf" ~label:"_Pdf" ~callback:(File.export "pdf"); item "Ps" ~label:"_Ps" ~callback:(File.export "ps"); ]; menu edit_menu [ item "Edit" ~label:"_Edit"; item "Undo" ~accel:"u" ~stock:`UNDO ~callback:(cb_on_current_term (fun t -> t.script#undo ())); item "Redo" ~stock:`REDO ~callback:(cb_on_current_term (fun t -> t.script#redo ())); item "Cut" ~stock:`CUT ~callback:(fun _ -> emit_to_focus w GtkText.View.S.cut_clipboard); item "Copy" ~stock:`COPY ~callback:(fun _ -> emit_to_focus w GtkText.View.S.copy_clipboard); item "Paste" ~stock:`PASTE ~callback:(fun _ -> emit_to_focus w GtkText.View.S.paste_clipboard); item "Find" ~stock:`FIND ~label:"Find / Replace" ~callback:(cb_on_current_term (fun t -> t.finder#show ())); item "Find Next" ~label:"Find _Next" ~stock:`GO_DOWN ~accel:"F3" ~callback:(cb_on_current_term (fun t -> t.finder#find_forward ())); item "Find Previous" ~label:"Find _Previous" ~stock:`GO_UP ~accel:"F3" ~callback:(cb_on_current_term (fun t -> t.finder#find_backward ())); item "External editor" ~label:"External editor" ~stock:`EDIT ~callback:(External.editor ~parent:w); item "Preferences" ~accel:"comma" ~stock:`PREFERENCES ~callback:(fun _ -> begin try Preferences.configure ~apply:refresh_notebook_pos w with e -> flash_info ("Editing preferences failed (" ^ Printexc.to_string e ^ ")") end; reset_revert_timer ()); ]; menu view_menu [ item "View" ~label:"_View"; item "Previous tab" ~label:"_Previous tab" ~accel:"Left" ~stock:`GO_BACK ~callback:(fun _ -> notebook#previous_page ()); item "Next tab" ~label:"_Next tab" ~accel:"Right" ~stock:`GO_FORWARD ~callback:(fun _ -> notebook#next_page ()); item "Zoom in" ~label:"_Zoom in" ~accel:("plus") ~stock:`ZOOM_IN ~callback:(fun _ -> let ft = Pango.Font.from_string text_font#get in Pango.Font.set_size ft (Pango.Font.get_size ft + Pango.scale); text_font#set (Pango.Font.to_string ft); save_pref ()); item "Zoom out" ~label:"_Zoom out" ~accel:("minus") ~stock:`ZOOM_OUT ~callback:(fun _ -> let ft = Pango.Font.from_string text_font#get in Pango.Font.set_size ft (Pango.Font.get_size ft - Pango.scale); text_font#set (Pango.Font.to_string ft); save_pref ()); item "Zoom fit" ~label:"_Zoom fit" ~accel:("0") ~stock:`ZOOM_FIT ~callback:(cb_on_current_term MiscMenu.zoom_fit); toggle_item "Show Toolbar" ~label:"Show _Toolbar" ~active:(show_toolbar#get) ~callback:(fun _ -> show_toolbar#set (not show_toolbar#get)); item "Query Pane" ~label:"_Query Pane" ~accel:"F1" ~callback:(cb_on_current_term MiscMenu.show_hide_query_pane); GAction.group_radio_actions ~init_value:( let v = diffs#get in List.iter (fun o -> Opt.set o v) Opt.diff_item.Opt.opts; if v = "on" then 1 else if v = "removed" then 2 else 0) ~callback:begin fun n -> (match n with | 0 -> List.iter (fun o -> Opt.set o "off"; diffs#set "off") Opt.diff_item.Opt.opts | 1 -> List.iter (fun o -> Opt.set o "on"; diffs#set "on") Opt.diff_item.Opt.opts | 2 -> List.iter (fun o -> Opt.set o "removed"; diffs#set "removed") Opt.diff_item.Opt.opts | _ -> assert false); send_to_coq (fun sn -> sn.coqops#show_goals) end [ radio "Unset diff" 0 ~label:"_Don't show diffs"; radio "Set diff" 1 ~label:"Show diffs: only _added"; radio "Set removed diff" 2 ~label:"Show diffs: added and _removed"; ]; ]; toggle_items view_menu Coq.PrintOpt.bool_items; let navitem (text, label, stock, callback, tooltip, accel) = let accel = modifier_for_navigation#get ^ accel in item text ~label ~stock ~callback ~tooltip ~accel in menu navigation_menu begin [ (fun e -> item "Navigation" ~label:"_Navigation" e); ] @ List.map navitem [ ("Forward", "_Forward", `GO_DOWN, Nav.forward_one, "Forward one command", "Down"); ("Backward", "_Backward", `GO_UP, Nav.backward_one, "Backward one command", "Up"); ("Go to", "_Go to", `JUMP_TO, Nav.goto, "Go to cursor", "Right"); ("Start", "_Start", `GOTO_TOP, Nav.restart, "Restart coq", "Home"); ("End", "_End", `GOTO_BOTTOM, Nav.goto_end, "Go to end", "End"); ("Interrupt", "_Interrupt", `STOP, Nav.interrupt, "Interrupt computations", "Break"); (* wait for this available in GtkSourceView ! ("Hide", "_Hide", `MISSING_IMAGE, ~callback:(fun _ -> let sess = notebook#current_term in toggle_proof_visibility sess.buffer sess.analyzed_view#get_insert), "Hide proof", "h"); *) ("Previous", "_Previous", `GO_BACK, Nav.previous_occ, "Previous occurrence", "less"); ("Next", "_Next", `GO_FORWARD, Nav.next_occ, "Next occurrence", "greater"); ("Force", "_Force", `EXECUTE, Nav.join_document, "Fully check the document", "f"); ] end; menu tactics_menu [ item "Tactics" ~label:"_Tactics"; ]; alpha_items tactics_menu "Tactic" Coq_commands.tactics; menu templates_menu [ item "Templates" ~label:"Te_mplates"; template_item ("Lemma new_lemma : .\nProof.\n\nQed.\n", 6,9, "J"); template_item ("Theorem new_theorem : .\nProof.\n\nQed.\n", 8,11, "T"); template_item ("Definition ident := .\n", 11,5, "E"); template_item ("Inductive ident : :=\n | : .\n", 10,5, "I"); template_item ("Fixpoint ident (_ : _) {struct _} : _ :=\n.\n", 9,5, "F"); template_item ("Scheme new_scheme := Induction for _ Sort _\n" ^ "with _ := Induction for _ Sort _.\n", 7,10, "S"); item "match" ~label:"match ..." ~accel:(modifier_for_templates#get^"M") ~callback:match_callback ]; alpha_items templates_menu "Template" Coq_commands.commands; let qitem s sc = let query = s ^ "..." in item s ~label:("_"^s) ~accel:(modifier_for_queries#get^sc) ~callback:(Query.query query) in menu queries_menu [ item "Queries" ~label:"_Queries"; qitem "Search" "K"; qitem "Check" "C"; qitem "Print" "P"; qitem "About" "A"; qitem "Locate" "L"; qitem "Print Assumptions" "N"; ]; user_queries_items queries_menu "User-Query" user_queries#get; menu tools_menu [ item "Tools" ~label:"_Tools"; item "Comment" ~label:"_Comment" ~accel:"D" ~callback:MiscMenu.comment; item "Uncomment" ~label:"_Uncomment" ~accel:"D" ~callback:MiscMenu.uncomment; item "Coqtop arguments" ~label:"Coqtop _arguments" ~callback:MiscMenu.coqtop_arguments; item "LaTeX-to-unicode" ~label:"_LaTeX-to-unicode" ~accel:"space" ~callback:MiscMenu.apply_unicode_binding; ]; menu compile_menu [ item "Compile" ~label:"_Compile"; item "Compile buffer" ~label:"_Compile buffer" ~callback:External.compile; item "Make" ~label:"_Make" ~accel:"F6" ~callback:External.make; item "Next error" ~label:"_Next error" ~accel:"F7" ~callback:External.next_error; item "Make makefile" ~label:"Make makefile" ~callback:External.coq_makefile; ]; menu windows_menu [ item "Windows" ~label:"_Windows"; item "Detach View" ~label:"Detach _View" ~callback:MiscMenu.detach_view ]; menu help_menu [ item "Help" ~label:"_Help"; item "Browse Coq Manual" ~label:"Browse Coq _Manual" ~callback:(fun _ -> browse notebook#current_term.messages#default_route#add_string Coq_config.wwwrefman); item "Browse Coq Library" ~label:"Browse Coq _Library" ~callback:(fun _ -> browse notebook#current_term.messages#default_route#add_string Coq_config.wwwstdlib); item "Help for keyword" ~label:"Help for _keyword" ~stock:`HELP ~callback:(fun _ -> on_current_term (fun sn -> browse_keyword sn.messages#default_route#add_string (get_current_word sn))); item "Help for μPG mode" ~label:"Help for μPG mode" ~callback:(fun _ -> on_current_term (fun sn -> sn.messages#default_route#clear; sn.messages#default_route#add_string (MicroPG.get_documentation ()))); item "About Coq" ~label:"_About" ~stock:`ABOUT ~callback:MiscMenu.about ]; Coqide_ui.init (); Coqide_ui.ui_m#insert_action_group file_menu 0; Coqide_ui.ui_m#insert_action_group export_menu 0; Coqide_ui.ui_m#insert_action_group edit_menu 0; Coqide_ui.ui_m#insert_action_group view_menu 0; Coqide_ui.ui_m#insert_action_group navigation_menu 0; Coqide_ui.ui_m#insert_action_group tactics_menu 0; Coqide_ui.ui_m#insert_action_group templates_menu 0; Coqide_ui.ui_m#insert_action_group tools_menu 0; Coqide_ui.ui_m#insert_action_group queries_menu 0; Coqide_ui.ui_m#insert_action_group compile_menu 0; Coqide_ui.ui_m#insert_action_group windows_menu 0; Coqide_ui.ui_m#insert_action_group help_menu 0; w#add_accel_group Coqide_ui.ui_m#get_accel_group ; GtkMain.Rc.parse_string "gtk-can-change-accels = 1"; if Coq_config.gtk_platform <> `QUARTZ then vbox#pack (Coqide_ui.ui_m#get_widget "/CoqIde MenuBar"); (* Toolbar *) let tbar = GtkButton.Toolbar.cast ((Coqide_ui.ui_m#get_widget "/CoqIde ToolBar")#as_widget) in let () = GtkButton.Toolbar.set ~orientation:`HORIZONTAL ~style:`ICONS tbar in let toolbar = new GButton.toolbar tbar in let () = vbox#pack toolbar#coerce in (* Emacs/PG mode *) MicroPG.init w notebook all_menus; (* On tab switch, reset, update location *) let _ = notebook#connect#switch_page ~callback:(fun n -> let _ = if reset_on_tab_switch#get then Nav.restart () in try let session = notebook#get_nth_term n in let ins = session.buffer#get_iter_at_mark `INSERT in Ideutils.display_location ins with _ -> ()) in (* Vertical Separator between Scripts and Goals *) let () = vbox#pack ~expand:true notebook#coerce in let () = refresh_notebook_pos () in let lower_hbox = GPack.hbox ~homogeneous:false ~packing:vbox#pack () in let () = lower_hbox#pack ~expand:true status#coerce in (* Location display *) let l = GMisc.label ~text:"Line: 1 Char: 1" ~packing:lower_hbox#pack () in let () = l#coerce#misc#set_name "location" in let () = set_location := l#set_text in (* Progress Bar *) let pbar = GRange.progress_bar ~pulse_step:0.1 () in let () = lower_hbox#pack pbar#coerce in let ready () = pbar#set_fraction 0.0; pbar#set_text "Coq is ready" in let pulse sn = if Coq.is_computing sn.coqtop then (pbar#set_text "Coq is computing"; pbar#pulse ()) else ready () in let callback () = on_current_term pulse; true in let _ = Glib.Timeout.add ~ms:300 ~callback in (* Pending proofs. It should be with a GtkSpinner... not bound *) let slaveinfo = GMisc.label ~xalign:0.5 ~width:50 () in let () = lower_hbox#pack slaveinfo#coerce in let () = slaveinfo#misc#set_has_tooltip true in let () = slaveinfo#misc#set_tooltip_markup "Proofs to be checked / Errors" in let update sn = let processed, to_process, jobs = sn.coqops#get_slaves_status in let missing = to_process - processed in let n_err = sn.coqops#get_n_errors in if n_err > 0 then slaveinfo#set_text (Printf.sprintf "%d / %d" missing n_err) else slaveinfo#set_text (Printf.sprintf "%d / %d" missing n_err); slaveinfo#set_use_markup true; sn.errpage#update sn.coqops#get_errors; sn.jobpage#update (Util.pi3 sn.coqops#get_slaves_status) in let callback () = on_current_term update; true in let _ = Glib.Timeout.add ~ms:300 ~callback in (* Initializing hooks *) let refresh_style style = let style = style_manager#style_scheme style in let iter_session v = v.script#source_buffer#set_style_scheme style in List.iter iter_session notebook#pages in let refresh_language lang = let lang = lang_manager#language lang in let iter_session v = v.script#source_buffer#set_language lang in List.iter iter_session notebook#pages in let refresh_toolbar b = if b then toolbar#misc#show () else toolbar#misc#hide () in stick show_toolbar toolbar refresh_toolbar; let _ = source_style#connect#changed ~callback:refresh_style in let _ = source_language#connect#changed ~callback:refresh_language in (* Showtime ! *) w#show (); w (** {2 Coqide main function } *) let make_file_buffer f = let f = if Filename.check_suffix f ".v" then f else f^".v" in FileAux.load_file ~maycreate:true f let make_scratch_buffer () = let session = create_session None in let _ = notebook#append_term session in () let main files = let w = build_ui () in reset_revert_timer (); reset_autosave_timer (); (match files with | [] -> make_scratch_buffer () | _ -> List.iter make_file_buffer files); notebook#goto_page 0; MiscMenu.initial_about (); on_current_term (fun t -> t.script#misc#grab_focus ()); Minilib.log "End of Coqide.main"; w (** {2 Argument parsing } *) (** By default, the coqtop we try to launch is exactly the current coqide full name, with the last occurrence of "coqide" replaced by "coqtop". This should correctly handle the ".opt", ".byte", ".exe" situations. If the replacement fails, we default to "coqtop", hoping it's somewhere in the path. Note that the -coqtop option to coqide overrides this default coqtop path *) let read_coqide_args argv = let rec filter_coqtop coqtop project_files bindings_files out = function |"-unicode-bindings" :: sfilenames :: args -> let filenames = Str.split (Str.regexp ",") sfilenames in filter_coqtop coqtop project_files (filenames @ bindings_files) out args |"-coqtop" :: prog :: args -> if coqtop = None then filter_coqtop (Some prog) project_files bindings_files out args else (output_string stderr "Error: multiple -coqtop options"; exit 1) |"-f" :: file :: args -> if project_files <> None then (output_string stderr "Error: multiple -f options"; exit 1); let d = CUnix.canonical_path_name (Filename.dirname file) in let warning_fn x = Format.eprintf "%s@\n%!" x in let p = CoqProject_file.read_project_file ~warning_fn file in filter_coqtop coqtop (Some (d,p)) out bindings_files args |"-f" :: [] -> output_string stderr "Error: missing project file name"; exit 1 |"-coqtop" :: [] -> output_string stderr "Error: missing argument after -coqtop"; exit 1 |"-debug"::args -> Minilib.debug := true; Flags.debug := true; Backtrace.record_backtrace true; filter_coqtop coqtop project_files bindings_files ("-debug"::out) args |"-coqtop-flags" :: flags :: args-> Coq.ideslave_coqtop_flags := Some flags; filter_coqtop coqtop project_files bindings_files out args |arg::args when out = [] && CString.is_prefix "-psn_" arg -> (* argument added by MacOS during .app launch *) filter_coqtop coqtop project_files bindings_files out args |arg::args -> filter_coqtop coqtop project_files bindings_files (arg::out) args |[] -> (coqtop,project_files,bindings_files,List.rev out) in let coqtop,project_files,bindings_files,argv = filter_coqtop None None [] [] argv in Ideutils.custom_coqtop := coqtop; custom_project_file := project_files; Unicode_bindings.load_files bindings_files; argv (** {2 Signal handling } *) (** The Ctrl-C (sigint) is handled as a interactive quit. For most of the other catchable signals we launch an emergency save of opened files and then exit. *) let signals_to_crash = [Sys.sigabrt; Sys.sigalrm; Sys.sigfpe; Sys.sighup; Sys.sigill; Sys.sigpipe; Sys.sigquit; Sys.sigusr1; Sys.sigusr2] let set_signal_handlers ?parent () = try Sys.set_signal Sys.sigint (Sys.Signal_handle (File.quit ?parent)); List.iter (fun i -> Sys.set_signal i (Sys.Signal_handle FileAux.crash_save)) signals_to_crash with _ -> Minilib.log "Signal ignored (normal if Win32)" coq-8.11.0/ide/document.ml0000644000175000017500000001412113612664533015176 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* ('a list * 'a list) option -> unit) -> unit method pushed : callback:('a -> ('a list * 'a list) option -> unit) -> unit end class ['a] signal () = object val mutable attached : ('a -> unit) list = [] method call (x : 'a) = let iter f = try f x with _ -> () in List.iter iter attached method connect f = attached <- f :: attached end type 'a document = { mutable stack : 'a sentence list; mutable context : ('a sentence list * 'a sentence list) option; pushed_sig : ('a * ('a list * 'a list) option) signal; popped_sig : ('a * ('a list * 'a list) option) signal; } let connect d : 'a signals = object method pushed ~callback = d.pushed_sig#connect (fun (x, ctx) -> callback x ctx) method popped ~callback = d.popped_sig#connect (fun (x, ctx) -> callback x ctx) end let create () = { stack = []; context = None; pushed_sig = new signal (); popped_sig = new signal (); } let repr_context s = match s.context with | None -> None | Some (cl, cr) -> let map s = s.data in Some (List.map map cl, List.map map cr) (* Invariant, only the tip is a allowed to have state_id = None *) let invariant l = l = [] || (List.hd l).state_id <> None let tip = function | { stack = [] } -> raise Empty | { stack = { state_id = Some id }::_ } -> id | { stack = { state_id = None }::_ } -> invalid_arg "tip" let tip_data = function | { stack = [] } -> raise Empty | { stack = { data }::_ } -> data let push d x = assert(invariant d.stack); d.stack <- { data = x; state_id = None } :: d.stack; d.pushed_sig#call (x, repr_context d) let pop = function | { stack = [] } -> raise Empty | { stack = { data }::xs } as s -> s.stack <- xs; s.popped_sig#call (data, repr_context s); data let focus d ~cond_top:c_start ~cond_bot:c_stop = assert(invariant d.stack); if d.context <> None then invalid_arg "focus"; let rec aux (a,s,b) grab = function | [] -> invalid_arg "focus" | { state_id = Some id; data } as x :: xs when not grab -> if c_start id data then aux (a,s,b) true (x::xs) else aux (x::a,s,b) grab xs | { state_id = Some id; data } as x :: xs -> if c_stop id data then List.rev a, List.rev (x::s), xs else aux (a,x::s,b) grab xs | _ -> assert false in let a, s, b = aux ([],[],[]) false d.stack in d.stack <- s; d.context <- Some (a, b) let unfocus = function | { context = None } -> invalid_arg "unfocus" | { context = Some (a,b); stack } as d -> assert(invariant stack); d.context <- None; d.stack <- a @ stack @ b let focused { context } = context <> None let to_lists = function | { context = None; stack = s } -> [],s,[] | { context = Some (a,b); stack = s } -> a,s,b let flat f b = fun x -> f b x.state_id x.data let find d f = let a, s, b = to_lists d in ( try List.find (flat f false) a with Not_found -> try List.find (flat f true) s with Not_found -> List.find (flat f false) b ).data let find_map d f = let a, s, b = to_lists d in try CList.find_map (flat f false) a with Not_found -> try CList.find_map (flat f true) s with Not_found -> CList.find_map (flat f false) b let is_empty = function | { stack = []; context = None } -> true | _ -> false let context d = let top, _, bot = to_lists d in let pair _ x y = try Option.get x, y with Option.IsNone -> assert false in List.map (flat pair true) top, List.map (flat pair true) bot let stateid_opt_equal = Option.equal Stateid.equal let is_in_focus d id = let _, focused, _ = to_lists d in List.exists (fun { state_id } -> stateid_opt_equal state_id (Some id)) focused let print d f = let top, mid, bot = to_lists d in let open Pp in v 0 (List.fold_right (fun i acc -> acc ++ cut() ++ flat f false i) top (List.fold_right (fun i acc -> acc ++ cut() ++ flat f true i) mid (List.fold_right (fun i acc -> acc ++ cut() ++ flat f false i) bot (mt())))) let assign_tip_id d id = match d with | { stack = { state_id = None } as top :: _ } -> top.state_id <- Some id | _ -> invalid_arg "assign_tip_id" let cut_at d id = let aux (n, zone) { state_id; data } = if stateid_opt_equal state_id (Some id) then CSig.Stop (n, zone) else CSig.Cont (n + 1, data :: zone) in let n, zone = CList.fold_left_until aux (0, []) d.stack in for _i = 1 to n do ignore(pop d) done; List.rev zone let find_id d f = let top, focus, bot = to_lists d in let pred = function | { state_id = Some id; data } when f id data -> Some id | _ -> None in try CList.find_map pred top, true with Not_found -> try CList.find_map pred focus, false with Not_found -> CList.find_map pred bot, true let before_tip d = let _, focused, rest = to_lists d in match focused with | _:: { state_id = Some id } :: _ -> id, false | _:: { state_id = None } :: _ -> assert false | [] -> raise Not_found | [_] -> match rest with | { state_id = Some id } :: _ -> id, true | { state_id = None } :: _ -> assert false | [] -> raise Not_found let fold_all d a f = let top, focused, bot = to_lists d in let a = List.fold_left (fun a -> flat (f a) false) a top in let a = List.fold_left (fun a -> flat (f a) true) a focused in let a = List.fold_left (fun a -> flat (f a) false) a bot in a coq-8.11.0/ide/coq_commands.ml0000644000175000017500000001715413612664533016034 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* "; "dependent rewrite <-"; "destruct"; "discriminate"; "do"; "double induction"; ]; [ "eapply"; "eauto"; "eauto with"; "eexact"; "elim"; "elim __ using"; "elim __ with"; "elimtype"; "exact"; "exists"; ]; [ "fail"; "field"; "first"; "firstorder"; "firstorder using"; "firstorder with"; "fix"; "fix __ with"; "fold"; "fold __ in"; "functional induction"; ]; [ "generalize"; "generalize dependent"; ]; [ "hnf"; ]; [ "idtac"; "induction"; "info"; "injection"; "instantiate (__:=__)"; "intro"; "intro after"; "intro __ after"; "intros"; "intros until"; "intuition"; "inversion"; "inversion __ in"; "inversion __ using"; "inversion __ using __ in"; "inversion__clear"; "inversion__clear __ in"; ]; [ "jp "; "jp"; ]; [ "lapply"; "lazy"; "lazy in"; "left"; ]; [ "move __ after"; ]; [ "omega"; ]; [ "pattern"; "pose"; "pose __:=__)"; "progress"; ]; [ "quote"; ]; [ "red"; "red in"; "refine"; "reflexivity"; "rename __ into"; "repeat"; "replace __ with"; "rewrite"; "rewrite __ in"; "rewrite <-"; "rewrite <- __ in"; "right"; "ring"; ]; [ "set"; "set (__:=__)"; "setoid__replace"; "setoid__rewrite"; "simpl"; "simpl __ in"; "simple destruct"; "simple induction"; "simple inversion"; "simplify__eq"; "solve"; "split"; (* "split__Rabs"; "split__Rmult"; *) "subst"; "symmetry"; "symmetry in"; ]; [ "tauto"; "transitivity"; "trivial"; "try"; ]; [ "unfold"; "unfold __ in"; ]; ] coq-8.11.0/ide/coq2.ico0000644000175000017500000001114613612664533014372 0ustar treinentreinen(Fhn   (  p ( @ʦ """)))UUUMMMBBB999|PP3f3333f333ff3fffff3f3f̙f3333f3333333333f3333333f3f33ff3f3f3f3333f3333333f3̙33333f333ff3ffffff3f33f3ff3f3f3ffff3fffffffff3fffffff3f̙ffff3ff333f3ff33fff33f3ff̙3f3f3333f333ff3fffff̙̙3̙f̙̙̙3f̙3f3f3333f333ff3fffff3f3f̙3ffffffffff!___www*       ( @ wwww www www ww wxxx x x p pp? pp   p pp( @ʦ """)))UUUMMMBBB999|PP3f3333f333ff3fffff3f3f̙f3333f3333333333f3333333f3f33ff3f3f3f3333f3333333f3̙33333f333ff3ffffff3f33f3ff3f3f3ffff3fffffffff3fffffff3f̙ffff3ff333f3ff33fff33f3ff̙3f3f3333f333ff3fffff̙̙3̙f̙̙̙3f̙3f3f3333f333ff3fffff3f3f̙3ffffffffff!___www      22222222222222222222222222**                  pcoq-8.11.0/ide/ideutils.ml0000644000175000017500000004310013612664533015201 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* incr size; ignore (status_context#push s)), (fun () -> decr size; status_context#pop ()), (fun () -> for _i = 1 to !size do status_context#pop () done; size := 0) type 'a mlist = Nil | Cons of { hd : 'a ; mutable tl : 'a mlist } let enqueue a x = let rec aux x = match x with | Nil -> assert false | Cons p -> match p.tl with | Nil -> p.tl <- Cons { hd = a ; tl = Nil } | _ -> aux p.tl in match !x with | Nil -> x := Cons { hd = a ; tl = Nil } | _ -> aux !x let pop = function | Cons p -> p.tl | Nil -> assert false let flash_info = let queue = ref Nil in let flash_context = status#new_context ~name:"Flash" in let rec process () = match !queue with | Cons { hd = (delay,text) } -> let msg = flash_context#push text in ignore (Glib.Timeout.add ~ms:delay ~callback:(fun () -> flash_context#remove msg; queue := pop !queue; process (); false)) | Nil -> () in fun ?(delay=5000) text -> let processing = !queue <> Nil in enqueue (delay,text) queue; if not processing then process () (* Note: Setting the same attribute with two separate tags appears to use the first value applied and not the second. I saw this trying to set the background color on Windows. A clean fix, if ever needed, would be to combine the attributes of the tags into a single composite tag before applying. This is left as an exercise for the reader. *) let insert_with_tags (buf : #GText.buffer_skel) mark rmark tags text = (* FIXME: LablGTK2 does not export the C insert_with_tags function, so that it has to reimplement its own helper function. Unluckily, it relies on a slow algorithm, so that we have to have our own quicker version here. Alas, it is still much slower than the native version... *) if CList.is_empty tags then buf#insert ~iter:(buf#get_iter_at_mark mark) text else let it = buf#get_iter_at_mark mark in let () = buf#move_mark rmark ~where:it in let () = buf#insert ~iter:(buf#get_iter_at_mark mark) text in let start = buf#get_iter_at_mark mark in let stop = buf#get_iter_at_mark rmark in let iter tag = buf#apply_tag tag ~start ~stop in List.iter iter (List.rev tags) let nl_white_regex = Str.regexp "^\\( *\n *\\)" let diff_regex = Str.regexp "^diff." let insert_xml ?(mark = `INSERT) ?(tags = []) (buf : #GText.buffer_skel) msg = let open Xml_datatype in let dtags = ref [] in let tag name = match GtkText.TagTable.lookup buf#tag_table name with | None -> raise Not_found | Some tag -> new GText.tag tag in let rmark = `MARK (buf#create_mark buf#start_iter) in (* insert the string, but don't apply diff highlights to white space at the begin/end of line *) let rec insert_str tags s = let etags = try List.hd !dtags :: tags with hd -> tags in try let start = Str.search_forward nl_white_regex s 0 in insert_with_tags buf mark rmark etags (String.sub s 0 start); insert_with_tags buf mark rmark tags (Str.matched_group 1 s); let mend = Str.match_end () in insert_str tags (String.sub s mend (String.length s - mend)) with Not_found -> insert_with_tags buf mark rmark etags s in let rec insert tags = function | PCData s -> insert_str tags s | Element (t, _, children) -> let (pfx, tname) = Pp.split_tag t in let is_diff = try let _ = Str.search_forward diff_regex tname 0 in true with Not_found -> false in let (tags, have_tag) = try let t = tag tname in if is_diff && pfx <> Pp.end_pfx then dtags := t :: !dtags; if pfx = "" then ((if is_diff then tags else t :: tags), true) else (tags, true) with Not_found -> (tags, false) in List.iter (fun xml -> insert tags xml) children; if have_tag && is_diff && pfx <> Pp.start_pfx then dtags := (try List.tl !dtags with tl -> []); in let () = try insert tags msg with _ -> () in buf#delete_mark rmark let set_location = ref (function s -> failwith "not ready") let display_location ins = let line = ins#line + 1 in let off = ins#line_offset + 1 in let msg = Printf.sprintf "Line: %5d Char: %3d" line off in !set_location msg (** A utf8 char is either a single byte (ascii char, 0xxxxxxx) or multi-byte (with a leading byte 11xxxxxx and extra bytes 10xxxxxx) *) let is_extra_byte c = ((Char.code c) lsr 6 = 2) (** For a string buffer that may contain utf8 chars, we convert a byte offset into a char offset by only counting char-starting bytes. Normally the string buffer starts with a char-starting byte (buffer produced by a [#get_text]) *) let byte_offset_to_char_offset s byte_offset = let extra_bytes = ref 0 in for i = 0 to min byte_offset (String.length s - 1) do if is_extra_byte s.[i] then incr extra_bytes done; byte_offset - !extra_bytes let glib_utf8_pos_to_offset s ~off = byte_offset_to_char_offset s off let do_convert s = let from_loc () = let _,char_set = Glib.Convert.get_charset () in flash_info ("Converting from locale ("^char_set^")"); Glib.Convert.convert_with_fallback ~to_codeset:"UTF-8" ~from_codeset:char_set s in let from_manual enc = flash_info ("Converting from "^ enc); Glib.Convert.convert s ~to_codeset:"UTF-8" ~from_codeset:enc in let s = if Glib.Utf8.validate s then (Minilib.log "Input is UTF-8"; s) else match encoding#get with |Preferences.Eutf8 | Preferences.Elocale -> from_loc () |Emanual enc -> try from_manual enc with _ -> from_loc () in Utf8_convert.f s let try_convert s = try do_convert s with _ -> "(* Fatal error: wrong encoding in input. \ Please choose a correct encoding in the preference panel.*)";; let export file_name s = let oc = open_out_bin file_name in let ending = line_ending#get in let is_windows = ref false in for i = 0 to String.length s - 1 do match s.[i] with | '\r' -> is_windows := true | '\n' -> begin match ending with | `DEFAULT -> if !is_windows then (output_char oc '\r'; output_char oc '\n') else output_char oc '\n' | `WINDOWS -> output_char oc '\r'; output_char oc '\n' | `UNIX -> output_char oc '\n' end | c -> output_char oc c done; close_out oc let try_export file_name s = let s = try match encoding#get with |Eutf8 -> Minilib.log "UTF-8 is enforced" ; s |Elocale -> let is_unicode,char_set = Glib.Convert.get_charset () in if is_unicode then (Minilib.log "Locale is UTF-8" ; s) else (Minilib.log ("Locale is "^char_set); Glib.Convert.convert_with_fallback ~from_codeset:"UTF-8" ~to_codeset:char_set s) |Emanual enc -> (Minilib.log ("Manual charset is "^ enc); Glib.Convert.convert_with_fallback ~from_codeset:"UTF-8" ~to_codeset:enc s) with e -> let str = Printexc.to_string e in Minilib.log ("Error ("^str^") in transcoding: falling back to UTF-8"); s in try export file_name s; true with e -> Minilib.log (Printexc.to_string e);false type timer = { run : ms:int -> callback:(unit->bool) -> unit; kill : unit -> unit } let mktimer () = let timer = ref None in { run = (fun ~ms ~callback -> timer := Some (GMain.Timeout.add ~ms ~callback)); kill = (fun () -> match !timer with | None -> () | Some id -> (try GMain.Timeout.remove id with Glib.GError _ -> ()); timer := None) } let filter_all_files () = GFile.filter ~name:"All" ~patterns:["*"] () let filter_coq_files () = GFile.filter ~name:"Coq source code" ~patterns:[ "*.v"] () let current_dir () = match project_path#get with | None -> "" | Some dir -> dir let select_file_for_open ~title ?(filter=true) ?parent ?filename () = let file_chooser = GWindow.file_chooser_dialog ~action:`OPEN ~modal:true ~title ?parent () in file_chooser#add_button_stock `CANCEL `CANCEL ; file_chooser#add_select_button_stock `OPEN `OPEN ; if filter then begin file_chooser#add_filter (filter_coq_files ()); file_chooser#add_filter (filter_all_files ()) end; file_chooser#set_default_response `OPEN; let dir = match filename with | None -> current_dir () | Some f -> Filename.dirname f in ignore (file_chooser#set_current_folder dir); let file = match file_chooser#run () with | `OPEN -> begin match file_chooser#filename with | None -> None | Some _ as f -> project_path#set file_chooser#current_folder; f end | `DELETE_EVENT | `CANCEL -> None in file_chooser#destroy (); file let select_file_for_save ~title ?parent ?filename () = let file = ref None in let file_chooser = GWindow.file_chooser_dialog ~action:`SAVE ~modal:true ~title ?parent () in file_chooser#add_button_stock `CANCEL `CANCEL ; file_chooser#add_select_button_stock `SAVE `SAVE ; file_chooser#add_filter (filter_coq_files ()); file_chooser#add_filter (filter_all_files ()); file_chooser#set_do_overwrite_confirmation true; file_chooser#set_default_response `SAVE; let dir,filename = match filename with |None -> current_dir (), "" |Some f -> Filename.dirname f, Filename.basename f in ignore (file_chooser#set_current_folder dir); ignore (file_chooser#set_current_name filename); begin match file_chooser#run () with | `SAVE -> begin file := file_chooser#filename; match !file with None -> () | Some s -> project_path#set file_chooser#current_folder end | `DELETE_EVENT | `CANCEL -> () end ; file_chooser#destroy (); !file let find_tag_start (tag :GText.tag) (it:GText.iter) = let it = it#copy in let tag = Some tag in while not (it#begins_tag tag) && it#nocopy#backward_char do () done; it let find_tag_stop (tag :GText.tag) (it:GText.iter) = let it = it#copy in let tag = Some tag in while not (it#ends_tag tag) && it#nocopy#forward_char do () done; it let find_tag_limits (tag :GText.tag) (it:GText.iter) = (find_tag_start tag it , find_tag_stop tag it) let stock_to_widget ?(size=`BUTTON) s = let img = GMisc.image () in (match size with | `CUSTOM(width,height) -> let opb = img#misc#render_icon ~size:`BUTTON s in let pb = GdkPixbuf.create ~width ~height ~bits:(GdkPixbuf.get_bits_per_sample opb) ~has_alpha:(GdkPixbuf.get_has_alpha opb) () in GdkPixbuf.scale ~width ~height ~dest:pb opb; img#set_pixbuf pb | #Gtk.Tags.icon_size as icon_size -> img#set_stock s; img#set_icon_size icon_size); img#coerce let custom_coqtop = ref None let coqtop_path () = let file = match !custom_coqtop with | Some s -> s | None -> match cmd_coqtop#get with | Some s -> s | None -> try let new_prog = System.get_toplevel_path "coqidetop" in (* The file exists or it is to be found by path *) if Sys.file_exists new_prog || CString.equal Filename.(basename new_prog) new_prog then new_prog else let in_macos_bundle = Filename.concat (Filename.dirname new_prog) (Filename.concat "../Resources/bin" (Filename.basename new_prog)) in if Sys.file_exists in_macos_bundle then in_macos_bundle else "coqidetop.opt" with Not_found -> "coqidetop.opt" in file (* In win32, when a command-line is to be executed via cmd.exe (i.e. Sys.command, Unix.open_process, ...), it cannot contain several quoted "..." zones otherwise some quotes are lost. Solution: we re-quote everything. Reference: http://ss64.com/nt/cmd.html *) let requote cmd = if Sys.os_type = "Win32" then "\""^cmd^"\"" else cmd let textview_width (view : #GText.view_skel) = let rect = view#visible_rect in let pixel_width = Gdk.Rectangle.width rect in let metrics = view#misc#pango_context#get_metrics () in let char_width = GPango.to_pixels metrics#approx_char_width in pixel_width / char_width type logger = Feedback.level -> Pp.t -> unit let default_logger level message = let level = match level with | Feedback.Debug -> `DEBUG | Feedback.Info -> `INFO | Feedback.Notice -> `NOTICE | Feedback.Warning -> `WARNING | Feedback.Error -> `ERROR in Minilib.log_pp ~level message (** {6 File operations} *) (** A customized [stat] function. Exceptions are caught. *) type stats = MTime of float | NoSuchFile | OtherError let stat f = try MTime (Unix.stat f).Unix.st_mtime with | Unix.Unix_error (Unix.ENOENT,_,_) -> NoSuchFile | _ -> OtherError (** I/O utilities Note: In a mono-thread coqide, we use the same buffer for different read operations *) let maxread = 4096 let read_string = Bytes.create maxread let read_buffer = Buffer.create maxread (** Read the content of file [f] and add it to buffer [b]. I/O Exceptions are propagated. *) let read_file name buf = let ic = Util.open_utf8_file_in name in let len = ref 0 in try while len := input ic read_string 0 maxread; !len > 0 do Buffer.add_subbytes buf read_string 0 !len done; close_in ic with e -> close_in ic; raise e (** Read what is available on a gtk channel. This channel should have been set as non-blocking. When there's nothing more to read, the inner loop will be exited via a GError exception concerning a EAGAIN unix error. Anyway, any other exception also stops the read. *) let io_read_all chan = Buffer.clear read_buffer; let read_once () = let len = Glib.Io.read_chars ~buf:read_string ~pos:0 ~len:maxread chan in Buffer.add_subbytes read_buffer read_string 0 len in begin try while true do read_once () done with Glib.GError _ -> () end; Buffer.contents read_buffer (** Run an external command asynchronously *) let run_command display finally cmd = let cin = Unix.open_process_in cmd in let fd = Unix.descr_of_in_channel cin in let () = Unix.set_nonblock fd in let io_chan = Glib.Io.channel_of_descr fd in let all_conds = [`ERR; `HUP; `IN; `NVAL; `PRI] in (* all except `OUT *) let rec has_errors = function | [] -> false | (`IN | `PRI) :: conds -> has_errors conds | e :: _ -> true in let handle_end () = finally (Unix.close_process_in cin); false in let handle_input conds = if has_errors conds then handle_end () else let s = io_read_all io_chan in if s = "" then handle_end () else (display (try_convert s); true) in ignore (Glib.Io.add_watch ~cond:all_conds ~callback:handle_input io_chan) (** Web browsing *) let browse prerr url = let com = Util.subst_command_placeholder cmd_browse#get url in let finally = function | Unix.WEXITED 127 -> prerr ("Could not execute:\n"^com^"\n"^ "check your preferences for setting a valid browser command\n") | _ -> () in run_command (fun _ -> ()) finally com let url_for_keyword = let ht = Hashtbl.create 97 in lazy ( begin try let cin = try let index_urls = Filename.concat (List.find (fun x -> Sys.file_exists (Filename.concat x "index_urls.txt")) (Minilib.coqide_data_dirs ())) "index_urls.txt" in open_in index_urls with Not_found -> raise Exit in try while true do let s = input_line cin in try let i = String.index s ',' in let k = String.sub s 0 i in let u = String.sub s (i + 1) (String.length s - i - 1) in Hashtbl.add ht k u with _ -> Minilib.log "Warning: Cannot parse documentation index file." done with End_of_file -> close_in cin with _ -> Minilib.log "Warning: Cannot find documentation index file." end; Hashtbl.find ht : string -> string) let browse_keyword prerr text = try let u = Lazy.force url_for_keyword text in browse prerr (Coq_config.wwwrefman ^ u) with Not_found -> prerr ("No documentation found for \""^text^"\".\n") let rec is_valid (s : Pp.t) = match Pp.repr s with | Pp.Ppcmd_empty | Pp.Ppcmd_print_break _ | Pp.Ppcmd_force_newline -> true | Pp.Ppcmd_glue l -> List.for_all is_valid l | Pp.Ppcmd_string s -> Glib.Utf8.validate s | Pp.Ppcmd_box (_,s) | Pp.Ppcmd_tag (_,s) -> is_valid s | Pp.Ppcmd_comment s -> List.for_all Glib.Utf8.validate s let validate s = if is_valid s then s else Pp.str "This error massage can't be printed." coq-8.11.0/ide/coq.mli0000644000175000017500000001567513612664533014332 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 'a task (** Monadic return of values as tasks. *) val bind : 'a task -> ('a -> 'b task) -> 'b task (** Monadic binding of tasks *) val lift : (unit -> 'a) -> 'a task (** Return the imperative computation waiting to be processed. *) val seq : unit task -> 'a task -> 'a task (** Sequential composition *) (** {5 Coqtop process management} *) type reset_kind = Planned | Unexpected (** A reset may occur accidentally or voluntarily, so we discriminate between these. *) val is_computing : coqtop -> bool (** Check if coqtop is computing, i.e. already has a current task *) val spawn_coqtop : string list -> coqtop (** Create a coqtop process with some command-line arguments. *) val set_reset_handler : coqtop -> unit task -> unit (** Register a handler called when a coqtop dies (badly or on purpose) *) val set_feedback_handler : coqtop -> (Feedback.feedback -> unit) -> unit (** Register a handler called when coqtop sends a feedback message *) val init_coqtop : coqtop -> unit task -> unit (** Finish initializing a freshly spawned coqtop, by running a first task on it. The task should run its inner continuation at the end. *) val break_coqtop : coqtop -> string list -> unit (** Interrupt the current computation of coqtop or the worker if coqtop it not running. *) val close_coqtop : coqtop -> unit (** Close coqtop. Subsequent requests will be discarded. Hook ignored. *) val reset_coqtop : coqtop -> unit (** Reset coqtop. Pending requests will be discarded. The reset handler of coqtop will be called with [Planned] as first argument *) val get_arguments : coqtop -> string list (** Get the current arguments used by coqtop. *) val set_arguments : coqtop -> string list -> unit (** Set process arguments. This also forces a planned reset. *) (** In win32, sockets are not like regular files *) val gio_channel_of_descr_socket : (Unix.file_descr -> Glib.Io.channel) ref (** {5 Task processing} *) val try_grab : coqtop -> unit task -> (unit -> unit) -> unit (** Try to schedule a task on a coqtop. If coqtop is available, the task callback is run (asynchronously), otherwise the [(unit->unit)] callback is triggered. - If coqtop ever dies during the computation, this function restarts coqtop and calls the restart hook with the fresh coqtop. - If the argument function raises an exception, a coqtop reset occurs. - The task may be discarded if a [close_coqtop] or [reset_coqtop] occurs before its completion. - The task callback should run its inner continuation at the end. *) (** {5 Atomic calls to coqtop} *) (** These atomic calls can be combined to form arbitrary multi-call tasks. They correspond to the protocol calls (cf [Serialize] for more details). Note that each call is asynchronous: it will return immediately, but the inner callback will be executed later to handle the call answer when this answer is available. Except for interp, we use the default logger for any call. *) type 'a query = 'a Interface.value task (** A type abbreviation for coqtop specific answers *) val add : Interface.add_sty -> Interface.add_rty query val edit_at : Interface.edit_at_sty -> Interface.edit_at_rty query val query : Interface.query_sty -> Interface.query_rty query val status : Interface.status_sty -> Interface.status_rty query val goals : Interface.goals_sty -> Interface.goals_rty query val evars : Interface.evars_sty -> Interface.evars_rty query val hints : Interface.hints_sty -> Interface.hints_rty query val mkcases : Interface.mkcases_sty -> Interface.mkcases_rty query val search : Interface.search_sty -> Interface.search_rty query val init : Interface.init_sty -> Interface.init_rty query val stop_worker: Interface.stop_worker_sty-> Interface.stop_worker_rty query (** A specialized version of [raw_interp] dedicated to set/unset options. *) module PrintOpt : sig type 'a t (** Representation of an option *) type 'a descr = { opts : 'a t list; init : 'a; label : string } val bool_items : bool descr list val diff_item : string descr val set : 'a t -> 'a -> unit val printing_unfocused: unit -> bool (** [enforce] transmits to coq the current option values. It is also called by [goals] and [evars] above. *) val enforce : unit task end (** {5 Miscellaneous} *) val short_version : unit -> string (** Return a short phrase identifying coqtop version and date of compilation, as given by the [configure] script. *) val version : unit -> string (** More verbose description, including details about libraries and architecture. *) val filter_coq_opts : string list -> string list (** * Launch a test coqtop processes, ask for a correct coqtop if it fails. @return the list of arguments that coqtop did not understand (the files probably ..). This command may terminate coqide in case of trouble. *) val check_connection : string list -> unit (** Launch a coqtop with the user args in order to be sure that it works, checking in particular that Prelude.vo is found. This command may terminate coqide in case of trouble *) val interrupter : (int -> unit) ref val save_all : (unit -> unit) ref (* Flags to be used for ideslave *) val ideslave_coqtop_flags : string option ref coq-8.11.0/ide/wg_Command.mli0000644000175000017500000000174313612664533015612 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Coq.coqtop -> CoqOps.coqops -> Wg_RoutedMessageViews.message_views_router -> object method new_query : ?command:string -> ?term:string -> unit -> unit method pack_in : (GObj.widget -> unit) -> unit method show : unit method hide : unit method visible : bool end coq-8.11.0/ide/coqide_main.ml0000644000175000017500000000530313612664533015632 0ustar treinentreinen(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* 0 in if level_is `ERROR then `FATAL else if level_is `CRITICAL then `ERROR else if level_is `DEBUG then `DEBUG else if level_is `WARNING then `WARNING else if level_is `MESSAGE then `NOTICE else `INFO in let handler ~level msg = let header = "Coqide internal error: " in match log_level level with |`FATAL -> let () = GToolbox.message_box ~title:"Error" (header ^ msg) in Coqide.crash_save 1 |`ERROR -> if !Flags.debug then GToolbox.message_box ~title:"Error" (header ^ msg) else Printf.eprintf "%s\n" (header ^ msg) |`DEBUG -> Minilib.log msg |level when Sys.os_type = "Win32" -> Minilib.log ~level msg |_ -> Printf.eprintf "%s\n" msg in let catch domain = ignore (Glib.Message.set_log_handler ~domain ~levels:all_levels handler) in List.iter catch ["GLib";"Gtk";"Gdk";"Pango"] let () = catch_gtk_messages () let load_prefs () = Preferences.load_pref ~warn:(fun ~delay -> Ideutils.flash_info ~delay) let () = Ideutils.push_info ("Ready"^ if Preferences.microPG#get then ", [μPG]" else ""); load_prefs (); let argl = List.tl (Array.to_list Sys.argv) in let argl = Coqide.read_coqide_args argl in let files = Coq.filter_coq_opts argl in let args = List.filter (fun x -> not (List.mem x files)) argl in Coq.check_connection args; Coqide.sup_args := args; let w = Coqide.main files in Coqide.set_signal_handlers ~parent:w (); Coqide_os_specific.init (); try GMain.main (); failwith "Gtk loop ended" with e -> Minilib.log ("CoqIde unexpected error:" ^ Printexc.to_string e); Coqide.crash_save 127 coq-8.11.0/ide/coq-ssreflect.lang0000644000175000017500000002303413612664533016446 0ustar treinentreinen *.v \(\* \*\)