(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * Copyright INRIA, CNRS and contributors *) (* (* in case the inductive has a type elimination, generates only one induction scheme, the other ones share the same code with the appropriate type *) let sigma, cte = Evd.fresh_constant_instance env sigma cte in let c = mkConstU cte in let t = type_of_constant_in (Global.env()) cte in let (mib,mip) = Global.lookup_inductive ind in let npars = (* if a constructor of [ind] contains a recursive call, the scheme is generalized only wrt recursively uniform parameters *) if (Inductiveops.mis_is_recursive_subset [snd ind] mip.mind_recargs) then mib.mind_nparams_rec else mib.mind_nparams in let sigma, sort = Evd.fresh_sort_in_family sigma sort in let sigma, t', c' = weaken_sort_scheme env sigma false sort npars c t in let sigma = Evd.minimize_universes sigma in (Evarutil.nf_evars_universes sigma c', Evd.evar_universe_context sigma) | None -> let sigma, pind = Evd.fresh_inductive_instance ~rigid:UState.univ_rigid env sigma ind in let sigma, c = build_induction_scheme env sigma pind dep sort in (c, Evd.evar_universe_context sigma) let build_induction_scheme_in_type dep sort ind = let env = Global.env () in let sigma = Evd.from_env env in let sigma, pind = Evd.fresh_inductive_instance ~rigid:UState.univ_rigid env sigma ind in let sigma, c = build_induction_scheme env sigma pind dep sort in c, Evd.evar_universe_context sigma let declare_individual_scheme_object name ?aux f = let f : individual_scheme_object_function = fun _ ind -> f ind in declare_individual_scheme_object name ?aux f let rect_scheme_kind_from_type = declare_individual_scheme_object "_rect_nodep" (fun x -> build_induction_scheme_in_type false InType x) let rect_scheme_kind_from_prop = declare_individual_scheme_object "_rect" ~aux:"_rect_from_prop" (fun x -> build_induction_scheme_in_type false InType x) let rect_dep_scheme_kind_from_type = declare_individual_scheme_object "_rect" ~aux:"_rect_from_type" (fun x -> build_induction_scheme_in_type true InType x) let rec_scheme_kind_from_type = declare_individual_scheme_object "_rec_nodep" ~aux:"_rec_nodep_from_type" (optimize_non_type_induction_scheme rect_scheme_kind_from_type false InSet) let rec_scheme_kind_from_prop = declare_individual_scheme_object "_rec" ~aux:"_rec_from_prop" (optimize_non_type_induction_scheme rect_scheme_kind_from_prop false InSet) let rec_dep_scheme_kind_from_type = declare_individual_scheme_object "_rec" ~aux:"_rec_from_type" (optimize_non_type_induction_scheme rect_dep_scheme_kind_from_type true InSet) let ind_scheme_kind_from_type = declare_individual_scheme_object "_ind_nodep" (optimize_non_type_induction_scheme rec_scheme_kind_from_type false InProp) let sind_scheme_kind_from_type = declare_individual_scheme_object "_sind_nodep" (fun x -> build_induction_scheme_in_type false InSProp x) let ind_dep_scheme_kind_from_type = declare_individual_scheme_object "_ind" ~aux:"_ind_from_type" (optimize_non_type_induction_scheme rec_dep_scheme_kind_from_type true InProp) let sind_dep_scheme_kind_from_type = declare_individual_scheme_object "_sind" ~aux:"_sind_from_type" (fun x -> build_induction_scheme_in_type true InSProp x) let ind_scheme_kind_from_prop = declare_individual_scheme_object "_ind" ~aux:"_ind_from_prop" (optimize_non_type_induction_scheme rec_scheme_kind_from_prop false InProp) let sind_scheme_kind_from_prop = declare_individual_scheme_object "_sind" ~aux:"_sind_from_prop" (fun x -> build_induction_scheme_in_type false InSProp x) let nondep_elim_scheme from_kind to_kind = match from_kind, to_kind with | InProp, InProp -> ind_scheme_kind_from_prop | InProp, InSProp -> sind_scheme_kind_from_prop | InProp, InSet -> rec_scheme_kind_from_prop | InProp, InType -> rect_scheme_kind_from_prop | _ , InProp -> ind_scheme_kind_from_type | _ , InSProp -> sind_scheme_kind_from_type | _ , InSet -> rec_scheme_kind_from_type | _ , InType -> rect_scheme_kind_from_type (* Case analysis *) let build_case_analysis_scheme_in_type dep sort ind = let env = Global.env () in let sigma = Evd.from_env env in let (sigma, indu) = Evd.fresh_inductive_instance env sigma ind in let (sigma, c) = build_case_analysis_scheme env sigma indu dep sort in c, Evd.evar_universe_context sigma let case_scheme_kind_from_type = declare_individual_scheme_object "_case_nodep" (fun x -> build_case_analysis_scheme_in_type false InType x) let case_scheme_kind_from_prop = declare_individual_scheme_object "_case" ~aux:"_case_from_prop" (fun x -> build_case_analysis_scheme_in_type false InType x) let case_dep_scheme_kind_from_type = declare_individual_scheme_object "_case" ~aux:"_case_from_type" (fun x -> build_case_analysis_scheme_in_type true InType x) let case_dep_scheme_kind_from_type_in_prop = declare_individual_scheme_object "_casep_dep" (fun x -> build_case_analysis_scheme_in_type true InProp x) let case_dep_scheme_kind_from_prop = declare_individual_scheme_object "_case_dep" (fun x -> build_case_analysis_scheme_in_type true InType x) let case_dep_scheme_kind_from_prop_in_prop = declare_individual_scheme_object "_casep" (fun x -> build_case_analysis_scheme_in_type true InProp x)