@@ -50,6 +50,8 @@ Inductive type_error :=
5050| NotEnoughFuel (n : nat)
5151| NotSupported (s : string).
5252
53+
54+
5355Definition string_of_type_error (e : type_error) : string :=
5456 match e with
5557 | UnboundRel n => "Unboound rel " ^ string_of_nat n
@@ -106,27 +108,23 @@ Section Lookups.
106108 end .
107109
108110 Definition lookup_constant_type cst u :=
109- match lookup_env Σ cst with
110- | Some (ConstantDecl {| cst_type := ty; cst_universes := uctx |}) =>
111- ret (subst_instance u ty)
111+ match lookup_constant_type Σ cst u with
112+ | Some res =>
113+ ret res
112114 | _ => raise (UndeclaredConstant cst)
113115 end .
114116
115117 Definition lookup_constant_type_cstrs cst u :=
116- match lookup_env Σ cst with
117- | Some (ConstantDecl {| cst_type := ty; cst_universes := uctx |}) =>
118- let cstrs := polymorphic_constraints uctx in
119- ret (subst_instance u ty, subst_instance_cstrs u cstrs)
120- | _ => raise (UndeclaredConstant cst)
118+ match lookup_constant_type_cstrs Σ cst u with
119+ | Some res =>
120+ ret res
121+ | _ => raise (UndeclaredConstant cst)
121122 end .
122123
123124 Definition lookup_ind_decl ind i :=
124- match lookup_env Σ ind with
125- | Some (InductiveDecl mdecl) =>
126- match nth_error mdecl.(ind_bodies) i with
127- | Some body => ret (mdecl, body)
128- | None => raise (UndeclaredInductive (mkInd ind i))
129- end
125+ match lookup_ind_decl Σ ind i with
126+ | Some res =>
127+ ret res
130128 | _ => raise (UndeclaredInductive (mkInd ind i))
131129 end .
132130
@@ -262,80 +260,13 @@ Section Reduce.
262260 res <- reduce_stack Γ n c [] ;;
263261 ret (zip res).
264262
265- Definition fix_decls (l : mfixpoint term) :=
266- let fix aux acc ds :=
267- match ds with
268- | nil => acc
269- | d :: ds => aux (vass d.(dname) d.(dtype) :: acc) ds
270- end
271- in aux [] l.
272-
273- Definition rebuild_case_predicate_ctx ind (p : predicate term) : context :=
274- match lookup_ind_decl Σ (inductive_mind ind) (inductive_ind ind) with
275- | TypeError _ => []
276- | Checked (mib, oib) => case_predicate_context ind mib oib p
277- end .
278-
279- Definition map_context_with_binders (f : context -> term -> term) (c : context) Γ : context :=
280- fold_left (fun acc decl => map_decl (f (Γ ,,, acc)) decl :: acc) (List.rev c) [].
281-
282- Definition map_predicate_with_binders (f : context -> term -> term) Γ ind (p : predicate term) :=
283- let pctx := rebuild_case_predicate_ctx ind p in
284- let Γparams := map_context_with_binders f pctx Γ in
285- {| pparams := map (f Γ) p.(pparams);
286- puinst := p.(puinst);
287- pcontext := p.(pcontext);
288- preturn := f (Γparams ++ Γ) (preturn p) |}.
289-
290- Definition rebuild_case_branch_ctx ind i p br :=
291- match lookup_constructor_decl Σ (inductive_mind ind) (inductive_ind ind) i with
292- | TypeError _ => []
293- | Checked (mib, cdecl) => case_branch_context ind mib cdecl p br
294- end .
295-
296- Definition map_case_branch_with_binders ind i (f : context -> term -> term) Γ p br :=
297- let ctx := rebuild_case_branch_ctx ind i p br in
298- map_branch (f (Γ ,,, ctx)) br.
299-
300- Definition map_constr_with_binders (f : context -> term -> term) Γ (t : term) : term :=
301- match t with
302- | tRel i => t
303- | tEvar ev args => tEvar ev (List.map (f Γ) args)
304- | tLambda na T M =>
305- let T' := f Γ T in
306- tLambda na T' (f (Γ,, vass na T') M)
307- | tApp u v => tApp (f Γ u) (List.map (f Γ) v)
308- | tProd na A B =>
309- let A' := f Γ A in
310- tProd na A' (f (Γ ,, vass na A') B)
311- | tCast c kind t => tCast (f Γ c) kind (f Γ t)
312- | tLetIn na b t c =>
313- let b' := f Γ b in
314- let t' := f Γ t in
315- tLetIn na b' t' (f (Γ ,, vdef na b' t') c)
316- | tCase ci p c brs =>
317- let p' := map_predicate_with_binders f Γ ci.(ci_ind) p in
318- let brs' := mapi (fun i x => map_case_branch_with_binders ci.(ci_ind) i f Γ p' x) brs in
319- tCase ci p' (f Γ c) brs'
320- | tProj p c => tProj p (f Γ c)
321- | tFix mfix idx =>
322- let Γ' := fix_decls mfix ++ Γ in
323- let mfix' := List.map (map_def (f Γ) (f Γ')) mfix in
324- tFix mfix' idx
325- | tCoFix mfix k =>
326- let Γ' := fix_decls mfix ++ Γ in
327- let mfix' := List.map (map_def (f Γ) (f Γ')) mfix in
328- tCoFix mfix' k
329- | x => x
330- end .
331-
332263 Fixpoint reduce_opt Γ n c :=
333264 match n with
334265 | 0 => None
335266 | S n =>
336267 match reduce_stack_term Γ n c with
337268 | Some c' =>
338- Some (map_constr_with_binders
269+ Some (map_term_with_context Σ
339270 (fun Γ t => match reduce_opt Γ n t with
340271 | Some t => t
341272 | None => t end ) Γ c')
@@ -776,7 +707,7 @@ Section Typecheck.
776707 (** TODO check branches *)
777708 let '(ind, u, args) := indargs in
778709 if eq_inductive ind ci.(ci_ind) then
779- let pctx := rebuild_case_predicate_ctx Σ ind p in
710+ let pctx := rebuild_case_predicate_ctx_with_context Σ ind p in
780711 let ptm := it_mkLambda_or_LetIn pctx p.(preturn) in
781712 ret (tApp ptm (List.skipn ci.(ci_npar) args ++ [c]))
782713 else
0 commit comments