aboutsummaryrefslogtreecommitdiff
path: root/kernel/inductive.mli
diff options
context:
space:
mode:
authorbarras2006-05-12 18:50:21 +0000
committerbarras2006-05-12 18:50:21 +0000
commit74a9510f976ed99b19d1081799e79aad09c27cdc (patch)
treeb1224a4e261cbc595359a404ed52f7840d8bc4ad /kernel/inductive.mli
parent041d3b49fc7ca9d0a70f43259e2b099ff21cb9ab (diff)
correction bugs de condition de garde (fix + cofix)
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@8810 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'kernel/inductive.mli')
-rw-r--r--kernel/inductive.mli28
1 files changed, 24 insertions, 4 deletions
diff --git a/kernel/inductive.mli b/kernel/inductive.mli
index a0ff0cefba..b4adbf0933 100644
--- a/kernel/inductive.mli
+++ b/kernel/inductive.mli
@@ -62,10 +62,6 @@ val type_case_branches :
given inductive type. *)
val check_case_info : env -> inductive -> case_info -> unit
-(* Find the ultimate inductive in the [mind_equiv] chain *)
-
-val scrape_mind : env -> mutual_inductive -> mutual_inductive
-
(*s Guard conditions for fix and cofix-points. *)
val check_fix : env -> fixpoint -> unit
val check_cofix : env -> cofixpoint -> unit
@@ -83,3 +79,27 @@ val find_inductive_level : env -> mind_specif -> inductive ->
val is_small_inductive : mind_specif -> bool
val max_inductive_sort : sorts array -> universe
+
+(***************************************************************)
+(* 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;
+ (* inductive of recarg of each fixpoint *)
+ inds : inductive array;
+ (* the recarg information of inductive family *)
+ recvec : wf_paths array;
+ (* dB of variables denoting subterms *)
+ genv : subterm_spec list;
+ }
+
+val subterm_specif : guard_env -> constr -> subterm_spec
+val case_branches_specif : guard_env -> constr -> inductive ->
+ constr array -> (guard_env * constr) array