aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2007-12-31 13:11:55 +0000
committerGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2007-12-31 13:11:55 +0000
commit5aab6b96318d440f818fdf2f5bea69ad5dcda431 (patch)
tree0d0689ab04ffbc471b5e046c670ffe9de21641c5
parent932d9dbc302b2c530aef8f1da6c7b36e228aa1f9 (diff)
Merged revisions 10358-10362,10365,10371-10373,10377,10383-10384,10394-10395,10398,10403-10408 via svnmerge from
svn+ssh://msozeau@scm.gforge.inria.fr/svn/coq/branches/TypeClasses ........ r10358 | msozeau | 2007-12-10 15:42:53 +0100 (Mon, 10 Dec 2007) | 1 line Comment grammar error ........ r10359 | msozeau | 2007-12-10 16:04:09 +0100 (Mon, 10 Dec 2007) | 7 lines The initial Type Classes patch. This patch introduces type classes and instance definitions a la Haskell. Technically, it uses the implicit arguments mechanism which was extended a bit. The patch also introduces a notation for explicitely marking implicit, maximally inserted parameters. It includes the tactic redefinition code too (Ltac tac ::= foo redefines tac). ........ r10360 | msozeau | 2007-12-10 16:14:30 +0100 (Mon, 10 Dec 2007) | 1 line Fix interface ........ r10361 | msozeau | 2007-12-10 16:28:19 +0100 (Mon, 10 Dec 2007) | 1 line Fix more xlate code ........ r10362 | msozeau | 2007-12-11 02:00:53 +0100 (Tue, 11 Dec 2007) | 3 lines Update coqdoc for type classes, fix proof state not being displayed on Next Obligation. ........ r10365 | msozeau | 2007-12-11 14:22:35 +0100 (Tue, 11 Dec 2007) | 3 lines Bug fixes in Instance decls. ........ r10371 | msozeau | 2007-12-12 21:17:30 +0100 (Wed, 12 Dec 2007) | 3 lines Streamline typeclass context implementation, prepare for class binders in proof statements. ........ r10372 | msozeau | 2007-12-12 22:03:38 +0100 (Wed, 12 Dec 2007) | 1 line Minor cosmetic fixes: allow sorts as typeclass param instances without parens and infer more types in class definitions ........ r10373 | msozeau | 2007-12-13 00:35:09 +0100 (Thu, 13 Dec 2007) | 2 lines Better names in g_vernac, binders in Lemmas and Context [] to introduce a typeclass context. ........ r10377 | msozeau | 2007-12-13 18:34:33 +0100 (Thu, 13 Dec 2007) | 1 line Stupid bug ........ r10383 | msozeau | 2007-12-16 00:04:48 +0100 (Sun, 16 Dec 2007) | 1 line Bug fixes in name handling and implicits, new syntax for using implicit mode in typeclass constraints ........ r10384 | msozeau | 2007-12-16 15:53:24 +0100 (Sun, 16 Dec 2007) | 1 line Streamlined implementation of instances again, the produced typeclass is a typeclass constraint. Added corresponding implicit/explicit behaviors ........ r10394 | msozeau | 2007-12-18 23:42:56 +0100 (Tue, 18 Dec 2007) | 4 lines Various fixes for implicit arguments, new "Enriching" kw to just enrich existing sets of impl args. New syntax !a to force an argument, even if not dependent. New tactic clrewrite using a setoid typeclass implementation to do setoid_rewrite under compatible morphisms... very experimental. Other bugs related to naming in typeclasses fixed. ........ r10395 | msozeau | 2007-12-19 17:11:55 +0100 (Wed, 19 Dec 2007) | 3 lines Progress on setoids using type classes, recognize setoid equalities in hyps better. Streamline implementation to return more information when resolving setoids (return the results setoid). ........ r10398 | msozeau | 2007-12-20 10:18:19 +0100 (Thu, 20 Dec 2007) | 1 line Syntax change, more like Coq ........ r10403 | msozeau | 2007-12-21 22:30:35 +0100 (Fri, 21 Dec 2007) | 1 line Add right-to-left rewriting in class_setoid, fix some discharge/substitution bug, adapt test-suite to latest syntax ........ r10404 | msozeau | 2007-12-24 21:47:58 +0100 (Mon, 24 Dec 2007) | 2 lines Work on type classes based rewrite tactic. ........ r10405 | msozeau | 2007-12-27 18:51:32 +0100 (Thu, 27 Dec 2007) | 2 lines Better evar handling in pretyping, reorder theories/Program and add some tactics for dealing with subsets. ........ r10406 | msozeau | 2007-12-27 18:52:05 +0100 (Thu, 27 Dec 2007) | 1 line Forgot to add a file ........ r10407 | msozeau | 2007-12-29 17:19:54 +0100 (Sat, 29 Dec 2007) | 4 lines Generalize usage of implicit arguments in terms, up to rawconstr. Binders are decorated with binding info, either Implicit or Explicit for rawconstr. Factorizes code for typeclasses, topconstrs decorations are Default (impl|expl) or TypeClass (impl|expl) and implicit quantification is resolve at internalization time, getting rid of the arbitrary prenex restriction on contexts. ........ r10408 | msozeau | 2007-12-31 00:58:50 +0100 (Mon, 31 Dec 2007) | 4 lines Fix parsing of subset binders, bugs in subtac_cases and handling of mutual defs obligations. Add useful tactics to Program.Subsets. ........ git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@10410 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--Makefile.common35
-rw-r--r--contrib/first-order/instances.ml4
-rw-r--r--contrib/funind/indfun.ml34
-rw-r--r--contrib/funind/indfun_common.ml4
-rw-r--r--contrib/funind/indfun_main.ml42
-rw-r--r--contrib/funind/merge.ml8
-rw-r--r--contrib/funind/rawterm_to_relation.ml16
-rw-r--r--contrib/funind/rawtermops.ml62
-rw-r--r--contrib/interface/name_to_ast.ml6
-rw-r--r--contrib/interface/xlate.ml26
-rw-r--r--contrib/recdef/recdef.ml42
-rw-r--r--contrib/subtac/eterm.ml5
-rw-r--r--contrib/subtac/eterm.mli6
-rw-r--r--contrib/subtac/g_subtac.ml421
-rw-r--r--contrib/subtac/subtac.ml52
-rw-r--r--contrib/subtac/subtac_cases.ml14
-rw-r--r--contrib/subtac/subtac_classes.ml187
-rw-r--r--contrib/subtac/subtac_classes.mli39
-rw-r--r--contrib/subtac/subtac_command.ml28
-rw-r--r--contrib/subtac/subtac_obligations.ml103
-rw-r--r--contrib/subtac/subtac_obligations.mli13
-rw-r--r--contrib/subtac/subtac_pretyping.ml46
-rw-r--r--contrib/subtac/subtac_pretyping.mli10
-rw-r--r--contrib/subtac/subtac_pretyping_F.ml43
-rw-r--r--contrib/subtac/subtac_utils.ml22
-rw-r--r--contrib/xml/xmlcommand.ml5
-rw-r--r--dev/include2
-rw-r--r--interp/constrextern.ml44
-rw-r--r--interp/constrintern.ml106
-rw-r--r--interp/constrintern.mli2
-rw-r--r--interp/implicit_quantifiers.ml215
-rw-r--r--interp/implicit_quantifiers.mli51
-rw-r--r--interp/reserve.ml6
-rw-r--r--interp/topconstr.ml78
-rw-r--r--interp/topconstr.mli25
-rw-r--r--kernel/term.mli4
-rw-r--r--lib/util.ml8
-rw-r--r--lib/util.mli2
-rw-r--r--library/decl_kinds.ml3
-rw-r--r--library/impargs.ml134
-rw-r--r--library/impargs.mli12
-rw-r--r--parsing/g_constr.ml466
-rw-r--r--parsing/g_ltac.ml413
-rw-r--r--parsing/g_tactic.ml44
-rw-r--r--parsing/g_vernac.ml477
-rw-r--r--parsing/g_xml.ml44
-rw-r--r--parsing/pcoq.ml43
-rw-r--r--parsing/pcoq.mli5
-rw-r--r--parsing/ppconstr.ml105
-rw-r--r--parsing/pptactic.ml9
-rw-r--r--parsing/ppvernac.ml66
-rw-r--r--parsing/prettyp.ml19
-rw-r--r--parsing/prettyp.mli5
-rw-r--r--parsing/q_constr.ml46
-rw-r--r--parsing/q_coqast.ml415
-rw-r--r--pretyping/cases.ml4
-rw-r--r--pretyping/detyping.ml30
-rw-r--r--pretyping/detyping.mli2
-rw-r--r--pretyping/indrec.ml8
-rw-r--r--pretyping/pattern.ml6
-rw-r--r--pretyping/pretyping.ml34
-rw-r--r--pretyping/rawterm.ml28
-rw-r--r--pretyping/rawterm.mli8
-rw-r--r--pretyping/typeclasses.ml311
-rw-r--r--pretyping/typeclasses.mli70
-rw-r--r--pretyping/typeclasses_errors.ml42
-rw-r--r--pretyping/typeclasses_errors.mli40
-rw-r--r--proofs/refiner.mli1
-rw-r--r--tactics/class_setoid.ml4224
-rw-r--r--tactics/decl_interp.ml20
-rw-r--r--tactics/tacinterp.ml63
-rw-r--r--tactics/tacinterp.mli2
-rw-r--r--tactics/tactics.ml5
-rw-r--r--tactics/tactics.mli7
-rw-r--r--test-suite/typeclasses/NewSetoid.v74
-rw-r--r--theories/Arith/Even.v2
-rw-r--r--theories/Classes/Init.v1
-rw-r--r--theories/Classes/Init.v.d1
-rw-r--r--theories/Classes/Setoid.v192
-rw-r--r--theories/Classes/SetoidTactics.v37
-rw-r--r--theories/Lists/List.v2
-rw-r--r--theories/Program/Equality.v47
-rw-r--r--theories/Program/FunctionalExtensionality.v12
-rw-r--r--theories/Program/Program.v1
-rw-r--r--theories/Program/Subset.v141
-rw-r--r--theories/Program/Tactics.v24
-rw-r--r--theories/Program/Utils.v18
-rw-r--r--theories/Program/Wf.v16
-rw-r--r--theories/Setoids/Setoid.v1
-rw-r--r--theories/Setoids/Setoid_tac.v1
-rw-r--r--tools/coqdoc/output.ml4
-rw-r--r--tools/coqdoc/pretty.mll12
-rw-r--r--toplevel/cerrors.ml2
-rw-r--r--toplevel/classes.ml505
-rw-r--r--toplevel/classes.mli65
-rw-r--r--toplevel/command.ml61
-rw-r--r--toplevel/command.mli12
-rw-r--r--toplevel/coqinit.ml1
-rw-r--r--toplevel/himsg.ml33
-rw-r--r--toplevel/himsg.mli3
-rw-r--r--toplevel/record.ml3
-rw-r--r--toplevel/record.mli2
-rw-r--r--toplevel/vernacentries.ml44
-rw-r--r--toplevel/vernacexpr.ml30
-rw-r--r--toplevel/whelp.ml46
105 files changed, 3489 insertions, 651 deletions
diff --git a/Makefile.common b/Makefile.common
index bc583c8a2..84a15bf83 100644
--- a/Makefile.common
+++ b/Makefile.common
@@ -105,6 +105,7 @@ PRETYPING:=\
pretyping/retyping.cmo pretyping/cbv.cmo \
pretyping/pretype_errors.cmo pretyping/recordops.cmo pretyping/typing.cmo \
pretyping/tacred.cmo pretyping/evarutil.cmo pretyping/evarconv.cmo \
+ pretyping/typeclasses_errors.cmo pretyping/typeclasses.cmo \
pretyping/classops.cmo pretyping/coercion.cmo \
pretyping/unification.cmo pretyping/clenv.cmo \
pretyping/rawterm.cmo pretyping/pattern.cmo \
@@ -115,7 +116,7 @@ INTERP:=\
parsing/lexer.cmo interp/topconstr.cmo interp/ppextend.cmo \
interp/notation.cmo \
interp/genarg.cmo interp/syntax_def.cmo interp/reserve.cmo \
- library/impargs.cmo interp/constrintern.cmo \
+ library/impargs.cmo interp/implicit_quantifiers.cmo interp/constrintern.cmo \
interp/modintern.cmo interp/constrextern.cmo interp/coqlib.cmo \
toplevel/discharge.cmo library/declare.cmo
@@ -154,11 +155,11 @@ TACTICS:=\
tactics/decl_interp.cmo tactics/decl_proof_instr.cmo
TOPLEVEL:=\
- toplevel/himsg.cmo toplevel/cerrors.cmo toplevel/class.cmo \
- toplevel/vernacexpr.cmo toplevel/metasyntax.cmo \
+ toplevel/himsg.cmo toplevel/cerrors.cmo \
+ toplevel/class.cmo toplevel/vernacexpr.cmo toplevel/metasyntax.cmo \
toplevel/auto_ind_decl.cmo \
toplevel/command.cmo toplevel/record.cmo \
- parsing/ppvernac.cmo \
+ parsing/ppvernac.cmo toplevel/classes.cmo \
toplevel/vernacinterp.cmo toplevel/mltop.cmo \
toplevel/vernacentries.cmo toplevel/whelp.cmo toplevel/vernac.cmo \
toplevel/line_oriented_parser.cmo toplevel/protectedtoplevel.cmo \
@@ -167,7 +168,7 @@ TOPLEVEL:=\
HIGHTACTICS:=\
tactics/refine.cmo tactics/extraargs.cmo \
- tactics/extratactics.cmo tactics/eauto.cmo
+ tactics/extratactics.cmo tactics/eauto.cmo tactics/class_setoid.cmo
SPECTAC:= tactics/tauto.ml4 tactics/eqdecide.ml4
USERTAC:= $(SPECTAC)
@@ -251,10 +252,10 @@ SUBTACCMO:=contrib/subtac/subtac_utils.cmo contrib/subtac/eterm.cmo \
contrib/subtac/subtac_errors.cmo contrib/subtac/subtac_coercion.cmo \
contrib/subtac/subtac_obligations.cmo contrib/subtac/subtac_cases.cmo \
contrib/subtac/subtac_pretyping_F.cmo contrib/subtac/subtac_pretyping.cmo \
- contrib/subtac/subtac_command.cmo contrib/subtac/subtac.cmo \
+ contrib/subtac/subtac_command.cmo contrib/subtac/subtac_classes.cmo \
+ contrib/subtac/subtac.cmo \
contrib/subtac/g_subtac.cmo
-
RTAUTOCMO:=contrib/rtauto/proof_search.cmo contrib/rtauto/refl_tauto.cmo \
contrib/rtauto/g_rtauto.cmo
@@ -415,15 +416,15 @@ PRINTERSCMO:=\
pretyping/reductionops.cmo pretyping/inductiveops.cmo \
pretyping/retyping.cmo pretyping/cbv.cmo \
pretyping/pretype_errors.cmo pretyping/recordops.cmo pretyping/typing.cmo \
- pretyping/evarutil.cmo pretyping/evarconv.cmo \
- pretyping/tacred.cmo pretyping/classops.cmo pretyping/detyping.cmo \
- pretyping/indrec.cmo pretyping/coercion.cmo \
+ pretyping/evarutil.cmo pretyping/evarconv.cmo pretyping/tacred.cmo \
+ pretyping/classops.cmo pretyping/typeclasses_errors.cmo pretyping/typeclasses.cmo \
+ pretyping/detyping.cmo pretyping/indrec.cmo pretyping/coercion.cmo \
pretyping/unification.cmo pretyping/cases.cmo \
pretyping/pretyping.cmo pretyping/clenv.cmo pretyping/pattern.cmo \
parsing/lexer.cmo interp/ppextend.cmo interp/genarg.cmo \
interp/topconstr.cmo interp/notation.cmo interp/reserve.cmo \
- library/impargs.cmo\
- interp/constrextern.cmo interp/syntax_def.cmo interp/constrintern.cmo \
+ library/impargs.cmo interp/constrextern.cmo \
+ interp/syntax_def.cmo interp/implicit_quantifiers.cmo interp/constrintern.cmo \
proofs/proof_trees.cmo proofs/logic.cmo proofs/refiner.cmo \
proofs/tacexpr.cmo \
proofs/evar_refiner.cmo proofs/pfedit.cmo proofs/tactic_debug.cmo \
@@ -432,8 +433,8 @@ PRINTERSCMO:=\
parsing/printer.cmo parsing/pptactic.cmo \
parsing/ppdecl_proof.cmo \
parsing/tactic_printer.cmo \
- parsing/egrammar.cmo toplevel/himsg.cmo \
- toplevel/cerrors.cmo toplevel/vernacexpr.cmo toplevel/vernacinterp.cmo \
+ parsing/egrammar.cmo toplevel/himsg.cmo toplevel/cerrors.cmo \
+ toplevel/vernacexpr.cmo toplevel/vernacinterp.cmo \
dev/top_printers.cmo
###########################################################################
@@ -705,11 +706,13 @@ SETOIDSVO:= theories/Setoids/Setoid_tac.vo theories/Setoids/Setoid_Prop.vo theor
UNICODEVO:= theories/Unicode/Utf8.vo
+CLASSESVO:= theories/Classes/Init.vo theories/Classes/Setoid.vo theories/Classes/SetoidTactics.v
+
THEORIESVO:=\
$(INITVO) $(LOGICVO) $(ARITHVO) $(BOOLVO) $(NARITHVO) $(ZARITHVO) \
$(SETOIDSVO) $(LISTSVO) $(STRINGSVO) $(SETSVO) $(FSETSVO) $(INTMAPVO) \
$(RELATIONSVO) $(WELLFOUNDEDVO) $(REALSVO) $(SORTINGVO) $(QARITHVO) \
- $(INTSVO) $(NUMBERSVO) $(UNICODEVO)
+ $(INTSVO) $(NUMBERSVO) $(UNICODEVO) $(CLASSESVO)
THEORIESLIGHTVO:= $(INITVO) $(LOGICVO) $(ARITHVO)
@@ -758,7 +761,7 @@ CCVO:=
DPVO:=contrib/dp/Dp.vo
-SUBTACVO:=theories/Program/Tactics.vo theories/Program/Equality.vo \
+SUBTACVO:=theories/Program/Tactics.vo theories/Program/Equality.vo theories/Program/Subset.vo \
theories/Program/Utils.vo theories/Program/Wf.vo theories/Program/Program.vo \
theories/Program/FunctionalExtensionality.vo
diff --git a/contrib/first-order/instances.ml b/contrib/first-order/instances.ml
index 8eeb8b642..56cea8e07 100644
--- a/contrib/first-order/instances.ml
+++ b/contrib/first-order/instances.ml
@@ -125,9 +125,9 @@ let mk_open_instance id gl m t=
let rec raux n t=
if n=0 then t else
match t with
- RLambda(loc,name,_,t0)->
+ RLambda(loc,name,k,_,t0)->
let t1=raux (n-1) t0 in
- RLambda(loc,name,RHole (dummy_loc,Evd.BinderType name),t1)
+ RLambda(loc,name,k,RHole (dummy_loc,Evd.BinderType name),t1)
| _-> anomaly "can't happen" in
let ntt=try
Pretyping.Default.understand evmap env (raux m rawt)
diff --git a/contrib/funind/indfun.ml b/contrib/funind/indfun.ml
index 612be05e3..97f7c1d97 100644
--- a/contrib/funind/indfun.ml
+++ b/contrib/funind/indfun.ml
@@ -139,8 +139,8 @@ type newfixpoint_expr =
let rec abstract_rawconstr c = function
| [] -> c
| Topconstr.LocalRawDef (x,b)::bl -> Topconstr.mkLetInC(x,b,abstract_rawconstr c bl)
- | Topconstr.LocalRawAssum (idl,t)::bl ->
- List.fold_right (fun x b -> Topconstr.mkLambdaC([x],t,b)) idl
+ | Topconstr.LocalRawAssum (idl,k,t)::bl ->
+ List.fold_right (fun x b -> Topconstr.mkLambdaC([x],k,t,b)) idl
(abstract_rawconstr c bl)
let interp_casted_constr_with_implicits sigma env impls c =
@@ -213,7 +213,7 @@ let rec is_rec names =
| RRec _ -> error "RRec not handled"
| RIf(_,b,_,lhs,rhs) ->
(lookup names b) || (lookup names lhs) || (lookup names rhs)
- | RLetIn(_,na,t,b) | RLambda(_,na,t,b) | RProd(_,na,t,b) ->
+ | RLetIn(_,na,t,b) | RLambda(_,na,_,t,b) | RProd(_,na,_,t,b) ->
lookup names t || lookup (Nameops.name_fold Idset.remove na names) b
| RLetTuple(_,nal,_,t,b) -> lookup names t ||
lookup
@@ -434,7 +434,7 @@ let register_mes fname rec_impls wf_mes_expr wf_arg using_lemmas args ret_type b
| None ->
begin
match args with
- | [Topconstr.LocalRawAssum ([(_,Name x)],t)] -> t,x
+ | [Topconstr.LocalRawAssum ([(_,Name x)],k,t)] -> t,x
| _ -> error "Recursive argument must be specified"
end
| Some wf_args ->
@@ -442,7 +442,7 @@ let register_mes fname rec_impls wf_mes_expr wf_arg using_lemmas args ret_type b
match
List.find
(function
- | Topconstr.LocalRawAssum(l,t) ->
+ | Topconstr.LocalRawAssum(l,k,t) ->
List.exists
(function (_,Name id) -> id = wf_args | _ -> false)
l
@@ -450,7 +450,7 @@ let register_mes fname rec_impls wf_mes_expr wf_arg using_lemmas args ret_type b
)
args
with
- | Topconstr.LocalRawAssum(_,t) -> t,wf_args
+ | Topconstr.LocalRawAssum(_,k,t) -> t,wf_args
| _ -> assert false
with Not_found -> assert false
in
@@ -462,7 +462,7 @@ let register_mes fname rec_impls wf_mes_expr wf_arg using_lemmas args ret_type b
let fun_from_mes =
let applied_mes =
Topconstr.mkAppC(wf_mes_expr,[Topconstr.mkIdentC wf_arg]) in
- Topconstr.mkLambdaC ([(dummy_loc,Name wf_arg)],wf_arg_type,applied_mes)
+ Topconstr.mkLambdaC ([(dummy_loc,Name wf_arg)],Topconstr.default_binder_kind,wf_arg_type,applied_mes)
in
let wf_rel_from_mes =
Topconstr.mkAppC(Topconstr.mkRefC ltof,[wf_arg_type;fun_from_mes])
@@ -570,11 +570,11 @@ let rec add_args id new_args b =
CArrow(loc,add_args id new_args b1, add_args id new_args b2)
| CProdN(loc,nal,b1) ->
CProdN(loc,
- List.map (fun (nal,b2) -> (nal,add_args id new_args b2)) nal,
+ List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
add_args id new_args b1)
| CLambdaN(loc,nal,b1) ->
CLambdaN(loc,
- List.map (fun (nal,b2) -> (nal,add_args id new_args b2)) nal,
+ List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
add_args id new_args b1)
| CLetIn(loc,na,b1,b2) ->
CLetIn(loc,na,add_args id new_args b1,add_args id new_args b2)
@@ -644,13 +644,15 @@ let rec chop_n_arrow n t =
let new_n =
let rec aux (n:int) = function
[] -> n
- | (nal,t'')::nal_ta' ->
+ | (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' = Topconstr.CProdN(dummy_loc,((snd (list_chop n nal)),t'')::nal_ta',t')
+ let new_t' =
+ Topconstr.CProdN(dummy_loc,
+ ((snd (list_chop n nal)),k,t'')::nal_ta',t')
in
raise (Stop new_t')
in
@@ -668,12 +670,12 @@ let rec get_args b t : Topconstr.local_binder list *
| Topconstr.CLambdaN (loc, (nal_ta), b') ->
begin
let n =
- (List.fold_left (fun n (nal,_) ->
+ (List.fold_left (fun n (nal,_,_) ->
n+List.length nal) 0 nal_ta )
in
let nal_tas,b'',t'' = get_args b' (chop_n_arrow n t) in
- (List.map (fun (nal,ta) ->
- (Topconstr.LocalRawAssum (nal,ta))) nal_ta)@nal_tas, b'',t''
+ (List.map (fun (nal,k,ta) ->
+ (Topconstr.LocalRawAssum (nal,k,ta))) nal_ta)@nal_tas, b'',t''
end
| _ -> [],b,t
@@ -716,7 +718,7 @@ let make_graph (f_ref:global_reference) =
(List.map
(function
| Topconstr.LocalRawDef (na,_)-> []
- | Topconstr.LocalRawAssum (nal,_) -> nal
+ | Topconstr.LocalRawAssum (nal,_,_) -> nal
)
bl
)
@@ -730,7 +732,7 @@ let make_graph (f_ref:global_reference) =
(List.map
(function
| Topconstr.LocalRawDef (na,_)-> []
- | Topconstr.LocalRawAssum (nal,_) ->
+ | Topconstr.LocalRawAssum (nal,_,_) ->
List.map
(fun (loc,n) ->
CRef(Libnames.Ident(loc, Nameops.out_name n)))
diff --git a/contrib/funind/indfun_common.ml b/contrib/funind/indfun_common.ml
index 5ad4632e4..78e2c4408 100644
--- a/contrib/funind/indfun_common.ml
+++ b/contrib/funind/indfun_common.ml
@@ -76,7 +76,7 @@ let chop_rlambda_n =
then List.rev acc,rt
else
match rt with
- | Rawterm.RLambda(_,name,t,b) -> chop_lambda_n ((name,t,false)::acc) (n-1) b
+ | Rawterm.RLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,false)::acc) (n-1) b
| Rawterm.RLetIn(_,name,v,b) -> chop_lambda_n ((name,v,true)::acc) (n-1) b
| _ ->
raise (Util.UserError("chop_rlambda_n",
@@ -90,7 +90,7 @@ let chop_rprod_n =
then List.rev acc,rt
else
match rt with
- | Rawterm.RProd(_,name,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b
+ | Rawterm.RProd(_,name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b
| _ -> raise (Util.UserError("chop_rprod_n",str "chop_rprod_n: Not enough products"))
in
chop_prod_n []
diff --git a/contrib/funind/indfun_main.ml4 b/contrib/funind/indfun_main.ml4
index b6398a329..4b3492b12 100644
--- a/contrib/funind/indfun_main.ml4
+++ b/contrib/funind/indfun_main.ml4
@@ -148,7 +148,7 @@ END
VERNAC ARGUMENT EXTEND binder2
[ "(" ne_ident_list(idl) ":" lconstr(c) ")"] ->
[
- LocalRawAssum (List.map (fun id -> (Util.dummy_loc,Name id)) idl,c) ]
+ LocalRawAssum (List.map (fun id -> (Util.dummy_loc,Name id)) idl,Topconstr.default_binder_kind,c) ]
END
diff --git a/contrib/funind/merge.ml b/contrib/funind/merge.ml
index 44df08598..adec67e36 100644
--- a/contrib/funind/merge.ml
+++ b/contrib/funind/merge.ml
@@ -830,7 +830,7 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
let _ = prNamedRConstr (string_of_name nme) tp in
let _ = prstr " ; " in
let typ = rawterm_to_constr_expr tp in
- LocalRawAssum ([(dummy_loc,nme)] , typ) :: acc)
+ LocalRawAssum ([(dummy_loc,nme)], Topconstr.default_binder_kind, typ) :: acc)
[] params in
let concl = Constrextern.extern_constr false (Global.env()) concl in
let arity,_ =
@@ -838,7 +838,7 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
(fun (acc,env) (nm,_,c) ->
let typ = Constrextern.extern_constr false env c in
let newenv = Environ.push_rel (nm,None,c) env in
- CProdN (dummy_loc, [[(dummy_loc,nm)],typ] , acc) , newenv)
+ CProdN (dummy_loc, [[(dummy_loc,nm)],Topconstr.default_binder_kind,typ] , acc) , newenv)
(concl,Global.env())
(shift.funresprms2 @ shift.funresprms1
@ shift.args2 @ shift.args1 @ shift.otherprms2 @ shift.otherprms1) in
@@ -867,7 +867,7 @@ let mkProd_reldecl (rdecl:rel_declaration) (t2:rawconstr) =
match rdecl with
| (nme,None,t) ->
let traw = Detyping.detype false [] [] t in
- RProd (dummy_loc,nme,traw,t2)
+ RProd (dummy_loc,nme,Explicit,traw,t2)
| (_,Some _,_) -> assert false
@@ -877,7 +877,7 @@ let mkProd_reldecl (rdecl:rel_declaration) (t2:rawconstr) =
match rdecl with
| (nme,None,t) ->
let traw = Detyping.detype false [] [] t in
- RProd (dummy_loc,nme,traw,t2)
+ RProd (dummy_loc,nme,Explicit,traw,t2)
| (_,Some _,_) -> assert false
diff --git a/contrib/funind/rawterm_to_relation.ml b/contrib/funind/rawterm_to_relation.ml
index 2f6f5914d..9bea4f1e0 100644
--- a/contrib/funind/rawterm_to_relation.ml
+++ b/contrib/funind/rawterm_to_relation.ml
@@ -586,7 +586,7 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
| RProd _ -> error "Cannot apply a type"
end (* end of the application treatement *)
- | RLambda(_,n,t,b) ->
+ | RLambda(_,n,_,t,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the type
@@ -601,7 +601,7 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
let new_env = raw_push_named (new_n,None,t) env in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_lam new_n) t_res b_res
- | RProd(_,n,t,b) ->
+ | RProd(_,n,_,t,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the type
@@ -865,7 +865,7 @@ let is_res id =
*)
let rec rebuild_cons nb_args relname args crossed_types depth rt =
match rt with
- | RProd(_,n,t,b) ->
+ | RProd(_,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
@@ -928,7 +928,7 @@ let rec rebuild_cons nb_args relname args crossed_types depth rt =
(Idset.filter not_free_in_t id_to_exclude)
| _ -> mkRProd(n,t,new_b),Idset.filter not_free_in_t id_to_exclude
end
- | RLambda(_,n,t,b) ->
+ | RLambda(_,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
@@ -944,7 +944,7 @@ let rec rebuild_cons nb_args relname args crossed_types depth rt =
then
new_b, Idset.remove id (Idset.filter not_free_in_t id_to_exclude)
else
- RProd(dummy_loc,n,t,new_b),Idset.filter not_free_in_t id_to_exclude
+ RProd(dummy_loc,n,k,t,new_b),Idset.filter not_free_in_t id_to_exclude
| _ -> anomaly "Should not have an anonymous function here"
(* We have renamed all the anonymous functions during alpha_renaming phase *)
@@ -1016,7 +1016,7 @@ let rec compute_cst_params relnames params = function
compute_cst_params_from_app [] (params,rtl)
| RApp(_,f,args) ->
List.fold_left (compute_cst_params relnames) params (f::args)
- | RLambda(_,_,t,b) | RProd(_,_,t,b) | RLetIn(_,_,t,b) | RLetTuple(_,_,_,t,b) ->
+ | RLambda(_,_,_,t,b) | RProd(_,_,_,t,b) | RLetIn(_,_,t,b) | RLetTuple(_,_,_,t,b) ->
let t_params = compute_cst_params relnames params t in
compute_cst_params relnames t_params b
| RCases _ -> params (* If there is still cases at this point they can only be
@@ -1153,7 +1153,7 @@ let do_build_inductive
else
Topconstr.CProdN
(dummy_loc,
- [[(dummy_loc,n)],Constrextern.extern_rawconstr Idset.empty t],
+ [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_rawconstr Idset.empty t],
acc
)
)
@@ -1173,7 +1173,7 @@ let do_build_inductive
Topconstr.LocalRawDef((dummy_loc,n), Constrextern.extern_rawconstr Idset.empty t)
else
Topconstr.LocalRawAssum
- ([(dummy_loc,n)], Constrextern.extern_rawconstr Idset.empty t)
+ ([(dummy_loc,n)], Topconstr.default_binder_kind, Constrextern.extern_rawconstr Idset.empty t)
)
rels_params
in
diff --git a/contrib/funind/rawtermops.ml b/contrib/funind/rawtermops.ml
index e8cce47ad..0c03c1e61 100644
--- a/contrib/funind/rawtermops.ml
+++ b/contrib/funind/rawtermops.ml
@@ -12,8 +12,8 @@ let idmap_is_empty m = m = Idmap.empty
let mkRRef ref = RRef(dummy_loc,ref)
let mkRVar id = RVar(dummy_loc,id)
let mkRApp(rt,rtl) = RApp(dummy_loc,rt,rtl)
-let mkRLambda(n,t,b) = RLambda(dummy_loc,n,t,b)
-let mkRProd(n,t,b) = RProd(dummy_loc,n,t,b)
+let mkRLambda(n,t,b) = RLambda(dummy_loc,n,Explicit,t,b)
+let mkRProd(n,t,b) = RProd(dummy_loc,n,Explicit,t,b)
let mkRLetIn(n,t,b) = RLetIn(dummy_loc,n,t,b)
let mkRCases(rto,l,brl) = RCases(dummy_loc,rto,l,brl)
let mkRSort s = RSort(dummy_loc,s)
@@ -26,7 +26,7 @@ let mkRCast(b,t) = RCast(dummy_loc,b,CastConv (Term.DEFAULTcast,t))
*)
let raw_decompose_prod =
let rec raw_decompose_prod args = function
- | RProd(_,n,t,b) ->
+ | RProd(_,n,k,t,b) ->
raw_decompose_prod ((n,t)::args) b
| rt -> args,rt
in
@@ -34,7 +34,7 @@ let raw_decompose_prod =
let raw_decompose_prod_or_letin =
let rec raw_decompose_prod args = function
- | RProd(_,n,t,b) ->
+ | RProd(_,n,k,t,b) ->
raw_decompose_prod ((n,None,Some t)::args) b
| RLetIn(_,n,t,b) ->
raw_decompose_prod ((n,Some t,None)::args) b
@@ -58,7 +58,7 @@ let raw_decompose_prod_n n =
if i<=0 then args,c
else
match c with
- | RProd(_,n,t,b) ->
+ | RProd(_,n,_,t,b) ->
raw_decompose_prod (i-1) ((n,t)::args) b
| rt -> args,rt
in
@@ -70,7 +70,7 @@ let raw_decompose_prod_or_letin_n n =
if i<=0 then args,c
else
match c with
- | RProd(_,n,t,b) ->
+ | RProd(_,n,_,t,b) ->
raw_decompose_prod (i-1) ((n,None,Some t)::args) b
| RLetIn(_,n,t,b) ->
raw_decompose_prod (i-1) ((n,Some t,None)::args) b
@@ -135,15 +135,17 @@ let change_vars =
change_vars mapping rt',
List.map (change_vars mapping) rtl
)
- | RLambda(loc,name,t,b) ->
+ | RLambda(loc,name,k,t,b) ->
RLambda(loc,
name,
+ k,
change_vars mapping t,
change_vars (remove_name_from_mapping mapping name) b
)
- | RProd(loc,name,t,b) ->
+ | RProd(loc,name,k,t,b) ->
RProd(loc,
name,
+ k,
change_vars mapping t,
change_vars (remove_name_from_mapping mapping name) b
)
@@ -261,21 +263,21 @@ let rec alpha_rt excluded rt =
let new_rt =
match rt with
| RRef _ | RVar _ | REvar _ | RPatVar _ -> rt
- | RLambda(loc,Anonymous,t,b) ->
+ | RLambda(loc,Anonymous,k,t,b) ->
let new_id = Nameops.next_ident_away (id_of_string "_x") 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
- RLambda(loc,Name new_id,new_t,new_b)
- | RProd(loc,Anonymous,t,b) ->
+ RLambda(loc,Name new_id,k,new_t,new_b)
+ | RProd(loc,Anonymous,k,t,b) ->
let new_t = alpha_rt excluded t in
let new_b = alpha_rt excluded b in
- RProd(loc,Anonymous,new_t,new_b)
+ RProd(loc,Anonymous,k,new_t,new_b)
| RLetIn(loc,Anonymous,t,b) ->
let new_t = alpha_rt excluded t in
let new_b = alpha_rt excluded b in
RLetIn(loc,Anonymous,new_t,new_b)
- | RLambda(loc,Name id,t,b) ->
+ | RLambda(loc,Name id,k,t,b) ->
let new_id = Nameops.next_ident_away id excluded in
let t,b =
if new_id = id
@@ -287,8 +289,8 @@ let rec alpha_rt excluded rt =
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
- RLambda(loc,Name new_id,new_t,new_b)
- | RProd(loc,Name id,t,b) ->
+ RLambda(loc,Name new_id,k,new_t,new_b)
+ | RProd(loc,Name id,k,t,b) ->
let new_id = Nameops.next_ident_away id excluded in
let new_excluded = new_id::excluded in
let t,b =
@@ -300,7 +302,7 @@ let rec alpha_rt excluded rt =
in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- RProd(loc,Name new_id,new_t,new_b)
+ RProd(loc,Name new_id,k,new_t,new_b)
| RLetIn(loc,Name id,t,b) ->
let new_id = Nameops.next_ident_away id excluded in
let t,b =
@@ -389,7 +391,7 @@ let is_free_in id =
| REvar _ -> false
| RPatVar _ -> false
| RApp(_,rt,rtl) -> List.exists is_free_in (rt::rtl)
- | RLambda(_,n,t,b) | RProd(_,n,t,b) | RLetIn(_,n,t,b) ->
+ | RLambda(_,n,_,t,b) | RProd(_,n,_,t,b) | RLetIn(_,n,t,b) ->
let check_in_b =
match n with
| Name id' -> id_ord id' id <> 0
@@ -460,17 +462,19 @@ let replace_var_by_term x_id term =
replace_var_by_pattern rt',
List.map replace_var_by_pattern rtl
)
- | RLambda(_,Name id,_,_) when id_ord id x_id == 0 -> rt
- | RLambda(loc,name,t,b) ->
+ | RLambda(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
+ | RLambda(loc,name,k,t,b) ->
RLambda(loc,
name,
+ k,
replace_var_by_pattern t,
replace_var_by_pattern b
)
- | RProd(_,Name id,_,_) when id_ord id x_id == 0 -> rt
- | RProd(loc,name,t,b) ->
+ | RProd(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
+ | RProd(loc,name,k,t,b) ->
RProd(loc,
name,
+ k,
replace_var_by_pattern t,
replace_var_by_pattern b
)
@@ -590,8 +594,8 @@ let ids_of_rawterm c =
| RVar (_,id) -> id::acc
| RApp (loc,g,args) ->
ids_of_rawterm [] g @ List.flatten (List.map (ids_of_rawterm []) args) @ acc
- | RLambda (loc,na,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
- | RProd (loc,na,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
+ | RLambda (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
+ | RProd (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
| RLetIn (loc,na,b,c) -> idof na :: ids_of_rawterm [] b @ ids_of_rawterm [] c @ acc
| RCast (loc,c,CastConv(k,t)) -> ids_of_rawterm [] c @ ids_of_rawterm [] t @ acc
| RCast (loc,c,CastCoerce) -> ids_of_rawterm [] c @ acc
@@ -622,15 +626,17 @@ let zeta_normalize =
zeta_normalize_term rt',
List.map zeta_normalize_term rtl
)
- | RLambda(loc,name,t,b) ->
+ | RLambda(loc,name,k,t,b) ->
RLambda(loc,
name,
+ k,
zeta_normalize_term t,
zeta_normalize_term b
)
- | RProd(loc,name,t,b) ->
+ | RProd(loc,name,k,t,b) ->
RProd(loc,
- name,
+ name,
+ k,
zeta_normalize_term t,
zeta_normalize_term b
)
@@ -691,8 +697,8 @@ let expand_as =
with Not_found -> rt
end
| RApp(loc,f,args) -> RApp(loc,expand_as map f,List.map (expand_as map) args)
- | RLambda(loc,na,t,b) -> RLambda(loc,na,expand_as map t, expand_as map b)
- | RProd(loc,na,t,b) -> RProd(loc,na,expand_as map t, expand_as map b)
+ | RLambda(loc,na,k,t,b) -> RLambda(loc,na,k,expand_as map t, expand_as map b)
+ | RProd(loc,na,k,t,b) -> RProd(loc,na,k,expand_as map t, expand_as map b)
| RLetIn(loc,na,v,b) -> RLetIn(loc,na, expand_as map v,expand_as map b)
| RLetTuple(loc,nal,(na,po),v,b) ->
RLetTuple(loc,nal,(na,Option.map (expand_as map) po),
diff --git a/contrib/interface/name_to_ast.ml b/contrib/interface/name_to_ast.ml
index b0a34cfd1..172c7479c 100644
--- a/contrib/interface/name_to_ast.ml
+++ b/contrib/interface/name_to_ast.ml
@@ -28,7 +28,7 @@ let convert_env =
let convert_binder env (na, b, c) =
match b with
| Some b -> LocalRawDef ((dummy_loc,na), extern_constr true env b)
- | None -> LocalRawAssum ([dummy_loc,na], extern_constr true env c) in
+ | None -> LocalRawAssum ([dummy_loc,na], default_binder_kind, extern_constr true env c) in
let rec cvrec env = function
[] -> []
| b::rest -> (convert_binder env b)::(cvrec (push_rel b env) rest) in
@@ -140,8 +140,8 @@ let make_variable_ast name typ implicits =
let make_definition_ast name c typ implicits =
- VernacDefinition ((Global,false,Definition), (dummy_loc,name), DefineBody ([], None,
- (constr_to_ast c), Some (constr_to_ast typ)),
+ VernacDefinition ((Global,false,Definition), (dummy_loc,name),
+ DefineBody ([], None, constr_to_ast c, Some (constr_to_ast typ)),
(fun _ _ -> ()))
::(implicits_to_ast_list implicits);;
diff --git a/contrib/interface/xlate.ml b/contrib/interface/xlate.ml
index 0e4f66072..05c227de0 100644
--- a/contrib/interface/xlate.ml
+++ b/contrib/interface/xlate.ml
@@ -312,7 +312,7 @@ let make_fix_struct (n,bl) =
let rec xlate_binder = function
- (l,t) -> CT_binder(xlate_id_opt_ne_list l, xlate_formula t)
+ (l,k,t) -> CT_binder(xlate_id_opt_ne_list l, xlate_formula t)
and xlate_return_info = function
| (Some Anonymous, None) | (None, None) ->
CT_coerce_NONE_to_RETURN_INFO CT_none
@@ -325,7 +325,7 @@ and xlate_formula_opt =
| Some e -> CT_coerce_FORMULA_to_FORMULA_OPT (xlate_formula e)
and xlate_binder_l = function
- LocalRawAssum(l,t) -> CT_binder(xlate_id_opt_ne_list l, xlate_formula t)
+ LocalRawAssum(l,_,t) -> CT_binder(xlate_id_opt_ne_list l, xlate_formula t)
| LocalRawDef(n,v) -> CT_coerce_DEF_to_BINDER(CT_def(xlate_id_opt n,
xlate_formula v))
and
@@ -459,7 +459,7 @@ and xlate_matched_formula = function
CT_coerce_FORMULA_to_MATCHED_FORMULA(xlate_formula f)
and xlate_formula_expl = function
(a, None) -> xlate_formula a
- | (a, Some (_,ExplByPos i)) ->
+ | (a, Some (_,ExplByPos (i, _))) ->
xlate_error "explicitation of implicit by rank not supported"
| (a, Some (_,ExplByName i)) ->
CT_labelled_arg(CT_ident (string_of_id i), xlate_formula a)
@@ -1609,7 +1609,7 @@ let rec xlate_vernac =
| VernacDeclareTacticDefinition (true, tacs) ->
(match List.map
(function
- ((_, id), body) ->
+ ((_, id), _, body) ->
CT_tac_def(CT_ident (string_of_id id), xlate_tactic body))
tacs with
[] -> assert false
@@ -1834,6 +1834,10 @@ let rec xlate_vernac =
xlate_error "TODO: Print Canonical Structures"
| PrintAssumptions _ ->
xlate_error "TODO: Print Needed Assumptions"
+ | PrintInstances _ ->
+ xlate_error "TODO: Print Instances"
+ | PrintTypeClasses ->
+ xlate_error "TODO: Print TypeClasses"
| PrintInspect n -> CT_inspect (CT_int n)
| PrintUniverses opt_s -> CT_print_universes(ctf_STRING_OPT opt_s)
| PrintSetoids -> CT_print_setoids
@@ -2075,6 +2079,12 @@ let rec xlate_vernac =
| Local -> CT_local in
CT_coercion (local_opt, id_opt, xlate_ident id1,
xlate_class id2, xlate_class id3)
+
+ (* TypeClasses *)
+ | VernacSetInstantiationTactic _|VernacDeclareInstance _|VernacContext _|
+ VernacInstance (_, _, _)|VernacClass (_, _, _, _, _) ->
+ xlate_error "TODO: Type Classes commands"
+
| VernacResetName id -> CT_reset (xlate_ident (snd id))
| VernacResetInitial -> CT_restore_state (CT_ident "Initial")
| VernacExtend (s, l) ->
@@ -2088,7 +2098,7 @@ let rec xlate_vernac =
| VernacNop -> CT_proof_no_op
| VernacComments l ->
CT_scomments(CT_scomment_content_list (List.map xlate_comment l))
- | VernacDeclareImplicits(true, id, opt_positions) ->
+ | VernacDeclareImplicits(true, id, _, opt_positions) ->
CT_implicits
(reference_to_ct_ID id,
match opt_positions with
@@ -2097,11 +2107,11 @@ let rec xlate_vernac =
CT_coerce_ID_LIST_to_ID_LIST_OPT
(CT_id_list
(List.map
- (function ExplByPos x, _
+ (function ExplByPos (x,_), _, _
-> xlate_error
"explication argument by rank is obsolete"
- | ExplByName id, _ -> CT_ident (string_of_id id)) l)))
- | VernacDeclareImplicits(false, id, opt_positions) ->
+ | ExplByName id, _, _ -> CT_ident (string_of_id id)) l)))
+ | VernacDeclareImplicits(false, id, _, opt_positions) ->
xlate_error "TODO: Implicit Arguments Global"
| VernacReserve((_,a)::l, f) ->
CT_reserve(CT_id_ne_list(xlate_ident a,
diff --git a/contrib/recdef/recdef.ml4 b/contrib/recdef/recdef.ml4
index 26b6b3094..87c4d9bc7 100644
--- a/contrib/recdef/recdef.ml4
+++ b/contrib/recdef/recdef.ml4
@@ -1061,7 +1061,7 @@ let (value_f:constr list -> global_reference -> constr) =
List.fold_left2
(fun acc x_id a ->
RLambda
- (d0, Name x_id, RDynamic(d0, constr_in a),
+ (d0, Name x_id, Explicit, RDynamic(d0, constr_in a),
acc
)
)
diff --git a/contrib/subtac/eterm.ml b/contrib/subtac/eterm.ml
index d1b4b50eb..9cb65f76e 100644
--- a/contrib/subtac/eterm.ml
+++ b/contrib/subtac/eterm.ml
@@ -121,7 +121,7 @@ let rec chop_product n t =
| Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (Termops.pop b) else None
| _ -> None
-let eterm_obligations env name isevars evm fs t tycon =
+let eterm_obligations env name isevars evm fs t ty =
(* 'Serialize' the evars, we assume that the types of the existentials
refer to previous existentials in the list only *)
trace (str " In eterm: isevars: " ++ my_print_evardefs isevars);
@@ -165,6 +165,7 @@ let eterm_obligations env name isevars evm fs t tycon =
let t', _, transparent = (* Substitute evar refs in the term by variables *)
subst_evar_constr evts 0 t
in
+ let ty, _, _ = subst_evar_constr evts 0 ty in
let evars =
List.map (fun (_, ((_, name), _, opaque, typ, deps)) -> name, typ, not (opaque = None) && not (Idset.mem name transparent), deps) evts
in
@@ -177,7 +178,7 @@ let eterm_obligations env name isevars evm fs t tycon =
Termops.print_constr_env (Global.env ()) typ))
evars);
with _ -> ());
- Array.of_list (List.rev evars), t'
+ Array.of_list (List.rev evars), t', ty
let mkMetas n = list_tabulate (fun _ -> Evarutil.mk_new_meta ()) n
diff --git a/contrib/subtac/eterm.mli b/contrib/subtac/eterm.mli
index a2582c5ca..e23fd535c 100644
--- a/contrib/subtac/eterm.mli
+++ b/contrib/subtac/eterm.mli
@@ -19,9 +19,9 @@ val mkMetas : int -> constr list
(* val eterm_term : evar_map -> constr -> types option -> constr * types option * (identifier * types) list *)
(* env, id, evars, number of
- function prototypes to try to clear from evars contexts, object and optional type *)
-val eterm_obligations : env -> identifier -> evar_defs -> evar_map -> int -> constr -> types option ->
- (identifier * types * bool * Intset.t) array * constr
+ function prototypes to try to clear from evars contexts, object and type *)
+val eterm_obligations : env -> identifier -> evar_defs -> evar_map -> int -> constr -> types ->
+ (identifier * types * bool * Intset.t) array * constr * types
(* Obl. name, type as product, opacity (true = opaque) and dependencies as indexes into the array *)
val etermtac : open_constr -> tactic
diff --git a/contrib/subtac/g_subtac.ml4 b/contrib/subtac/g_subtac.ml4
index c2691c781..49e312aeb 100644
--- a/contrib/subtac/g_subtac.ml4
+++ b/contrib/subtac/g_subtac.ml4
@@ -44,17 +44,20 @@ struct
let subtac_nameopt : identifier option Gram.Entry.e = gec "subtac_nameopt"
end
+open Rawterm
open SubtacGram
open Util
open Pcoq
-
+open Prim
+open Constr
let sigref = mkRefC (Qualid (dummy_loc, Libnames.qualid_of_string "Coq.Init.Specif.sig"))
GEXTEND Gram
- GLOBAL: subtac_gallina_loc Constr.binder_let Constr.binder subtac_nameopt;
+ GLOBAL: subtac_gallina_loc typeclass_constraint Constr.binder_let Constr.binder subtac_nameopt;
subtac_gallina_loc:
- [ [ g = Vernac.gallina -> loc, g ] ]
+ [ [ g = Vernac.gallina -> loc, g
+ | g = Vernac.gallina_ext -> loc, g ] ]
;
subtac_nameopt:
@@ -63,18 +66,18 @@ GEXTEND Gram
;
Constr.binder_let:
- [ [ "("; id=Prim.name; ":"; t=Constr.lconstr; "|"; c=Constr.lconstr; ")" ->
- let typ = mkAppC (sigref, [mkLambdaC ([id], t, c)]) in
- LocalRawAssum ([id], typ)
+ [[ "("; id=Prim.name; ":"; t=Constr.lconstr; "|"; c=Constr.lconstr; ")" ->
+ let typ = mkAppC (sigref, [mkLambdaC ([id], default_binder_kind, t, c)]) in
+ LocalRawAssum ([id], default_binder_kind, typ)
] ];
Constr.binder:
[ [ "("; id=Prim.name; ":"; c=Constr.lconstr; "|"; p=Constr.lconstr; ")" ->
- ([id],mkAppC (sigref, [mkLambdaC ([id], c, p)]))
+ ([id],default_binder_kind, mkAppC (sigref, [mkLambdaC ([id], default_binder_kind, c, p)]))
| "("; id=Prim.name; ":"; c=Constr.lconstr; ")" ->
- ([id],c)
+ ([id],default_binder_kind, c)
| "("; id=Prim.name; lid=LIST1 Prim.name; ":"; c=Constr.lconstr; ")" ->
- (id::lid,c)
+ (id::lid,default_binder_kind, c)
] ];
END
diff --git a/contrib/subtac/subtac.ml b/contrib/subtac/subtac.ml
index 4f5e302c2..56ebc4f52 100644
--- a/contrib/subtac/subtac.ml
+++ b/contrib/subtac/subtac.ml
@@ -49,6 +49,22 @@ open Decl_kinds
open Tacinterp
open Tacexpr
+let solve_tccs_in_type env id isevars evm c typ =
+ if not (evm = Evd.empty) then
+ let stmt_id = Nameops.add_suffix id "_stmt" in
+ let obls, c', t' = eterm_obligations env stmt_id !isevars evm 0 c typ in
+ (** Make all obligations transparent so that real dependencies can be sorted out by the user *)
+ let obls = Array.map (fun (id, t, op, d) -> (id, t, false, d)) obls in
+ match Subtac_obligations.add_definition stmt_id c' typ obls with
+ Subtac_obligations.Defined cst -> constant_value (Global.env()) cst
+ | _ ->
+ errorlabstrm "start_proof"
+ (str "The statement obligations could not be resolved automatically, " ++ spc () ++
+ str "write a statement definition first.")
+ else
+ let _ = Typeops.infer_type env c in c
+
+
let start_proof_com env isevars sopt kind (bl,t) hook =
let id = match sopt with
| Some id ->
@@ -60,21 +76,11 @@ let start_proof_com env isevars sopt kind (bl,t) hook =
next_global_ident_away false (id_of_string "Unnamed_thm")
(Pfedit.get_all_proof_names ())
in
- let evm, c, typ =
+ let evm, c, typ, _imps =
Subtac_pretyping.subtac_process env isevars id [] (Command.generalize_constr_expr t bl) None
in
- if not (evm = Evd.empty) then
- let stmt_id = Nameops.add_suffix id "_stmt" in
- let obls, c' = eterm_obligations env stmt_id !isevars evm 0 c (Some typ) in
- match Subtac_obligations.add_definition stmt_id c' typ obls with
- Subtac_obligations.Defined cst -> Command.start_proof id kind (constant_value (Global.env()) cst) hook
- | _ ->
- errorlabstrm "start_proof"
- (str "The statement obligations could not be resolved automatically, " ++ spc () ++
- str "write a statement definition first.")
- else
- let _ = Typeops.infer_type env c in
- Command.start_proof id kind c hook
+ let c = solve_tccs_in_type env id isevars evm c typ in
+ Command.start_proof id kind c hook
let print_subgoals () = Flags.if_verbose (fun () -> msg (Printer.pr_open_subgoals ())) ()
@@ -88,10 +94,12 @@ let assumption_message id =
Flags.if_verbose message ((string_of_id id) ^ " is assumed")
let declare_assumption env isevars idl is_coe k bl c nl =
- if not (Pfedit.refining ()) then
- let evm, c, typ =
- Subtac_pretyping.subtac_process env isevars (snd (List.hd idl)) [] (Command.generalize_constr_expr c bl) None
+ if not (Pfedit.refining ()) then
+ let id = snd (List.hd idl) in
+ let evm, c, typ, imps =
+ Subtac_pretyping.subtac_process env isevars id [] (Command.generalize_constr_expr c bl) None
in
+ let c = solve_tccs_in_type env id isevars evm c typ in
List.iter (Command.declare_one_assumption is_coe k c nl) idl
else
errorlabstrm "Command.Assumption"
@@ -114,8 +122,13 @@ let subtac (loc, command) =
match command with
VernacDefinition (defkind, (locid, id), expr, hook) ->
(match expr with
- ProveBody (bl, c) -> ignore(Subtac_pretyping.subtac_proof env isevars id bl c None)
- | DefineBody (bl, _, c, tycon) ->
+ | ProveBody (bl, t) ->
+ if Lib.is_modtype () then
+ errorlabstrm "Subtac_command.StartProof"
+ (str "Proof editing mode not supported in module types");
+ start_proof_and_print env isevars (Some id) (Global, DefinitionBody Definition) (bl,t)
+ (fun _ _ -> ())
+ | DefineBody (bl, _, c, tycon) ->
ignore(Subtac_pretyping.subtac_proof env isevars id bl c tycon))
| VernacFixpoint (l, b) ->
let _ = trace (str "Building fixpoint") in
@@ -135,6 +148,9 @@ let subtac (loc, command) =
| VernacAssumption (stre,nl,l) ->
vernac_assumption env isevars stre l nl
+ | VernacInstance (sup, is, props) ->
+ Subtac_classes.new_instance sup is props
+
(* | VernacCoFixpoint (l, b) -> *)
(* let _ = trace (str "Building cofixpoint") in *)
(* ignore(Subtac_command.build_recursive l b) *)
diff --git a/contrib/subtac/subtac_cases.ml b/contrib/subtac/subtac_cases.ml
index 34c398289..e46ca822e 100644
--- a/contrib/subtac/subtac_cases.ml
+++ b/contrib/subtac/subtac_cases.ml
@@ -1395,9 +1395,9 @@ let set_arity_signature dep n arsign tomatchl pred x =
let rec decomp_lam_force n avoid l p =
if n = 0 then (List.rev l,p,avoid) else
match p with
- | RLambda (_,(Name id as na),_,c) ->
+ | RLambda (_,(Name id as na),k,_,c) ->
decomp_lam_force (n-1) (id::avoid) (na::l) c
- | RLambda (_,(Anonymous as na),_,c) -> decomp_lam_force (n-1) avoid (na::l) c
+ | RLambda (_,(Anonymous as na),k,_,c) -> decomp_lam_force (n-1) avoid (na::l) c
| _ ->
let x = next_ident_away (id_of_string "x") avoid in
decomp_lam_force (n-1) (x::avoid) (Name x :: l)
@@ -2091,13 +2091,12 @@ let compile_cases loc (typing_fun, isevars) (tycon : Evarutil.type_constraint) e
(* let env = nf_evars_env !isevars (env : env) in *)
(* trace (str "Evars : " ++ my_print_evardefs !isevars); *)
(* trace (str "Env : " ++ my_print_env env); *)
-
- let tomatchs, tomatchs_lets = abstract_tomatch env tomatchs in
- let tomatchs_len = List.length tomatchs_lets in
- let tycon = lift_tycon tomatchs_len tycon in
- let env = push_rel_context tomatchs_lets env in
let _isdep = List.exists (fun (x, y) -> is_dependent_ind y) tomatchs in
if predopt = None then
+ let tomatchs, tomatchs_lets = abstract_tomatch env tomatchs in
+ let tomatchs_len = List.length tomatchs_lets in
+ let tycon = lift_tycon tomatchs_len tycon in
+ let env = push_rel_context tomatchs_lets env in
let len = List.length eqns in
let sign, allnames, signlen, eqs, neqs, args =
(* The arity signature *)
@@ -2185,7 +2184,6 @@ let compile_cases loc (typing_fun, isevars) (tycon : Evarutil.type_constraint) e
let j = compile pb in
(* We check for unused patterns *)
List.iter (check_unused_pattern env) matx;
- let j = { j with uj_val = it_mkLambda_or_LetIn j.uj_val tomatchs_lets } in
inh_conv_coerce_to_tycon loc env isevars j tycon
end
diff --git a/contrib/subtac/subtac_classes.ml b/contrib/subtac/subtac_classes.ml
new file mode 100644
index 000000000..da91b0406
--- /dev/null
+++ b/contrib/subtac/subtac_classes.ml
@@ -0,0 +1,187 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.mli 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+open Pretyping
+open Evd
+open Environ
+open Term
+open Rawterm
+open Topconstr
+open Names
+open Libnames
+open Pp
+open Vernacexpr
+open Constrintern
+open Subtac_command
+open Typeclasses
+open Typeclasses_errors
+open Termops
+open Decl_kinds
+open Entries
+
+module SPretyping = Subtac_pretyping.Pretyping
+
+let interp_binder_evars evdref env na t =
+ let t = Constrintern.intern_gen true (Evd.evars_of !evdref) env t in
+ SPretyping.understand_tcc_evars evdref env IsType t
+
+let interp_binders_evars isevars env avoid l =
+ List.fold_left
+ (fun (env, ids, params) ((loc, i), t) ->
+ let n = Name i in
+ let t' = interp_binder_evars isevars env n t in
+ let d = (i,None,t') in
+ (push_named d env, i :: ids, d::params))
+ (env, avoid, []) l
+
+let interp_typeclass_context_evars isevars env avoid l =
+ List.fold_left
+ (fun (env, ids, params) (iid, bk, cl) ->
+ let t' = interp_binder_evars isevars env (snd iid) cl in
+ let i = match snd iid with
+ | Anonymous -> Nameops.next_name_away (Termops.named_hd env t' Anonymous) ids
+ | Name id -> id
+ in
+ let d = (i,None,t') in
+ (push_named d env, i :: ids, d::params))
+ (env, avoid, []) l
+
+let interp_constrs_evars isevars env avoid l =
+ List.fold_left
+ (fun (env, ids, params) t ->
+ let t' = interp_binder_evars isevars env Anonymous t in
+ let id = Nameops.next_name_away (Termops.named_hd env t' Anonymous) ids in
+ let d = (id,None,t') in
+ (push_named d env, id :: ids, d::params))
+ (env, avoid, []) l
+
+let type_ctx_instance isevars env ctx inst subst =
+ List.fold_left2
+ (fun (subst, instctx) (na, _, t) ce ->
+ let t' = replace_vars subst t in
+ let c = interp_casted_constr_evars isevars env ce t' in
+ let d = na, Some c, t' in
+ (na, c) :: subst, d :: instctx)
+ (subst, []) (List.rev ctx) inst
+
+let substitution_of_constrs ctx cstrs =
+ List.fold_right2 (fun c (na, _, _) acc -> (na, c) :: acc) cstrs ctx []
+
+let new_instance sup (instid, bk, cl) props =
+ let id, par = Implicit_quantifiers.destClassApp cl in
+ let env = Global.env() in
+ let isevars = ref (Evd.create_evar_defs Evd.empty) in
+ let avoid = Termops.ids_of_context env in
+ let k =
+ try class_info (snd id)
+ with Not_found -> unbound_class env id
+ in
+ let gen_ctx, sup = Implicit_quantifiers.resolve_class_binders (vars_of_env env) sup in
+ let gen_ctx =
+ let is_free ((_, x), _) =
+ let f l = not (List.exists (fun ((_, y), _) -> y = x) l) in
+ let g l = not (List.exists (fun ((_, y), _, _) -> y = Name x) l) in
+ f gen_ctx && g sup
+ in
+ let parbinders = Implicit_quantifiers.compute_constrs_freevars_binders (vars_of_env env) par in
+ let parbinders' = List.filter is_free parbinders in
+ gen_ctx @ parbinders'
+ in
+ let env', avoid, genctx = interp_binders_evars isevars env avoid gen_ctx in
+ let env', avoid, supctx = interp_typeclass_context_evars isevars env' avoid sup in
+
+ let subst =
+ match bk with
+ | Explicit ->
+ if List.length par <> List.length k.cl_context + List.length k.cl_params then
+ Classes.mismatched_params env par (k.cl_params @ k.cl_context);
+ let len = List.length k.cl_context in
+ let ctx, par = Util.list_chop len par in
+ let subst, _ = type_ctx_instance isevars env' k.cl_context ctx [] in
+ let subst = Typeclasses.substitution_of_named_context isevars env' k.cl_name len subst
+ k.cl_super
+ in
+ let subst, par = type_ctx_instance isevars env' k.cl_params par subst in subst
+
+ | Implicit ->
+ let t' = interp_type_evars isevars env' (Topconstr.mkAppC (CRef (Ident id), par)) in
+ match kind_of_term t' with
+ App (c, args) ->
+ substitution_of_constrs (k.cl_params @ k.cl_super @ k.cl_context)
+ (List.rev (Array.to_list args))
+ | _ -> assert false
+ in
+ isevars := Evarutil.nf_evar_defs !isevars;
+ let sigma = Evd.evars_of !isevars in
+ let substctx = Typeclasses.nf_substitution sigma subst in
+ let subst, _propsctx =
+ let props =
+ List.map (fun (x, l, d) ->
+ x, Topconstr.abstract_constr_expr d (Classes.binders_of_lidents l))
+ props
+ in
+ if List.length props > List.length k.cl_props then
+ Classes.mismatched_props env' (List.map snd props) k.cl_props;
+ let props, rest =
+ List.fold_left
+ (fun (props, rest) (id,_,_) ->
+ try
+ let (_, c) = List.find (fun ((_,id'), c) -> id' = id) rest in
+ let rest' = List.filter (fun ((_,id'), c) -> id' <> id) rest in
+ c :: props, rest'
+ with Not_found -> (CHole Util.dummy_loc :: props), rest)
+ ([], props) k.cl_props
+ in
+ if rest <> [] then
+ unbound_method env k.cl_name (fst (List.hd rest))
+ else
+ type_ctx_instance isevars env' k.cl_props props substctx
+ in
+ let app =
+ applistc (mkConstruct (k.cl_impl, 1)) (List.rev_map snd subst)
+ in
+ let term = it_mkNamedLambda_or_LetIn (it_mkNamedLambda_or_LetIn app supctx) genctx in
+ isevars := Evarutil.nf_evar_defs !isevars;
+ let term = Evarutil.nf_isevar !isevars term in
+ let termtype =
+ let app = applistc (mkInd (k.cl_impl)) (List.rev_map snd substctx) in
+ let t = it_mkNamedProd_or_LetIn (it_mkNamedProd_or_LetIn app supctx) genctx in
+ Evarutil.nf_isevar !isevars t
+ in
+ isevars := undefined_evars !isevars;
+ let id =
+ match snd instid with
+ Name id -> id
+ | Anonymous ->
+ let i = Nameops.add_suffix (snd id) "_instance_" in
+ Termops.next_global_ident_away false i (Termops.ids_of_context env)
+ in
+ let imps =
+ Util.list_map_i
+ (fun i (na, b, t) -> ExplByPos (i, Some na), (true, true))
+ 1 (genctx @ List.rev supctx)
+ in
+ let hook cst =
+ let inst =
+ { is_class = k;
+ is_name = id;
+(* is_params = paramsctx; *)
+(* is_super = superctx; *)
+ is_impl = cst;
+(* is_add_hint = (fun () -> Classes.add_instance_hint id); *)
+ }
+ in
+ Classes.add_instance_hint id;
+ Impargs.declare_manual_implicits false (ConstRef cst) false imps;
+ Typeclasses.add_instance inst
+ in
+ let evm = Subtac_utils.evars_of_term (Evd.evars_of !isevars) Evd.empty term in
+ let obls, constr, typ = Eterm.eterm_obligations env id !isevars evm 0 term termtype in
+ ignore (Subtac_obligations.add_definition id constr typ ~kind:Instance ~hook obls)
diff --git a/contrib/subtac/subtac_classes.mli b/contrib/subtac/subtac_classes.mli
new file mode 100644
index 000000000..ce4b32cad
--- /dev/null
+++ b/contrib/subtac/subtac_classes.mli
@@ -0,0 +1,39 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.mli 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Topconstr
+open Util
+open Typeclasses
+open Implicit_quantifiers
+open Classes
+(*i*)
+
+val type_ctx_instance : Evd.evar_defs ref ->
+ Environ.env ->
+ (Names.identifier * 'a * Term.constr) list ->
+ Topconstr.constr_expr list ->
+ (Names.identifier * Term.constr) list ->
+ (Names.identifier * Term.constr) list *
+ (Names.identifier * Term.constr option * Term.constr) list
+
+val new_instance :
+ typeclass_context ->
+ typeclass_constraint ->
+ binder_def_list ->
+ unit
diff --git a/contrib/subtac/subtac_command.ml b/contrib/subtac/subtac_command.ml
index a9a30c57a..9afa6f067 100644
--- a/contrib/subtac/subtac_command.ml
+++ b/contrib/subtac/subtac_command.ml
@@ -98,11 +98,11 @@ let interp_binder sigma env na t =
let interp_context sigma env params =
List.fold_left
(fun (env,params) d -> match d with
- | LocalRawAssum ([_,na],(CHole _ as t)) ->
+ | LocalRawAssum ([_,na],k,(CHole _ as t)) ->
let t = interp_binder sigma env na t in
let d = (na,None,t) in
(push_rel d env, d::params)
- | LocalRawAssum (nal,t) ->
+ | LocalRawAssum (nal,k,t) ->
let t = interp_type sigma env t in
let ctx = list_map_i (fun i (_,na) -> (na,None,lift i t)) 0 nal in
let ctx = List.rev ctx in
@@ -152,7 +152,7 @@ let collect_non_rec env =
let list_of_local_binders l =
let rec aux acc = function
Topconstr.LocalRawDef (n, c) :: tl -> aux ((n, Some c, None) :: acc) tl
- | Topconstr.LocalRawAssum (nl, c) :: tl ->
+ | Topconstr.LocalRawAssum (nl, k, c) :: tl ->
aux (List.fold_left (fun acc n -> (n, None, Some c) :: acc) acc nl) tl
| [] -> List.rev acc
in aux [] l
@@ -294,12 +294,11 @@ let build_wellfounded (recname, n, bl,arityc,body) r measure notation boxed =
lift lift_cst prop ;
lift lift_cst intern_body_lam |])
| Some f ->
- lift (succ after_length)
- (mkApp (constr_of_global (Lazy.force fix_measure_sub_ref),
- [| argtyp ;
- f ;
- lift lift_cst prop ;
- lift lift_cst intern_body_lam |]))
+ mkApp (constr_of_global (Lazy.force fix_measure_sub_ref),
+ [| lift lift_cst argtyp ;
+ lift lift_cst f ;
+ lift lift_cst prop ;
+ lift lift_cst intern_body_lam |])
in
let def_appl = applist (fix_def, gen_rels (after_length + 1)) in
let def = it_mkLambda_or_LetIn def_appl binders_rel in
@@ -316,13 +315,13 @@ let build_wellfounded (recname, n, bl,arityc,body) r measure notation boxed =
let evm = non_instanciated_map env isevars evm in
(* let _ = try trace (str "Non instanciated evars map: " ++ Evd.pr_evar_map evm) with _ -> () in *)
- let evars, evars_def = Eterm.eterm_obligations env recname !isevars evm 0 fullcoqc (Some fullctyp) in
+ let evars, evars_def, evars_typ = Eterm.eterm_obligations env recname !isevars evm 0 fullcoqc fullctyp in
(* (try trace (str "Generated obligations : "); *)
(* Array.iter *)
(* (fun (n, t, _) -> trace (str "Evar " ++ str (string_of_id n) ++ spc () ++ my_print_constr env t)) *)
(* evars; *)
(* with _ -> ()); *)
- Subtac_obligations.add_definition recname evars_def fullctyp evars
+ Subtac_obligations.add_definition recname evars_def evars_typ evars
let nf_evar_context isevars ctx =
List.map (fun (n, b, t) ->
@@ -412,11 +411,12 @@ let build_mutrec lnameargsardef boxed =
(* Generalize by the recursive prototypes *)
let def =
Termops.it_mkNamedLambda_or_LetIn def rec_sign
- and typ =
+ and typ =
Termops.it_mkNamedProd_or_LetIn typ rec_sign
in
- let evm = Subtac_utils.evars_of_term evm Evd.empty def in
- let evars, def = Eterm.eterm_obligations env id isevars evm recdefs def (Some typ) in
+ let evm' = Subtac_utils.evars_of_term evm Evd.empty def in
+ let evm' = Subtac_utils.evars_of_term evm evm' typ in
+ let evars, def, typ = Eterm.eterm_obligations env id isevars evm' recdefs def typ in
collect_evars (succ i) ((id, def, typ, evars) :: acc)
else acc
in
diff --git a/contrib/subtac/subtac_obligations.ml b/contrib/subtac/subtac_obligations.ml
index e3cf7a57a..c184169ac 100644
--- a/contrib/subtac/subtac_obligations.ml
+++ b/contrib/subtac/subtac_obligations.ml
@@ -12,6 +12,8 @@ open Decl_kinds
open Util
open Evd
+type definition_hook = constant -> unit
+
let pperror cmd = Util.errorlabstrm "Program" cmd
let error s = pperror (str s)
@@ -40,6 +42,9 @@ type program_info = {
prg_obligations: obligations;
prg_deps : identifier list;
prg_nvrec : int array;
+ prg_implicits : (Topconstr.explicitation * (bool * bool)) list;
+ prg_kind : definition_object_kind;
+ prg_hook : definition_hook;
}
let assumption_message id =
@@ -106,7 +111,7 @@ let progmap_union = ProgMap.fold ProgMap.add
let cache (_, (infos, tac)) =
from_prg := infos;
- default_tactic_expr := tac
+ set_default_tactic tac
let (input,output) =
declare_object
@@ -125,24 +130,30 @@ let rec intset_to = function
let subst_body prg =
let obls, _ = prg.prg_obligations in
- subst_deps obls (intset_to (pred (Array.length obls))) prg.prg_body
-
+ let ints = intset_to (pred (Array.length obls)) in
+ subst_deps obls ints prg.prg_body,
+ subst_deps obls ints (Termops.refresh_universes prg.prg_type)
+
let declare_definition prg =
- let body = subst_body prg in
+ let body, typ = subst_body prg in
(try trace (str "Declaring: " ++ Ppconstr.pr_id prg.prg_name ++ spc () ++
my_print_constr (Global.env()) body ++ str " : " ++
my_print_constr (Global.env()) prg.prg_type);
with _ -> ());
let ce =
{ const_entry_body = body;
- const_entry_type = Some (Termops.refresh_universes prg.prg_type);
+ const_entry_type = Some typ;
const_entry_opaque = false;
const_entry_boxed = false}
in
let c = Declare.declare_constant
- prg.prg_name (DefinitionEntry ce,IsDefinition Definition)
+ prg.prg_name (DefinitionEntry ce,IsDefinition prg.prg_kind)
in
+ let gr = ConstRef c in
+ if Impargs.is_implicit_args () || prg.prg_implicits <> [] then
+ Impargs.declare_manual_implicits false gr (Impargs.is_implicit_args ()) prg.prg_implicits;
print_message (definition_message c);
+ prg.prg_hook c;
c
open Pp
@@ -151,14 +162,18 @@ open Ppconstr
let declare_mutual_definition l =
let len = List.length l in
let namerec = Array.of_list (List.map (fun x -> x.prg_name) l) in
- let arrec =
- Array.of_list (List.map (fun x -> snd (decompose_prod_n len x.prg_type)) l)
- in
- let recvec =
- Array.of_list
- (List.map (fun x ->
- let subs = (subst_body x) in
- snd (decompose_lam_n len subs)) l)
+(* let arrec = *)
+(* Array.of_list (List.map (fun x -> snd (decompose_prod_n len x.prg_type)) l) *)
+(* in *)
+ let recvec, arrec =
+ let l, l' =
+ List.split
+ (List.map (fun x ->
+ let subs, typ = (subst_body x) in
+ snd (decompose_lam_n len subs),
+ snd (decompose_prod_n len typ)) l)
+ in
+ Array.of_list l, Array.of_list l'
in
let nvrec = (List.hd l).prg_nvrec in
let recdecls = (Array.map (fun id -> Name id) namerec, arrec, recvec) in
@@ -205,7 +220,7 @@ let try_tactics obls =
let red = Reductionops.nf_betaiota
-let init_prog_info n b t deps nvrec obls =
+let init_prog_info n b t deps nvrec obls impls kind hook =
let obls' =
Array.mapi
(fun i (n, t, o, d) ->
@@ -216,7 +231,7 @@ let init_prog_info n b t deps nvrec obls =
obls
in
{ prg_name = n ; prg_body = b; prg_type = red t; prg_obligations = (obls', Array.length obls');
- prg_deps = deps; prg_nvrec = nvrec; }
+ prg_deps = deps; prg_nvrec = nvrec; prg_implicits = impls; prg_kind = kind; prg_hook = hook; }
let get_prog name =
let prg_infos = !from_prg in
@@ -327,7 +342,8 @@ let rec solve_obligation prg num =
| _ -> ());
trace (str "Started obligation " ++ int user_num ++ str " proof: " ++
Subtac_utils.my_print_constr (Global.env ()) obl.obl_type);
- Pfedit.by !default_tactic
+ Pfedit.by !default_tactic;
+ Flags.if_verbose (fun () -> msg (Printer.pr_open_subgoals ())) ()
| l -> pperror (str "Obligation " ++ int user_num ++ str " depends on obligation(s) "
++ str (string_of_list ", " (fun x -> string_of_int (succ x)) l))
@@ -393,9 +409,22 @@ and auto_solve_obligations n : progress =
Flags.if_verbose msgnl (str "Solving obligations automatically...");
try solve_obligations n !default_tactic with NoObligations _ -> Dependent
-let add_definition n b t obls =
+open Pp
+let show_obligations ?(msg=true) n =
+ let prg = get_prog_err n in
+ let n = prg.prg_name in
+ let obls, rem = prg.prg_obligations in
+ if msg then msgnl (int rem ++ str " obligation(s) remaining: ");
+ Array.iteri (fun i x ->
+ match x.obl_body with
+ None -> msgnl (str "Obligation" ++ spc() ++ int (succ i) ++ spc () ++ str "of" ++ spc() ++ str (string_of_id n) ++ str ":" ++ spc () ++
+ my_print_constr (Global.env ()) x.obl_type ++ str "." ++ fnl ())
+ | Some _ -> ())
+ obls
+
+let add_definition n b t ?(implicits=[]) ?(kind=Definition) ?(hook=fun x -> ()) obls =
Flags.if_verbose pp (str (string_of_id n) ++ str " has type-checked");
- let prg = init_prog_info n b t [] (Array.make 0 0) obls in
+ let prg = init_prog_info n b t [] (Array.make 0 0) obls implicits kind hook in
let obls,_ = prg.prg_obligations in
if Array.length obls = 0 then (
Flags.if_verbose ppnl (str ".");
@@ -406,19 +435,30 @@ let add_definition n b t obls =
let len = Array.length obls in
let _ = Flags.if_verbose ppnl (str ", generating " ++ int len ++ str " obligation(s)") in
from_prg := ProgMap.add n prg !from_prg;
- auto_solve_obligations (Some n))
+ let res = auto_solve_obligations (Some n) in
+ match res with
+ | Remain rem when rem < 5 -> Flags.if_verbose (fun () -> show_obligations ~msg:false (Some n)) (); res
+ | _ -> res)
-let add_mutual_definitions l nvrec =
+let add_mutual_definitions l ?(implicits=[]) ?(kind=Definition) nvrec =
let deps = List.map (fun (n, b, t, obls) -> n) l in
let upd = List.fold_left
(fun acc (n, b, t, obls) ->
- let prg = init_prog_info n b t deps nvrec obls in
+ let prg = init_prog_info n b t deps nvrec obls implicits kind (fun x -> ()) in
ProgMap.add n prg acc)
!from_prg l
in
from_prg := upd;
- List.iter (fun x -> ignore(auto_solve_obligations (Some x))) deps
-
+ let _defined =
+ List.fold_left (fun finished x ->
+ if finished then finished
+ else
+ match auto_solve_obligations (Some x) with
+ Defined _ -> (* If one definition is turned into a constant, the whole block is defined. *) true
+ | _ -> false)
+ false deps
+ in ()
+
let admit_obligations n =
let prg = get_prog_err n in
let obls, rem = prg.prg_obligations in
@@ -447,18 +487,5 @@ let next_obligation n =
array_find (fun x -> x.obl_body = None && deps_remaining obls x.obl_deps = [])
obls
in solve_obligation prg i
-
-open Pp
-let show_obligations n =
- let prg = get_prog_err n in
- let n = prg.prg_name in
- let obls, rem = prg.prg_obligations in
- msgnl (int rem ++ str " obligation(s) remaining: ");
- Array.iteri (fun i x ->
- match x.obl_body with
- None -> msgnl (str "Obligation" ++ spc() ++ int (succ i) ++ spc () ++ str "of" ++ spc() ++ str (string_of_id n) ++ str ":" ++ spc () ++
- my_print_constr (Global.env ()) x.obl_type ++ str "." ++ fnl ())
- | Some _ -> ())
- obls
-
+
let default_tactic () = !default_tactic
diff --git a/contrib/subtac/subtac_obligations.mli b/contrib/subtac/subtac_obligations.mli
index cd70d5233..30fbd0284 100644
--- a/contrib/subtac/subtac_obligations.mli
+++ b/contrib/subtac/subtac_obligations.mli
@@ -12,11 +12,18 @@ type progress = (* Resolution status of a program *)
val set_default_tactic : Tacexpr.glob_tactic_expr -> unit
val default_tactic : unit -> Proof_type.tactic
+type definition_hook = constant -> unit
+
val add_definition : Names.identifier -> Term.constr -> Term.types ->
- obligation_info -> progress
+ ?implicits:(Topconstr.explicitation * (bool * bool)) list ->
+ ?kind:Decl_kinds.definition_object_kind ->
+ ?hook:definition_hook -> obligation_info -> progress
val add_mutual_definitions :
- (Names.identifier * Term.constr * Term.types * obligation_info) list -> int array -> unit
+ (Names.identifier * Term.constr * Term.types * obligation_info) list ->
+ ?implicits:(Topconstr.explicitation * (bool * bool)) list ->
+ ?kind:Decl_kinds.definition_object_kind ->
+ int array -> unit
val subtac_obligation : int * Names.identifier option * Topconstr.constr_expr option -> unit
@@ -31,7 +38,7 @@ val try_solve_obligation : int -> Names.identifier option -> Proof_type.tactic -
val try_solve_obligations : Names.identifier option -> Proof_type.tactic -> unit
-val show_obligations : Names.identifier option -> unit
+val show_obligations : ?msg:bool -> Names.identifier option -> unit
val admit_obligations : Names.identifier option -> unit
diff --git a/contrib/subtac/subtac_pretyping.ml b/contrib/subtac/subtac_pretyping.ml
index a0e7e6951..0703bcb83 100644
--- a/contrib/subtac/subtac_pretyping.ml
+++ b/contrib/subtac/subtac_pretyping.ml
@@ -70,7 +70,12 @@ let merge_evms x y =
let interp env isevars c tycon =
let j = pretype tycon env isevars ([],[]) c in
- let evm = evars_of !isevars in
+ let _ = isevars := Evarutil.nf_evar_defs !isevars in
+ let evd,_ = consider_remaining_unif_problems env !isevars in
+ let unevd = undefined_evars evd in
+ let unevd' = Typeclasses.resolve_typeclasses env (Evd.evars_of unevd) evd in
+ let evm = evars_of unevd' in
+ isevars := unevd';
nf_evar evm j.uj_val, nf_evar evm j.uj_type
let find_with_index x l =
@@ -98,7 +103,7 @@ let env_with_binders env isevars l =
let coqdef, deftyp = interp env isevars rawdef empty_tycon in
let reldecl = (name, Some coqdef, deftyp) in
aux (push_rel reldecl env, reldecl :: rels) tl
- | Topconstr.LocalRawAssum (bl, typ) :: tl ->
+ | Topconstr.LocalRawAssum (bl, k, typ) :: tl ->
let rawtyp = coqintern_type !isevars env typ in
let coqtyp, typtyp = interp env isevars rawtyp empty_tycon in
let acc =
@@ -111,45 +116,28 @@ let env_with_binders env isevars l =
| [] -> acc
in aux (env, []) l
-let subtac_process env isevars id l c tycon =
- let c = Command.abstract_constr_expr c l in
-(* let env_binders, binders_rel = env_with_binders env isevars l in *)
+let subtac_process env isevars id bl c tycon =
+(* let bl = Implicit_quantifiers.ctx_of_class_binders (vars_of_env env) cbl @ l in *)
+ let imps = Implicit_quantifiers.implicits_of_binders bl in
+ let c = Command.abstract_constr_expr c bl in
let tycon =
match tycon with
None -> empty_tycon
| Some t ->
- let t = Command.generalize_constr_expr t l in
+ let t = Command.generalize_constr_expr t bl in
let t = coqintern_type !isevars env t in
let coqt, ttyp = interp env isevars t empty_tycon in
mk_tycon coqt
in
let c = coqintern_constr !isevars env c in
let coqc, ctyp = interp env isevars c tycon in
-(* let _ = try trace (str "Interpreted term: " ++ my_print_constr env coqc ++ spc () ++ *)
-(* str "Coq type: " ++ my_print_constr env ctyp) *)
-(* with _ -> () *)
-(* in *)
-(* let _ = try trace (str "Original evar map: " ++ Evd.pr_evar_map (evars_of !isevars)) with _ -> () in *)
-
-(* let fullcoqc = it_mkLambda_or_LetIn coqc binders_rel *)
-(* and fullctyp = it_mkProd_or_LetIn ctyp binders_rel *)
-(* in *)
- let fullcoqc = Evarutil.nf_evar (evars_of !isevars) coqc in
- let fullctyp = Evarutil.nf_evar (evars_of !isevars) ctyp in
-(* let evm = evars_of_term (evars_of !isevars) Evd.empty fullctyp in *)
-(* let evm = evars_of_term (evars_of !isevars) evm fullcoqc in *)
-(* let _ = try trace (str "After evar normalization remain: " ++ spc () ++ *)
-(* Evd.pr_evar_map evm) *)
-(* with _ -> () *)
-(* in *)
let evm = non_instanciated_map env isevars (evars_of !isevars) in
-(* let _ = try trace (str "Non instanciated evars map: " ++ Evd.pr_evar_map evm) with _ -> () in *)
- evm, fullcoqc, fullctyp
+ evm, coqc, ctyp, imps
open Subtac_obligations
-let subtac_proof env isevars id l c tycon =
- let evm, coqc, coqt = subtac_process env isevars id l c tycon in
+let subtac_proof env isevars id bl c tycon =
+ let evm, coqc, coqt, imps = subtac_process env isevars id bl c tycon in
let evm = Subtac_utils.evars_of_term evm Evd.empty coqc in
- let evars, def = Eterm.eterm_obligations env id !isevars evm 0 coqc (Some coqt) in
- add_definition id def coqt evars
+ let evars, def, ty = Eterm.eterm_obligations env id !isevars evm 0 coqc coqt in
+ add_definition id def ty ~implicits:imps evars
diff --git a/contrib/subtac/subtac_pretyping.mli b/contrib/subtac/subtac_pretyping.mli
index afe554c87..4af37043f 100644
--- a/contrib/subtac/subtac_pretyping.mli
+++ b/contrib/subtac/subtac_pretyping.mli
@@ -5,11 +5,19 @@ open Sign
open Evd
open Global
open Topconstr
+open Implicit_quantifiers
+open Impargs
module Pretyping : Pretyping.S
+val interp :
+ Environ.env ->
+ Evd.evar_defs ref ->
+ Rawterm.rawconstr ->
+ Evarutil.type_constraint -> Term.constr * Term.constr
+
val subtac_process : env -> evar_defs ref -> identifier -> local_binder list ->
- constr_expr -> constr_expr option -> evar_map * constr * types
+ constr_expr -> constr_expr option -> evar_map * constr * types * manual_explicitation list
val subtac_proof : env -> evar_defs ref -> identifier -> local_binder list ->
constr_expr -> constr_expr option -> Subtac_obligations.progress
diff --git a/contrib/subtac/subtac_pretyping_F.ml b/contrib/subtac/subtac_pretyping_F.ml
index 0ed0632c6..a2d5be66c 100644
--- a/contrib/subtac/subtac_pretyping_F.ml
+++ b/contrib/subtac/subtac_pretyping_F.ml
@@ -200,11 +200,11 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
| RRec (loc,fixkind,names,bl,lar,vdef) ->
let rec type_bl env ctxt = function
[] -> ctxt
- | (na,None,ty)::bl ->
+ | (na,k,None,ty)::bl ->
let ty' = pretype_type empty_valcon env isevars lvar ty in
let dcl = (na,None,ty'.utj_val) in
type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl
- | (na,Some bd,ty)::bl ->
+ | (na,k,Some bd,ty)::bl ->
let ty' = pretype_type empty_valcon env isevars lvar ty in
let bd' = pretype (mk_tycon ty'.utj_val) env isevars lvar ty in
let dcl = (na,Some bd'.uj_val,ty'.utj_val) in
@@ -326,7 +326,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
| _ -> resj in
inh_conv_coerce_to_tycon loc env isevars resj tycon
- | RLambda(loc,name,c1,c2) ->
+ | RLambda(loc,name,k,c1,c2) ->
let (name',dom,rng) = evd_comb1 (split_tycon loc env) isevars tycon in
let dom_valcon = valcon_of_tycon dom in
let j = pretype_type dom_valcon env isevars lvar c1 in
@@ -334,7 +334,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
let j' = pretype rng (push_rel var env) isevars lvar c2 in
judge_of_abstraction env name j j'
- | RProd(loc,name,c1,c2) ->
+ | RProd(loc,name,k,c1,c2) ->
let j = pretype_type empty_valcon env isevars lvar c1 in
let var = (name,j.utj_val) in
let env' = push_rel_assum var env in
@@ -557,7 +557,14 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
(pretype tycon env isevars lvar c).uj_val
| IsType ->
(pretype_type empty_valcon env isevars lvar c).utj_val in
- nf_evar (evars_of !isevars) c'
+ let evd,_ = consider_remaining_unif_problems env !isevars in
+ let evd = nf_evar_defs evd in
+ let c' = nf_evar (evars_of evd) c' in
+(* let evd = undefined_evars evd in *)
+ let evd = Typeclasses.resolve_typeclasses env (evars_of evd) evd in
+ let c' = nf_evar (evars_of evd) c' in
+ isevars := evd;
+ c'
(* TODO: comment faire remonter l'information si le typage a resolu des
variables du sigma original. il faudrait que la fonction de typage
@@ -585,11 +592,16 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
let ise_pretype_gen fail_evar sigma env lvar kind c =
let isevars = ref (Evd.create_evar_defs sigma) in
let c = pretype_gen isevars env lvar kind c in
- let isevars,_ = consider_remaining_unif_problems env !isevars in
- let c = nf_evar (evars_of isevars) c in
- if fail_evar then check_evars env sigma isevars c;
- isevars, c
-
+(* let evd,_ = consider_remaining_unif_problems env !isevars in *)
+(* let evd = nf_evar_defs evd in *)
+(* let c = nf_evar (evars_of evd) c in *)
+(* let evd = undefined_evars evd in *)
+(* let evd = Typeclasses.resolve_typeclasses env sigma evd in *)
+(* let c = nf_evar (evars_of evd) c in *)
+ let evd = !isevars in
+ if fail_evar then check_evars env (Evd.evars_of evd) evd c;
+ evd, c
+
(** Entry points of the high-level type synthesis algorithm *)
let understand_gen kind sigma env c =
@@ -604,8 +616,15 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
let understand_ltac sigma env lvar kind c =
ise_pretype_gen false sigma env lvar kind c
- let understand_tcc_evars isevars env kind c =
- pretype_gen isevars env ([],[]) kind c
+ let understand_tcc_evars evdref env kind c =
+ pretype_gen evdref env ([],[]) kind c
+
+(* evdref := nf_evar_defs !evdref; *)
+(* let c = nf_evar (evars_of !evdref) c in *)
+(* let evd = undefined_evars !evdref in *)
+(* let evd = Typeclasses.resolve_typeclasses env (evars_of evd) !evdref in *)
+(* evdref := evd; *)
+(* nf_evar (evars_of evd) c *)
let understand_tcc sigma env ?expected_type:exptyp c =
let ev, t = ise_pretype_gen false sigma env ([],[]) (OfType exptyp) c in
diff --git a/contrib/subtac/subtac_utils.ml b/contrib/subtac/subtac_utils.ml
index d5df9adc9..46a8bd203 100644
--- a/contrib/subtac/subtac_utils.ml
+++ b/contrib/subtac/subtac_utils.ml
@@ -357,19 +357,31 @@ let print_message m =
(* Solve an obligation using tactics, return the corresponding proof term *)
let solve_by_tac evi t =
- debug 2 (str "Solving goal using tactics: " ++ Evd.pr_evar_info evi);
let id = id_of_string "H" in
- try
+ try
Pfedit.start_proof id goal_kind evi.evar_hyps evi.evar_concl
(fun _ _ -> ());
- debug 2 (str "Started proof");
Pfedit.by (tclCOMPLETE t);
- let _,(const,_,_) = Pfedit.cook_proof () in
+ let _,(const,_,_) = Pfedit.cook_proof () in
Pfedit.delete_current_proof (); const.Entries.const_entry_body
- with e ->
+ with e ->
Pfedit.delete_current_proof();
raise Exit
+(* let apply_tac t goal = t goal *)
+
+(* let solve_by_tac evi t = *)
+(* let ev = 1 in *)
+(* let evm = Evd.add Evd.empty ev evi in *)
+(* let goal = {it = evi; sigma = evm } in *)
+(* let (res, valid) = apply_tac t goal in *)
+(* if res.it = [] then *)
+(* let prooftree = valid [] in *)
+(* let proofterm, obls = Refiner.extract_open_proof res.sigma prooftree in *)
+(* if obls = [] then proofterm *)
+(* else raise Exit *)
+(* else raise Exit *)
+
let rec string_of_list sep f = function
[] -> ""
| x :: [] -> f x
diff --git a/contrib/xml/xmlcommand.ml b/contrib/xml/xmlcommand.ml
index 82d7c19e5..3a0fdfb9b 100644
--- a/contrib/xml/xmlcommand.ml
+++ b/contrib/xml/xmlcommand.ml
@@ -491,7 +491,10 @@ let kind_of_constant kn =
| DK.IsDefinition DK.IdentityCoercion ->
Pp.warning "IdentityCoercion not supported in dtd (used Definition instead)";
"DEFINITION","Definition"
- | DK.IsProof (DK.Theorem|DK.Lemma|DK.Corollary|DK.Fact|DK.Remark as thm) ->
+ | DK.IsDefinition DK.Instance ->
+ Pp.warning "Instance not supported in dtd (used Definition instead)";
+ "DEFINITION","Definition"
+ | DK.IsProof (DK.Theorem|DK.Lemma|DK.Corollary|DK.Fact|DK.Remark as thm) ->
"THEOREM",DK.string_of_theorem_kind thm
| DK.IsProof _ ->
Pp.warning "Unsupported theorem kind (used Theorem instead)";
diff --git a/dev/include b/dev/include
index 7d0bfdca5..3011bcedf 100644
--- a/dev/include
+++ b/dev/include
@@ -2,7 +2,7 @@
(* File to include to install the pretty-printers in the ocaml toplevel *)
(* clflags.cmi (a ocaml compilation by-product) must be in the library path *)
-Clflags.recursive_types := true;;
+(* Clflags.recursive_types := true;;*)
#cd ".";;
#use "base_include";;
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 1117d2507..7e288f311 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -216,7 +216,7 @@ let rec check_same_type ty1 ty2 =
| _ when ty1=ty2 -> ()
| _ -> failwith "not same type"
-and check_same_binder (nal1,e1) (nal2,e2) =
+and check_same_binder (nal1,_,e1) (nal2,_,e2) =
List.iter2 (fun (_,na1) (_,na2) ->
if na1<>na2 then failwith "not same name") nal1 nal2;
check_same_type e1 e2
@@ -224,10 +224,10 @@ and check_same_binder (nal1,e1) (nal2,e2) =
and check_same_fix_binder bl1 bl2 =
List.iter2 (fun b1 b2 ->
match b1,b2 with
- LocalRawAssum(nal1,ty1), LocalRawAssum(nal2,ty2) ->
- check_same_binder (nal1,ty1) (nal2,ty2)
+ LocalRawAssum(nal1,k,ty1), LocalRawAssum(nal2,k',ty2) ->
+ check_same_binder (nal1,k,ty1) (nal2,k',ty2)
| LocalRawDef(na1,def1), LocalRawDef(na2,def2) ->
- check_same_binder ([na1],def1) ([na2],def2)
+ check_same_binder ([na1],default_binder_kind,def1) ([na2],default_binder_kind,def2)
| _ -> failwith "not same binder") bl1 bl2
let same c d = try check_same_type c d; true with _ -> false
@@ -255,10 +255,10 @@ let rec same_raw c d =
| RPatVar(_,pv1), RPatVar(_,pv2) -> if pv1<>pv2 then failwith "RPatVar"
| RApp(_,f1,a1), RApp(_,f2,a2) ->
List.iter2 same_raw (f1::a1) (f2::a2)
- | RLambda(_,na1,t1,m1), RLambda(_,na2,t2,m2) ->
+ | RLambda(_,na1,bk1,t1,m1), RLambda(_,na2,bk2,t2,m2) ->
if na1 <> na2 then failwith "RLambda";
same_raw t1 t2; same_raw m1 m2
- | RProd(_,na1,t1,m1), RProd(_,na2,t2,m2) ->
+ | RProd(_,na1,bk1,t1,m1), RProd(_,na2,bk2,t2,m2) ->
if na1 <> na2 then failwith "RProd";
same_raw t1 t2; same_raw m1 m2
| RLetIn(_,na1,t1,m1), RLetIn(_,na2,t2,m2) ->
@@ -283,7 +283,7 @@ let rec same_raw c d =
| RRec(_,fk1,na1,bl1,ty1,def1), RRec(_,fk2,na2,bl2,ty2,def2) ->
if fk1 <> fk2 || na1 <> na2 then failwith "RRec";
array_iter2
- (List.iter2 (fun (na1,bd1,ty1) (na2,bd2,ty2) ->
+ (List.iter2 (fun (na1,bk1,bd1,ty1) (na2,bk2,bd2,ty2) ->
if na1<>na2 then failwith "RRec";
Option.iter2 same_raw bd1 bd2;
same_raw ty1 ty2)) bl1 bl2;
@@ -582,7 +582,7 @@ let rec rename_rawconstr_var id0 id1 = function
let rec share_fix_binders n rbl ty def =
match ty,def with
- RProd(_,na0,t0,b), RLambda(_,na1,t1,m) ->
+ RProd(_,na0,bk0,t0,b), RLambda(_,na1,bk1,t1,m) ->
if not(same_rawconstr t0 t1) then List.rev rbl, ty, def
else
let (na,b,m) =
@@ -672,7 +672,7 @@ let rec extern inctx scopes vars r =
explicitize loc inctx [] (None,sub_extern false scopes vars f)
(List.map (sub_extern true scopes vars) args))
- | RProd (loc,Anonymous,t,c) ->
+ | RProd (loc,Anonymous,_,t,c) ->
(* Anonymous product are never factorized *)
CArrow (loc,extern_typ scopes vars t, extern_typ scopes vars c)
@@ -680,15 +680,15 @@ let rec extern inctx scopes vars r =
CLetIn (loc,(loc,na),sub_extern false scopes vars t,
extern inctx scopes (add_vname vars na) c)
- | RProd (loc,na,t,c) ->
+ | RProd (loc,na,bk,t,c) ->
let t = extern_typ scopes vars (anonymize_if_reserved na t) in
let (idl,c) = factorize_prod scopes (add_vname vars na) t c in
- CProdN (loc,[(dummy_loc,na)::idl,t],c)
+ CProdN (loc,[(dummy_loc,na)::idl,Default bk,t],c)
- | RLambda (loc,na,t,c) ->
+ | RLambda (loc,na,bk,t,c) ->
let t = extern_typ scopes vars (anonymize_if_reserved na t) in
let (idl,c) = factorize_lambda inctx scopes (add_vname vars na) t c in
- CLambdaN (loc,[(dummy_loc,na)::idl,t],c)
+ CLambdaN (loc,[(dummy_loc,na)::idl,Default bk,t],c)
| RCases (loc,rtntypopt,tml,eqns) ->
let vars' =
@@ -775,7 +775,7 @@ and factorize_prod scopes vars aty c =
if !Flags.raw_print or !print_no_symbol then raise No_match;
([],extern_symbol scopes vars c (uninterp_notations c))
with No_match -> match c with
- | RProd (loc,(Name id as na),ty,c)
+ | RProd (loc,(Name id as na),bk,ty,c)
when same aty (extern_typ scopes vars (anonymize_if_reserved na ty))
& not (occur_var_constr_expr id aty) (* avoid na in ty escapes scope *)
-> let (nal,c) = factorize_prod scopes (Idset.add id vars) aty c in
@@ -787,7 +787,7 @@ and factorize_lambda inctx scopes vars aty c =
if !Flags.raw_print or !print_no_symbol then raise No_match;
([],extern_symbol scopes vars c (uninterp_notations c))
with No_match -> match c with
- | RLambda (loc,na,ty,c)
+ | RLambda (loc,na,bk,ty,c)
when same aty (extern_typ scopes vars (anonymize_if_reserved na ty))
& not (occur_name na aty) (* To avoid na in ty' escapes scope *)
-> let (nal,c) =
@@ -797,24 +797,24 @@ and factorize_lambda inctx scopes vars aty c =
and extern_local_binder scopes vars = function
[] -> ([],[])
- | (na,Some bd,ty)::l ->
+ | (na,bk,Some bd,ty)::l ->
let (ids,l) =
extern_local_binder scopes (name_fold Idset.add na vars) l in
(na::ids,
LocalRawDef((dummy_loc,na), extern false scopes vars bd) :: l)
- | (na,None,ty)::l ->
+ | (na,bk,None,ty)::l ->
let ty = extern_typ scopes vars (anonymize_if_reserved na ty) in
(match extern_local_binder scopes (name_fold Idset.add na vars) l with
- (ids,LocalRawAssum(nal,ty')::l)
+ (ids,LocalRawAssum(nal,k,ty')::l)
when same ty ty' &
match na with Name id -> not (occur_var_constr_expr id ty')
| _ -> true ->
(na::ids,
- LocalRawAssum((dummy_loc,na)::nal,ty')::l)
+ LocalRawAssum((dummy_loc,na)::nal,k,ty')::l)
| (ids,l) ->
(na::ids,
- LocalRawAssum([(dummy_loc,na)],ty) :: l))
+ LocalRawAssum([(dummy_loc,na)],Default bk,ty) :: l))
and extern_eqn inctx scopes vars (loc,ids,pl,c) =
(loc,[List.map (extern_cases_pattern_in_scope scopes vars) pl],
@@ -927,11 +927,11 @@ let rec raw_of_pat env = function
RApp (loc,RPatVar (loc,(true,n)),
List.map (raw_of_pat env) args)
| PProd (na,t,c) ->
- RProd (loc,na,raw_of_pat env t,raw_of_pat (na::env) c)
+ RProd (loc,na,Explicit,raw_of_pat env t,raw_of_pat (na::env) c)
| PLetIn (na,t,c) ->
RLetIn (loc,na,raw_of_pat env t, raw_of_pat (na::env) c)
| PLambda (na,t,c) ->
- RLambda (loc,na,raw_of_pat env t, raw_of_pat (na::env) c)
+ RLambda (loc,na,Explicit,raw_of_pat env t, raw_of_pat (na::env) c)
| PIf (c,b1,b2) ->
RIf (loc, raw_of_pat env c, (Anonymous,None),
raw_of_pat env b1, raw_of_pat env b2)
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index a41825346..3214ca6c4 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -86,7 +86,7 @@ let explain_bad_patterns_number n1 n2 =
let explain_bad_explicitation_number n po =
match n with
- | ExplByPos n ->
+ | ExplByPos (n,_id) ->
let s = match po with
| None -> str "a regular argument"
| Some p -> int p in
@@ -683,7 +683,7 @@ let extract_explicit_arg imps args =
user_err_loc (loc,"",str "Argument name " ++ pr_id id
++ str " occurs more than once");
id
- | ExplByPos p ->
+ | ExplByPos (p,_id) ->
let id =
try
let imp = List.nth imps (p-1) in
@@ -775,6 +775,16 @@ let set_type_scope (ids,tmp_scope,scopes) =
let reset_tmp_scope (ids,tmp_scope,scopes) =
(ids,None,scopes)
+let rec it_mkRProd env body =
+ match env with
+ (na, bk, _, t) :: tl -> it_mkRProd tl (RProd (dummy_loc, na, bk, t, body))
+ | [] -> body
+
+let rec it_mkRLambda env body =
+ match env with
+ (na, bk, _, t) :: tl -> it_mkRLambda tl (RLambda (dummy_loc, na, bk, t, body))
+ | [] -> body
+
(**********************************************************************)
(* Main loop *)
@@ -844,15 +854,15 @@ let internalise sigma globalenv env allow_patvar lvar c =
Array.map (fun (_,ty,_) -> ty) idl,
Array.map (fun (_,_,bd) -> bd) idl)
| CArrow (loc,c1,c2) ->
- RProd (loc, Anonymous, intern_type env c1, intern_type env c2)
+ RProd (loc, Anonymous, Explicit, intern_type env c1, intern_type env c2)
| CProdN (loc,[],c2) ->
intern_type env c2
- | CProdN (loc,(nal,ty)::bll,c2) ->
- iterate_prod loc env ty (CProdN (loc, bll, c2)) nal
+ | CProdN (loc,(nal,bk,ty)::bll,c2) ->
+ iterate_prod loc env bk ty (CProdN (loc, bll, c2)) nal
| CLambdaN (loc,[],c2) ->
intern env c2
- | CLambdaN (loc,(nal,ty)::bll,c2) ->
- iterate_lam loc (reset_tmp_scope env) ty (CLambdaN (loc, bll, c2)) nal
+ | CLambdaN (loc,(nal,bk,ty)::bll,c2) ->
+ iterate_lam loc (reset_tmp_scope env) bk ty (CLambdaN (loc, bll, c2)) nal
| CLetIn (loc,(_,na),c1,c2) ->
RLetIn (loc, na, intern (reset_tmp_scope env) c1,
intern (push_name_env lvar env na) c2)
@@ -934,17 +944,20 @@ let internalise sigma globalenv env allow_patvar lvar c =
and intern_type env = intern (set_type_scope env)
and intern_local_binder ((ids,ts,sc as env),bl) = function
- | LocalRawAssum(nal,ty) ->
- let (loc,na) = List.hd nal in
- (* TODO: fail if several names with different implicit types *)
- let ty = locate_if_isevar loc na (intern_type env ty) in
- List.fold_left
- (fun ((ids,ts,sc),bl) (_,na) ->
- ((name_fold Idset.add na ids,ts,sc), (na,None,ty)::bl))
- (env,bl) nal
+ | LocalRawAssum(nal,bk,ty) ->
+ (match bk with
+ | Default k ->
+ let (loc,na) = List.hd nal in
+ (* TODO: fail if several names with different implicit types *)
+ let ty = locate_if_isevar loc na (intern_type env ty) in
+ List.fold_left
+ (fun ((ids,ts,sc),bl) (_,na) ->
+ ((name_fold Idset.add na ids,ts,sc), (na,k,None,ty)::bl))
+ (env,bl) nal
+ | TypeClass b -> anomaly ("TODO: intern_local_binder TypeClass"))
| LocalRawDef((loc,na),def) ->
((name_fold Idset.add na ids,ts,sc),
- (na,Some(intern env def),RHole(loc,Evd.BinderType na))::bl)
+ (na,Explicit,Some(intern env def),RHole(loc,Evd.BinderType na))::bl)
(* Expands a multiple pattern into a disjunction of multiple patterns *)
and intern_multiple_pattern scopes pl =
@@ -1004,22 +1017,50 @@ let internalise sigma globalenv env allow_patvar lvar c =
| _, None -> Anonymous
| _, Some na -> na in
(tm',(na,typ)), na::ids
+
+ and intern_typeclass_binders env bl =
+ List.fold_left
+ (fun ((ids,ts,sc) as env,bl) ((loc, na), bk, ty) ->
+ let ty = locate_if_isevar loc na (intern_type env ty) in
+ ((name_fold Idset.add na ids,ts,sc), (na,bk,None,ty)::bl))
+ env bl
- and iterate_prod loc2 env ty body = function
+ and iterate_prod loc2 env bk ty body nal =
+ let rec default env bk = function
| (loc1,na)::nal ->
if nal <> [] then check_capture loc1 ty na;
- let body = iterate_prod loc2 (push_name_env lvar env na) ty body nal in
+ let body = default (push_name_env lvar env na) bk nal in
let ty = locate_if_isevar loc1 na (intern_type env ty) in
- RProd (join_loc loc1 loc2, na, ty, body)
+ RProd (join_loc loc1 loc2, na, bk, ty, body)
| [] -> intern_type env body
-
- and iterate_lam loc2 env ty body = function
- | (loc1,na)::nal ->
- if nal <> [] then check_capture loc1 ty na;
- let body = iterate_lam loc2 (push_name_env lvar env na) ty body nal in
- let ty = locate_if_isevar loc1 na (intern_type env ty) in
- RLambda (join_loc loc1 loc2, na, ty, body)
- | [] -> intern env body
+ in
+ match bk with
+ | Default b -> default env b nal
+ | TypeClass b ->
+ let ctx = (List.hd nal, b, ty) in
+ let (fvs, bind) = Implicit_quantifiers.generalize_class_binders_raw (pi1 env) [ctx] in
+ let env, ifvs = intern_typeclass_binders (env,[]) fvs in
+ let env, ibind = intern_typeclass_binders (env,ifvs) bind in
+ let body = intern_type env body in
+ it_mkRProd ibind body
+
+ and iterate_lam loc2 env bk ty body nal =
+ let rec default env bk = function
+ | (loc1,na)::nal ->
+ if nal <> [] then check_capture loc1 ty na;
+ let body = default (push_name_env lvar env na) bk nal in
+ let ty = locate_if_isevar loc1 na (intern_type env ty) in
+ RLambda (join_loc loc1 loc2, na, bk, ty, body)
+ | [] -> intern env body
+ in match bk with
+ | Default b -> default env b nal
+ | TypeClass b ->
+ let ctx = (List.hd nal, b, ty) in
+ let (fvs, bind) = Implicit_quantifiers.generalize_class_binders_raw (pi1 env) [ctx] in
+ let env, ifvs = intern_typeclass_binders (env,[]) fvs in
+ let env, ibind = intern_typeclass_binders (env,ifvs) bind in
+ let body = intern env body in
+ it_mkRLambda ibind body
and intern_impargs c env l subscopes args =
let eargs, rargs = extract_explicit_arg l args in
@@ -1046,7 +1087,7 @@ let internalise sigma globalenv env allow_patvar lvar c =
| (imp::impl', []) ->
if eargs <> [] then
(let (id,(loc,_)) = List.hd eargs in
- user_err_loc (loc,"",str "Not enough non implicit
+ user_err_loc (loc,"",str "Not enough non implicit
arguments to accept the argument bound to " ++ pr_id id));
[]
| ([], rargs) ->
@@ -1123,7 +1164,6 @@ let interp_open_constr sigma env c =
let interp_constr_judgment sigma env c =
Default.understand_judgment sigma env (intern_constr sigma env c)
-
let interp_constr_evars_gen evdref env ?(impls=([],[])) kind c =
Default.understand_tcc_evars evdref env kind
(intern_gen (kind=IsType) ~impls (Evd.evars_of !evdref) env c)
@@ -1175,11 +1215,11 @@ open Term
let interp_context sigma env params =
List.fold_left
(fun (env,params) d -> match d with
- | LocalRawAssum ([_,na],(CHole _ as t)) ->
+ | LocalRawAssum ([_,na],k,(CHole _ as t)) ->
let t = interp_binder sigma env na t in
let d = (na,None,t) in
(push_rel d env, d::params)
- | LocalRawAssum (nal,t) ->
+ | LocalRawAssum (nal,k,t) ->
let t = interp_type sigma env t in
let ctx = list_map_i (fun i (_,na) -> (na,None,lift i t)) 0 nal in
let ctx = List.rev ctx in
@@ -1193,11 +1233,11 @@ let interp_context sigma env params =
let interp_context_evars evdref env params =
List.fold_left
(fun (env,params) d -> match d with
- | LocalRawAssum ([_,na],(CHole _ as t)) ->
+ | LocalRawAssum ([_,na],k,(CHole _ as t)) ->
let t = interp_binder_evars evdref env na t in
let d = (na,None,t) in
(push_rel d env, d::params)
- | LocalRawAssum (nal,t) ->
+ | LocalRawAssum (nal,k,t) ->
let t = interp_type_evars evdref env t in
let ctx = list_map_i (fun i (_,na) -> (na,None,lift i t)) 0 nal in
let ctx = List.rev ctx in
diff --git a/interp/constrintern.mli b/interp/constrintern.mli
index eac01a92a..f4272a2b1 100644
--- a/interp/constrintern.mli
+++ b/interp/constrintern.mli
@@ -103,6 +103,8 @@ val interp_reference : ltac_sign -> reference -> rawconstr
val interp_binder : evar_map -> env -> name -> constr_expr -> types
+val interp_binder_evars : evar_defs ref -> env -> name -> constr_expr -> types
+
(* Interpret contexts: returns extended env and context *)
val interp_context : evar_map -> env -> local_binder list -> env * rel_context
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
new file mode 100644
index 000000000..d9113644c
--- /dev/null
+++ b/interp/implicit_quantifiers.ml
@@ -0,0 +1,215 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.ml 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Util
+open Rawterm
+open Topconstr
+open Libnames
+open Typeclasses
+open Typeclasses_errors
+(*i*)
+
+(* Auxilliary functions for the inference of implicitly quantified variables. *)
+
+let free_vars_of_constr_expr c ?(bound=Idset.empty) l =
+ let found id bdvars l = if Idset.mem id bdvars then l else if List.mem id l then l else id :: l in
+ let rec aux bdvars l c = match c with
+ | CRef (Ident (_,id)) -> found id bdvars l
+ | CNotation (_, "{ _ : _ | _ }", (CRef (Ident (_, id))) :: _) when not (Idset.mem id bdvars) ->
+ fold_constr_expr_with_binders (fun a l -> Idset.add a l) aux (Idset.add id bdvars) l c
+ | c -> fold_constr_expr_with_binders (fun a l -> Idset.add a l) aux bdvars l c
+ in aux bound l c
+
+
+let locate_reference qid =
+ match Nametab.extended_locate qid with
+ | TrueGlobal ref -> ref
+ | SyntacticDef kn ->
+ match Syntax_def.search_syntactic_definition dummy_loc kn with
+ | Rawterm.RRef (_,ref) -> ref
+ | _ -> raise Not_found
+
+let is_global id =
+ try
+ let _ = locate_reference (make_short_qualid id) in true
+ with Not_found ->
+ false
+
+let is_freevar ids env x =
+ try
+ if Idset.mem x ids then false
+ else
+ try ignore(Environ.lookup_named x env) ; false
+ with _ -> not (is_global x)
+ with _ -> true
+
+let freevars_of_ids env ids =
+ List.filter (is_freevar env (Global.env())) ids
+
+let compute_constrs_freevars env constrs =
+ let ids =
+ List.rev (List.fold_left
+ (fun acc x -> free_vars_of_constr_expr x acc)
+ [] constrs)
+ in freevars_of_ids env ids
+
+(* let compute_context_freevars env ctx = *)
+(* let ids = *)
+(* List.rev *)
+(* (List.fold_left *)
+(* (fun acc (_,i,x) -> free_vars_of_constr_expr x acc) *)
+(* [] constrs) *)
+(* in freevars_of_ids ids *)
+
+let compute_constrs_freevars_binders env constrs =
+ let elts = compute_constrs_freevars env constrs in
+ List.map (fun id -> (dummy_loc, id), CHole dummy_loc) elts
+
+let ids_of_named_context_avoiding avoid l =
+ List.fold_left (fun (ids, avoid) id ->
+ let id' = Nameops.next_ident_away_from id avoid in id' :: ids, id' :: avoid)
+ ([], avoid) (Termops.ids_of_named_context l)
+
+let combine_params avoid applied needed =
+ let rec aux ids app need =
+ match app, need with
+ [], need ->
+ let need', avoid = ids_of_named_context_avoiding avoid need in
+ List.rev ids @ (List.map mkIdentC need'), avoid
+ | x :: app, _ :: need -> aux (x :: ids) app need
+ | _ :: _, [] -> failwith "combine_params: overly applied typeclass"
+ in aux [] applied needed
+
+let compute_context_vars env l =
+ List.fold_left (fun l (iid, _, c) ->
+ (match snd iid with Name i -> [i] | Anonymous -> []) @ free_vars_of_constr_expr c ~bound:env l)
+ [] l
+
+let destClassApp cl =
+ match cl with
+ | CApp (loc, (None,CRef (Ident f)), l) -> f, List.map fst l
+ | _ -> raise Not_found
+
+let full_class_binders env l =
+ let avoid = compute_context_vars env l in
+ let l', avoid =
+ List.fold_left (fun (l', avoid) (iid, bk, cl as x) ->
+ match bk with
+ Explicit ->
+ let (id, l) = destClassApp cl in
+ (try
+ let c = class_info (snd id) in
+ let args, avoid = combine_params avoid l (List.rev c.cl_context @ List.rev c.cl_super @ List.rev c.cl_params) in
+ (iid, bk, CAppExpl (fst id, (None, Ident id), args)) :: l', avoid
+ with Not_found -> unbound_class (Global.env ()) id)
+ | Implicit -> (x :: l', avoid))
+ ([], avoid) l
+ in List.rev l'
+
+let constr_expr_of_constraint (kind, id) l =
+ match kind with
+ | Explicit -> CAppExpl (fst id, (None, Ident id), l)
+ | Implicit -> CApp (fst id, (None, CRef (Ident id)),
+ List.map (fun x -> x, None) l)
+
+(* | CApp of loc * (proj_flag * constr_expr) * *)
+(* (constr_expr * explicitation located option) list *)
+
+
+let constrs_of_context l =
+ List.map (fun (_, id, l) -> constr_expr_of_constraint id l) l
+
+let compute_context_freevars env ctx =
+ let bound, ids =
+ List.fold_left
+ (fun (bound, acc) (oid, id, x) ->
+ let bound = match snd oid with Name n -> Idset.add n bound | Anonymous -> bound in
+ bound, free_vars_of_constr_expr x ~bound acc)
+ (env,[]) ctx
+ in freevars_of_ids env (List.rev ids)
+
+let resolve_class_binders env l =
+ let ctx = full_class_binders env l in
+ let fv_ctx =
+ let elts = compute_context_freevars env ctx in
+ List.map (fun id -> (dummy_loc, id), CHole dummy_loc) elts
+ in
+ fv_ctx, ctx
+
+let generalize_class_binders env l =
+ let fv_ctx, cstrs = resolve_class_binders env l in
+ List.map (fun ((loc, id), t) -> LocalRawAssum ([loc, Name id], Default Implicit, t)) fv_ctx,
+ List.map (fun (iid, bk, c) -> LocalRawAssum ([iid], Default Implicit, c))
+ cstrs
+
+let generalize_class_binders_raw env l =
+ let fv_ctx, cstrs = resolve_class_binders env l in
+ List.map (fun ((loc, id), t) -> ((loc, Name id), Implicit, t)) fv_ctx,
+ List.map (fun (iid, bk, c) -> (iid, Implicit, c)) cstrs
+
+let ctx_of_class_binders env l =
+ let (x, y) = generalize_class_binders env l in x @ y
+
+let implicits_of_binders l =
+ let rec aux i l =
+ match l with
+ [] -> []
+ | hd :: tl ->
+ let res, reslen =
+ match hd with
+ LocalRawAssum (nal, Default Implicit, t) ->
+ list_map_i (fun i (_,id) ->
+ let name =
+ match id with
+ Name id -> Some id
+ | Anonymous -> None
+ in ExplByPos (i, name), (true, true))
+ i nal, List.length nal
+ | LocalRawAssum (nal, _, _) -> [], List.length nal
+ | LocalRawDef _ -> [], 1
+ in res @ (aux (i + reslen) tl)
+ in aux 1 l
+
+let implicits_of_rawterm l =
+ let rec aux i c =
+ match c with
+ RProd (loc, na, bk, t, b) | RLambda (loc, na, bk, t, b) ->
+ let rest = aux (succ i) b in
+ if bk = Implicit then
+ let name =
+ match na with
+ Name id -> Some id
+ | Anonymous -> None
+ in
+ (ExplByPos (i, name), (true, true)) :: rest
+ else rest
+ | RLetIn (loc, na, t, b) -> aux i b
+ | _ -> []
+ in aux 1 l
+
+let nf_named_context sigma ctx =
+ Sign.map_named_context (Reductionops.nf_evar sigma) ctx
+
+let nf_rel_context sigma ctx =
+ Sign.map_rel_context (Reductionops.nf_evar sigma) ctx
+
+let nf_env sigma env =
+ let nc' = nf_named_context sigma (Environ.named_context env) in
+ let rel' = nf_rel_context sigma (Environ.rel_context env) in
+ push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env)
diff --git a/interp/implicit_quantifiers.mli b/interp/implicit_quantifiers.mli
new file mode 100644
index 000000000..a61dbcadf
--- /dev/null
+++ b/interp/implicit_quantifiers.mli
@@ -0,0 +1,51 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.ml 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Rawterm
+open Topconstr
+open Util
+open Typeclasses
+(*i*)
+
+val destClassApp : constr_expr -> identifier located * constr_expr list
+
+val free_vars_of_constr_expr : Topconstr.constr_expr ->
+ ?bound:Idset.t ->
+ Names.identifier list -> Names.identifier list
+
+val compute_constrs_freevars : Idset.t -> constr_expr list -> identifier list
+val compute_constrs_freevars_binders : Idset.t -> constr_expr list -> (identifier located * constr_expr) list
+val resolve_class_binders : Idset.t -> typeclass_context ->
+ (identifier located * constr_expr) list * typeclass_context
+
+val full_class_binders : Idset.t -> typeclass_context -> typeclass_context
+
+val generalize_class_binders_raw : Idset.t -> typeclass_context ->
+ (name located * binding_kind * constr_expr) list * (name located * binding_kind * constr_expr) list
+
+val ctx_of_class_binders : Idset.t -> typeclass_context -> local_binder list
+
+val implicits_of_binders : local_binder list -> (Topconstr.explicitation * (bool * bool)) list
+
+val implicits_of_rawterm : Rawterm.rawconstr -> (Topconstr.explicitation * (bool * bool)) list
+
+val nf_named_context : evar_map -> named_context -> named_context
+val nf_rel_context : evar_map -> rel_context -> rel_context
+val nf_env : evar_map -> env -> env
+
diff --git a/interp/reserve.ml b/interp/reserve.ml
index 02e15f069..f3c3506b5 100644
--- a/interp/reserve.ml
+++ b/interp/reserve.ml
@@ -54,8 +54,8 @@ open Rawterm
let rec unloc = function
| RVar (_,id) -> RVar (dummy_loc,id)
| RApp (_,g,args) -> RApp (dummy_loc,unloc g, List.map unloc args)
- | RLambda (_,na,ty,c) -> RLambda (dummy_loc,na,unloc ty,unloc c)
- | RProd (_,na,ty,c) -> RProd (dummy_loc,na,unloc ty,unloc c)
+ | RLambda (_,na,bk,ty,c) -> RLambda (dummy_loc,na,bk,unloc ty,unloc c)
+ | RProd (_,na,bk,ty,c) -> RProd (dummy_loc,na,bk,unloc ty,unloc c)
| RLetIn (_,na,b,c) -> RLetIn (dummy_loc,na,unloc b,unloc c)
| RCases (_,rtntypopt,tml,pl) ->
RCases (dummy_loc,
@@ -69,7 +69,7 @@ let rec unloc = function
| RRec (_,fk,idl,bl,tyl,bv) ->
RRec (dummy_loc,fk,idl,
Array.map (List.map
- (fun (na,obd,ty) -> (na,Option.map unloc obd, unloc ty)))
+ (fun (na,k,obd,ty) -> (na,k,Option.map unloc obd, unloc ty)))
bl,
Array.map unloc tyl,
Array.map unloc bv)
diff --git a/interp/topconstr.ml b/interp/topconstr.ml
index e29f17210..2994bc3ae 100644
--- a/interp/topconstr.ml
+++ b/interp/topconstr.ml
@@ -74,9 +74,9 @@ let rawconstr_of_aconstr_with_binders loc g f e = function
let outerl = (ldots_var,inner)::(if swap then [x,RVar(loc,y)] else []) in
subst_rawvars outerl it
| ALambda (na,ty,c) ->
- let e,na = name_fold_map g e na in RLambda (loc,na,f e ty,f e c)
+ let e,na = name_fold_map g e na in RLambda (loc,na,Explicit,f e ty,f e c)
| AProd (na,ty,c) ->
- let e,na = name_fold_map g e na in RProd (loc,na,f e ty,f e c)
+ let e,na = name_fold_map g e na in RProd (loc,na,Explicit,f e ty,f e c)
| ALetIn (na,b,c) ->
let e,na = name_fold_map g e na in RLetIn (loc,na,f e b,f e c)
| ACases (rtntypopt,tml,eqnl) ->
@@ -131,9 +131,9 @@ let compare_rawconstr f t1 t2 = match t1,t2 with
| RRef (_,r1), RRef (_,r2) -> r1 = r2
| RVar (_,v1), RVar (_,v2) -> v1 = v2
| RApp (_,f1,l1), RApp (_,f2,l2) -> f f1 f2 & List.for_all2 f l1 l2
- | RLambda (_,na1,ty1,c1), RLambda (_,na2,ty2,c2) when na1 = na2 ->
+ | RLambda (_,na1,bk1,ty1,c1), RLambda (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 ->
f ty1 ty2 & f c1 c2
- | RProd (_,na1,ty1,c1), RProd (_,na2,ty2,c2) when na1 = na2 ->
+ | RProd (_,na1,bk1,ty1,c1), RProd (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 ->
f ty1 ty2 & f c1 c2
| RHole _, RHole _ -> true
| RSort (_,s1), RSort (_,s2) -> s1 = s2
@@ -180,8 +180,8 @@ let aconstr_and_vars_of_rawconstr a =
found := ldots_var :: !found; assert lassoc;
AList (x,y,AApp (AVar ldots_var,[AVar x]),aux t,lassoc)
| RApp (_,g,args) -> AApp (aux g, List.map aux args)
- | RLambda (_,na,ty,c) -> add_name found na; ALambda (na,aux ty,aux c)
- | RProd (_,na,ty,c) -> add_name found na; AProd (na,aux ty,aux c)
+ | RLambda (_,na,bk,ty,c) -> add_name found na; ALambda (na,aux ty,aux c)
+ | RProd (_,na,bk,ty,c) -> add_name found na; AProd (na,aux ty,aux c)
| RLetIn (_,na,b,c) -> add_name found na; ALetIn (na,aux b,aux c)
| RCases (_,rtntypopt,tml,eqnl) ->
let f (_,idl,pat,rhs) = found := idl@(!found); (pat,aux rhs) in
@@ -377,7 +377,7 @@ let abstract_return_type_context pi mklam tml rtno =
let abstract_return_type_context_rawconstr =
abstract_return_type_context (fun (_,_,_,nal) -> nal)
- (fun na c -> RLambda(dummy_loc,na,RHole(dummy_loc,Evd.InternalHole),c))
+ (fun na c -> RLambda(dummy_loc,na,Explicit,RHole(dummy_loc,Evd.InternalHole),c))
let abstract_return_type_context_aconstr =
abstract_return_type_context pi3
@@ -440,9 +440,9 @@ let rec match_ alp metas sigma a1 a2 = match (a1,a2) with
| RApp (_,f1,l1), AList (x,_,(AApp (f2,l2) as iter),termin,lassoc)
when List.length l1 = List.length l2 ->
match_alist alp metas sigma (f1::l1) (f2::l2) x iter termin lassoc
- | RLambda (_,na1,t1,b1), ALambda (na2,t2,b2) ->
+ | RLambda (_,na1,_,t1,b1), ALambda (na2,t2,b2) ->
match_binders alp metas na1 na2 (match_ alp metas sigma t1 t2) b1 b2
- | RProd (_,na1,t1,b1), AProd (na2,t2,b2) ->
+ | RProd (_,na1,_,t1,b1), AProd (na2,t2,b2) ->
match_binders alp metas na1 na2 (match_ alp metas sigma t1 t2) b1 b2
| RLetIn (_,na1,t1,b1), ALetIn (na2,t2,b2) ->
match_binders alp metas na1 na2 (match_ alp metas sigma t1 t2) b1 b2
@@ -530,7 +530,9 @@ let match_aconstr c (metas_scl,pat) =
type notation = string
-type explicitation = ExplByPos of int | ExplByName of identifier
+type explicitation = ExplByPos of int * identifier option | ExplByName of identifier
+
+type binder_kind = Default of binding_kind | TypeClass of binding_kind
type proj_flag = int option (* [Some n] = proj of the n-th visible argument *)
@@ -550,8 +552,8 @@ type constr_expr =
| CFix of loc * identifier located * fixpoint_expr list
| CCoFix of loc * identifier located * cofixpoint_expr list
| CArrow of loc * constr_expr * constr_expr
- | CProdN of loc * (name located list * constr_expr) list * constr_expr
- | CLambdaN of loc * (name located list * constr_expr) list * constr_expr
+ | CProdN of loc * (name located list * binder_kind * constr_expr) list * constr_expr
+ | CLambdaN of loc * (name located list * binder_kind * constr_expr) list * constr_expr
| CLetIn of loc * name located * constr_expr * constr_expr
| CAppExpl of loc * (proj_flag * reference) * constr_expr list
| CApp of loc * (proj_flag * constr_expr) *
@@ -579,7 +581,11 @@ and fixpoint_expr =
and local_binder =
| LocalRawDef of name located * constr_expr
- | LocalRawAssum of name located list * constr_expr
+ | LocalRawAssum of name located list * binder_kind * constr_expr
+
+and typeclass_constraint = name located * binding_kind * constr_expr
+
+and typeclass_context = typeclass_constraint list
and cofixpoint_expr =
identifier * local_binder list * constr_expr * constr_expr
@@ -592,21 +598,23 @@ and recursion_order_expr =
(***********************)
(* For binders parsing *)
+let default_binder_kind = Default Explicit
+
let rec local_binders_length = function
| [] -> 0
| LocalRawDef _::bl -> 1 + local_binders_length bl
- | LocalRawAssum (idl,_)::bl -> List.length idl + local_binders_length bl
+ | LocalRawAssum (idl,_,_)::bl -> List.length idl + local_binders_length bl
let rec local_assums_length = function
| [] -> 0
| LocalRawDef _::bl -> local_binders_length bl
- | LocalRawAssum (idl,_)::bl -> List.length idl + local_binders_length bl
+ | LocalRawAssum (idl,_,_)::bl -> List.length idl + local_binders_length bl
let names_of_local_assums bl =
- List.flatten (List.map (function LocalRawAssum(l,_)->l|_->[]) bl)
+ List.flatten (List.map (function LocalRawAssum(l,_,_)->l|_->[]) bl)
let names_of_local_binders bl =
- List.flatten (List.map (function LocalRawAssum(l,_)->l|LocalRawDef(l,_)->[l]) bl)
+ List.flatten (List.map (function LocalRawAssum(l,_,_)->l|LocalRawDef(l,_)->[l]) bl)
(**********************************************************************)
(* Functions on constr_expr *)
@@ -684,7 +692,7 @@ let ids_of_pattern_list =
Idset.empty
let rec fold_constr_expr_binders g f n acc b = function
- | (nal,t)::l ->
+ | (nal,bk,t)::l ->
let nal = snd (List.split nal) in
let n' = List.fold_right (name_fold g) nal n in
f n (fold_constr_expr_binders g f n' acc b l) t
@@ -692,7 +700,7 @@ let rec fold_constr_expr_binders g f n acc b = function
f n acc b
let rec fold_local_binders g f n acc b = function
- | LocalRawAssum (nal,t)::l ->
+ | LocalRawAssum (nal,bk,t)::l ->
let nal = snd (List.split nal) in
let n' = List.fold_right (name_fold g) nal n in
f n (fold_local_binders g f n' acc b l) t
@@ -706,7 +714,7 @@ let fold_constr_expr_with_binders g f n acc = function
| CAppExpl (loc,(_,_),l) -> List.fold_left (f n) acc l
| CApp (loc,(_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l)
| CProdN (_,l,b) | CLambdaN (_,l,b) -> fold_constr_expr_binders g f n acc b l
- | CLetIn (_,na,a,b) -> fold_constr_expr_binders g f n acc b [[na],a]
+ | CLetIn (_,na,a,b) -> fold_constr_expr_binders g f n acc b [[na],default_binder_kind,a]
| CCast (loc,a,CastConv(_,b)) -> f n (f n acc a) b
| CCast (loc,a,CastCoerce) -> f n acc a
| CNotation (_,_,l) -> List.fold_left (f n) acc l
@@ -746,40 +754,40 @@ let mkIdentC id = CRef (Ident (dummy_loc, id))
let mkRefC r = CRef r
let mkAppC (f,l) = CApp (dummy_loc, (None,f), List.map (fun x -> (x,None)) l)
let mkCastC (a,k) = CCast (dummy_loc,a,k)
-let mkLambdaC (idl,a,b) = CLambdaN (dummy_loc,[idl,a],b)
+let mkLambdaC (idl,bk,a,b) = CLambdaN (dummy_loc,[idl,bk,a],b)
let mkLetInC (id,a,b) = CLetIn (dummy_loc,id,a,b)
-let mkProdC (idl,a,b) = CProdN (dummy_loc,[idl,a],b)
+let mkProdC (idl,bk,a,b) = CProdN (dummy_loc,[idl,bk,a],b)
let rec mkCProdN loc bll c =
match bll with
- | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll ->
- CProdN (loc,[idl,t],mkCProdN (join_loc loc1 loc) bll c)
+ | LocalRawAssum ((loc1,_)::_ as idl,bk,t) :: bll ->
+ CProdN (loc,[idl,bk,t],mkCProdN (join_loc loc1 loc) bll c)
| LocalRawDef ((loc1,_) as id,b) :: bll ->
CLetIn (loc,id,b,mkCProdN (join_loc loc1 loc) bll c)
| [] -> c
- | LocalRawAssum ([],_) :: bll -> mkCProdN loc bll c
+ | LocalRawAssum ([],_,_) :: bll -> mkCProdN loc bll c
let rec mkCLambdaN loc bll c =
match bll with
- | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll ->
- CLambdaN (loc,[idl,t],mkCLambdaN (join_loc loc1 loc) bll c)
+ | LocalRawAssum ((loc1,_)::_ as idl,bk,t) :: bll ->
+ CLambdaN (loc,[idl,bk,t],mkCLambdaN (join_loc loc1 loc) bll c)
| LocalRawDef ((loc1,_) as id,b) :: bll ->
CLetIn (loc,id,b,mkCLambdaN (join_loc loc1 loc) bll c)
| [] -> c
- | LocalRawAssum ([],_) :: bll -> mkCLambdaN loc bll c
+ | LocalRawAssum ([],_,_) :: bll -> mkCLambdaN loc bll c
let rec abstract_constr_expr c = function
| [] -> c
| LocalRawDef (x,b)::bl -> mkLetInC(x,b,abstract_constr_expr c bl)
- | LocalRawAssum (idl,t)::bl ->
- List.fold_right (fun x b -> mkLambdaC([x],t,b)) idl
+ | LocalRawAssum (idl,bk,t)::bl ->
+ List.fold_right (fun x b -> mkLambdaC([x],bk,t,b)) idl
(abstract_constr_expr c bl)
let rec prod_constr_expr c = function
| [] -> c
| LocalRawDef (x,b)::bl -> mkLetInC(x,b,prod_constr_expr c bl)
- | LocalRawAssum (idl,t)::bl ->
- List.fold_right (fun x b -> mkProdC([x],t,b)) idl
+ | LocalRawAssum (idl,bk,t)::bl ->
+ List.fold_right (fun x b -> mkProdC([x],bk,t,b)) idl
(prod_constr_expr c bl)
let coerce_to_id = function
@@ -794,15 +802,15 @@ let map_binder g e nal = List.fold_right (fun (_,na) -> name_fold g na) nal e
let map_binders f g e bl =
(* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
- let h (e,bl) (nal,t) = (map_binder g e nal,(nal,f e t)::bl) in
+ let h (e,bl) (nal,bk,t) = (map_binder g e nal,(nal,bk,f e t)::bl) in
let (e,rbl) = List.fold_left h (e,[]) bl in
(e, List.rev rbl)
let map_local_binders f g e bl =
(* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
let h (e,bl) = function
- LocalRawAssum(nal,ty) ->
- (map_binder g e nal, LocalRawAssum(nal,f e ty)::bl)
+ LocalRawAssum(nal,k,ty) ->
+ (map_binder g e nal, LocalRawAssum(nal,k,f e ty)::bl)
| LocalRawDef((loc,na),ty) ->
(name_fold g na e, LocalRawDef((loc,na),f e ty)::bl) in
let (e,rbl) = List.fold_left h (e,[]) bl in
diff --git a/interp/topconstr.mli b/interp/topconstr.mli
index 3d928bbb4..608cedb3c 100644
--- a/interp/topconstr.mli
+++ b/interp/topconstr.mli
@@ -90,7 +90,9 @@ val match_aconstr : rawconstr -> interpretation ->
type notation = string
-type explicitation = ExplByPos of int | ExplByName of identifier
+type explicitation = ExplByPos of int * identifier option | ExplByName of identifier
+
+type binder_kind = Default of binding_kind | TypeClass of binding_kind
type proj_flag = int option (* [Some n] = proj of the n-th visible argument *)
@@ -110,8 +112,8 @@ type constr_expr =
| CFix of loc * identifier located * fixpoint_expr list
| CCoFix of loc * identifier located * cofixpoint_expr list
| CArrow of loc * constr_expr * constr_expr
- | CProdN of loc * (name located list * constr_expr) list * constr_expr
- | CLambdaN of loc * (name located list * constr_expr) list * constr_expr
+ | CProdN of loc * (name located list * binder_kind * constr_expr) list * constr_expr
+ | CLambdaN of loc * (name located list * binder_kind * constr_expr) list * constr_expr
| CLetIn of loc * name located * constr_expr * constr_expr
| CAppExpl of loc * (proj_flag * reference) * constr_expr list
| CApp of loc * (proj_flag * constr_expr) *
@@ -146,7 +148,11 @@ and recursion_order_expr =
and local_binder =
| LocalRawDef of name located * constr_expr
- | LocalRawAssum of name located list * constr_expr
+ | LocalRawAssum of name located list * binder_kind * constr_expr
+
+type typeclass_constraint = name located * binding_kind * constr_expr
+
+and typeclass_context = typeclass_constraint list
(**********************************************************************)
(* Utilities on constr_expr *)
@@ -161,6 +167,8 @@ val replace_vars_constr_expr :
val free_vars_of_constr_expr : constr_expr -> Idset.t
val occur_var_constr_expr : identifier -> constr_expr -> bool
+val default_binder_kind : binder_kind
+
(* Specific function for interning "in indtype" syntax of "match" *)
val ids_of_cases_indtype : constr_expr -> identifier list
@@ -168,9 +176,9 @@ val mkIdentC : identifier -> constr_expr
val mkRefC : reference -> constr_expr
val mkAppC : constr_expr * constr_expr list -> constr_expr
val mkCastC : constr_expr * constr_expr cast_type -> constr_expr
-val mkLambdaC : name located list * constr_expr * constr_expr -> constr_expr
+val mkLambdaC : name located list * binder_kind * constr_expr * constr_expr -> constr_expr
val mkLetInC : name located * constr_expr * constr_expr -> constr_expr
-val mkProdC : name located list * constr_expr * constr_expr -> constr_expr
+val mkProdC : name located list * binder_kind * constr_expr * constr_expr -> constr_expr
val coerce_to_id : constr_expr -> identifier located
@@ -195,6 +203,11 @@ val names_of_local_assums : local_binder list -> name located list
(* With let binders *)
val names_of_local_binders : local_binder list -> name located list
+(* Used in typeclasses *)
+
+val fold_constr_expr_with_binders : (identifier -> 'a -> 'a) ->
+ ('a -> 'b -> constr_expr -> 'b) -> 'a -> 'b -> constr_expr -> 'b
+
(* Used in correctness and interface; absence of var capture not guaranteed *)
(* in pattern-matching clauses and in binders of the form [x,y:T(x)] *)
diff --git a/kernel/term.mli b/kernel/term.mli
index b5304b651..f71207c5e 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -473,10 +473,10 @@ val replace_vars : (identifier * constr) list -> constr -> constr
val subst_var : identifier -> constr -> constr
(* [subst_vars [id1;...;idn] t] substitute [VAR idj] by [Rel j] in [t]
- if two names are identical, the one of least indice is keeped *)
+ if two names are identical, the one of least indice is kept *)
val subst_vars : identifier list -> constr -> constr
(* [substn_vars n [id1;...;idn] t] substitute [VAR idj] by [Rel j+n-1] in [t]
- if two names are identical, the one of least indice is keeped *)
+ if two names are identical, the one of least indice is kept *)
val substn_vars : int -> identifier list -> constr -> constr
diff --git a/lib/util.ml b/lib/util.ml
index 586365dd8..6dd0a792b 100644
--- a/lib/util.ml
+++ b/lib/util.ml
@@ -207,6 +207,14 @@ let list_map3 f l1 l2 l3 =
in
map (l1,l2,l3)
+let list_map4 f l1 l2 l3 l4 =
+ let rec map = function
+ | ([], [], [], []) -> []
+ | ((h1::t1), (h2::t2), (h3::t3), (h4::t4)) -> let v = f h1 h2 h3 h4 in v::map (t1,t2,t3,t4)
+ | (_, _, _, _) -> invalid_arg "map4"
+ in
+ map (l1,l2,l3,l4)
+
let list_index x =
let rec index_x n = function
| y::l -> if x = y then n else index_x (succ n) l
diff --git a/lib/util.mli b/lib/util.mli
index 88b246a72..d52db650a 100644
--- a/lib/util.mli
+++ b/lib/util.mli
@@ -101,6 +101,8 @@ val list_map2_i :
(int -> 'a -> 'b -> 'c) -> int -> 'a list -> 'b list -> 'c list
val list_map3 :
('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
+val list_map4 :
+ ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
(* [list_index] returns the 1st index of an element in a list (counting from 1) *)
val list_index : 'a -> 'a list -> int
(* [list_unique_index x l] returns [Not_found] if [x] doesn't occur exactly once *)
diff --git a/library/decl_kinds.ml b/library/decl_kinds.ml
index c33738644..8ea183509 100644
--- a/library/decl_kinds.ml
+++ b/library/decl_kinds.ml
@@ -38,6 +38,7 @@ type definition_object_kind =
| Scheme
| StructureComponent
| IdentityCoercion
+ | Instance
type strength = locality_flag (* For compatibility *)
@@ -97,5 +98,5 @@ let string_of_definition_kind (l,boxed,d) =
| Global, Example -> "Example"
| Local, (CanonicalStructure|Example) ->
anomaly "Unsupported local definition kind"
- | _, (StructureComponent|Scheme|CoFixpoint|Fixpoint|IdentityCoercion)
+ | _, (StructureComponent|Scheme|CoFixpoint|Fixpoint|IdentityCoercion|Instance)
-> anomaly "Internal definition kind"
diff --git a/library/impargs.ml b/library/impargs.ml
index 479936c88..4272f413a 100644
--- a/library/impargs.ml
+++ b/library/impargs.ml
@@ -117,7 +117,7 @@ type implicit_explanation =
| DepRigid of argument_position
| DepFlex of argument_position
| DepFlexAndRigid of (*flex*) argument_position * (*rig*) argument_position
- | Manual
+ | Manual of bool
let argument_less = function
| Hyp n, Hyp n' -> n<n'
@@ -137,7 +137,7 @@ let update pos rig (na,st) =
| Some (DepFlex fpos) ->
if argument_less (pos,fpos) or pos=fpos then DepRigid pos
else DepFlexAndRigid (fpos,pos)
- | Some Manual -> assert false
+ | Some (Manual _) -> assert false
else
match st with
| None -> DepFlex pos
@@ -147,7 +147,7 @@ let update pos rig (na,st) =
if argument_less (pos,fpos) then DepFlexAndRigid (pos,rpos) else x
| Some (DepFlex fpos as x) ->
if argument_less (pos,fpos) then DepFlex pos else x
- | Some Manual -> assert false
+ | Some (Manual _) -> assert false
in na, Some e
(* modified is_rigid_reference with a truncated env *)
@@ -197,12 +197,20 @@ let add_free_rels_until strict strongly_strict revpat bound env m pos acc =
(* calcule la liste des arguments implicites *)
-let compute_implicits_gen strict strongly_strict revpat contextual env t =
+let concrete_name avoid_flags l env_names n all c =
+ if n = Anonymous & noccurn 1 c then
+ (Anonymous,l)
+ else
+ let fresh_id = next_name_not_occuring avoid_flags n l env_names c in
+ let idopt = if not all && noccurn 1 c then Anonymous else Name fresh_id in
+ (idopt, fresh_id::l)
+
+let compute_implicits_gen strict strongly_strict revpat contextual all env t =
let rec aux env avoid n names t =
let t = whd_betadeltaiota env t in
match kind_of_term t with
| Prod (na,a,b) ->
- let na',avoid' = Termops.concrete_name None avoid names na b in
+ let na',avoid' = concrete_name None avoid names na all b in
add_free_rels_until strict strongly_strict revpat n env a (Hyp (n+1))
(aux (push_rel (na',None,a) env) avoid' (n+1) (na'::names) b)
| _ ->
@@ -214,7 +222,7 @@ let compute_implicits_gen strict strongly_strict revpat contextual env t =
in
match kind_of_term (whd_betadeltaiota env t) with
| Prod (na,a,b) ->
- let na',avoid = Termops.concrete_name None [] [] na b in
+ let na',avoid = concrete_name None [] [] na all b in
let v = aux (push_rel (na',None,a) env) avoid 1 [na'] b in
Array.to_list v
| _ -> []
@@ -227,54 +235,79 @@ let rec prepare_implicits f = function
Some (id,imp,set_maximality imps' f.maximal) :: imps'
| _::imps -> None :: prepare_implicits f imps
+let compute_implicits_flags env f all t =
+ compute_implicits_gen
+ f.strict f.strongly_strict f.reversible_pattern f.contextual all env t
+
let compute_implicits_auto env f t =
- let l =
- compute_implicits_gen
- f.strict f.strongly_strict f.reversible_pattern f.contextual env t in
- prepare_implicits f l
+ let l = compute_implicits_flags env f false t in
+ prepare_implicits f l
let compute_implicits env t = compute_implicits_auto env !implicit_args t
let set_implicit id imp insmax =
- (id,(match imp with None -> Manual | Some imp -> imp),insmax)
+ (id,(match imp with None -> Manual false | Some imp -> imp),insmax)
+
+let merge_imps f = function
+ None -> Some (Manual f)
+ | x -> x
-let compute_manual_implicits flags ref l =
+let rec assoc_by_pos k = function
+ (ExplByPos (k', x), b) :: tl when k = k' -> (x,b), tl
+ | hd :: tl -> let (x, tl) = assoc_by_pos k tl in x, hd :: tl
+ | [] -> raise Not_found
+
+let compute_manual_implicits flags ref enriching l =
let t = Global.type_of_global ref in
+ let env = Global.env () in
let autoimps =
- compute_implicits_gen false false false true (Global.env()) t in
+ if enriching then compute_implicits_flags env flags true t
+ else compute_implicits_gen false false false true true env t in
let n = List.length autoimps in
+ let try_forced k l =
+ try
+ let (id, (b, f)), l' = assoc_by_pos k l in
+ if f then
+ let id = match id with Some id -> id | None -> id_of_string ("arg_" ^ string_of_int k) in
+ l', Some (id,Manual f,b)
+ else l, None
+ with Not_found -> l, None
+ in
if not (list_distinct l) then
error ("Some parameters are referred more than once");
(* Compare with automatic implicits to recover printing data and names *)
let rec merge k l = function
| (Name id,imp)::imps ->
- let l',m =
+ let l',imp,m =
try
- let b = List.assoc (ExplByName id) l in
- List.remove_assoc (ExplByName id) l, Some b
+ let (b, f) = List.assoc (ExplByName id) l in
+ List.remove_assoc (ExplByName id) l, merge_imps f imp,Some b
with Not_found ->
try
- let b = List.assoc (ExplByPos k) l in
- List.remove_assoc (ExplByPos k) l, Some b
+ let (id, (b, f)), l' = assoc_by_pos k l in
+ l', merge_imps f imp,Some b
with Not_found ->
- l, None in
+ l,imp, if enriching && imp <> None then Some flags.maximal else None
+ in
let imps' = merge (k+1) l' imps in
let m = Option.map (set_maximality imps') m in
Option.map (set_implicit id imp) m :: imps'
- | (Anonymous,_imp)::imps ->
- None :: merge (k+1) l imps
+ | (Anonymous,imp)::imps ->
+ let l', forced = try_forced k l in
+ forced :: merge (k+1) l' imps
| [] when l = [] -> []
- | _ ->
+ | [] ->
match List.hd l with
- | ExplByName id,_ ->
- error ("Wrong or not dependent implicit argument name: "^(string_of_id id))
- | ExplByPos i,_ ->
- if i<1 or i>n then
- error ("Bad implicit argument number: "^(string_of_int i))
- else
- errorlabstrm ""
- (str "Cannot set implicit argument number " ++ int i ++
- str ": it has no name") in
+ | ExplByName id,b ->
+ error ("Wrong or not dependent implicit argument name: "^(string_of_id id))
+ | ExplByPos (i,_id),_t ->
+ if i<1 or i>n then
+ error ("Bad implicit argument number: "^(string_of_int i))
+ else
+ errorlabstrm ""
+ (str "Cannot set implicit argument number " ++ int i ++
+ str ": it has no name")
+ in
merge 1 l autoimps
type maximal_insertion = bool (* true = maximal contextual insertion *)
@@ -297,7 +330,11 @@ let maximal_insertion_of = function
| Some (_,_,b) -> b
| None -> anomaly "Not an implicit argument"
-(* [in_ctx] means we now the expected type, [n] is the index of the argument *)
+let forced_implicit = function
+ | Some (_,Manual b,_) -> b
+ | _ -> false
+
+(* [in_ctx] means we know the expected type, [n] is the index of the argument *)
let is_inferable_implicit in_ctx n = function
| None -> false
| Some (_,DepRigid (Hyp p),_) -> in_ctx or n >= p
@@ -306,7 +343,7 @@ let is_inferable_implicit in_ctx n = function
| Some (_,DepRigid Conclusion,_) -> in_ctx
| Some (_,DepFlex Conclusion,_) -> false
| Some (_,DepFlexAndRigid (_,Conclusion),_) -> in_ctx
- | Some (_,Manual,_) -> true
+ | Some (_,Manual _,_) -> true
let positions_of_implicits =
let rec aux n = function
@@ -368,11 +405,24 @@ let compute_global_implicits flags = function
| ConstructRef ((kn,i),j) ->
let (_,cimps) = (compute_mib_implicits flags kn).(i) in snd cimps.(j-1)
+(* Merge a manual explicitation with an implicit_status list *)
+
+let list_split_at index l =
+ let rec aux i acc = function
+ tl when i = index -> (List.rev acc), tl
+ | hd :: tl -> aux (succ i) (hd :: acc) tl
+ | [] -> failwith "list_split_at: Invalid argument"
+ in aux 0 [] l
+
+let merge_impls oimpls impls =
+ let oimpls, _ = list_split_at (List.length oimpls - List.length impls) oimpls in
+ oimpls @ impls
+
(* Caching implicits *)
type implicit_interactive_request =
| ImplAuto
- | ImplManual of (explicitation * bool) list
+ | ImplManual of implicit_status list
type implicit_discharge_request =
| ImplNoDischarge
@@ -420,9 +470,10 @@ let rebuild_implicits (req,l) =
| ImplInteractive (ref,flags,o) ->
match o with
| ImplAuto -> [ref,compute_global_implicits flags ref]
- | ImplManual l ->
- error "Discharge of global manually given implicit arguments not implemented" in
- (req,l')
+ | ImplManual l ->
+ let auto = compute_global_implicits flags ref in
+ let l' = merge_impls auto l in [ref,l']
+ in (req,l')
let (inImplicits, _) =
@@ -464,15 +515,16 @@ let declare_mib_implicits kn =
(inImplicits (ImplMutualInductive (kn,flags),List.flatten imps))
(* Declare manual implicits *)
+type manual_explicitation = Topconstr.explicitation * (bool * bool)
-let declare_manual_implicits local ref l =
+let declare_manual_implicits local ref enriching l =
let flags = !implicit_args in
- let l' = compute_manual_implicits flags ref l in
+ let l' = compute_manual_implicits flags ref enriching l in
let req =
if local or isVarRef ref then ImplNoDischarge
- else ImplInteractive(ref,flags,ImplManual l)
+ else ImplInteractive(ref,flags,ImplManual l')
in
- add_anonymous_leaf (inImplicits (req,[ref,l']))
+ add_anonymous_leaf (inImplicits (req,[ref,l']))
(*s Registration as global tables *)
diff --git a/library/impargs.mli b/library/impargs.mli
index 8e6e25194..22c1ea23c 100644
--- a/library/impargs.mli
+++ b/library/impargs.mli
@@ -41,10 +41,13 @@ val with_implicits : implicits_flags -> ('a -> 'b) -> 'a -> 'b
type implicit_status
type implicits_list = implicit_status list
+type implicit_explanation
+
val is_status_implicit : implicit_status -> bool
val is_inferable_implicit : bool -> int -> implicit_status -> bool
val name_of_implicit : implicit_status -> identifier
val maximal_insertion_of : implicit_status -> bool
+val forced_implicit : implicit_status -> bool
val positions_of_implicits : implicits_list -> int list
@@ -60,8 +63,13 @@ val declare_mib_implicits : mutual_inductive -> unit
val declare_implicits : bool -> global_reference -> unit
+(* A [manual_explicitation] is a tuple of a positional or named explicitation with
+ maximal insertion and forcing flags. *)
+type manual_explicitation = Topconstr.explicitation * (bool * bool)
+
(* Manual declaration of which arguments are expected implicit *)
-val declare_manual_implicits : bool -> global_reference ->
- (Topconstr.explicitation * bool) list -> unit
+val declare_manual_implicits : bool -> global_reference -> bool ->
+ manual_explicitation list -> unit
val implicits_of_global : global_reference -> implicits_list
+
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index 4627d2305..c649b5846 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -37,7 +37,7 @@ let mk_lam = function
| (bl,c) -> CLambdaN(constr_loc c, bl,c)
let loc_of_binder_let = function
- | LocalRawAssum ((loc,_)::_,_)::_ -> loc
+ | LocalRawAssum ((loc,_)::_,_,_)::_ -> loc
| LocalRawDef ((loc,_),_)::_ -> loc
| _ -> dummy_loc
@@ -99,10 +99,10 @@ let lpar_id_coloneq =
| _ -> raise Stream.Failure)
| _ -> raise Stream.Failure)
-
GEXTEND Gram
GLOBAL: binder_constr lconstr constr operconstr sort global
- constr_pattern lconstr_pattern Constr.ident binder binder_let pattern;
+ constr_pattern lconstr_pattern Constr.ident
+ binder binder_let binders_let typeclass_constraint typeclass_param pattern;
Constr.ident:
[ [ id = Prim.ident -> id
@@ -307,28 +307,70 @@ GEXTEND Gram
;
binder_list:
[ [ idl=LIST1 name; bl=LIST0 binder_let ->
- LocalRawAssum (idl,CHole loc)::bl
+ LocalRawAssum (idl,Default Explicit,CHole loc)::bl
| idl=LIST1 name; ":"; c=lconstr ->
- [LocalRawAssum (idl,c)]
- | "("; idl=LIST1 name; ":"; c=lconstr; ")"; bl=LIST0 binder_let ->
- LocalRawAssum (idl,c)::bl ] ]
+ [LocalRawAssum (idl,Default Explicit,c)]
+ | cl = binders_let -> cl
+ ] ]
+ ;
+ binders_let:
+ [ [ "["; ctx = LIST1 typeclass_constraint_binder SEP ","; "]"; bl=binders_let ->
+ ctx @ bl
+ | cl = LIST0 binder_let -> cl
+ ] ]
;
binder_let:
[ [ id=name ->
- LocalRawAssum ([id],CHole loc)
+ LocalRawAssum ([id],Default Explicit,CHole loc)
| "("; id=name; idl=LIST1 name; ":"; c=lconstr; ")" ->
- LocalRawAssum (id::idl,c)
+ LocalRawAssum (id::idl,Default Explicit,c)
| "("; id=name; ":"; c=lconstr; ")" ->
- LocalRawAssum ([id],c)
+ LocalRawAssum ([id],Default Explicit,c)
+ | "`"; id=name; "`" ->
+ LocalRawAssum ([id],Default Implicit,CHole loc)
+ | "`"; id=name; idl=LIST1 name; ":"; c=lconstr; "`" ->
+ LocalRawAssum (id::idl,Default Implicit,c)
+ | "`"; id=name; ":"; c=lconstr; "`" ->
+ LocalRawAssum ([id],Default Implicit,c)
+ | "`"; id=name; idl=LIST1 name; "`" ->
+ LocalRawAssum (id::idl,Default Implicit,CHole loc)
| "("; id=name; ":="; c=lconstr; ")" ->
LocalRawDef (id,c)
| "("; id=name; ":"; t=lconstr; ":="; c=lconstr; ")" ->
LocalRawDef (id,CCast (join_loc (constr_loc t) loc,c, CastConv (DEFAULTcast,t)))
+ | "["; tc = typeclass_constraint_binder; "]" -> tc
] ]
;
binder:
- [ [ id=name -> ([id],CHole loc)
- | "("; idl=LIST1 name; ":"; c=lconstr; ")" -> (idl,c) ] ]
+ [ [ id=name -> ([id],Default Explicit,CHole loc)
+ | "("; idl=LIST1 name; ":"; c=lconstr; ")" -> (idl,Default Explicit,c)
+ | "`"; idl=LIST1 name; ":"; c=lconstr; "`" -> (idl,Default Implicit,c)
+ ] ]
+ ;
+ typeclass_constraint_binder:
+ [ [ tc = typeclass_constraint ->
+ let (n,bk,t) = tc in
+ LocalRawAssum ([n], TypeClass bk, t)
+ ] ]
+ ;
+ typeclass_constraint:
+ [ [ id=identref ; cl = LIST1 typeclass_param ->
+ (loc, Anonymous), Explicit, mkAppC (mkIdentC (snd id), cl)
+ | "?" ; id=identref ; cl = LIST1 typeclass_param ->
+ (loc, Anonymous), Implicit, mkAppC (mkIdentC (snd id), cl)
+ | iid=identref ; ":" ; id=typeclass_name ; cl = LIST1 typeclass_param ->
+ (fst iid, Name (snd iid)), (fst id), mkAppC (mkIdentC (snd (snd id)), cl)
+ ] ]
+ ;
+ typeclass_name:
+ [ [ id=identref -> (Explicit, id)
+ | "?"; id = identref -> (Implicit, id)
+ ] ]
+ ;
+ typeclass_param:
+ [ [ id = identref -> CRef (Libnames.Ident id)
+ | c = sort -> CSort (loc, c)
+ | "("; c = lconstr; ")" -> c ] ]
;
type_cstr:
[ [ c=OPT [":"; c=lconstr -> c] -> (loc,c) ] ]
diff --git a/parsing/g_ltac.ml4 b/parsing/g_ltac.ml4
index 3c5c88e89..79e88f8cd 100644
--- a/parsing/g_ltac.ml4
+++ b/parsing/g_ltac.ml4
@@ -215,12 +215,17 @@ GEXTEND Gram
| n = integer -> MsgInt n ] ]
;
+ ltac_def_kind:
+ [ [ ":=" -> false
+ | "::=" -> true ] ]
+ ;
+
(* Definitions for tactics *)
tacdef_body:
- [ [ name = identref; it=LIST1 input_fun; ":="; body = tactic_expr ->
- (name, TacFun (it, body))
- | name = identref; ":="; body = tactic_expr ->
- (name, body) ] ]
+ [ [ name = identref; it=LIST1 input_fun; redef = ltac_def_kind; body = tactic_expr ->
+ (name, redef, TacFun (it, body))
+ | name = identref; redef = ltac_def_kind; body = tactic_expr ->
+ (name, redef, body) ] ]
;
tactic:
[ [ tac = tactic_expr -> tac ] ]
diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4
index 23e4ba621..f305ebd69 100644
--- a/parsing/g_tactic.ml4
+++ b/parsing/g_tactic.ml4
@@ -95,9 +95,9 @@ open Tactic
let mk_fix_tac (loc,id,bl,ann,ty) =
let n =
match bl,ann with
- [([_],_)], None -> 1
+ [([_],_,_)], None -> 1
| _, Some x ->
- let ids = List.map snd (List.flatten (List.map fst bl)) in
+ let ids = List.map snd (List.flatten (List.map pi1 bl)) in
(try list_index (snd x) ids
with Not_found -> error "no such fix variable")
| _ -> error "cannot guess decreasing argument of fix" in
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index 729693aa7..ee2036167 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -46,6 +46,7 @@ let noedit_mode = Gram.Entry.create "vernac:noedit_command"
let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr"
let thm_token = Gram.Entry.create "vernac:thm_token"
let def_body = Gram.Entry.create "vernac:def_body"
+let typeclass_context = Gram.Entry.create "vernac:typeclass_context"
let get_command_entry () =
match Decl_mode.get_current_mode () with
@@ -116,11 +117,11 @@ let no_coercion loc (c,x) =
(* Gallina declarations *)
GEXTEND Gram
- GLOBAL: gallina gallina_ext thm_token def_body;
+ GLOBAL: gallina gallina_ext thm_token def_body typeclass_context typeclass_constraint;
gallina:
(* Definition, Theorem, Variable, Axiom, ... *)
- [ [ thm = thm_token; id = identref; bl = LIST0 binder_let; ":";
+ [ [ thm = thm_token; id = identref; bl = binders_let; ":";
c = lconstr ->
VernacStartTheoremProof (thm, id, (bl, c), false, no_hook)
| stre = assumption_token; nl = inline; bl = assum_list ->
@@ -164,6 +165,10 @@ GEXTEND Gram
*)
] ]
;
+ typeclass_context:
+ [ [ "["; l=LIST1 typeclass_constraint SEP ","; "]" -> l
+ | -> [] ] ]
+ ;
thm_token:
[ [ "Theorem" -> Theorem
| IDENT "Lemma" -> Lemma
@@ -209,14 +214,14 @@ GEXTEND Gram
;
(* Simple definitions *)
def_body:
- [ [ bl = LIST0 binder_let; ":="; red = reduce; c = lconstr ->
+ [ [ bl = binders_let; ":="; red = reduce; c = lconstr ->
(match c with
CCast(_,c, Rawterm.CastConv (k,t)) -> DefineBody (bl, red, c, Some t)
| _ -> DefineBody (bl, red, c, None))
- | bl = LIST0 binder_let; ":"; t = lconstr; ":="; red = reduce; c = lconstr ->
- DefineBody (bl, red, c, Some t)
- | bl = LIST0 binder_let; ":"; t = lconstr ->
- ProveBody (bl, t) ] ]
+ | bl = binders_let; ":"; t = lconstr; ":="; red = reduce; c = lconstr ->
+ DefineBody (bl, red, c, Some t)
+ | bl = binders_let; ":"; t = lconstr ->
+ ProveBody (bl, t) ] ]
;
reduce:
[ [ IDENT "Eval"; r = Tactic.red_expr; "in" -> Some r
@@ -431,7 +436,7 @@ END
(* Extensions: implicits, coercions, etc. *)
GEXTEND Gram
- GLOBAL: gallina_ext;
+ GLOBAL: gallina_ext typeclass_param;
gallina_ext:
[ [ (* Transparent and Opaque *)
@@ -467,22 +472,68 @@ GEXTEND Gram
t = class_rawexpr ->
VernacCoercion (Global, qid, s, t)
+ (* Type classes *)
+ | IDENT "Class"; sup = OPT [ l = typeclass_context; "=>" -> l ];
+ qid = identref; pars = LIST0 typeclass_param_type;
+ s = [ ":"; c = sort -> loc, c | -> (loc,Rawterm.RType None) ];
+ props = typeclass_field_types ->
+ VernacClass (qid, pars, s, (match sup with None -> [] | Some l -> l), props)
+
+ | IDENT "Context"; c = typeclass_context ->
+ VernacContext c
+
+ | IDENT "Instance"; sup = OPT [ l = typeclass_context ; "=>" -> l ];
+ is = typeclass_constraint ; props = typeclass_field_defs ->
+ let sup = match sup with None -> [] | Some l -> l in
+ VernacInstance (sup, is, props)
+
+ | IDENT "Existing"; IDENT "Instance"; is = identref -> VernacDeclareInstance is
+
+ | IDENT "Instantiation"; IDENT "Tactic"; ":="; tac = Tactic.tactic ->
+ VernacSetInstantiationTactic tac
+
(* Implicit *)
- | IDENT "Implicit"; IDENT "Arguments";
+ | IDENT "Implicit"; IDENT "Arguments"; enrich = [ IDENT "Enriching" -> true | -> false ];
(local,qid,pos) =
[ local = [ IDENT "Global" -> false | IDENT "Local" -> true ];
qid = global -> (local,qid,None)
| qid = global;
l = OPT [ "["; l = LIST0 implicit_name; "]" ->
- List.map (fun (id,b) -> (ExplByName id,b)) l ] ->
+ List.map (fun (id,b,f) -> (ExplByName id,b,f)) l ] ->
(true,qid,l) ] ->
- VernacDeclareImplicits (local,qid,pos)
+ VernacDeclareImplicits (local,qid,enrich,pos)
| IDENT "Implicit"; ["Type" | IDENT "Types"];
idl = LIST1 identref; ":"; c = lconstr -> VernacReserve (idl,c) ] ]
;
+
+(* typeclass_ctx: *)
+(* [ [ sup = LIST1 operconstr SEP "->"; "=>" -> sup *)
+(* ] ] *)
+(* ; *)
implicit_name:
- [ [ id = ident -> (id,false) | "["; id = ident; "]" -> (id,true) ] ]
+ [ [ "!"; id = ident -> (id, false, true)
+ | id = ident -> (id,false,false)
+ | "["; "!"; id = ident; "]" -> (id,true,true)
+ | "["; id = ident; "]" -> (id,true, false) ] ]
+ ;
+ typeclass_param_type:
+ [ [ "(" ; id = identref; ":"; t = lconstr ; ")" -> id, t
+ | id = identref -> id, CHole loc ] ]
+ ;
+ typeclass_field_type:
+ [ [ id = identref; ":"; t = lconstr -> id, t ] ]
+ ;
+ typeclass_field_def:
+ [ [ id = identref; params = LIST0 identref; ":="; t = lconstr -> id, params, t ] ]
+ ;
+ typeclass_field_types:
+ [ [ ":="; l = LIST1 typeclass_field_type SEP ";" -> l
+ | -> [] ] ]
+ ;
+ typeclass_field_defs:
+ [ [ ":="; l = LIST1 typeclass_field_def SEP ";" -> l
+ | -> [] ] ]
;
END
@@ -614,6 +665,8 @@ GEXTEND Gram
| IDENT "ML"; IDENT "Modules" -> PrintMLModules
| IDENT "Graph" -> PrintGraph
| IDENT "Classes" -> PrintClasses
+ | IDENT "TypeClasses" -> PrintTypeClasses
+ | IDENT "Instances"; qid = global -> PrintInstances qid
| IDENT "Ltac"; qid = global -> PrintLtac qid
| IDENT "Coercions" -> PrintCoercions
| IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr
diff --git a/parsing/g_xml.ml4 b/parsing/g_xml.ml4
index cd929e5af..47725cbd6 100644
--- a/parsing/g_xml.ml4
+++ b/parsing/g_xml.ml4
@@ -141,12 +141,12 @@ let rec interp_xml_constr = function
| XmlTag (loc,"LAMBDA",al,(_::_ as xl)) ->
let body,decls = list_sep_last xl in
let ctx = List.map interp_xml_decl decls in
- List.fold_right (fun (na,t) b -> RLambda (loc, na, t, b))
+ List.fold_right (fun (na,t) b -> RLambda (loc, na, Explicit, t, b))
ctx (interp_xml_target body)
| XmlTag (loc,"PROD",al,(_::_ as xl)) ->
let body,decls = list_sep_last xl in
let ctx = List.map interp_xml_decl decls in
- List.fold_right (fun (na,t) b -> RProd (loc, na, t, b))
+ List.fold_right (fun (na,t) b -> RProd (loc, na, Explicit, t, b))
ctx (interp_xml_target body)
| XmlTag (loc,"LETIN",al,(_::_ as xl)) ->
let body,defs = list_sep_last xl in
diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4
index d2380dad6..56d547cb5 100644
--- a/parsing/pcoq.ml4
+++ b/parsing/pcoq.ml4
@@ -434,6 +434,9 @@ module Constr =
let lconstr_pattern = gec_constr "lconstr_pattern"
let binder = Gram.Entry.create "constr:binder"
let binder_let = Gram.Entry.create "constr:binder_let"
+ let binders_let = Gram.Entry.create "constr:binders_let"
+ let typeclass_constraint = Gram.Entry.create "constr:typeclass_constraint"
+ let typeclass_param = Gram.Entry.create "constr:typeclass_param"
end
module Module =
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index 651554486..eaf298e06 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -161,8 +161,11 @@ module Constr :
val annot : constr_expr Gram.Entry.e
val constr_pattern : constr_expr Gram.Entry.e
val lconstr_pattern : constr_expr Gram.Entry.e
- val binder : (name located list * constr_expr) Gram.Entry.e
+ val binder : (name located list * binder_kind * constr_expr) Gram.Entry.e
val binder_let : local_binder Gram.Entry.e
+ val binders_let : local_binder list Gram.Entry.e
+ val typeclass_constraint : (name located * binding_kind * constr_expr) Gram.Entry.e
+ val typeclass_param : constr_expr Gram.Entry.e
end
module Module :
diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml
index 858c6685f..85bf11806 100644
--- a/parsing/ppconstr.ml
+++ b/parsing/ppconstr.ml
@@ -129,7 +129,7 @@ let pr_qualid = pr_qualid
let pr_expl_args pr (a,expl) =
match expl with
| None -> pr (lapp,L) a
- | Some (_,ExplByPos n) ->
+ | Some (_,ExplByPos (n,_id)) ->
anomaly("Explicitation by position not implemented")
| Some (_,ExplByName id) ->
str "(" ++ pr_id id ++ str ":=" ++ pr ltop a ++ str ")"
@@ -206,30 +206,41 @@ let pr_eqn pr (loc,pl,rhs) =
let begin_of_binder = function
LocalRawDef((loc,_),_) -> fst (unloc loc)
- | LocalRawAssum((loc,_)::_,_) -> fst (unloc loc)
+ | LocalRawAssum((loc,_)::_,_,_) -> fst (unloc loc)
| _ -> assert false
let begin_of_binders = function
| b::_ -> begin_of_binder b
| _ -> 0
-let pr_binder many pr (nal,t) =
+let surround_binder k p =
+ match k with
+ Default Explicit -> hov 1 (str"(" ++ p ++ str")")
+ | Default Implicit -> hov 1 (str"`" ++ p ++ str"`")
+ | TypeClass b -> hov 1 (str"[" ++ p ++ str"]")
+
+let surround_implicit k p =
+ match k with
+ Default Explicit -> p
+ | Default Implicit -> (str"`" ++ p ++ str"`")
+ | TypeClass b -> (str"[" ++ p ++ str"]")
+
+let pr_binder many pr (nal,k,t) =
match t with
| CHole _ -> prlist_with_sep spc pr_lname nal
| _ ->
let s = prlist_with_sep spc pr_lname nal ++ str" : " ++ pr t in
- hov 1 (if many then surround s else s)
+ hov 1 (if many then surround_binder k s else surround_implicit k s)
let pr_binder_among_many pr_c = function
- | LocalRawAssum (nal,t) ->
- pr_binder true pr_c (nal,t)
+ | LocalRawAssum (nal,k,t) ->
+ pr_binder true pr_c (nal,k,t)
| LocalRawDef (na,c) ->
let c,topt = match c with
| CCast(_,c, CastConv (_,t)) -> c, t
| _ -> c, CHole dummy_loc in
- hov 1 (surround
- (pr_lname na ++ pr_opt_type pr_c topt ++
- str":=" ++ cut() ++ pr_c c))
+ hov 1 (pr_lname na ++ pr_opt_type pr_c topt ++
+ str":=" ++ cut() ++ pr_c c)
let pr_undelimited_binders pr_c =
prlist_with_sep spc (pr_binder_among_many pr_c)
@@ -237,8 +248,8 @@ let pr_undelimited_binders pr_c =
let pr_delimited_binders kw pr_c bl =
let n = begin_of_binders bl in
match bl with
- | [LocalRawAssum (nal,t)] ->
- pr_com_at n ++ kw() ++ pr_binder false pr_c (nal,t)
+ | [LocalRawAssum (nal,k,t)] ->
+ pr_com_at n ++ kw() ++ pr_binder false pr_c (nal,k,t)
| LocalRawAssum _ :: _ as bdl ->
pr_com_at n ++ kw() ++ pr_undelimited_binders pr_c bdl
| _ -> assert false
@@ -249,9 +260,9 @@ let rec extract_prod_binders = function
if bl = [] then [], x else LocalRawDef (na,b) :: bl, c*)
| CProdN (loc,[],c) ->
extract_prod_binders c
- | CProdN (loc,(nal,t)::bl,c) ->
+ | CProdN (loc,(nal,bk,t)::bl,c) ->
let bl,c = extract_prod_binders (CProdN(loc,bl,c)) in
- LocalRawAssum (nal,t) :: bl, c
+ LocalRawAssum (nal,bk,t) :: bl, c
| c -> [], c
let rec extract_lam_binders = function
@@ -260,15 +271,15 @@ let rec extract_lam_binders = function
if bl = [] then [], x else LocalRawDef (na,b) :: bl, c*)
| CLambdaN (loc,[],c) ->
extract_lam_binders c
- | CLambdaN (loc,(nal,t)::bl,c) ->
+ | CLambdaN (loc,(nal,bk,t)::bl,c) ->
let bl,c = extract_lam_binders (CLambdaN(loc,bl,c)) in
- LocalRawAssum (nal,t) :: bl, c
+ LocalRawAssum (nal,bk,t) :: bl, c
| c -> [], c
let split_lambda = function
- | CLambdaN (loc,[[na],t],c) -> (na,t,c)
- | CLambdaN (loc,([na],t)::bl,c) -> (na,t,CLambdaN(loc,bl,c))
- | CLambdaN (loc,(na::nal,t)::bl,c) -> (na,t,CLambdaN(loc,(nal,t)::bl,c))
+ | CLambdaN (loc,[[na],bk,t],c) -> (na,t,c)
+ | CLambdaN (loc,([na],bk,t)::bl,c) -> (na,t,CLambdaN(loc,bl,c))
+ | CLambdaN (loc,(na::nal,bk,t)::bl,c) -> (na,t,CLambdaN(loc,(nal,bk,t)::bl,c))
| _ -> anomaly "ill-formed fixpoint body"
let rename na na' t c =
@@ -279,13 +290,13 @@ let rename na na' t c =
let split_product na' = function
| CArrow (loc,t,c) -> (na',t,c)
- | CProdN (loc,[[na],t],c) -> rename na na' t c
- | CProdN (loc,([na],t)::bl,c) -> rename na na' t (CProdN(loc,bl,c))
- | CProdN (loc,(na::nal,t)::bl,c) ->
- rename na na' t (CProdN(loc,(nal,t)::bl,c))
+ | CProdN (loc,[[na],bk,t],c) -> rename na na' t c
+ | CProdN (loc,([na],bk,t)::bl,c) -> rename na na' t (CProdN(loc,bl,c))
+ | CProdN (loc,(na::nal,bk,t)::bl,c) ->
+ rename na na' t (CProdN(loc,(nal,bk,t)::bl,c))
| _ -> anomaly "ill-formed fixpoint body"
-let merge_binders (na1,ty1) cofun (na2,ty2) codom =
+let merge_binders (na1,bk1,ty1) cofun (na2,bk2,ty2) codom =
let na =
match snd na1, snd na2 with
Anonymous, Name id ->
@@ -306,42 +317,42 @@ let merge_binders (na1,ty1) cofun (na2,ty2) codom =
| _ ->
Constrextern.check_same_type ty1 ty2;
ty2 in
- (LocalRawAssum ([na],ty), codom)
+ (LocalRawAssum ([na],bk1,ty), codom)
let rec strip_domain bvar cofun c =
match c with
| CArrow(loc,a,b) ->
- merge_binders bvar cofun ((dummy_loc,Anonymous),a) b
- | CProdN(loc,[([na],ty)],c') ->
- merge_binders bvar cofun (na,ty) c'
- | CProdN(loc,([na],ty)::bl,c') ->
- merge_binders bvar cofun (na,ty) (CProdN(loc,bl,c'))
- | CProdN(loc,(na::nal,ty)::bl,c') ->
- merge_binders bvar cofun (na,ty) (CProdN(loc,(nal,ty)::bl,c'))
+ merge_binders bvar cofun ((dummy_loc,Anonymous),default_binder_kind,a) b
+ | CProdN(loc,[([na],bk,ty)],c') ->
+ merge_binders bvar cofun (na,bk,ty) c'
+ | CProdN(loc,([na],bk,ty)::bl,c') ->
+ merge_binders bvar cofun (na,bk,ty) (CProdN(loc,bl,c'))
+ | CProdN(loc,(na::nal,bk,ty)::bl,c') ->
+ merge_binders bvar cofun (na,bk,ty) (CProdN(loc,(nal,bk,ty)::bl,c'))
| _ -> failwith "not a product"
(* Note: binder sharing is lost *)
-let rec strip_domains (nal,ty) cofun c =
+let rec strip_domains (nal,bk,ty) cofun c =
match nal with
[] -> assert false
| [na] ->
- let bnd, c' = strip_domain (na,ty) cofun c in
+ let bnd, c' = strip_domain (na,bk,ty) cofun c in
([bnd],None,c')
| na::nal ->
- let f = CLambdaN(dummy_loc,[(nal,ty)],cofun) in
- let bnd, c1 = strip_domain (na,ty) f c in
+ let f = CLambdaN(dummy_loc,[(nal,bk,ty)],cofun) in
+ let bnd, c1 = strip_domain (na,bk,ty) f c in
(try
- let bl, rest, c2 = strip_domains (nal,ty) cofun c1 in
+ let bl, rest, c2 = strip_domains (nal,bk,ty) cofun c1 in
(bnd::bl, rest, c2)
- with Failure _ -> ([bnd],Some (nal,ty), c1))
+ with Failure _ -> ([bnd],Some (nal,bk,ty), c1))
(* Re-share binders *)
let rec factorize_binders = function
| ([] | [_] as l) -> l
- | LocalRawAssum (nal,ty) as d :: (LocalRawAssum (nal',ty')::l as l') ->
+ | LocalRawAssum (nal,k,ty) as d :: (LocalRawAssum (nal',k',ty')::l as l') ->
(try
let _ = Constrextern.check_same_type ty ty' in
- factorize_binders (LocalRawAssum (nal@nal',ty)::l)
+ factorize_binders (LocalRawAssum (nal@nal',k,ty)::l)
with _ ->
d :: factorize_binders l')
| d :: l -> d :: factorize_binders l
@@ -370,7 +381,7 @@ let rec split_fix n typ def =
let (na,_,def) = split_lambda def in
let (na,t,typ) = split_product na typ in
let (bl,typ,def) = split_fix (n-1) typ def in
- (LocalRawAssum ([na],t)::bl,typ,def)
+ (LocalRawAssum ([na],default_binder_kind,t)::bl,typ,def)
let pr_recursive_decl pr pr_dangling dangling_with_for id bl annot t c =
let pr_body =
@@ -601,27 +612,27 @@ let rec strip_context n iscast t =
if n = 0 then
[], if iscast then match t with CCast (_,c,_) -> c | _ -> t else t
else match t with
- | CLambdaN (loc,(nal,t)::bll,c) ->
+ | CLambdaN (loc,(nal,bk,t)::bll,c) ->
let n' = List.length nal in
if n' > n then
let nal1,nal2 = list_chop n nal in
- [LocalRawAssum (nal1,t)], CLambdaN (loc,(nal2,t)::bll,c)
+ [LocalRawAssum (nal1,bk,t)], CLambdaN (loc,(nal2,bk,t)::bll,c)
else
let bl', c = strip_context (n-n') iscast
(if bll=[] then c else CLambdaN (loc,bll,c)) in
- LocalRawAssum (nal,t) :: bl', c
- | CProdN (loc,(nal,t)::bll,c) ->
+ LocalRawAssum (nal,bk,t) :: bl', c
+ | CProdN (loc,(nal,bk,t)::bll,c) ->
let n' = List.length nal in
if n' > n then
let nal1,nal2 = list_chop n nal in
- [LocalRawAssum (nal1,t)], CProdN (loc,(nal2,t)::bll,c)
+ [LocalRawAssum (nal1,bk,t)], CProdN (loc,(nal2,bk,t)::bll,c)
else
let bl', c = strip_context (n-n') iscast
(if bll=[] then c else CProdN (loc,bll,c)) in
- LocalRawAssum (nal,t) :: bl', c
+ LocalRawAssum (nal,bk,t) :: bl', c
| CArrow (loc,t,c) ->
let bl', c = strip_context (n-1) iscast c in
- LocalRawAssum ([loc,Anonymous],t) :: bl', c
+ LocalRawAssum ([loc,Anonymous],default_binder_kind,t) :: bl', c
| CCast (_,c,_) -> strip_context n false c
| CLetIn (_,na,b,c) ->
let bl', c = strip_context (n-1) iscast c in
diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml
index 019063527..7d83cffb6 100644
--- a/parsing/pptactic.ml
+++ b/parsing/pptactic.ml
@@ -302,9 +302,10 @@ let strip_prod_binders_expr n ty =
match ty with
Topconstr.CProdN(_,bll,a) ->
let nb =
- List.fold_left (fun i (nal,_) -> i + List.length nal) 0 bll in
- if nb >= n then (List.rev (bll@acc), a)
- else strip_ty (bll@acc) (n-nb) a
+ List.fold_left (fun i (nal,_,_) -> i + List.length nal) 0 bll in
+ let bll = List.map (fun (x, _, y) -> x, y) bll in
+ if nb >= n then (List.rev (bll@acc)), a
+ else strip_ty (bll@acc) (n-nb) a
| Topconstr.CArrow(_,a,b) ->
if n=1 then
(List.rev (([(dummy_loc,Anonymous)],a)::acc), b)
@@ -978,7 +979,7 @@ let strip_prod_binders_rawterm n (ty,_) =
let rec strip_ty acc n ty =
if n=0 then (List.rev acc, (ty,None)) else
match ty with
- Rawterm.RProd(loc,na,a,b) ->
+ Rawterm.RProd(loc,na,Explicit,a,b) ->
strip_ty (([dummy_loc,na],(a,None))::acc) (n-1) b
| _ -> error "Cannot translate fix tactic: not enough products" in
strip_ty [] n ty
diff --git a/parsing/ppvernac.ml b/parsing/ppvernac.ml
index 1161c3b61..1470f6902 100644
--- a/parsing/ppvernac.ml
+++ b/parsing/ppvernac.ml
@@ -146,11 +146,12 @@ let pr_search a b pr_p = match a with
let pr_locality local = if local then str "Local " else str ""
let pr_non_globality local = if local then str "" else str "Global "
-let pr_explanation (e,b) =
+let pr_explanation (e,b,f) =
let a = match e with
- | ExplByPos n -> anomaly "No more supported"
+ | ExplByPos (n,_) -> anomaly "No more supported"
| ExplByName id -> pr_id id in
- if b then str "[" ++ a ++ str "]" else a
+ let a = if f then str"!" ++ a else a in
+ if b then str "[" ++ a ++ str "]" else a
let pr_class_rawexpr = function
| FunClass -> str"Funclass"
@@ -394,6 +395,18 @@ let make_pr_vernac pr_constr pr_lconstr =
let pr_constrarg c = spc () ++ pr_constr c in
let pr_lconstrarg c = spc () ++ pr_lconstr c in
let pr_intarg n = spc () ++ int n in
+let pr_lident_constr sep (i,c) = pr_lident i ++ sep ++ pr_constrarg c in
+let pr_lname_lident_constr (oi,bk,a) =
+ (match snd oi with Anonymous -> mt () | Name id -> pr_lident (fst oi, id) ++ spc () ++ str":" ++ spc ())
+ ++ pr_lconstr a in
+let pr_typeclass_context l =
+ match l with
+ [] -> mt ()
+ | _ -> str"[" ++ spc () ++ prlist_with_sep (fun () -> str"," ++ spc()) pr_lname_lident_constr l
+ ++ spc () ++ str"]" ++ spc ()
+in
+let pr_instance_def sep (i,l,c) = pr_lident i ++ prlist_with_sep spc pr_lident l
+ ++ sep ++ pr_constrarg c in
let rec pr_vernac = function
@@ -565,7 +578,7 @@ let rec pr_vernac = function
| VernacFixpoint (recs,b) ->
let name_of_binder = function
- | LocalRawAssum (nal,_) -> nal
+ | LocalRawAssum (nal,_,_) -> nal
| LocalRawDef (_,_) -> [] in
let pr_onerec = function
| (id,(n,ro),bl,type_,def),ntn ->
@@ -679,6 +692,39 @@ let rec pr_vernac = function
spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++ spc() ++ str">->" ++
spc() ++ pr_class_rawexpr c2)
+
+ | VernacClass (id, par, ar, sup, props) ->
+ hov 1 (
+ str"Class" ++ spc () ++ pr_lident id ++
+ prlist_with_sep (spc) (pr_lident_constr (spc() ++ str ":" ++ spc())) par ++
+ spc () ++ str":" ++ spc() ++ pr_rawsort (snd ar) ++
+ spc () ++ str"where" ++ spc () ++
+ prlist_with_sep (fun () -> str";" ++ spc()) (pr_lident_constr (spc () ++ str":" ++ spc())) props )
+
+
+ | VernacInstance (sup, (instid, bk, cl), props) ->
+ hov 1 (
+ str"Instance" ++ spc () ++
+ pr_typeclass_context sup ++
+ str"=>" ++ spc () ++
+ (match snd instid with Name id -> pr_lident (fst instid, id) ++ spc () ++ str":" ++ spc () | Anonymous -> mt ()) ++
+ pr_constr_expr cl ++ spc () ++
+ spc () ++ str"where" ++ spc () ++
+ prlist_with_sep (fun () -> str";" ++ spc()) (pr_instance_def (spc () ++ str":=" ++ spc())) props)
+
+ | VernacContext l ->
+ hov 1 (
+ str"Context" ++ spc () ++ str"[" ++ spc () ++
+ prlist_with_sep (fun () -> str"," ++ spc()) pr_lname_lident_constr l ++
+ spc () ++ str "]")
+
+
+ | VernacDeclareInstance id ->
+ hov 1 (str"Instance" ++ spc () ++ pr_lident id)
+
+ | VernacSetInstantiationTactic tac ->
+ hov 1 (str"Instantiation Tactic :=" ++ spc () ++ pr_raw_tactic tac)
+
(* Modules and Module Types *)
| VernacDefineModule (export,m,bl,ty,bd) ->
let b = pr_module_binders_list bl pr_lconstr in
@@ -738,7 +784,7 @@ let rec pr_vernac = function
(* Commands *)
| VernacDeclareTacticDefinition (rc,l) ->
- let pr_tac_body (id, body) =
+ let pr_tac_body (id, redef, body) =
let idl, body =
match body with
| Tacexpr.TacFun (idl,b) -> idl,b
@@ -746,10 +792,10 @@ let rec pr_vernac = function
pr_located pr_ltac_id id ++
prlist (function None -> str " _"
| Some id -> spc () ++ pr_id id) idl
- ++ str" :=" ++ brk(1,1) ++
+ ++ (if redef then str" ::=" else str" :=") ++ brk(1,1) ++
let idl = List.map Option.get (List.filter (fun x -> not (x=None)) idl)in
pr_raw_tactic_env
- (idl @ List.map snd (List.map fst l))
+ (idl @ List.map snd (List.map (fun (x, _, _) -> x) l))
(Global.env())
body in
hov 1
@@ -765,9 +811,9 @@ let rec pr_vernac = function
(str"Notation " ++ pr_locality local ++ pr_id id ++ str" :=" ++
pr_constrarg c ++
pr_syntax_modifiers (if onlyparsing then [SetOnlyParsing] else []))
- | VernacDeclareImplicits (local,q,None) ->
+ | VernacDeclareImplicits (local,q,e,None) ->
hov 2 (str"Implicit Arguments" ++ spc() ++ pr_reference q)
- | VernacDeclareImplicits (local,q,Some imps) ->
+ | VernacDeclareImplicits (local,q,e,Some imps) ->
hov 1 (str"Implicit Arguments" ++ pr_non_globality local ++
spc() ++ pr_reference q ++ spc() ++
str"[" ++ prlist_with_sep sep pr_explanation imps ++ str"]")
@@ -810,6 +856,8 @@ let rec pr_vernac = function
| PrintMLModules -> str"Print ML Modules"
| PrintGraph -> str"Print Graph"
| PrintClasses -> str"Print Classes"
+ | PrintTypeClasses -> str"Print TypeClasses"
+ | PrintInstances qid -> str"Print Instances" ++ spc () ++ pr_reference qid
| PrintLtac qid -> str"Print Ltac" ++ spc() ++ pr_reference qid
| PrintCoercions -> str"Print Coercions"
| PrintCoercionPaths (s,t) -> str"Print Coercion Paths" ++ spc() ++ pr_class_rawexpr s ++ spc() ++ pr_class_rawexpr t
diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml
index e92db84fe..6fe4d80d4 100644
--- a/parsing/prettyp.ml
+++ b/parsing/prettyp.ml
@@ -746,3 +746,22 @@ let print_canonical_projections () =
(*************************************************************************)
+(*************************************************************************)
+(* Pretty-printing functions for type classes *)
+
+open Typeclasses
+
+let pr_typeclass env t =
+ gallina_print_inductive (fst t.cl_impl)
+
+let print_typeclasses () =
+ let env = Global.env () in
+ prlist_with_sep pr_spc (pr_typeclass env) (typeclasses ())
+
+let pr_instance env i =
+ gallina_print_constant_with_infos i.is_impl
+
+let print_instances r =
+ let env = Global.env () in
+ let inst = instances r in
+ prlist_with_sep pr_spc (pr_instance env) inst
diff --git a/parsing/prettyp.mli b/parsing/prettyp.mli
index df16e3526..35e23d923 100644
--- a/parsing/prettyp.mli
+++ b/parsing/prettyp.mli
@@ -58,6 +58,11 @@ val print_coercions : unit -> std_ppcmds
val print_path_between : Classops.cl_typ -> Classops.cl_typ -> std_ppcmds
val print_canonical_projections : unit -> std_ppcmds
+(* Pretty-printing functions for type classes and instances *)
+val print_typeclasses : unit -> std_ppcmds
+val print_instances : reference -> std_ppcmds
+
+
val inspect : int -> std_ppcmds
(* Locate *)
diff --git a/parsing/q_constr.ml4 b/parsing/q_constr.ml4
index d8ce0a570..497b8a3c6 100644
--- a/parsing/q_constr.ml4
+++ b/parsing/q_constr.ml4
@@ -49,9 +49,9 @@ EXTEND
constr:
[ "200" RIGHTA
[ LIDENT "forall"; id = ident; ":"; c1 = constr; ","; c2 = constr ->
- <:expr< Rawterm.RProd ($dloc$,Name $id$,$c1$,$c2$) >>
+ <:expr< Rawterm.RProd ($dloc$,Name $id$,Rawterm.Explicit,$c1$,$c2$) >>
| "fun"; id = ident; ":"; c1 = constr; "=>"; c2 = constr ->
- <:expr< Rawterm.RLambda ($dloc$,Name $id$,$c1$,$c2$) >>
+ <:expr< Rawterm.RLambda ($dloc$,Name $id$,Rawterm.Explicit,$c1$,$c2$) >>
| "let"; id = ident; ":="; c1 = constr; "in"; c2 = constr ->
<:expr< Rawterm.RLetin ($dloc$,Name $id$,$c1$,$c2$) >>
(* fix todo *)
@@ -61,7 +61,7 @@ EXTEND
<:expr< Rawterm.RCast($dloc$,$c1$,DEFAULTcast,$c2$) >> ]
| "90" RIGHTA
[ c1 = constr; "->"; c2 = SELF ->
- <:expr< Rawterm.RProd ($dloc$,Anonymous,$c1$,$c2$) >> ]
+ <:expr< Rawterm.RProd ($dloc$,Anonymous,Rawterm.Explicit,$c1$,$c2$) >> ]
| "75" RIGHTA
[ "~"; c = constr ->
apply_ref <:expr< coq_not_ref >> [c] ]
diff --git a/parsing/q_coqast.ml4 b/parsing/q_coqast.ml4
index fbd2e6e07..6fdfba23d 100644
--- a/parsing/q_coqast.ml4
+++ b/parsing/q_coqast.ml4
@@ -129,8 +129,16 @@ let mlexpr_of_red_flags {
let mlexpr_of_explicitation = function
| Topconstr.ExplByName id -> <:expr< Topconstr.ExplByName $mlexpr_of_ident id$ >>
- | Topconstr.ExplByPos n -> <:expr< Topconstr.ExplByPos $mlexpr_of_int n$ >>
+ | Topconstr.ExplByPos (n,_id) -> <:expr< Topconstr.ExplByPos $mlexpr_of_int n$ >>
+let mlexpr_of_binding_kind = function
+ | Rawterm.Implicit -> <:expr< Rawterm.Implicit >>
+ | Rawterm.Explicit -> <:expr< Rawterm.Explicit >>
+
+let mlexpr_of_binder_kind = function
+ | Topconstr.Default b -> <:expr< Topconstr.Default $mlexpr_of_binding_kind b$ >>
+ | Topconstr.TypeClass b -> <:expr< Topconstr.TypeClass $mlexpr_of_binding_kind b$ >>
+
let rec mlexpr_of_constr = function
| Topconstr.CRef (Libnames.Ident (loc,id)) when is_meta (string_of_id id) ->
anti loc (string_of_id id)
@@ -139,8 +147,9 @@ let rec mlexpr_of_constr = function
| Topconstr.CCoFix (loc,_,_) -> failwith "mlexpr_of_constr: TODO"
| Topconstr.CArrow (loc,a,b) ->
<:expr< Topconstr.CArrow $dloc$ $mlexpr_of_constr a$ $mlexpr_of_constr b$ >>
- | Topconstr.CProdN (loc,l,a) -> <:expr< Topconstr.CProdN $dloc$ $mlexpr_of_list (mlexpr_of_pair (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
- | Topconstr.CLambdaN (loc,l,a) -> <:expr< Topconstr.CLambdaN $dloc$ $mlexpr_of_list (mlexpr_of_pair (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
+ | Topconstr.CProdN (loc,l,a) -> <:expr< Topconstr.CProdN $dloc$ $mlexpr_of_list
+ (mlexpr_of_triple (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_binder_kind mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
+ | Topconstr.CLambdaN (loc,l,a) -> <:expr< Topconstr.CLambdaN $dloc$ $mlexpr_of_list (mlexpr_of_triple (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_binder_kind mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
| Topconstr.CLetIn (loc,_,_,_) -> failwith "mlexpr_of_constr: TODO"
| Topconstr.CAppExpl (loc,a,l) -> <:expr< Topconstr.CAppExpl $dloc$ $mlexpr_of_pair (mlexpr_of_option mlexpr_of_int) mlexpr_of_reference a$ $mlexpr_of_list mlexpr_of_constr l$ >>
| Topconstr.CApp (loc,a,l) -> <:expr< Topconstr.CApp $dloc$ $mlexpr_of_pair (mlexpr_of_option mlexpr_of_int) mlexpr_of_constr a$ $mlexpr_of_list (mlexpr_of_pair mlexpr_of_constr (mlexpr_of_option (mlexpr_of_located mlexpr_of_explicitation))) l$ >>
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 058f3d210..e71083c50 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -1408,9 +1408,9 @@ let set_arity_signature dep n arsign tomatchl pred x =
let rec decomp_lam_force n avoid l p =
if n = 0 then (List.rev l,p,avoid) else
match p with
- | RLambda (_,(Name id as na),_,c) ->
+ | RLambda (_,(Name id as na),_,_,c) ->
decomp_lam_force (n-1) (id::avoid) (na::l) c
- | RLambda (_,(Anonymous as na),_,c) -> decomp_lam_force (n-1) avoid (na::l) c
+ | RLambda (_,(Anonymous as na),_,_,c) -> decomp_lam_force (n-1) avoid (na::l) c
| _ ->
let x = next_ident_away (id_of_string "x") avoid in
decomp_lam_force (n-1) (x::avoid) (Name x :: l)
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 2cfa7076a..a9e550bf7 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -268,7 +268,7 @@ let is_nondep_branch c n =
let extract_nondep_branches test c b n =
let rec strip n r = if n=0 then r else
match r with
- | RLambda (_,_,_,t) -> strip (n-1) t
+ | RLambda (_,_,_,_,t) -> strip (n-1) t
| RLetIn (_,_,_,t) -> strip (n-1) t
| _ -> assert false in
if test c n then Some (strip n b) else None
@@ -276,7 +276,7 @@ let extract_nondep_branches test c b n =
let it_destRLambda_or_LetIn_names n c =
let rec aux n nal c =
if n=0 then (List.rev nal,c) else match c with
- | RLambda (_,na,_,c) -> aux (n-1) (na::nal) c
+ | RLambda (_,na,_,_,c) -> aux (n-1) (na::nal) c
| RLetIn (_,na,_,c) -> aux (n-1) (na::nal) c
| _ ->
(* eta-expansion *)
@@ -308,7 +308,7 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| Some p ->
let nl,typ = it_destRLambda_or_LetIn_names k p in
let n,typ = match typ with
- | RLambda (_,x,t,c) -> x, c
+ | RLambda (_,x,_,t,c) -> x, c
| _ -> Anonymous, typ in
let aliastyp =
if List.for_all ((=) Anonymous) nl then None
@@ -444,14 +444,14 @@ and share_names isgoal n l avoid env c t =
let t = detype isgoal avoid env t in
let id = next_name_away na avoid in
let avoid = id::avoid and env = add_name (Name id) env in
- share_names isgoal (n-1) ((Name id,None,t)::l) avoid env c c'
+ share_names isgoal (n-1) ((Name id,Explicit,None,t)::l) avoid env c c'
(* May occur for fix built interactively *)
| LetIn (na,b,t',c), _ when n > 0 ->
let t' = detype isgoal avoid env t' in
let b = detype isgoal avoid env b in
let id = next_name_away na avoid in
let avoid = id::avoid and env = add_name (Name id) env in
- share_names isgoal n ((Name id,Some b,t')::l) avoid env c t
+ share_names isgoal n ((Name id,Explicit,Some b,t')::l) avoid env c t
(* Only if built with the f/n notation or w/o let-expansion in types *)
| _, LetIn (_,b,_,t) when n > 0 ->
share_names isgoal n l avoid env c (subst1 b t)
@@ -461,7 +461,7 @@ and share_names isgoal n l avoid env c t =
let id = next_name_away na' avoid in
let avoid = id::avoid and env = add_name (Name id) env in
let appc = mkApp (lift 1 c,[|mkRel 1|]) in
- share_names isgoal (n-1) ((Name id,None,t')::l) avoid env appc c'
+ share_names isgoal (n-1) ((Name id,Explicit,None,t')::l) avoid env appc c'
(* If built with the f/n notation: we renounce to share names *)
| _ ->
if n>0 then warning "Detyping.detype: cannot factorize fix enough";
@@ -524,8 +524,8 @@ and detype_binder isgoal bk avoid env na ty c =
concrete_name (avoid_flag isgoal) avoid env na c in
let r = detype isgoal avoid' (add_name na' env) c in
match bk with
- | BProd -> RProd (dl, na',detype isgoal avoid env ty, r)
- | BLambda -> RLambda (dl, na',detype isgoal avoid env ty, r)
+ | BProd -> RProd (dl, na',Explicit,detype isgoal avoid env ty, r)
+ | BLambda -> RLambda (dl, na',Explicit,detype isgoal avoid env ty, r)
| BLetIn -> RLetIn (dl, na',detype isgoal avoid env ty, r)
let rec detype_rel_context where avoid env sign =
@@ -541,7 +541,7 @@ let rec detype_rel_context where avoid env sign =
else concrete_name None avoid env na c in
let b = Option.map (detype false avoid env) b in
let t = detype false avoid env t in
- (na',b,t) :: aux avoid' (add_name na' env) rest
+ (na',Explicit,b,t) :: aux avoid' (add_name na' env) rest
in aux avoid env (List.rev sign)
(**********************************************************************)
@@ -573,15 +573,15 @@ let rec subst_rawconstr subst raw =
if r' == r && rl' == rl then raw else
RApp(loc,r',rl')
- | RLambda (loc,n,r1,r2) ->
+ | RLambda (loc,n,bk,r1,r2) ->
let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RLambda (loc,n,r1',r2')
+ RLambda (loc,n,bk,r1',r2')
- | RProd (loc,n,r1,r2) ->
+ | RProd (loc,n,bk,r1,r2) ->
let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RProd (loc,n,r1',r2')
+ RProd (loc,n,bk,r1',r2')
| RLetIn (loc,n,r1,r2) ->
let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
@@ -629,10 +629,10 @@ let rec subst_rawconstr subst raw =
let ra1' = array_smartmap (subst_rawconstr subst) ra1
and ra2' = array_smartmap (subst_rawconstr subst) ra2 in
let bl' = array_smartmap
- (list_smartmap (fun (na,obd,ty as dcl) ->
+ (list_smartmap (fun (na,k,obd,ty as dcl) ->
let ty' = subst_rawconstr subst ty in
let obd' = Option.smartmap (subst_rawconstr subst) obd in
- if ty'==ty & obd'==obd then dcl else (na,obd',ty')))
+ if ty'==ty & obd'==obd then dcl else (na,k,obd',ty')))
bl in
if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else
RRec (loc,fix,ida,bl',ra1',ra2')
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index 588bc8c84..72379dfcf 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -41,7 +41,7 @@ val detype_case :
val detype_sort : sorts -> rawsort
val detype_rel_context : constr option -> identifier list -> names_context ->
- rel_context -> (name * rawconstr option * rawconstr) list
+ rel_context -> rawdecl list
(* look for the index of a named var or a nondep var as it is renamed *)
val lookup_name_as_renamed : env -> constr -> identifier -> int option
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index b894a9aae..37d60e470 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -520,11 +520,11 @@ let build_mutual_indrec env sigma = function
(List.map
(function (mind',mibi',mipi',dep',s') ->
let (sp',_) = mind' in
- if sp=sp' then
+ if sp=sp' then
let (mibi',mipi') = lookup_mind_specif env mind' in
- (mind',mibi',mipi',dep',s')
- else
- raise (RecursionSchemeError (NotMutualInScheme (mind,mind'))))
+ (mind',mibi',mipi',dep',s')
+ else
+ raise (RecursionSchemeError (NotMutualInScheme (mind,mind'))))
lrecspec)
in
let _ = check_arities listdepkind in
diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml
index 63dfbb2d9..845f9fae1 100644
--- a/pretyping/pattern.ml
+++ b/pretyping/pattern.ml
@@ -220,10 +220,10 @@ let rec pat_of_raw metas vars = function
| RApp (_,c,cl) ->
PApp (pat_of_raw metas vars c,
Array.of_list (List.map (pat_of_raw metas vars) cl))
- | RLambda (_,na,c1,c2) ->
+ | RLambda (_,na,bk,c1,c2) ->
PLambda (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | RProd (_,na,c1,c2) ->
+ | RProd (_,na,bk,c1,c2) ->
PProd (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
| RLetIn (_,na,c1,c2) ->
@@ -241,7 +241,7 @@ let rec pat_of_raw metas vars = function
PIf (pat_of_raw metas vars c,
pat_of_raw metas vars b1,pat_of_raw metas vars b2)
| RLetTuple (loc,nal,(_,None),b,c) ->
- let mkRLambda c na = RLambda (loc,na,RHole (loc,Evd.InternalHole),c) in
+ let mkRLambda c na = RLambda (loc,na,Explicit,RHole (loc,Evd.InternalHole),c) in
let c = List.fold_left mkRLambda c nal in
PCase ((LetStyle,[|1|],None,None),PMeta None,pat_of_raw metas vars b,
[|pat_of_raw metas vars c|])
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 9ecd7e251..3689ae174 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -321,11 +321,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| RRec (loc,fixkind,names,bl,lar,vdef) ->
let rec type_bl env ctxt = function
[] -> ctxt
- | (na,None,ty)::bl ->
+ | (na,bk,None,ty)::bl ->
let ty' = pretype_type empty_valcon env evdref lvar ty in
let dcl = (na,None,ty'.utj_val) in
type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl
- | (na,Some bd,ty)::bl ->
+ | (na,bk,Some bd,ty)::bl ->
let ty' = pretype_type empty_valcon env evdref lvar ty in
let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar ty in
let dcl = (na,Some bd'.uj_val,ty'.utj_val) in
@@ -420,7 +420,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| _ -> resj in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLambda(loc,name,c1,c2) ->
+ | RLambda(loc,name,bk,c1,c2) ->
let (name',dom,rng) = evd_comb1 (split_tycon loc env) evdref tycon in
let dom_valcon = valcon_of_tycon dom in
let j = pretype_type dom_valcon env evdref lvar c1 in
@@ -428,7 +428,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let j' = pretype rng (push_rel var env) evdref lvar c2 in
judge_of_abstraction env (orelse_name name name') j j'
- | RProd(loc,name,c1,c2) ->
+ | RProd(loc,name,bk,c1,c2) ->
let j = pretype_type empty_valcon env evdref lvar c1 in
let var = (name,j.utj_val) in
let env' = push_rel_assum var env in
@@ -652,7 +652,13 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(pretype tycon env evdref lvar c).uj_val
| IsType ->
(pretype_type empty_valcon env evdref lvar c).utj_val in
- nf_evar (evars_of !evdref) c'
+ let evd,_ = consider_remaining_unif_problems env !evdref in
+ let evd = nf_evar_defs evd in
+ let c' = nf_evar (evars_of evd) c' in
+ let evd = Typeclasses.resolve_typeclasses env (evars_of evd) evd in
+ let c' = nf_evar (evars_of evd) c' in
+ evdref := evd;
+ c'
(* TODO: comment faire remonter l'information si le typage a resolu des
variables du sigma original. il faudrait que la fonction de typage
@@ -680,10 +686,15 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let ise_pretype_gen fail_evar sigma env lvar kind c =
let evdref = ref (Evd.create_evar_defs sigma) in
let c = pretype_gen evdref env lvar kind c in
+(* let evd,_ = consider_remaining_unif_problems env !evdref in *)
+(* let evd = nf_evar_defs evd in *)
+(* let c = nf_evar (evars_of evd) c in *)
+(* let evd = undefined_evars evd in *)
+(* let evd = Typeclasses.resolve_typeclasses env sigma evd in *)
+(* let c = nf_evar (evars_of evd) c in *)
let evd,_ = consider_remaining_unif_problems env !evdref in
- let c = nf_evar (evars_of evd) c in
- if fail_evar then check_evars env sigma evd c;
- evd, c
+ if fail_evar then check_evars env (Evd.evars_of evd) evd c;
+ evd, c
(** Entry points of the high-level type synthesis algorithm *)
@@ -701,6 +712,13 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let understand_tcc_evars evdref env kind c =
pretype_gen evdref env ([],[]) kind c
+(* let c = pretype_gen evdref env ([],[]) kind c in *)
+(* evdref := nf_evar_defs !evdref; *)
+(* let c = nf_evar (evars_of !evdref) c in *)
+(* let evd = undefined_evars !evdref in *)
+(* let evd = Typeclasses.resolve_typeclasses env (evars_of evd) !evdref in *)
+(* evdref := evd; *)
+(* nf_evar (evars_of evd) c *)
let understand_tcc sigma env ?expected_type:exptyp c =
let evd, t = ise_pretype_gen false sigma env ([],[]) (OfType exptyp) c in
diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml
index d2a347363..bf5cc2272 100644
--- a/pretyping/rawterm.ml
+++ b/pretyping/rawterm.ml
@@ -36,6 +36,8 @@ type rawsort = RProp of Term.contents | RType of Univ.universe option
type binder_kind = BProd | BLambda | BLetIn
+type binding_kind = Explicit | Implicit
+
type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list
@@ -57,8 +59,8 @@ type rawconstr =
| REvar of loc * existential_key * rawconstr list option
| RPatVar of loc * (bool * patvar) (* Used for patterns only *)
| RApp of loc * rawconstr * rawconstr list
- | RLambda of loc * name * rawconstr * rawconstr
- | RProd of loc * name * rawconstr * rawconstr
+ | RLambda of loc * name * binding_kind * rawconstr * rawconstr
+ | RProd of loc * name * binding_kind * rawconstr * rawconstr
| RLetIn of loc * name * rawconstr * rawconstr
| RCases of loc * rawconstr option * tomatch_tuple * cases_clauses
| RLetTuple of loc * name list * (name * rawconstr option) *
@@ -71,7 +73,7 @@ type rawconstr =
| RCast of loc * rawconstr * rawconstr cast_type
| RDynamic of loc * Dyn.t
-and rawdecl = name * rawconstr option * rawconstr
+and rawdecl = name * binding_kind * rawconstr option * rawconstr
and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr
@@ -101,13 +103,13 @@ let cases_predicate_names tml =
- boolean in POldCase means it is recursive
i*)
-let map_rawdecl f (na,obd,ty) = (na,Option.map f obd,f ty)
+let map_rawdecl f (na,k,obd,ty) = (na,k,Option.map f obd,f ty)
let map_rawconstr f = function
| RVar (loc,id) -> RVar (loc,id)
| RApp (loc,g,args) -> RApp (loc,f g, List.map f args)
- | RLambda (loc,na,ty,c) -> RLambda (loc,na,f ty,f c)
- | RProd (loc,na,ty,c) -> RProd (loc,na,f ty,f c)
+ | RLambda (loc,na,bk,ty,c) -> RLambda (loc,na,bk,f ty,f c)
+ | RProd (loc,na,bk,ty,c) -> RProd (loc,na,bk,f ty,f c)
| RLetIn (loc,na,b,c) -> RLetIn (loc,na,f b,f c)
| RCases (loc,rtntypopt,tml,pl) ->
RCases (loc,Option.map f rtntypopt,
@@ -166,8 +168,8 @@ let occur_rawconstr id =
let rec occur = function
| RVar (loc,id') -> id = id'
| RApp (loc,f,args) -> (occur f) or (List.exists occur args)
- | RLambda (loc,na,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
- | RProd (loc,na,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
+ | RLambda (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
+ | RProd (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
| RLetIn (loc,na,b,c) -> (occur b) or ((na <> Name id) & (occur c))
| RCases (loc,rtntypopt,tml,pl) ->
(occur_option rtntypopt)
@@ -182,7 +184,7 @@ let occur_rawconstr id =
not (array_for_all4 (fun fid bl ty bd ->
let rec occur_fix = function
[] -> not (occur ty) && (fid=id or not(occur bd))
- | (na,bbd,bty)::bl ->
+ | (na,k,bbd,bty)::bl ->
not (occur bty) &&
(match bbd with
Some bd -> not (occur bd)
@@ -211,7 +213,7 @@ let free_rawvars =
let rec vars bounded vs = function
| RVar (loc,id') -> if Idset.mem id' bounded then vs else Idset.add id' vs
| RApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args)
- | RLambda (loc,na,ty,c) | RProd (loc,na,ty,c) | RLetIn (loc,na,ty,c) ->
+ | RLambda (loc,na,_,ty,c) | RProd (loc,na,_,ty,c) | RLetIn (loc,na,ty,c) ->
let vs' = vars bounded vs ty in
let bounded' = add_name_to_ids bounded na in
vars bounded' vs' c
@@ -234,7 +236,7 @@ let free_rawvars =
let vars_fix i vs fid =
let vs1,bounded1 =
List.fold_left
- (fun (vs,bounded) (na,bbd,bty) ->
+ (fun (vs,bounded) (na,k,bbd,bty) ->
let vs' = vars_option bounded vs bbd in
let vs'' = vars bounded vs' bty in
let bounded' = add_name_to_ids bounded na in
@@ -272,8 +274,8 @@ let loc_of_rawconstr = function
| REvar (loc,_,_) -> loc
| RPatVar (loc,_) -> loc
| RApp (loc,_,_) -> loc
- | RLambda (loc,_,_,_) -> loc
- | RProd (loc,_,_,_) -> loc
+ | RLambda (loc,_,_,_,_) -> loc
+ | RProd (loc,_,_,_,_) -> loc
| RLetIn (loc,_,_,_) -> loc
| RCases (loc,_,_,_) -> loc
| RLetTuple (loc,_,_,_,_) -> loc
diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli
index c43c290e8..2a5dab6e4 100644
--- a/pretyping/rawterm.mli
+++ b/pretyping/rawterm.mli
@@ -40,6 +40,8 @@ type rawsort = RProp of Term.contents | RType of Univ.universe option
type binder_kind = BProd | BLambda | BLetIn
+type binding_kind = Explicit | Implicit
+
type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list
@@ -61,8 +63,8 @@ type rawconstr =
| REvar of loc * existential_key * rawconstr list option
| RPatVar of loc * (bool * patvar) (* Used for patterns only *)
| RApp of loc * rawconstr * rawconstr list
- | RLambda of loc * name * rawconstr * rawconstr
- | RProd of loc * name * rawconstr * rawconstr
+ | RLambda of loc * name * binding_kind * rawconstr * rawconstr
+ | RProd of loc * name * binding_kind * rawconstr * rawconstr
| RLetIn of loc * name * rawconstr * rawconstr
| RCases of loc * rawconstr option * tomatch_tuple * cases_clauses
| RLetTuple of loc * name list * (name * rawconstr option) *
@@ -75,7 +77,7 @@ type rawconstr =
| RCast of loc * rawconstr * rawconstr cast_type
| RDynamic of loc * Dyn.t
-and rawdecl = name * rawconstr option * rawconstr
+and rawdecl = name * binding_kind * rawconstr option * rawconstr
and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
new file mode 100644
index 000000000..b0e7cb147
--- /dev/null
+++ b/pretyping/typeclasses.ml
@@ -0,0 +1,311 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.ml 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Util
+open Typeclasses_errors
+(*i*)
+
+let mismatched_params env n m = mismatched_ctx_inst env Parameters n m
+(* let mismatched_defs env n m = mismatched_ctx_inst env Definitions n m *)
+let mismatched_props env n m = mismatched_ctx_inst env Properties n m
+
+type rels = constr list
+
+(* This module defines type-classes *)
+type typeclass = {
+ cl_name : identifier; (* Name of the class *)
+ cl_context : named_context; (* Context in which superclasses and params are typed (usually types) *)
+ cl_super : named_context; (* Superclasses applied to some of the params *)
+ cl_params : named_context; (* Context of the parameters (usually types) *)
+(* cl_defs : named_context; (\* Context of the definitions (usually functions), which may be shared *\) *)
+ cl_props : named_context; (* Context of the properties on defs, in Prop, will not be shared *)
+ cl_impl : inductive; (* The class implementation: a record parameterized by params and defs *)
+}
+
+type typeclasses = (identifier, typeclass) Gmap.t
+
+type instance = {
+ is_class: typeclass;
+ is_name: identifier; (* Name of the instance *)
+(* is_params: named_context; (\* Context of the parameters, instanciated *\) *)
+(* is_super: named_context; (\* The corresponding superclasses *\) *)
+(* is_defs: named_context; (\* Context of the definitions, instanciated *\) *)
+ is_impl: constant;
+(* is_add_hint : unit -> unit; (\* Hook to add an hint for the instance *\) *)
+}
+
+type instances = (identifier, instance list) Gmap.t
+
+let classes : typeclasses ref = ref Gmap.empty
+
+let methods : (identifier, identifier) Gmap.t ref = ref Gmap.empty
+
+let instances : instances ref = ref Gmap.empty
+
+let freeze () = !classes, !methods, !instances
+
+let unfreeze (cl,m,is) =
+ classes:=cl;
+ methods:=m;
+ instances:=is
+
+let init () =
+ classes:= Gmap.empty;
+ methods:= Gmap.empty;
+ instances:= Gmap.empty
+
+let _ =
+ Summary.declare_summary "classes_and_instances"
+ { Summary.freeze_function = freeze;
+ Summary.unfreeze_function = unfreeze;
+ Summary.init_function = init;
+ Summary.survive_module = true;
+ Summary.survive_section = true }
+
+let gmap_merge old ne =
+ Gmap.fold (fun k v acc -> Gmap.add k v acc) ne old
+
+let gmap_list_merge old ne =
+ let ne =
+ Gmap.fold (fun k v acc ->
+ let oldv = try Gmap.find k old with Not_found -> [] in
+ let v' =
+ List.fold_left (fun acc x ->
+ if not (List.exists (fun y -> y.is_name = x.is_name) v) then x :: acc else acc)
+ v oldv
+ in Gmap.add k v' acc)
+ ne Gmap.empty
+ in
+ Gmap.fold (fun k v acc ->
+ let newv = try Gmap.find k ne with Not_found -> [] in
+ let v' =
+ List.fold_left (fun acc x ->
+ if not (List.exists (fun y -> y.is_name = x.is_name) acc) then x :: acc else acc)
+ newv v
+ in Gmap.add k v' acc)
+ old ne
+
+let cache (_, (cl, m, inst)) =
+ classes := gmap_merge !classes cl;
+ methods := gmap_merge !methods m;
+ instances := gmap_list_merge !instances inst
+
+open Libobject
+
+let subst (_,subst,(cl,m,inst)) =
+ let do_subst_con c = fst (Mod_subst.subst_con subst c)
+ and do_subst c = Mod_subst.subst_mps subst c
+ and do_subst_ind (kn,i) = (Mod_subst.subst_kn subst kn,i)
+ in
+ let do_subst_named ctx =
+ List.map (fun (na, b, t) ->
+ (na, Option.smartmap do_subst b, do_subst t))
+ ctx
+ in
+ let subst_class cl =
+ let cl' = { cl with cl_impl = do_subst_ind cl.cl_impl;
+ cl_context = do_subst_named cl.cl_context;
+ cl_super = do_subst_named cl.cl_super;
+ cl_params = do_subst_named cl.cl_params;
+ cl_props = do_subst_named cl.cl_props; }
+ in if cl' = cl then cl else cl'
+ in
+ let subst_inst classes insts =
+ List.map (fun is ->
+ { is with is_class = Gmap.find is.is_class.cl_name classes;
+ is_impl = do_subst_con is.is_impl }) insts
+ in
+ let classes = Gmap.map subst_class cl in
+ let instances = Gmap.map (subst_inst classes) inst in
+ (classes, m, instances)
+
+let discharge (_,(cl,m,inst)) =
+ let subst_class cl =
+ { cl with cl_impl = Lib.discharge_inductive cl.cl_impl }
+ in
+ let subst_inst classes insts =
+ List.map (fun is -> { is with is_impl = Lib.discharge_con is.is_impl;
+ is_class = Gmap.find is.is_class.cl_name classes; }) insts
+ in
+ let classes = Gmap.map subst_class cl in
+ let instances = Gmap.map (subst_inst classes) inst in
+ Some (classes, m, instances)
+
+let (input,output) =
+ declare_object
+ { (default_object "type classes state") with
+ cache_function = cache;
+ load_function = (fun _ -> cache);
+ open_function = (fun _ -> cache);
+ classify_function = (fun (_,x) -> Substitute x);
+ discharge_function = discharge;
+ subst_function = subst;
+ export_function = (fun x -> Some x) }
+
+let update () =
+ Lib.add_anonymous_leaf (input (!classes, !methods, !instances))
+
+let class_methods cl =
+ List.map (function (x, _, _) -> x) cl.cl_props
+
+let add_class c =
+ classes := Gmap.add c.cl_name c !classes;
+ methods := List.fold_left (fun acc x -> Gmap.add x c.cl_name acc) !methods (class_methods c);
+ update ()
+
+let class_info c =
+ Gmap.find c !classes
+
+let class_of_inductive ind =
+ let res = Gmap.fold
+ (fun k v acc -> match acc with None -> if v.cl_impl = ind then Some v else acc | _ -> acc)
+ !classes None
+ in match res with
+ None -> raise Not_found
+ | Some cl -> cl
+
+let typeclasses () = Gmap.fold (fun _ l c -> l :: c) !classes []
+
+let gmapl_add x y m =
+ try
+ let l = Gmap.find x m in
+ Gmap.add x (y::Util.list_except y l) m
+ with Not_found ->
+ Gmap.add x [y] m
+
+let instances_of c =
+ try Gmap.find c.cl_name !instances with Not_found -> []
+
+let add_instance i =
+ instances := gmapl_add i.is_class.cl_name i !instances;
+ update ()
+
+open Libnames
+
+let id_of_reference r =
+ let (_, id) = repr_qualid (snd (qualid_of_reference r)) in id
+
+let instances r =
+ let id = id_of_reference r in
+ try let cl = class_info id in instances_of cl
+ with Not_found -> unbound_class (Global.env()) (loc_of_reference r, id)
+
+let solve_instanciation_problem = ref (fun _ _ _ _ -> assert false)
+
+let resolve_typeclass env ev evi (evd, defined as acc) =
+ if evi.evar_body = Evar_empty then
+ try
+ !solve_instanciation_problem env evd ev evi
+ with Exit -> acc
+ else acc
+
+let resolve_one_typeclass env types =
+ try
+ let evi = Evd.make_evar (Environ.named_context_val env) types in
+ let ev = 1 in
+ let evm = Evd.add Evd.empty ev evi in
+ let evd = Evd.create_evar_defs evm in
+ let evd', b = !solve_instanciation_problem env evd ev evi in
+ if b then
+ let evm' = Evd.evars_of evd' in
+ match Evd.evar_body (Evd.find evm' ev) with
+ Evar_empty -> raise Not_found
+ | Evar_defined c -> c
+ else raise Not_found
+ with Exit -> raise Not_found
+
+let resolve_one_typeclass_evd env evd types =
+ try
+ let ev = Evarutil.e_new_evar evd env types in
+ let (ev,_) = destEvar ev in
+ let evd', b =
+ !solve_instanciation_problem env !evd ev (Evd.find (Evd.evars_of !evd) ev)
+ in
+ evd := evd';
+ if b then
+ let evm' = Evd.evars_of evd' in
+ match Evd.evar_body (Evd.find evm' ev) with
+ Evar_empty -> raise Not_found
+ | Evar_defined c -> c
+ else raise Not_found
+ with Exit -> raise Not_found
+
+let method_typeclass ref =
+ match ref with
+ | ConstRef c ->
+ let (_, _, l) = Names.repr_con c in
+ class_info (Gmap.find (Names.id_of_label l) !methods)
+ | _ -> raise Not_found
+
+let is_class ind =
+ Gmap.fold (fun k v acc -> acc || v.cl_impl = ind) !classes false
+
+let is_implicit_arg k =
+ match k with
+ ImplicitArg (ref, (n, id)) -> true
+ | InternalHole -> true
+ | _ -> false
+
+let resolve_typeclasses ?(check=true) env sigma evd =
+ let evm = Evd.evars_of evd in
+ let tc_evars =
+ let f ev evi acc =
+ let (l, k) = Evd.evar_source ev evd in
+ if not check || is_implicit_arg k then
+ match kind_of_term evi.evar_concl with
+ | App(ki, args) when isInd ki ->
+ if is_class (destInd ki) then Evd.add acc ev evi
+ else acc
+ | _ -> acc
+ else acc
+ in Evd.fold f evm Evd.empty
+ in
+ let rec sat evars =
+ let (evars', progress) =
+ Evd.fold
+ (fun ev evi acc ->
+ if Evd.mem tc_evars ev then resolve_typeclass env ev evi acc else acc)
+ (Evd.evars_of evars) (evars, false)
+ in
+ if not progress then evars'
+ else
+ sat (Evarutil.nf_evar_defs evars')
+ in sat evd
+
+let class_of_constr c =
+ match kind_of_term c with
+ App (c, _) ->
+ (match kind_of_term c with
+ Ind ind -> (try Some (class_of_inductive ind) with Not_found -> None)
+ | _ -> None)
+ | _ -> None
+
+
+type substitution = (identifier * constr) list
+
+let substitution_of_named_context isevars env id n subst l =
+ List.fold_right
+ (fun (na, _, t) subst ->
+ let t' = replace_vars subst t in
+ let b = Evarutil.e_new_evar isevars env ~src:(dummy_loc, ImplicitArg (VarRef id, (n, Some na))) t' in
+ (na, b) :: subst)
+ l subst
+
+let nf_substitution sigma subst =
+ List.map (function (na, c) -> na, Reductionops.nf_evar sigma c) subst
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
new file mode 100644
index 000000000..fbe48e06a
--- /dev/null
+++ b/pretyping/typeclasses.mli
@@ -0,0 +1,70 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.mli 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Topconstr
+open Util
+(*i*)
+
+(* This module defines type-classes *)
+type typeclass = {
+ cl_name : identifier; (* Name of the class *)
+ cl_context : named_context; (* Context in which superclasses and params are typed (usually types) *)
+ cl_super : named_context; (* Superclasses applied to some of the params *)
+ cl_params : named_context; (* Context of the real parameters (types and operations) *)
+(* cl_defs : rel_context; (\* Context of the definitions (usually functions), which may be shared *\) *)
+ cl_props : named_context; (* Context of the properties on defs, in Prop, will not be shared *)
+ cl_impl : inductive; (* The class implementation: a record parameterized by params and defs *)
+}
+
+type instance = {
+ is_class: typeclass;
+ is_name: identifier; (* Name of the instance *)
+(* is_params: named_context; (\* Context of the parameters, instanciated *\) *)
+(* is_super: named_context; (\* The corresponding superclasses *\) *)
+ is_impl: constant;
+(* is_add_hint : unit -> unit; (\* Hook to add an hint for the instance *\) *)
+}
+
+val instances : Libnames.reference -> instance list
+val typeclasses : unit -> typeclass list
+val add_class : typeclass -> unit
+val add_instance : instance -> unit
+
+val class_info : identifier -> typeclass (* raises Not_found *)
+val class_of_inductive : inductive -> typeclass (* raises Not_found *)
+
+val resolve_one_typeclass : env -> types -> types (* Raises Not_found *)
+val resolve_one_typeclass_evd : env -> evar_defs ref -> types -> types (* Raises Not_found *)
+
+val is_class : inductive -> bool
+
+val class_of_constr : constr -> typeclass option
+
+val resolve_typeclass : env -> evar -> evar_info -> evar_defs * bool -> evar_defs * bool
+val resolve_typeclasses : ?check:bool -> env -> evar_map -> evar_defs -> evar_defs
+
+val solve_instanciation_problem : (env -> evar_defs -> existential_key -> evar_info -> evar_defs * bool) ref
+
+type substitution = (identifier * constr) list
+
+val substitution_of_named_context :
+ evar_defs ref -> env -> identifier -> int ->
+ substitution -> named_context -> substitution
+
+val nf_substitution : evar_map -> substitution -> substitution
diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml
new file mode 100644
index 000000000..980f327cb
--- /dev/null
+++ b/pretyping/typeclasses_errors.ml
@@ -0,0 +1,42 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.ml 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Topconstr
+open Util
+(*i*)
+
+type contexts = Parameters | Properties
+
+type typeclass_error =
+ | UnboundClass of identifier located
+ | UnboundMethod of identifier * identifier located (* Class name, method *)
+ | NoInstance of identifier located * constr list
+ | MismatchedContextInstance of contexts * constr_expr list * named_context (* found, expected *)
+
+exception TypeClassError of env * typeclass_error
+
+let typeclass_error env err = raise (TypeClassError (env, err))
+
+let unbound_class env id = typeclass_error env (UnboundClass id)
+
+let unbound_method env cid id = typeclass_error env (UnboundMethod (cid, id))
+
+let no_instance env id args = typeclass_error env (NoInstance (id, args))
+
+let mismatched_ctx_inst env c n m = typeclass_error env (MismatchedContextInstance (c, n, m))
diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli
new file mode 100644
index 000000000..fbc2f2544
--- /dev/null
+++ b/pretyping/typeclasses_errors.mli
@@ -0,0 +1,40 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.mli 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Topconstr
+open Util
+(*i*)
+
+type contexts = Parameters | Properties
+
+type typeclass_error =
+ | UnboundClass of identifier located
+ | UnboundMethod of identifier * identifier located (* Class name, method *)
+ | NoInstance of identifier located * constr list
+ | MismatchedContextInstance of contexts * constr_expr list * named_context (* found, expected *)
+
+exception TypeClassError of env * typeclass_error
+
+val unbound_class : env -> identifier located -> 'a
+
+val unbound_method : env -> identifier -> identifier located -> 'a
+
+val no_instance : env -> identifier located -> constr list -> 'a
+
+val mismatched_ctx_inst : env -> contexts -> constr_expr list -> named_context -> 'a
diff --git a/proofs/refiner.mli b/proofs/refiner.mli
index ba092e89e..e79534b38 100644
--- a/proofs/refiner.mli
+++ b/proofs/refiner.mli
@@ -198,6 +198,7 @@ val solve_pftreestate : tactic -> pftreestate -> pftreestate
val weak_undo_pftreestate : pftreestate -> pftreestate
val mk_pftreestate : goal -> pftreestate
+val extract_open_proof : evar_map -> proof_tree -> constr * (int * types) list
val extract_open_pftreestate : pftreestate -> constr * Termops.metamap
val extract_pftreestate : pftreestate -> constr
val first_unproven : pftreestate -> pftreestate
diff --git a/tactics/class_setoid.ml4 b/tactics/class_setoid.ml4
new file mode 100644
index 000000000..0c8bdd298
--- /dev/null
+++ b/tactics/class_setoid.ml4
@@ -0,0 +1,224 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i camlp4deps: "parsing/grammar.cma" i*)
+
+(* $Id: eauto.ml4 10346 2007-12-05 21:11:19Z aspiwack $ *)
+
+open Pp
+open Util
+open Names
+open Nameops
+open Term
+open Termops
+open Sign
+open Reduction
+open Proof_type
+open Proof_trees
+open Declarations
+open Tacticals
+open Tacmach
+open Evar_refiner
+open Tactics
+open Pattern
+open Clenv
+open Auto
+open Rawterm
+open Hiddentac
+open Typeclasses
+open Typeclasses_errors
+
+let e_give_exact c gl =
+ let t1 = (pf_type_of gl c) and t2 = pf_concl gl in
+ if occur_existential t1 or occur_existential t2 then
+ tclTHEN (Clenvtac.unify t1) (exact_check c) gl
+ else exact_check c gl
+
+let assumption id = e_give_exact (mkVar id)
+
+let morphism_class = lazy (class_info (id_of_string "Morphism"))
+let morphism2_class = lazy (class_info (id_of_string "BinaryMorphism"))
+let morphism3_class = lazy (class_info (id_of_string "TernaryMorphism"))
+
+let get_respect cl =
+ Option.get (List.hd (Recordops.lookup_projections cl.cl_impl))
+
+let respect_proj = lazy (get_respect (Lazy.force morphism_class))
+let respect2_proj = lazy (get_respect (Lazy.force morphism2_class))
+let respect3_proj = lazy (get_respect (Lazy.force morphism3_class))
+
+let gen_constant dir s = Coqlib.gen_constant "Class_setoid" dir s
+let coq_proj1 = lazy(gen_constant ["Init"; "Logic"] "proj1")
+let coq_proj2 = lazy(gen_constant ["Init"; "Logic"] "proj2")
+let iff = lazy (gen_constant ["Init"; "Logic"] "iff")
+
+let iff_setoid = lazy (gen_constant ["Classes"; "Setoid"] "iff_setoid")
+let setoid_equiv = lazy (gen_constant ["Classes"; "Setoid"] "equiv")
+let setoid_morphism = lazy (gen_constant ["Classes"; "Setoid"] "setoid_morphism")
+let setoid_refl_proj = lazy (gen_constant ["Classes"; "Setoid"] "equiv_refl")
+
+let arrow_morphism a b =
+ mkLambda (Name (id_of_string "A"), a,
+ mkLambda (Name (id_of_string "B"), b,
+ mkProd (Anonymous, mkRel 2, mkRel 2)))
+
+let setoid_refl l sa x =
+ applistc (Lazy.force setoid_refl_proj) (l @ [sa ; x])
+
+let class_one = lazy (Lazy.force morphism_class, Lazy.force respect_proj)
+let class_two = lazy (Lazy.force morphism2_class, Lazy.force respect2_proj)
+let class_three = lazy (Lazy.force morphism3_class, Lazy.force respect3_proj)
+
+exception Found of (constr * constant * constr list * int * constr * constr array *
+ (constr * (constr * constr * constr * constr * constr)) option array)
+
+let resolve_morphism_evd env evd app =
+ let ev = Evarutil.e_new_evar evd env app in
+ let evd' = resolve_typeclasses ~check:false env (Evd.evars_of !evd) !evd in
+ let evm' = Evd.evars_of evd' in
+ match Evd.evar_body (Evd.find evm' (fst (destEvar ev))) with
+ Evd.Evar_empty -> raise Not_found
+ | Evd.Evar_defined c -> evd := Evarutil.nf_evar_defs evd'; c
+
+let is_equiv env sigma t =
+ isConst t && Reductionops.is_conv env sigma (Lazy.force setoid_equiv) t
+
+let resolve_morphism env sigma direction oldt m args args' =
+ let evars = ref (Evd.create_evar_defs Evd.empty) in
+ let morph_instance, proj, subst, len, m', args, args' =
+ if is_equiv env sigma m then
+ let params, rest = array_chop 3 args in
+ let a, r, s = params.(0), params.(1), params.(2) in
+ let params', rest' = array_chop 3 args' in
+ let inst = mkApp (Lazy.force setoid_morphism, params) in
+ (* Equiv gives a binary morphism *)
+ let (cl, proj) = Lazy.force class_two in
+ let ctxargs = [ a; r; a; r; mkProp; Lazy.force iff; s; s; Lazy.force iff_setoid; ] in
+ let m' = mkApp (m, [| a; r; s |]) in
+ inst, proj, ctxargs, 6, m', rest, rest'
+ else
+ let cls =
+ match Array.length args with
+ 1 -> [Lazy.force class_one, 1]
+ | 2 -> [Lazy.force class_two, 2; Lazy.force class_one, 1]
+ | 3 -> [Lazy.force class_three, 3; Lazy.force class_two, 2; Lazy.force class_one, 1]
+ | n -> [Lazy.force class_three, 3; Lazy.force class_two, 2; Lazy.force class_one, 1]
+ in
+ try
+ List.iter (fun ((cl, proj), n) ->
+ evars := Evd.create_evar_defs Evd.empty;
+ let ctxevs = substitution_of_named_context evars env cl.cl_name 0 [] cl.cl_context in
+ let len = List.length ctxevs in
+ let superevs = substitution_of_named_context evars env cl.cl_name len ctxevs cl.cl_super in
+ let morphargs, morphobjs = array_chop (Array.length args - n) args in
+ let morphargs', morphobjs' = array_chop (Array.length args - n) args' in
+ let args = List.rev_map (fun (_, c) -> c) superevs in
+ let appm = mkApp(m, morphargs) in
+ let appmtype = Typing.type_of env sigma appm in
+ let app = applistc (mkInd cl.cl_impl) (args @ [appm]) in
+ let mtype = replace_vars superevs (pi3 (List.hd cl.cl_params)) in
+ try
+ evars := Unification.w_unify true env CONV ~mod_delta:true appmtype mtype !evars;
+ evars := Evarutil.nf_evar_defs !evars;
+ let app = Evarutil.nf_isevar !evars app in
+ raise (Found (resolve_morphism_evd env evars app, proj, args, len, appm, morphobjs, morphobjs'))
+ with Not_found -> ()
+ | Stdpp.Exc_located (_, Pretype_errors.PretypeError _)
+ | Pretype_errors.PretypeError _ -> ())
+ cls;
+ raise Not_found
+ with Found x -> x
+ in
+ evars := Evarutil.nf_evar_defs !evars;
+ let evm = Evd.evars_of !evars in
+ let ctxargs = List.map (Reductionops.nf_evar evm) subst in
+ let ctx, sup = Util.list_chop len ctxargs in
+ let m' = Reductionops.nf_evar evm m' in
+ let appproj = applistc (mkConst proj) (ctxargs @ [m' ; morph_instance]) in
+ let projargs, respars, ressetoid, typeargs =
+ array_fold_left2
+ (fun (acc, ctx, sup, typeargs') x y ->
+ let par, ctx = list_chop 2 ctx in
+ let setoid, sup = List.hd sup, List.tl sup in
+ match y with
+ None ->
+ let refl_proof = setoid_refl par setoid x in
+ [ refl_proof ; x ; x ] @ acc, ctx, sup, x :: typeargs'
+ | Some (p, (_, _, _, _, t')) ->
+ if direction then
+ [ p ; t'; x ] @ acc, ctx, sup, t' :: typeargs'
+ else [ p ; x; t' ] @ acc, ctx, sup, t' :: typeargs')
+ ([], ctx, sup, []) args args'
+ in
+ let proof = applistc appproj (List.rev projargs) in
+ let newt = applistc m' (List.rev typeargs) in
+ match respars, ressetoid with
+ [ a ; r ], [ s ] -> (proof, (a, r, s, oldt, newt))
+ | _ -> assert(false)
+
+let build_new gl env setoid direction origt newt hyp hypinfo concl =
+ let rec aux t =
+ match kind_of_term t with
+ | _ when eq_constr t origt ->
+ Some (hyp, hypinfo)
+ | App (m, args) ->
+ let args' = Array.map aux args in
+ if array_for_all (fun x -> x = None) args' then None
+ else
+ (try Some (resolve_morphism env (project gl) direction t m args args')
+ with Not_found -> None)
+ | Prod (_, x, b) ->
+ let x', b' = aux x, aux b in
+ if x' = None && b' = None then None
+ else
+ (try Some (resolve_morphism env (project gl) direction t (arrow_morphism (pf_type_of gl x) (pf_type_of gl b)) [| x ; b |] [| x' ; b' |])
+ with Not_found -> None)
+
+ | _ -> None
+ in aux concl
+
+let decompose_setoid_eqhyp env sigma c dir t =
+ match kind_of_term t with
+ | App (equiv, [| a; r; s; x; y |]) ->
+ if dir then (c, (a, r, s, x, y))
+ else (c, (a, r, s, y, x))
+ | App (r, args) when Array.length args >= 2 ->
+ (try
+ let (p, (a, r, s, _, _)) = resolve_morphism env sigma dir t r args (Array.map (fun _ -> None) args) in
+ let _, args = array_chop (Array.length args - 2) args in
+ if dir then (c, (a, r, s, args.(0), args.(1)))
+ else (c, (a, r, s, args.(1), args.(0)))
+ with Not_found -> error "Not a (declared) setoid equality")
+ | _ -> error "Not a setoid equality"
+
+let cl_rewrite c left2right gl =
+ let env = pf_env gl in
+ let sigma = project gl in
+ let hyp = pf_type_of gl c in
+ let hypt, (typ, rel, setoid, origt, newt as hypinfo) = decompose_setoid_eqhyp env sigma c left2right hyp in
+ let concl = pf_concl gl in
+ let _concltyp = pf_type_of gl concl in
+ let eq = build_new gl env setoid left2right origt newt hypt hypinfo concl in
+ match eq with
+ Some (p, (_, _, _, _, t)) ->
+ let proj =
+ if left2right then
+ applistc (Lazy.force coq_proj2)
+ [ mkProd (Anonymous, concl, t) ; mkProd (Anonymous, t, concl) ; p ]
+ else
+ applistc (Lazy.force coq_proj1)
+ [ mkProd (Anonymous, t, concl) ; mkProd (Anonymous, concl, t) ; p ]
+ in
+ (Tactics.apply proj) gl
+ | None -> tclIDTAC gl
+
+open Extraargs
+
+TACTIC EXTEND class_rewrite
+| [ "clrewrite" orient(o) constr(c) ] -> [ cl_rewrite c o ]
+END
diff --git a/tactics/decl_interp.ml b/tactics/decl_interp.ml
index 771dbe736..b33fbde8a 100644
--- a/tactics/decl_interp.ml
+++ b/tactics/decl_interp.ml
@@ -182,17 +182,17 @@ let interp_constr_or_thesis check_sort sigma env = function
let type_tester_var body typ =
raw_app(dummy_loc,
- RLambda(dummy_loc,Anonymous,typ,
+ RLambda(dummy_loc,Anonymous,Explicit,typ,
RSort (dummy_loc,RProp Null)),body)
let abstract_one_hyp inject h raw =
match h with
Hvar (loc,(id,None)) ->
- RProd (dummy_loc,Name id, RHole (loc,Evd.BinderType (Name id)), raw)
+ RProd (dummy_loc,Name id, Explicit, RHole (loc,Evd.BinderType (Name id)), raw)
| Hvar (loc,(id,Some typ)) ->
- RProd (dummy_loc,Name id,fst typ, raw)
+ RProd (dummy_loc,Name id, Explicit, fst typ, raw)
| Hprop st ->
- RProd (dummy_loc,st.st_label,inject st.st_it, raw)
+ RProd (dummy_loc,st.st_label, Explicit, inject st.st_it, raw)
let rawconstr_of_hyps inject hyps head =
List.fold_right (abstract_one_hyp inject) hyps head
@@ -254,18 +254,18 @@ let rec raw_of_pat =
let prod_one_hyp = function
(loc,(id,None)) ->
(fun raw ->
- RProd (dummy_loc,Name id,
+ RProd (dummy_loc,Name id, Explicit,
RHole (loc,Evd.BinderType (Name id)), raw))
| (loc,(id,Some typ)) ->
(fun raw ->
- RProd (dummy_loc,Name id,fst typ, raw))
+ RProd (dummy_loc,Name id, Explicit, fst typ, raw))
let prod_one_id (loc,id) raw =
- RProd (dummy_loc,Name id,
+ RProd (dummy_loc,Name id, Explicit,
RHole (loc,Evd.BinderType (Name id)), raw)
let let_in_one_alias (id,pat) raw =
- RLetIn (dummy_loc,Name id,raw_of_pat pat, raw)
+ RLetIn (dummy_loc,Name id, raw_of_pat pat, raw)
let rec bind_primary_aliases map pat =
match pat with
@@ -417,11 +417,11 @@ let interp_casee sigma env = function
let abstract_one_arg = function
(loc,(id,None)) ->
(fun raw ->
- RLambda (dummy_loc,Name id,
+ RLambda (dummy_loc,Name id, Explicit,
RHole (loc,Evd.BinderType (Name id)), raw))
| (loc,(id,Some typ)) ->
(fun raw ->
- RLambda (dummy_loc,Name id,fst typ, raw))
+ RLambda (dummy_loc,Name id, Explicit, fst typ, raw))
let rawconstr_of_fun args body =
List.fold_right abstract_one_arg args (fst body)
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index 3211cc6cf..dbb4648bc 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -2579,22 +2579,32 @@ let bad_tactic_args s =
(* Declaration of the TAC-DEFINITION object *)
let add (kn,td) = mactab := Gmap.add kn td !mactab
+type tacdef_kind = | NewTac of identifier
+ | UpdateTac of ltac_constant
+
let load_md i ((sp,kn),defs) =
let dp,_ = repr_path sp in
let mp,dir,_ = repr_kn kn in
List.iter (fun (id,t) ->
- let sp = Libnames.make_path dp id in
- let kn = Names.make_kn mp dir (label_of_id id) in
- Nametab.push_tactic (Until i) sp kn;
- add (kn,t)) defs
-
+ match id with
+ NewTac id ->
+ let sp = Libnames.make_path dp id in
+ let kn = Names.make_kn mp dir (label_of_id id) in
+ Nametab.push_tactic (Until i) sp kn;
+ add (kn,t)
+ | UpdateTac kn ->
+ add (kn,t)) defs
+
let open_md i((sp,kn),defs) =
let dp,_ = repr_path sp in
let mp,dir,_ = repr_kn kn in
List.iter (fun (id,t) ->
- let sp = Libnames.make_path dp id in
- let kn = Names.make_kn mp dir (label_of_id id) in
- Nametab.push_tactic (Exactly i) sp kn) defs
+ match id with
+ NewTac id ->
+ let sp = Libnames.make_path dp id in
+ let kn = Names.make_kn mp dir (label_of_id id) in
+ Nametab.push_tactic (Exactly i) sp kn
+ | UpdateTac kn -> ()) defs
let cache_md x = load_md 1 x
@@ -2622,27 +2632,40 @@ let print_ltac id =
(pr_qualid id ++ spc() ++ str "is not a user defined tactic")
(* Adds a definition for tactics in the table *)
-let make_absolute_name (loc,id) =
- let kn = Lib.make_kn id in
- if Gmap.mem kn !mactab or is_atomic_kn kn then
+let make_absolute_name (loc,id) repl =
+ try
+ let kn = if repl then Nametab.locate_tactic (make_short_qualid id) else Lib.make_kn id in
+ if Gmap.mem kn !mactab then
+ if repl then kn
+ else
+ user_err_loc (loc,"Tacinterp.add_tacdef",
+ str "There is already an Ltac named " ++ pr_id id)
+ else if is_atomic_kn kn then
+ user_err_loc (loc,"Tacinterp.add_tacdef",
+ str "Reserved Ltac name " ++ pr_id id)
+ else kn
+ with Not_found ->
user_err_loc (loc,"Tacinterp.add_tacdef",
- str "There is already an Ltac named " ++ pr_id id);
- kn
-
+ str "There is no Ltac named " ++ pr_id id)
+
let add_tacdef isrec tacl =
(* let isrec = if !Flags.p1 then isrec else true in*)
- let rfun = List.map (fun ((loc,id as locid),_) -> (id,make_absolute_name locid)) tacl in
+ let rfun = List.map (fun ((loc,id as locid),b,_) -> (id,make_absolute_name locid b)) tacl in
let ist =
{(make_empty_glob_sign()) with ltacrecvars = if isrec then rfun else []} in
let gtacl =
- List.map (fun ((_,id),def) ->
- (id,Flags.with_option strict_check (intern_tactic ist) def))
- tacl in
+ List.map2 (fun ((_,id),b,def) (_, qid) ->
+ let k = if b then UpdateTac qid else NewTac id in
+ let t = Flags.with_option strict_check (intern_tactic ist) def in
+ (k, t))
+ tacl rfun in
let id0 = fst (List.hd rfun) in
let _ = Lib.add_leaf id0 (inMD gtacl) in
List.iter
- (fun (id,_) -> Flags.if_verbose msgnl (pr_id id ++ str " is defined"))
- rfun
+ (fun ((_,id),b,_) ->
+ if b then Flags.if_verbose msgnl (pr_id id ++ str " is redefined")
+ else Flags.if_verbose msgnl (pr_id id ++ str " is defined"))
+ tacl
(***************************************************************************)
(* Other entry points *)
diff --git a/tactics/tacinterp.mli b/tactics/tacinterp.mli
index 420ae8d74..6ea0505a0 100644
--- a/tactics/tacinterp.mli
+++ b/tactics/tacinterp.mli
@@ -63,7 +63,7 @@ val get_debug : unit -> debug_info
(* Adds a definition for tactics in the table *)
val add_tacdef :
- bool -> (identifier Util.located * raw_tactic_expr) list -> unit
+ bool -> (identifier Util.located * bool * raw_tactic_expr) list -> unit
val add_primitive_tactic : string -> glob_tactic_expr -> unit
(* Tactic extensions *)
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index f1f169394..dbd3aacbf 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -1753,7 +1753,10 @@ let abstract_args gl id =
let argty = pf_type_of gl arg in
let liftarg = lift (List.length ctx) arg in
let liftargty = lift (List.length ctx) argty in
- let convertible = Reductionops.is_conv ctxenv sigma ty liftargty in
+ let convertible =
+ Reductionops.is_conv ctxenv sigma
+ (Termops.refresh_universes ty) (Termops.refresh_universes liftargty)
+ in
match kind_of_term arg with
| Var _ | Rel _ when convertible ->
(subst1 arg arity, ctx, ctxenv, mkApp (c, [|arg|]), args, eqs, refls, (Anonymous, liftarg, liftarg) :: finalargs, env)
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index eb62f602a..db46c621f 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -327,4 +327,11 @@ val tclABSTRACT : identifier option -> tactic -> tactic
val admit_as_an_axiom : tactic
+val make_abstract_generalize : 'a ->
+ Names.identifier ->
+ Term.constr ->
+ Sign.rel_context ->
+ Term.types ->
+ Term.types list ->
+ Term.constr list -> Term.constr list -> Term.constr -> Term.constr
val abstract_generalize : identifier -> tactic
diff --git a/test-suite/typeclasses/NewSetoid.v b/test-suite/typeclasses/NewSetoid.v
new file mode 100644
index 000000000..8ad038809
--- /dev/null
+++ b/test-suite/typeclasses/NewSetoid.v
@@ -0,0 +1,74 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* Certified Haskell Prelude.
+ * Author: Matthieu Sozeau
+ * Institution: LRI, CNRS UMR 8623 - UniversitÃcopyright Paris Sud
+ * 91405 Orsay, France *)
+
+(* $Id: FSetAVL_prog.v 616 2007-08-08 12:28:10Z msozeau $ *)
+
+Require Import Coq.Program.Program.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+Require Import Coq.Classes.SetoidTactics.
+
+Goal not True == not (not False) -> ((not True -> True)) \/ True.
+ intros.
+ clrewrite H.
+ clrewrite <- H.
+ right ; auto.
+Defined.
+
+Definition reduced_thm := Eval compute in Unnamed_thm.
+
+(* Print reduced_thm. *)
+
+Lemma foo [ Setoid a R ] : True. (* forall x y, R x y -> x -> y. *)
+Proof.
+ intros.
+ Print respect2.
+ pose setoid_morphism.
+ pose (respect2 (b0:=b)).
+ simpl in b0.
+ unfold binary_respectful in b0.
+ pose (arrow_morphism R).
+ pose (respect2 (b0:=b1)).
+ unfold binary_respectful in b2.
+
+ pose (eq_morphism (A:=a)).
+ pose (respect2 (b0:=b3)).
+ unfold binary_respectful in b4.
+ exact I.
+Qed.
+
+Goal forall A B C (H : A <-> B) (H' : B <-> C), A /\ B <-> B /\ C.
+ intros.
+ Set Printing All.
+ Print iff_morphism.
+ clrewrite H.
+ clrewrite H'.
+ reflexivity.
+Defined.
+
+Goal forall A B C (H : A <-> B) (H' : B <-> C), A /\ B <-> B /\ C.
+ intros.
+ rewrite H.
+ rewrite H'.
+ reflexivity.
+Defined.
+
+Require Import Setoid_tac.
+Require Import Setoid_Prop.
+
+(* Print Unnamed_thm0. *)
+(* Print Unnamed_thm1. *)
+
+
diff --git a/theories/Arith/Even.v b/theories/Arith/Even.v
index c61d89a00..ee3d5fe0d 100644
--- a/theories/Arith/Even.v
+++ b/theories/Arith/Even.v
@@ -40,7 +40,7 @@ Proof.
induction n.
auto with arith.
elim IHn; auto with arith.
-Qed.
+Defined.
Lemma not_even_and_odd : forall n, even n -> odd n -> False.
Proof.
diff --git a/theories/Classes/Init.v b/theories/Classes/Init.v
new file mode 100644
index 000000000..307005382
--- /dev/null
+++ b/theories/Classes/Init.v
@@ -0,0 +1 @@
+Instantiation Tactic := eauto 50 with typeclass_instances || eauto.
diff --git a/theories/Classes/Init.v.d b/theories/Classes/Init.v.d
new file mode 100644
index 000000000..482ac0796
--- /dev/null
+++ b/theories/Classes/Init.v.d
@@ -0,0 +1 @@
+theories/Classes/Init.vo theories/Classes/Init.glob: theories/Classes/Init.v
diff --git a/theories/Classes/Setoid.v b/theories/Classes/Setoid.v
new file mode 100644
index 000000000..915a7a944
--- /dev/null
+++ b/theories/Classes/Setoid.v
@@ -0,0 +1,192 @@
+(* -*- coq-prog-args: ("-emacs-U" "-nois") -*- *)
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* Certified Haskell Prelude.
+ * Author: Matthieu Sozeau
+ * Institution: LRI, CNRS UMR 8623 - UniversitÃcopyright Paris Sud
+ * 91405 Orsay, France *)
+
+(* $Id: FSetAVL_prog.v 616 2007-08-08 12:28:10Z msozeau $ *)
+
+Require Import Coq.Program.Program.
+Require Import Coq.Classes.Init.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+(** We first define setoids on a carrier, it amounts to an equivalence relation.
+ Now [equiv] is overloaded for every [Setoid].
+*)
+
+Require Export Coq.Relations.Relations.
+
+Class Setoid (carrier : Type) (equiv : relation carrier) :=
+ equiv_prf : equivalence carrier equiv.
+
+(** Overloaded notation for setoid equivalence. Not to be confused with [eq] and [=]. *)
+
+Definition equiv [ Setoid A R ] : _ := R.
+
+Infix "==" := equiv (at level 70, no associativity).
+
+Definition equiv_refl [ s : Setoid A R ] : forall x : A, R x x := equiv_refl _ _ equiv_prf.
+Definition equiv_sym [ s : Setoid A R ] : forall x y : A, R x y -> R y x := equiv_sym _ _ equiv_prf.
+Definition equiv_trans [ s : Setoid A R ] : forall x y z : A, R x y -> R y z -> R x z := equiv_trans _ _ equiv_prf.
+
+
+Ltac refl :=
+ match goal with
+ | [ |- @equiv ?A ?R ?s ?X _ ] => apply (equiv_refl (A:=A) (R:=R) (s:=s) X)
+ | [ |- ?R ?X _ ] => apply (equiv_refl (R:=R) X)
+ | [ |- ?R ?A ?X _ ] => apply (equiv_refl (R:=R A) X)
+ | [ |- ?R ?A ?B ?X _ ] => apply (equiv_refl (R:=R A B) X)
+ | [ |- ?R ?A ?B ?C ?X _ ] => apply (equiv_refl (R:=R A B C) X)
+ end.
+
+Ltac sym :=
+ match goal with
+ | [ |- @equiv ?A ?R ?s ?X ?Y ] => apply (equiv_sym (A:=A) (R:=R) (s:=s) (x:=X) (y:=Y))
+ | [ |- ?R ?X ?Y ] => apply (equiv_sym (R:=R) (x:=Y) (y:=X))
+ | [ |- ?R ?A ?X ?Y ] => apply (equiv_sym (R:=R A) (x:=Y) (y:=X))
+ | [ |- ?R ?A ?B ?X ?Y ] => apply (equiv_sym (R:=R A B) (x:=Y) (y:=X))
+ | [ |- ?R ?A ?B ?C ?X ?Y ] => apply (equiv_sym (R:=R A B C) (x:=Y) (y:=X))
+ end.
+
+Ltac trans Y :=
+ match goal with
+ | [ |- @equiv ?A ?R ?s ?X ?Z ] => apply (equiv_trans (A:=A) (R:=R) (s:=s) (x:=X) (y:=Y) (z:=Z))
+ | [ |- ?R ?X ?Z ] => apply (equiv_trans (R:=R) (x:=X) (y:=Y) (z:=Z))
+ | [ |- ?R ?A ?X ?Z ] => apply (equiv_trans (R:=R A) (x:=X) (y:=Y) (z:=Z))
+ | [ |- ?R ?A ?B ?X ?Z ] => apply (equiv_trans (R:=R A B) (x:=X) (y:=Y) (z:=Z))
+ | [ |- ?R ?A ?B ?C ?X ?Z ] => apply (equiv_trans (R:=R A B C) (x:=X) (y:=Y) (z:=Z))
+ end.
+
+Definition respectful [ sa : Setoid a eqa, sb : Setoid b eqb ] (m : a -> b) : Prop :=
+ forall x y, eqa x y -> eqb (m x) (m y).
+
+Class [ Setoid a eqa, Setoid b eqb ] => Morphism (m : a -> b) :=
+ respect : respectful m.
+
+(** Here we build a setoid instance for functions which relates respectful ones only. *)
+
+Definition respecting [ Setoid a R, Setoid b R' ] : Type := { morph : a -> b | respectful morph }.
+
+Obligations Tactic := try red ; program_simpl ; unfold equiv in * ; try tauto.
+
+Program Instance [ sa : Setoid a R, sb : Setoid b R' ] => arrow_setoid :
+ Setoid ({ morph : a -> b | respectful morph })
+ (fun (f g : respecting) => forall (x y : a), R x y -> R' (`f x) (`g y)) :=
+ equiv_prf := Build_equivalence _ _ _ _ _.
+
+Next Obligation.
+Proof.
+ trans (y x0) ; eauto.
+ apply H.
+ refl.
+Qed.
+
+Next Obligation.
+Proof.
+ sym ; apply H.
+ sym ; auto.
+Qed.
+
+(** We redefine respect for binary and ternary morphims because we cannot get a satisfying instance of [Setoid (a -> b)] from
+ some arbitrary domain and codomain setoids. We can define it on respectful Coq functions though, see [arrow_setoid] above. *)
+
+Definition binary_respectful [ sa : Setoid a eqa, sb : Setoid b eqb, Setoid c eqc ] (m : a -> b -> c) : Prop :=
+ forall x y, eqa x y ->
+ forall z w, eqb z w -> eqc (m x z) (m y w).
+
+Class [ sa : Setoid a eqa, sb : Setoid b eqb, sc : Setoid c eqc ] => BinaryMorphism (m : a -> b -> c) :=
+ respect2 : binary_respectful m.
+
+Definition ternary_respectful [ sa : Setoid a eqa, sb : Setoid b eqb, sc : Setoid c eqc, Setoid d eqd ] (m : a -> b -> c -> d) : Prop :=
+ forall x y, eqa x y -> forall z w, eqb z w -> forall u v, eqc u v -> eqd (m x z u) (m y w v).
+
+Class [ sa : Setoid a eqa, sb : Setoid b eqb, sc : Setoid c eqc, sd : Setoid d eqd ] => TernaryMorphism (m : a -> b -> c -> d) :=
+ respect3 : ternary_respectful m.
+
+(** Definition of the usual morphisms in [Prop]. *)
+
+Program Instance iff_setoid : Setoid Prop iff :=
+ equiv_prf := @Build_equivalence _ _ iff_refl iff_trans iff_sym.
+
+Program Instance not_morphism : Morphism Prop iff Prop iff not.
+
+Program Instance and_morphism : ? BinaryMorphism iff_setoid iff_setoid iff_setoid and.
+
+(* We make the setoids implicit, they're always [iff] *)
+
+Implicit Arguments Enriching BinaryMorphism [[!sa] [!sb] [!sc]].
+
+Program Instance or_morphism : ? BinaryMorphism or.
+
+Definition impl (A B : Prop) := A -> B.
+
+Program Instance impl_morphism : ? BinaryMorphism impl.
+
+Next Obligation.
+Proof.
+ unfold impl. tauto.
+Qed.
+
+(** Every setoid relation gives rise to a morphism, in fact every partial setoid does. *)
+
+Program Instance [ Setoid a R ] => setoid_morphism : ? BinaryMorphism R.
+
+Next Obligation.
+Proof with auto.
+ split ; intros.
+ trans x. sym... trans z...
+ trans y... trans w... sym...
+Qed.
+
+Definition iff_morphism : BinaryMorphism iff := setoid_morphism.
+
+Existing Instance iff_morphism.
+
+Implicit Arguments eq [[A]].
+
+Program Instance eq_setoid : Setoid A eq :=
+ equiv_prf := Build_equivalence _ _ _ _ _.
+
+Program Instance eq_morphism : BinaryMorphism A eq A eq Prop iff eq.
+
+Program Instance arrow_morphism : BinaryMorphism A eq B eq C eq m.
+
+Implicit Arguments arrow_morphism [[A] [B] [C]].
+
+Program Instance type_setoid : Setoid Type (fun x y => x = y) :=
+ equiv_prf := Build_equivalence _ _ _ _ _.
+
+Lemma setoid_subst : forall (x y : Type), x == y -> x -> y.
+Proof.
+ intros.
+ rewrite <- H.
+ apply X.
+Qed.
+
+Lemma prop_setoid_subst : forall (x y : Prop), x == y -> x -> y.
+Proof.
+ intros.
+ clrewrite <- H.
+ apply H0.
+Qed.
+
+Program Instance [ sa : Setoid a eqa, sb : Setoid b eqb, sc : Setoid c eqc,
+ ? Morphism sb sc g, ? Morphism sa sb f ] =>
+ compose_morphism : ? Morphism sa sc (fun x => g (f x)).
+
+Next Obligation.
+Proof.
+ apply (respect (m0:=m)).
+ apply (respect (m0:=m0)).
+ assumption.
+Qed.
diff --git a/theories/Classes/SetoidTactics.v b/theories/Classes/SetoidTactics.v
new file mode 100644
index 000000000..8e037db1a
--- /dev/null
+++ b/theories/Classes/SetoidTactics.v
@@ -0,0 +1,37 @@
+(* -*- coq-prog-args: ("-emacs-U" "-nois") -*- *)
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* Certified Haskell Prelude.
+ * Author: Matthieu Sozeau
+ * Institution: LRI, CNRS UMR 8623 - UniversitÃcopyright Paris Sud
+ * 91405 Orsay, France *)
+
+(* $Id: FSetAVL_prog.v 616 2007-08-08 12:28:10Z msozeau $ *)
+
+Require Import Coq.Program.Program.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+Require Export Coq.Classes.Setoid.
+
+Ltac rew H := clrewrite H.
+
+Lemma setoideq_eq [ sa : Setoid a eqa ] : forall x y, eqa x y -> x = y.
+Proof.
+ admit.
+Qed.
+
+Implicit Arguments setoideq_eq [[a] [eqa] [sa]].
+
+Ltac setoideq :=
+ match goal with
+ [ |- @eq ?A ?X ?Y ] => apply (setoideq_eq (a:=A) X Y)
+ end.
+
diff --git a/theories/Lists/List.v b/theories/Lists/List.v
index d0bf5bee9..d9f5d10a1 100644
--- a/theories/Lists/List.v
+++ b/theories/Lists/List.v
@@ -9,7 +9,7 @@
(*i $Id$ i*)
Require Import Le Gt Minus Min Bool.
-Require Import Setoid.
+Require Import Coq.Setoids.Setoid.
Set Implicit Arguments.
diff --git a/theories/Program/Equality.v b/theories/Program/Equality.v
index 96c1c8f4d..85ed18891 100644
--- a/theories/Program/Equality.v
+++ b/theories/Program/Equality.v
@@ -85,17 +85,6 @@ Ltac abstract_eq_hyp H' p :=
end
end.
-(** Try to abstract a proof of equality, if no proof of the same equality is present in the context. *)
-
-Ltac abstract_any_hyp H' p :=
- match type of p with
- ?X =>
- match goal with
- | [ H : X |- _ ] => fail 1
- | _ => set (H':=p) ; try (change p with H') ; clearbody H' ; simpl in H'
- end
- end.
-
(** Apply the tactic tac to proofs of equality appearing as coercion arguments.
Just redefine this tactic (using [Ltac on_coerce_proof tac ::=]) to handle custom coercion operators.
*)
@@ -180,3 +169,39 @@ Ltac clear_eqs := repeat clear_eq.
Ltac simplify_eqs :=
simpl ; simpl_eqs ; clear_eq_ctx ; clear_refl_eqs ;
try subst ; simpl ; repeat simpl_uip ; rewrite_refl_id.
+
+(** A tactic to remove trivial equality guards in hypotheses. *)
+
+Ltac simpl_IH_eq H :=
+ let tac H' := clear H ; rename H' into H in
+ let H' := fresh "H" in
+ match type of H with
+ | JMeq _ _ -> _ =>
+ assert (H' := H (JMeq_refl _)) ; tac H'
+ | _ = _ -> _ =>
+ assert (H' := H (refl_equal _)) ; tac H'
+ end.
+
+Ltac simpl_IH_eqs H := repeat simpl_IH_eq H.
+
+Ltac simpl_IHs_eqs :=
+ match goal with
+ | [ H : JMeq _ _ -> _ |- _ ] => simpl_IH_eqs H
+ | [ H : _ = _ -> _ |- _ ] => simpl_IH_eqs H
+ end.
+
+Require Import Coq.Program.Tactics.
+
+(** The following tactics allow to do induction on an already instantiated inductive predicate
+ by first generalizing it and adding the proper equalities to the context, in a maner similar to
+ the BasicElim tactic of "Elimination with a motive" by Conor McBride. *)
+
+Tactic Notation "dependent" "induction" ident(H) :=
+ generalize_eqs H ; clear H ; (intros until 1 || intros until H) ;
+ induction H ; intros ; subst* ; try discriminates ; try simpl_IHs_eqs.
+
+(** This tactic also generalizes the goal by the given variables before the induction. *)
+
+Tactic Notation "dependent" "induction" ident(H) "generalizing" ne_hyp_list(l) :=
+ generalize_eqs H ; clear H ; (intros until 1 || intros until H) ;
+ generalize l ; clear l ; induction H ; intros ; subst* ; try discriminates ; try simpl_IHs_eqs.
diff --git a/theories/Program/FunctionalExtensionality.v b/theories/Program/FunctionalExtensionality.v
index 22cb93d56..382252ce2 100644
--- a/theories/Program/FunctionalExtensionality.v
+++ b/theories/Program/FunctionalExtensionality.v
@@ -48,13 +48,13 @@ Ltac apply_ext :=
(** For a function defined with Program using a well-founded order. *)
-Lemma fix_sub_eq_ext :
+Program Lemma fix_sub_eq_ext :
forall (A : Set) (R : A -> A -> Prop) (Rwf : well_founded R)
(P : A -> Set)
- (F_sub : forall x : A, (forall {y : A | R y x}, P (`y)) -> P x),
+ (F_sub : forall x : A, (forall (y : A | R y x), P y) -> P x),
forall x : A,
Fix_sub A R Rwf P F_sub x =
- F_sub x (fun {y : A | R y x}=> Fix A R Rwf P F_sub (`y)).
+ F_sub x (fun (y : A | R y x) => Fix A R Rwf P F_sub y).
Proof.
intros ; apply Fix_eq ; auto.
intros.
@@ -65,12 +65,12 @@ Qed.
(** For a function defined with Program using a measure. *)
-Lemma fix_sub_measure_eq_ext :
+Program Lemma fix_sub_measure_eq_ext :
forall (A : Type) (f : A -> nat) (P : A -> Type)
- (F_sub : forall x : A, (forall {y : A | f y < f x}, P (`y)) -> P x),
+ (F_sub : forall x : A, (forall (y : A | f y < f x), P y) -> P x),
forall x : A,
Fix_measure_sub A f P F_sub x =
- F_sub x (fun {y : A | f y < f x}=> Fix_measure_sub A f P F_sub (`y)).
+ F_sub x (fun (y : A | f y < f x) => Fix_measure_sub A f P F_sub y).
Proof.
intros ; apply Fix_measure_eq ; auto.
intros.
diff --git a/theories/Program/Program.v b/theories/Program/Program.v
index 39c5b7734..fb172db84 100644
--- a/theories/Program/Program.v
+++ b/theories/Program/Program.v
@@ -1,3 +1,4 @@
Require Export Coq.Program.Utils.
Require Export Coq.Program.Wf.
Require Export Coq.Program.Equality.
+Require Export Coq.Program.Subset.
diff --git a/theories/Program/Subset.v b/theories/Program/Subset.v
new file mode 100644
index 000000000..54d830c89
--- /dev/null
+++ b/theories/Program/Subset.v
@@ -0,0 +1,141 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Coq.Program.Utils.
+Require Import Coq.Program.Equality.
+
+(** Tactics related to subsets and proof irrelevance. *)
+
+(** Simplify dependent equality using sigmas to equality of the codomains if possible. *)
+
+Ltac simpl_existT :=
+ match goal with
+ [ H : existT _ ?x _ = existT _ ?x _ |- _ ] =>
+ let Hi := fresh H in assert(Hi:=inj_pairT2 _ _ _ _ _ H) ; clear H
+ end.
+
+Ltac simpl_existTs := repeat simpl_existT.
+
+(** The following tactics implement a poor-man's solution for proof-irrelevance: it tries to
+ factorize every proof of the same proposition in a goal so that equality of such proofs becomes trivial. *)
+
+Ltac on_subset_proof_aux tac T :=
+ match T with
+ | context [ exist ?P _ ?p ] => try on_subset_proof_aux tac P ; tac p
+ end.
+
+Ltac on_subset_proof tac :=
+ match goal with
+ [ |- ?T ] => on_subset_proof_aux tac T
+ end.
+
+Ltac abstract_any_hyp H' p :=
+ match type of p with
+ ?X =>
+ match goal with
+ | [ H : X |- _ ] => fail 1
+ | _ => set (H':=p) ; try (change p with H') ; clearbody H'
+ end
+ end.
+
+Ltac abstract_subset_proof :=
+ on_subset_proof ltac:(fun p => let H := fresh "eqH" in abstract_any_hyp H p ; simpl in H).
+
+Ltac abstract_subset_proofs := repeat abstract_subset_proof.
+
+Ltac pi_subset_proof_hyp p :=
+ match type of p with
+ ?X =>
+ match goal with
+ | [ H : X |- _ ] =>
+ match p with
+ | H => fail 2
+ | _ => rewrite (proof_irrelevance X p H)
+ end
+ | _ => fail " No hypothesis with same type "
+ end
+ end.
+
+Ltac pi_subset_proof := on_subset_proof pi_subset_proof_hyp.
+
+Ltac pi_subset_proofs := repeat pi_subset_proof.
+
+(** Clear duplicated hypotheses *)
+
+Ltac clear_dup :=
+ match goal with
+ | [ H : ?X |- _ ] =>
+ match goal with
+ | [ H' : X |- _ ] =>
+ match H' with
+ | H => fail 2
+ | _ => clear H' || clear H
+ end
+ end
+ end.
+
+Ltac clear_dups := repeat clear_dup.
+
+(** The two preceding tactics in sequence. *)
+
+Ltac clear_subset_proofs :=
+ abstract_subset_proofs ; simpl in * |- ; pi_subset_proofs ; clear_dups.
+
+Ltac pi := repeat progress f_equal ; apply proof_irrelevance.
+
+Lemma subset_eq : forall A (P : A -> Prop) (n m : sig P), n = m <-> (`n) = (`m).
+Proof.
+ induction n.
+ induction m.
+ simpl.
+ split ; intros ; subst.
+
+ inversion H.
+ reflexivity.
+
+ pi.
+Qed.
+
+(* Somewhat trivial definition, but not unfolded automatically hence we can match on [match_eq ?A ?B ?x ?f]
+ in tactics. *)
+
+Program Definition match_eq (A B : Type) (x : A) (fn : forall (y : A | y = x), B) : B :=
+ fn (exist _ x (refl_equal x)).
+
+(* This is what we want to be able to do: replace the originaly matched object by a new,
+ propositionally equal one. If [fn] works on [x] it should work on any [y | y = x]. *)
+
+Lemma match_eq_rewrite : forall (A B : Type) (x : A) (fn : forall (y : A | y = x), B)
+ (y : A | y = x),
+ match_eq A B x fn = fn y.
+Proof.
+ intros.
+ unfold match_eq.
+ f_equal.
+ destruct y.
+ (* uses proof-irrelevance *)
+ apply <- subset_eq.
+ symmetry. assumption.
+Qed.
+
+(** Now we make a tactic to be able to rewrite a term [t] which is applied to a [match_eq] using an arbitrary
+ equality [t = u], and [u] is now the subject of the [match]. *)
+
+Ltac rewrite_match_eq H :=
+ match goal with
+ [ |- ?T ] =>
+ match T with
+ context [ match_eq ?A ?B ?t ?f ] =>
+ rewrite (match_eq_rewrite A B t f (exist _ _ (sym_eq H)))
+ end
+ end.
+
+(** Otherwise we can simply unfold [match_eq] and the term trivially reduces to the original definition. *)
+
+Ltac simpl_match_eq := unfold match_eq ; simpl.
+
diff --git a/theories/Program/Tactics.v b/theories/Program/Tactics.v
index e4c60e14a..ac0f5cf71 100644
--- a/theories/Program/Tactics.v
+++ b/theories/Program/Tactics.v
@@ -50,6 +50,7 @@ Tactic Notation "destruct" "exist" ident(t) ident(Ht) := destruct t as [t Ht].
Tactic Notation "destruct" "or" ident(H) := destruct H as [H|H].
(** Discriminate that also work on a [x <> x] hypothesis. *)
+
Ltac discriminates :=
match goal with
| [ H : ?x <> ?x |- _ ] => elim H ; reflexivity
@@ -151,21 +152,12 @@ Ltac bang :=
Tactic Notation "contradiction" "by" constr(t) :=
let H := fresh in assert t as H by auto with * ; contradiction.
-(** The following tactics allow to do induction on an already instantiated inductive predicate
- by first generalizing it and adding the proper equalities to the context, in a manner similar to
- the BasicElim tactic of "Elimination with a motive" by Conor McBride. *)
-
-Tactic Notation "dependent" "induction" ident(H) :=
- generalize_eqs H ; clear H ; (intros until 1 || intros until H) ;
- induction H ; intros ; subst* ; try discriminates.
-
-(** This tactic also generalizes the goal by the given variables before the induction. *)
-
-Tactic Notation "dependent" "induction" ident(H) "generalizing" ne_hyp_list(l) :=
- generalize_eqs H ; clear H ; (intros until 1 || intros until H) ;
- generalize l ; clear l ; induction H ; intros ; subst* ; try discriminates.
+(** The default simplification tactic used by Program is defined by [program_simpl], sometimes [auto with *]
+ is overkill and slows things down, better rebind using [Obligations Tactic := tac] in this case,
+ possibly using [program_simplify] to use standard goal-cleaning tactics. *)
-(** The default simplification tactic used by Program. *)
+Ltac program_simplify :=
+ simpl ; intros ; destruct_conjs ; simpl proj1_sig in * ; subst* ;
+ try (solve [ red ; intros ; destruct_conjs ; discriminate ]).
-Ltac program_simpl := simpl ; intros ; destruct_conjs ; simpl in * ; try subst ;
- try (solve [ red ; intros ; destruct_conjs ; discriminate ]) ; auto with *.
+Ltac program_simpl := program_simplify ; auto with *.
diff --git a/theories/Program/Utils.v b/theories/Program/Utils.v
index a268f0afb..6af81a410 100644
--- a/theories/Program/Utils.v
+++ b/theories/Program/Utils.v
@@ -18,9 +18,9 @@ Notation " {{ x }} " := (tt : { y : unit | x }).
(** A simpler notation for subsets defined on a cartesian product. *)
-Notation "{ ( x , y ) : A | P }" :=
- (sig (fun anonymous : A => let (x,y) := anonymous in P))
- (x ident, y ident) : type_scope.
+(* Notation "{ ( x , y ) : A | P }" := *)
+(* (sig (fun anonymous : A => let (x,y) := anonymous in P)) *)
+(* (x ident, y ident, at level 10) : type_scope. *)
(** Generates an obligation to prove False. *)
@@ -45,13 +45,13 @@ Notation " x '`=' y " := ((x :>) = (y :>)) (at level 70).
(** Quantifying over subsets. *)
-Notation "'fun' { x : A | P } => Q" :=
- (fun x:{x:A|P} => Q)
- (at level 200, x ident, right associativity).
+(* Notation "'fun' ( x : A | P ) => Q" := *)
+(* (fun (x :A|P} => Q) *)
+(* (at level 200, x ident, right associativity). *)
-Notation "'forall' { x : A | P } , Q" :=
- (forall x:{x:A|P}, Q)
- (at level 200, x ident, right associativity).
+(* Notation "'forall' ( x : A | P ), Q" := *)
+(* (forall (x : A | P), Q) *)
+(* (at level 200, x ident, right associativity). *)
Require Import Coq.Bool.Sumbool.
diff --git a/theories/Program/Wf.v b/theories/Program/Wf.v
index 55784671f..70b1b1b5a 100644
--- a/theories/Program/Wf.v
+++ b/theories/Program/Wf.v
@@ -34,11 +34,11 @@ Section Well_founded.
Hypothesis
F_ext :
forall (x:A) (f g:forall y:{y:A | R y x}, P (`y)),
- (forall y:{ y:A | R y x}, f y = g y) -> F_sub x f = F_sub x g.
+ (forall (y : A | R y x), f y = g y) -> F_sub x f = F_sub x g.
Lemma Fix_F_eq :
forall (x:A) (r:Acc R x),
- F_sub x (fun (y:{y:A|R y x}) => Fix_F (`y) (Acc_inv r (proj1_sig y) (proj2_sig y))) = Fix_F x r.
+ F_sub x (fun (y:A|R y x) => Fix_F (`y) (Acc_inv r (proj1_sig y) (proj2_sig y))) = Fix_F x r.
Proof.
destruct r using Acc_inv_dep; auto.
Qed.
@@ -52,7 +52,7 @@ Section Well_founded.
rewrite (proof_irrelevance (Acc R x) r s) ; auto.
Qed.
- Lemma Fix_eq : forall x:A, Fix x = F_sub x (fun (y:{y:A|R y x}) => Fix (proj1_sig y)).
+ Lemma Fix_eq : forall x:A, Fix x = F_sub x (fun (y:A|R y x) => Fix (proj1_sig y)).
Proof.
intro x; unfold Fix in |- *.
rewrite <- (Fix_F_eq ).
@@ -64,7 +64,7 @@ Section Well_founded.
forall x : A,
Fix_sub P F_sub x =
let f_sub := F_sub in
- f_sub x (fun {y : A | R y x}=> Fix (`y)).
+ f_sub x (fun (y : A | R y x) => Fix (`y)).
exact Fix_eq.
Qed.
@@ -98,7 +98,7 @@ Section Well_founded_measure.
Section FixPoint.
Variable P : A -> Type.
- Variable F_sub : forall x:A, (forall y: { y : A | m y < m x }, P (proj1_sig y)) -> P x.
+ Variable F_sub : forall x:A, (forall (y : A | m y < m x), P (proj1_sig y)) -> P x.
Notation Fix_F := (Fix_measure_F_sub P F_sub) (only parsing). (* alias *)
@@ -106,8 +106,8 @@ Section Well_founded_measure.
Hypothesis
F_ext :
- forall (x:A) (f g:forall y:{y:A | m y < m x}, P (`y)),
- (forall y:{ y:A | m y < m x}, f y = g y) -> F_sub x f = F_sub x g.
+ forall (x:A) (f g:forall y : { y : A | m y < m x}, P (`y)),
+ (forall y : { y : A | m y < m x}, f y = g y) -> F_sub x f = F_sub x g.
Lemma Fix_measure_F_eq :
forall (x:A) (r:Acc lt (m x)),
@@ -137,7 +137,7 @@ Section Well_founded_measure.
forall x : A,
Fix_measure_sub P F_sub x =
let f_sub := F_sub in
- f_sub x (fun {y : A | m y < m x}=> Fix_measure (`y)).
+ f_sub x (fun (y : A | m y < m x) => Fix_measure (`y)).
exact Fix_measure_eq.
Qed.
diff --git a/theories/Setoids/Setoid.v b/theories/Setoids/Setoid.v
index 18e8b6e77..f93a12955 100644
--- a/theories/Setoids/Setoid.v
+++ b/theories/Setoids/Setoid.v
@@ -1,4 +1,3 @@
-
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
diff --git a/theories/Setoids/Setoid_tac.v b/theories/Setoids/Setoid_tac.v
index f979ec11c..664316805 100644
--- a/theories/Setoids/Setoid_tac.v
+++ b/theories/Setoids/Setoid_tac.v
@@ -1,4 +1,3 @@
-
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
diff --git a/tools/coqdoc/output.ml b/tools/coqdoc/output.ml
index 1cbee6046..647cc5da3 100644
--- a/tools/coqdoc/output.ml
+++ b/tools/coqdoc/output.ml
@@ -46,9 +46,11 @@ let is_keyword =
"Delimit"; "Bind"; "Open"; "Scope";
"Boxed"; "Unboxed";
"Arguments";
+ "Instance"; "Class"; "where"; "Instantiation";
(* Program *)
"Program Definition"; "Program Fixpoint"; "Program Lemma";
"Obligation"; "Obligations"; "Solve"; "using"; "Next Obligation"; "Next";
+ "Program Instance";
(*i (* correctness *)
"array"; "assert"; "begin"; "do"; "done"; "else"; "end"; "if";
"in"; "invariant"; "let"; "of"; "ref"; "state"; "then"; "variant";
@@ -366,7 +368,7 @@ module Html = struct
module_ref m s
| Ref (m,fullid) ->
ident_ref m fullid s
- | Mod _ | Ref _ ->
+ | Mod _ ->
raw_ident s)
with Not_found ->
raw_ident s
diff --git a/tools/coqdoc/pretty.mll b/tools/coqdoc/pretty.mll
index 8c46150a2..2af5bda09 100644
--- a/tools/coqdoc/pretty.mll
+++ b/tools/coqdoc/pretty.mll
@@ -207,6 +207,7 @@ let thm_token =
let def_token =
"Definition"
| "Let"
+ | "Class"
| "SubClass"
| "Example"
| "Local"
@@ -214,11 +215,10 @@ let def_token =
| "CoFixpoint"
| "Record"
| "Structure"
+ | "Instance"
| "Scheme"
| "Inductive"
| "CoInductive"
- | "Program" space+ "Definition"
- | "Program" space+ "Fixpoint"
let decl_token =
"Hypothesis"
@@ -263,12 +263,17 @@ let commands =
| "Check"
| "Type"
+ | "Section"
+ | "Chapter"
+ | "Variable" 's'?
+ | ("Hypothesis" | "Hypotheses")
+ | "End"
+
let extraction =
"Extraction"
| "Recursive" space+ "Extraction"
| "Extract"
-
let gallina_kw = thm_token | def_token | decl_token | gallina_ext | commands | extraction
let prog_kw =
@@ -291,7 +296,6 @@ let gallina_kw_to_hide =
| "Transparent"
| "Opaque"
| ("Declare" space+ ("Morphism" | "Step") )
- | "Chapter"
| ("Set" | "Unset") space+ "Printing" space+ "Coercions"
| "Declare" space+ ("Left" | "Right") space+ "Step"
diff --git a/toplevel/cerrors.ml b/toplevel/cerrors.ml
index 5a1b8b8b4..bad30a849 100644
--- a/toplevel/cerrors.ml
+++ b/toplevel/cerrors.ml
@@ -77,6 +77,8 @@ let rec explain_exn_default_aux anomaly_string report_fn = function
hov 0 (str "Error:" ++ spc () ++ Himsg.explain_type_error ctx te)
| PretypeError(ctx,te) ->
hov 0 (str "Error:" ++ spc () ++ Himsg.explain_pretype_error ctx te)
+ | Typeclasses_errors.TypeClassError(env, te) ->
+ hov 0 (str "Error:" ++ spc () ++ Himsg.explain_typeclass_error env te)
| InductiveError e ->
hov 0 (str "Error:" ++ spc () ++ Himsg.explain_inductive_error e)
| RecursionSchemeError e ->
diff --git a/toplevel/classes.ml b/toplevel/classes.ml
new file mode 100644
index 000000000..7ee4513d1
--- /dev/null
+++ b/toplevel/classes.ml
@@ -0,0 +1,505 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.ml 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Termops
+open Sign
+open Entries
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Util
+open Typeclasses_errors
+open Typeclasses
+open Libnames
+open Constrintern
+open Rawterm
+open Topconstr
+(*i*)
+
+let mismatched_params env n m = mismatched_ctx_inst env Parameters n m
+(* let mismatched_defs env n m = mismatched_ctx_inst env Definitions n m *)
+let mismatched_props env n m = mismatched_ctx_inst env Properties n m
+
+type binder_list = (identifier located * constr_expr) list
+
+let interp_binders_evars isevars env avoid l =
+ List.fold_left
+ (fun (env, ids, params) ((loc, i), t) ->
+ let n = Name i in
+ let t' = interp_binder_evars isevars env n t in
+ let d = (i,None,t') in
+ (push_named d env, i :: ids, d::params))
+ (env, avoid, []) l
+
+let interp_typeclass_context_evars isevars env avoid l =
+ List.fold_left
+ (fun (env, ids, params) (iid, bk, l) ->
+ let t' = interp_binder_evars isevars env (snd iid) l in
+ let i = match snd iid with
+ | Anonymous -> Nameops.next_name_away (Termops.named_hd env t' Anonymous) ids
+ | Name id -> id
+ in
+ let d = (i,None,t') in
+ (push_named d env, i :: ids, d::params))
+ (env, avoid, []) l
+
+let interp_constrs_evars isevars env avoid l =
+ List.fold_left
+ (fun (env, ids, params) t ->
+ let t' = interp_binder_evars isevars env Anonymous t in
+ let id = Nameops.next_name_away (Termops.named_hd env t' Anonymous) ids in
+ let d = (id,None,t') in
+ (push_named d env, id :: ids, d::params))
+ (env, avoid, []) l
+
+let raw_assum_of_binders k =
+ List.map (fun ((loc,i),t) -> LocalRawAssum ([(loc, Name i)], k, t))
+
+let raw_assum_of_constrs k =
+ List.map2 (fun t (n,_,_) -> LocalRawAssum ([(dummy_loc, Name n)], k, t))
+
+let raw_assum_of_anonymous_constrs k =
+ List.map (fun t -> LocalRawAssum ([(dummy_loc, Anonymous)], k, t))
+
+let decl_expr_of_binders =
+ List.map (fun ((loc,i),t) -> false, Vernacexpr.AssumExpr ((loc, Name i), t))
+
+let rec unfold n f acc =
+ match n with
+ | 0 -> f 0 acc
+ | n -> unfold (n - 1) f (f n acc)
+
+(* Declare everything in the parameters as implicit, and the class instance as well *)
+open Topconstr
+
+let declare_implicit_proj c proj =
+ let len = List.length c.cl_context + List.length c.cl_super + List.length c.cl_params in
+ let (ctx, _) = decompose_prod_n (len + 1) (Typeops.type_of_constant (Global.env()) proj) in
+ let expls =
+ let rec aux i expls = function
+ [] -> expls
+ | (Name n, _) :: tl ->
+ let impl = ExplByPos (i, Some n), (true, true) in
+ aux (succ i) (impl :: List.remove_assoc (ExplByName n) expls) tl
+ | (Anonymous,_) :: _ -> assert(false)
+ in
+ aux 1 [] ctx
+ in Impargs.declare_manual_implicits true (ConstRef proj) true expls
+
+let declare_implicits cl =
+ let projs = Recordops.lookup_projections cl.cl_impl in
+ List.iter
+ (fun c ->
+ match c with
+ | None -> assert(false)
+ | Some p -> declare_implicit_proj cl p)
+ projs;
+ let indimps =
+ list_map_i
+ (fun i (na, b, t) -> ExplByPos (i, Some na), (false, true))
+ 1 (List.rev cl.cl_context)
+ in
+ Impargs.declare_manual_implicits true (IndRef cl.cl_impl) false indimps
+
+let rel_of_named_context subst l =
+ List.fold_right
+ (fun (id, _, t) (ctx, acc) -> (Name id, None, subst_vars acc t) :: ctx, id :: acc)
+ l ([], subst)
+
+let degenerate_decl (na,b,t) =
+ let id = match na with
+ | Name id -> id
+ | Anonymous -> anomaly "Unnamed record variable" in
+ match b with
+ | None -> (id, Entries.LocalAssum t)
+ | Some b -> (id, Entries.LocalDef b)
+
+
+let declare_structure env id idbuild params arity fields =
+ let nparams = List.length params and nfields = List.length fields in
+ let args = extended_rel_list nfields params in
+ let ind = applist (mkRel (1+nparams+nfields), args) in
+ let type_constructor = it_mkProd_or_LetIn ind fields in
+ let mie_ind =
+ { mind_entry_typename = id;
+ mind_entry_arity = arity;
+ mind_entry_consnames = [idbuild];
+ mind_entry_lc = [type_constructor] } in
+ let mie =
+ { mind_entry_params = List.map degenerate_decl params;
+ mind_entry_record = true;
+ mind_entry_finite = true;
+ mind_entry_inds = [mie_ind] } in
+ let kn = Command.declare_mutual_with_eliminations true mie in
+ let rsp = (kn,0) in (* This is ind path of idstruc *)
+ let kinds,sp_projs = Record.declare_projections rsp (List.map (fun _ -> false) fields) fields in
+ let _build = ConstructRef (rsp,1) in
+ Recordops.declare_structure(rsp,idbuild,List.rev kinds,List.rev sp_projs);
+ rsp
+
+
+let mk_interning_data env na typ =
+ let impl =
+ if Impargs.is_implicit_args() then
+ Impargs.compute_implicits env typ
+ else []
+ in (na, ([], impl, Notation.compute_arguments_scope typ))
+
+let interp_fields_evars isevars env avoid l =
+ List.fold_left
+ (fun (env, ids, params, impls) ((loc, i), t) ->
+ let t' = interp_type_evars isevars env ~impls t in
+ let data = mk_interning_data env i t' in
+ let d = (i,None,t') in
+ (push_named d env, i :: ids, d::params, ([], data :: snd impls)))
+ (env, avoid, [], ([], [])) l
+
+let new_class id par ar sup props =
+ let env0 = Global.env() in
+ let isevars = ref (Evd.create_evar_defs Evd.empty) in
+ let avoid = Termops.ids_of_context env0 in
+
+ let env_params, avoid = env0, avoid in
+
+ (* Find the implicitly quantified variables *)
+ let gen_ctx, super = Implicit_quantifiers.resolve_class_binders (vars_of_env env0) sup in
+
+ let env_super_ctx, avoid, ctx_super_ctx = interp_binders_evars isevars env_params avoid gen_ctx in
+
+ (* Interpret the superclasses constraints *)
+ let env_super, avoid, ctx_super0 =
+ interp_typeclass_context_evars isevars env_super_ctx avoid super
+ in
+
+ let env_params, avoid, ctx_params = interp_binders_evars isevars env_super avoid par in
+
+ (* Interpret the arity *)
+ let arity = interp_type_evars isevars env_params (CSort (fst ar, snd ar)) in
+
+ (* let fullarity = it_mkProd_or_LetIn (it_mkProd_or_LetIn arity ctx_defs) ctx_params in*)
+
+ (* Interpret the definitions and propositions *)
+ let env_props, avoid, ctx_props, _ =
+ interp_fields_evars isevars env_params avoid props
+ in
+
+ (* Instantiate evars and check all are resolved *)
+ let isevars,_ = Evarconv.consider_remaining_unif_problems env_props !isevars in
+ let sigma = Evd.evars_of isevars in
+ let ctx_super_ctx = Implicit_quantifiers.nf_named_context sigma ctx_super_ctx in
+ let ctx_super = Implicit_quantifiers.nf_named_context sigma ctx_super0 in
+ let ctx_params = Implicit_quantifiers.nf_named_context sigma ctx_params in
+ let ctx_props = Implicit_quantifiers.nf_named_context sigma ctx_props in
+ let arity = Reductionops.nf_evar sigma arity in
+ let ce t = Evarutil.check_evars env0 Evd.empty isevars t in
+ let kn =
+ let idb = id_of_string ("Build_" ^ (string_of_id (snd id))) in
+ let params, subst = rel_of_named_context [] (ctx_params @ ctx_super @ ctx_super_ctx) in
+ let fields, _ = rel_of_named_context subst ctx_props in
+ List.iter (fun (_,c,t) -> ce t; match c with Some c -> ce c | None -> ()) (params @ fields);
+ declare_structure env0 (snd id) idb params arity fields
+ in
+ let ctx_context, ctx_super =
+ let class_binders, regular_binders =
+ List.partition (fun (na, b, t) ->
+ Typeclasses.class_of_constr t <> None)
+ ctx_super_ctx
+ in
+ if (ctx_super_ctx = class_binders @ regular_binders) then
+ regular_binders, ctx_super @ class_binders
+ else ctx_super_ctx, ctx_super
+ in
+ let k =
+ { cl_name = snd id;
+ cl_context = ctx_context;
+ cl_super = ctx_super;
+ cl_params = ctx_params;
+ cl_props = ctx_props;
+ cl_impl = kn }
+ in
+ declare_implicits k;
+ add_class k
+
+open Decl_kinds
+open Entries
+
+let hint_db = "typeclass_instances"
+
+let add_instance_hint inst =
+ Auto.add_hints false [hint_db] (Vernacexpr.HintsResolve [CRef (Ident (dummy_loc, inst))])
+
+let declare_instance (_,id) =
+ add_instance_hint id
+
+type binder_def_list = (identifier located * identifier located list * constr_expr) list
+
+let binders_of_lidents l =
+ List.map (fun (loc, id) -> LocalRawAssum ([loc, Name id], Default Rawterm.Implicit, CHole loc)) l
+
+let subst_ids_in_named_context subst l =
+ let x, _ =
+ List.fold_right
+ (fun (id, _, t) (ctx, k) -> (id, None, substn_vars k subst t) :: ctx, succ k)
+ l ([], 1)
+ in x
+
+let subst_one_named inst ids t =
+ substnl inst 0 (substn_vars 1 ids t)
+
+let subst_named inst subst ctx =
+ let ids = List.map (fun (id, _, _) -> id) subst in
+ let ctx' = subst_ids_in_named_context ids ctx in
+ let ctx', _ =
+ List.fold_right
+ (fun (id, _, t) (ctx, k) -> (id, None, substnl inst k t) :: ctx, succ k)
+ ctx' ([], 0)
+ in ctx'
+
+let push_named_context = List.fold_right push_named
+
+let destClass c =
+ match kind_of_term c with
+ App (c, args) ->
+ (match kind_of_term c with
+ | Ind ind -> (try class_of_inductive ind, args with _ -> assert false)
+ | _ -> assert false)
+ | _ -> assert false
+
+let infer_super_instances env params params_ctx super =
+ let super = subst_named params params_ctx super in
+ List.fold_right
+ (fun (na, _, t) (sups, ids, supctx) ->
+ let t = subst_one_named sups ids t in
+ let inst =
+ try resolve_one_typeclass env t
+ with Not_found ->
+ let cl, args = destClass t in
+ no_instance (push_named_context supctx env) (dummy_loc, cl.cl_name) (Array.to_list args)
+ in
+ let d = (na, Some inst, t) in
+ inst :: sups, na :: ids, d :: supctx)
+ super ([], [], [])
+
+(* let evars_of_context ctx id n env = *)
+(* List.fold_right (fun (na, _, t) (n, env, nc) -> *)
+(* let b = Evarutil.e_new_evar isevars env ~src:(dummy_loc, ImplicitArg (Ident id, (n * Some na))) t in *)
+(* let d = (na, Some b, t) in *)
+(* (succ n, push_named d env, d :: nc)) *)
+(* ctx (n, env, []) *)
+
+let type_ctx_instance isevars env ctx inst subst =
+ List.fold_left2
+ (fun (subst, instctx) (na, _, t) ce ->
+ let t' = replace_vars subst t in
+ let c = interp_casted_constr_evars isevars env ce t' in
+ let d = na, Some c, t' in
+ (na, c) :: subst, d :: instctx)
+ (subst, []) (List.rev ctx) inst
+
+let substitution_of_constrs ctx cstrs =
+ List.fold_right2 (fun c (na, _, _) acc -> (na, c) :: acc) cstrs ctx []
+
+let destClassApp cl =
+ match cl with
+ | CApp (loc, (None,CRef (Ident f)), l) -> f, List.map fst l
+ | _ -> raise Not_found
+
+let new_instance sup (instid, bk, cl) props =
+ let id, par = destClassApp cl in
+ let env = Global.env() in
+ let isevars = ref (Evd.create_evar_defs Evd.empty) in
+ let avoid = Termops.ids_of_context env in
+ let k =
+ try class_info (snd id)
+ with Not_found -> unbound_class env id
+ in
+ let gen_ctx, sup = Implicit_quantifiers.resolve_class_binders (vars_of_env env) sup in
+ let env', avoid, genctx = interp_binders_evars isevars env avoid gen_ctx in
+ let env', avoid, supctx = interp_typeclass_context_evars isevars env' avoid sup in
+ let subst =
+ match bk with
+ Explicit ->
+ if List.length par <> List.length k.cl_context + List.length k.cl_params then
+ mismatched_params env par (k.cl_params @ k.cl_context);
+ let len = List.length k.cl_context in
+ let ctx, par = Util.list_chop len par in
+ let subst, _ = type_ctx_instance isevars env' k.cl_context ctx [] in
+ let subst =
+ Typeclasses.substitution_of_named_context isevars env' k.cl_name len subst
+ k.cl_super
+ in
+ if List.length par <> List.length k.cl_params then
+ mismatched_params env par k.cl_params;
+ let subst, par = type_ctx_instance isevars env' k.cl_params par subst in subst
+
+ | Implicit ->
+ let t' = interp_type_evars isevars env (Topconstr.mkAppC (CRef (Ident id), par)) in
+ match kind_of_term t' with
+ App (c, args) ->
+ substitution_of_constrs (k.cl_params @ k.cl_super @ k.cl_context)
+ (List.rev (Array.to_list args))
+ | _ -> assert false
+ in
+ isevars := Evarutil.nf_evar_defs !isevars;
+ let sigma = Evd.evars_of !isevars in
+ let env' = Implicit_quantifiers.nf_env sigma env' in
+ let subst = Typeclasses.nf_substitution sigma subst in
+ let subst, propsctx =
+ let props =
+ List.map (fun (x, l, d) ->
+ Topconstr.abstract_constr_expr d (binders_of_lidents l))
+ props
+ in
+ if List.length props <> List.length k.cl_props then
+ mismatched_props env' props k.cl_props;
+ type_ctx_instance isevars env' k.cl_props props subst
+ in
+ let app =
+ applistc (mkConstruct (k.cl_impl, 1)) (List.rev_map snd subst)
+ in
+ let term = Termops.it_mkNamedLambda_or_LetIn (Termops.it_mkNamedLambda_or_LetIn app supctx) genctx in
+ isevars := Evarutil.nf_evar_defs !isevars;
+ let term = Evarutil.nf_isevar !isevars term in
+ let cdecl =
+ let kind = IsDefinition Instance in
+ let entry =
+ { const_entry_body = term;
+ const_entry_type = None;
+ const_entry_opaque = false;
+ const_entry_boxed = false }
+ in DefinitionEntry entry, kind
+ in
+ let id, cst =
+ let instid =
+ match snd instid with
+ Name id -> id
+ | Anonymous ->
+ let i = Nameops.add_suffix (snd id) "_instance_" in
+ Termops.next_global_ident_away false i (Termops.ids_of_context env)
+ in
+ instid, Declare.declare_constant instid cdecl
+ in
+ let inst =
+ { is_class = k;
+ is_name = id;
+(* is_params = paramsctx; (\* missing gen_ctx *\) *)
+(* is_super = superctx; *)
+ is_impl = cst;
+(* is_add_hint = (fun () -> add_instance_hint id); *)
+ }
+ in
+ add_instance_hint id;
+ add_instance inst
+
+let goal_kind = Decl_kinds.Global, Decl_kinds.DefinitionBody Decl_kinds.Definition
+
+let solve_by_tac env evd evar evi t =
+ let goal = {it = evi; sigma = (Evd.evars_of evd) } in
+ let (res, valid) = t goal in
+ if res.it = [] then
+ let prooftree = valid [] in
+ let proofterm, obls = Refiner.extract_open_proof res.sigma prooftree in
+ if obls = [] then
+ let evd' = evars_reset_evd res.sigma evd in
+ let evd' = evar_define evar proofterm evd' in
+ evd', true
+ else evd, false
+ else evd, false
+
+let context l =
+ let env = Global.env() in
+ let isevars = ref (Evd.create_evar_defs Evd.empty) in
+ let avoid = Termops.ids_of_context env in
+ let ctx, l = Implicit_quantifiers.resolve_class_binders (vars_of_env env) l in
+ let env', avoid, ctx = interp_binders_evars isevars env avoid ctx in
+ let env', avoid, l = interp_typeclass_context_evars isevars env' avoid l in
+ isevars := Evarutil.nf_evar_defs !isevars;
+ let sigma = Evd.evars_of !isevars in
+ let fullctx = Implicit_quantifiers.nf_named_context sigma (l @ ctx) in
+ List.iter (function (id,_,t) ->
+ Command.declare_one_assumption false (Local (* global *), Definitional) t false (* inline *) (dummy_loc, id))
+ (List.rev fullctx)
+
+(* let init () = hints := [] *)
+(* let freeze () = !hints *)
+(* let unfreeze fs = hints := fs *)
+
+(* let _ = Summary.declare_summary "hints db" *)
+(* { Summary.freeze_function = freeze; *)
+(* Summary.unfreeze_function = unfreeze; *)
+(* Summary.init_function = init; *)
+(* Summary.survive_module = false; *)
+(* Summary.survive_section = true } *)
+
+open Libobject
+
+(* let cache (_, db) := hints := db *)
+
+(* let (input,output) = *)
+(* declare_object *)
+(* { (default_object "hints db") with *)
+(* cache_function = cache; *)
+(* load_function = (fun _ -> cache); *)
+(* open_function = (fun _ -> cache); *)
+(* classify_function = (fun (_,x) -> Keep x); *)
+(* export_function = (fun x -> Some x) } *)
+
+let tactic = ref Tacticals.tclIDTAC
+let tactic_expr = ref (Obj.magic ())
+
+let set_instantiation_tactic t =
+ tactic := Tacinterp.interp t; tactic_expr := t
+
+let freeze () = !tactic_expr
+let unfreeze t = set_instantiation_tactic t
+let init () =
+ set_instantiation_tactic (Tacexpr.TacId[])
+
+let cache (_, tac) =
+ set_instantiation_tactic tac
+
+let _ =
+ Summary.declare_summary "typeclasses instantiation tactic"
+ { Summary.freeze_function = freeze;
+ Summary.unfreeze_function = unfreeze;
+ Summary.init_function = init;
+ Summary.survive_module = true;
+ Summary.survive_section = true }
+
+let (input,output) =
+ declare_object
+ { (default_object "type classes instantiation tactic state") with
+ cache_function = cache;
+ load_function = (fun _ -> cache);
+ open_function = (fun _ -> cache);
+ classify_function = (fun (_,x) -> Keep x);
+ export_function = (fun x -> Some x) }
+
+let set_instantiation_tactic t =
+ Lib.add_anonymous_leaf (input t)
+
+
+let initialize () =
+ if !tactic_expr = Tacexpr.TacId [] then
+ let qualid = (dummy_loc, qualid_of_dirpath (dirpath_of_string "Coq.Classes.Init")) in
+ Library.require_library [qualid] None
+
+let _ =
+ Typeclasses.solve_instanciation_problem :=
+ (fun env -> initialize ();
+ fun evd ev evi ->
+ solve_by_tac env evd ev evi !tactic)
+
+
diff --git a/toplevel/classes.mli b/toplevel/classes.mli
new file mode 100644
index 000000000..102f1b8b0
--- /dev/null
+++ b/toplevel/classes.mli
@@ -0,0 +1,65 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: classes.mli 6748 2005-02-18 22:17:50Z herbelin $ i*)
+
+(*i*)
+open Names
+open Decl_kinds
+open Term
+open Sign
+open Evd
+open Environ
+open Nametab
+open Mod_subst
+open Topconstr
+open Util
+open Typeclasses
+open Implicit_quantifiers
+(*i*)
+
+type binder_list = (identifier located * constr_expr) list
+type binder_def_list = (identifier located * identifier located list * constr_expr) list
+
+val binders_of_lidents : identifier located list -> local_binder list
+
+val declare_implicit_proj : typeclass -> constant -> unit
+
+val infer_super_instances : env -> constr list ->
+ named_context -> named_context -> types list * identifier list * named_context
+
+val new_class : identifier located ->
+ binder_list ->
+ Vernacexpr.sort_expr located ->
+ typeclass_context ->
+ binder_list -> unit
+
+val new_instance :
+ typeclass_context ->
+ typeclass_constraint ->
+ binder_def_list ->
+ unit
+
+val context : typeclass_context -> unit
+
+val add_instance_hint : identifier -> unit
+
+val declare_instance : identifier located -> unit
+
+val set_instantiation_tactic : Tacexpr.raw_tactic_expr -> unit
+
+val mismatched_params : env -> constr_expr list -> named_context -> 'a
+
+val mismatched_props : env -> constr_expr list -> named_context -> 'a
+
+val solve_by_tac : env ->
+ Evd.evar_defs ->
+ Evd.evar ->
+ evar_info ->
+ Proof_type.tactic ->
+ Evd.evar_defs * bool
diff --git a/toplevel/command.ml b/toplevel/command.ml
index 05ee50daf..445555251 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -46,21 +46,21 @@ open Goptions
open Mod_subst
open Evd
-let mkLambdaCit = List.fold_right (fun (x,a) b -> mkLambdaC(x,a,b))
-let mkProdCit = List.fold_right (fun (x,a) b -> mkProdC(x,a,b))
+let mkLambdaCit = List.fold_right (fun (x,a) b -> mkLambdaC(x,default_binder_kind,a,b))
+let mkProdCit = List.fold_right (fun (x,a) b -> mkProdC(x,default_binder_kind,a,b))
let rec abstract_constr_expr c = function
| [] -> c
| LocalRawDef (x,b)::bl -> mkLetInC(x,b,abstract_constr_expr c bl)
- | LocalRawAssum (idl,t)::bl ->
- List.fold_right (fun x b -> mkLambdaC([x],t,b)) idl
+ | LocalRawAssum (idl,k,t)::bl ->
+ List.fold_right (fun x b -> mkLambdaC([x],k,t,b)) idl
(abstract_constr_expr c bl)
let rec generalize_constr_expr c = function
| [] -> c
| LocalRawDef (x,b)::bl -> mkLetInC(x,b,generalize_constr_expr c bl)
- | LocalRawAssum (idl,t)::bl ->
- List.fold_right (fun x b -> mkProdC([x],t,b)) idl
+ | LocalRawAssum (idl,k,t)::bl ->
+ List.fold_right (fun x b -> mkProdC([x],k,t,b)) idl
(generalize_constr_expr c bl)
let rec under_binders env f n c =
@@ -103,10 +103,13 @@ let definition_message id =
let constant_entry_of_com (bl,com,comtypopt,opacity,boxed) =
let sigma = Evd.empty in
let env = Global.env() in
- match comtypopt with
+ match comtypopt with
None ->
let b = abstract_constr_expr com bl in
- let j = interp_constr_judgment sigma env b in
+ let ib = intern_constr sigma env b in
+ let imps = Implicit_quantifiers.implicits_of_rawterm ib in
+ let j = Default.understand_judgment sigma env ib in
+ imps,
{ const_entry_body = j.uj_val;
const_entry_type = None;
const_entry_opaque = opacity;
@@ -115,7 +118,10 @@ let constant_entry_of_com (bl,com,comtypopt,opacity,boxed) =
(* We use a cast to avoid troubles with evars in comtyp *)
(* that can only be resolved knowing com *)
let b = abstract_constr_expr (mkCastC (com, Rawterm.CastConv (DEFAULTcast,comtyp))) bl in
- let (body,typ) = destSubCast (interp_constr sigma env b) in
+ let ib = intern_gen false sigma env b in
+ let imps = Implicit_quantifiers.implicits_of_rawterm ib in
+ let (body,typ) = destSubCast (Default.understand_gen (OfType None) sigma env ib) in
+ imps,
{ const_entry_body = body;
const_entry_type = Some typ;
const_entry_opaque = opacity;
@@ -130,15 +136,18 @@ let red_constant_entry bl ce = function
(local_binders_length bl)
body }
-let declare_global_definition ident ce local =
+let declare_global_definition ident ce local imps =
let kn = declare_constant ident (DefinitionEntry ce,IsDefinition Definition) in
- if local = Local && Flags.is_verbose() then
- msg_warning (pr_id ident ++ str" is declared as a global definition");
- definition_message ident;
- ConstRef kn
+ let gr = ConstRef kn in
+ if Impargs.is_implicit_args () || imps <> [] then
+ declare_manual_implicits false gr (Impargs.is_implicit_args ()) imps;
+ if local = Local && Flags.is_verbose() then
+ msg_warning (pr_id ident ++ str" is declared as a global definition");
+ definition_message ident;
+ gr
let declare_definition ident (local,boxed,dok) bl red_option c typopt hook =
- let ce = constant_entry_of_com (bl,c,typopt,false,boxed) in
+ let imps, ce = constant_entry_of_com (bl,c,typopt,false,boxed) in
let ce' = red_constant_entry bl ce red_option in
let r = match local with
| Local when Lib.sections_are_opened () ->
@@ -152,7 +161,7 @@ let declare_definition ident (local,boxed,dok) bl red_option c typopt hook =
str" is not visible from current goals");
VarRef ident
| (Global|Local) ->
- declare_global_definition ident ce' local in
+ declare_global_definition ident ce' local imps in
hook local r
let syntax_definition ident c local onlyparse =
@@ -550,8 +559,8 @@ let eq_constr_expr c1 c2 =
(* Very syntactical equality *)
let eq_local_binder d1 d2 = match d1,d2 with
- | LocalRawAssum (nal1,c1), LocalRawAssum (nal2,c2) ->
- List.length nal1 = List.length nal2 &&
+ | LocalRawAssum (nal1,k1,c1), LocalRawAssum (nal2,k2,c2) ->
+ List.length nal1 = List.length nal2 && k1 = k2 &&
List.for_all2 (fun (_,na1) (_,na2) -> na1 = na2) nal1 nal2 &&
eq_constr_expr c1 c2
| LocalRawDef ((_,id1),c1), LocalRawDef ((_,id2),c2) ->
@@ -742,7 +751,7 @@ let interp_fix_body evdref env_rec impls (_,ctx) fix ccl =
let build_fix_type (_,ctx) ccl = it_mkProd_or_LetIn ccl ctx
-let declare_fix boxed kind f def t =
+let declare_fix boxed kind f def t imps =
let ce = {
const_entry_body = def;
const_entry_type = Some t;
@@ -750,7 +759,11 @@ let declare_fix boxed kind f def t =
const_entry_boxed = boxed
} in
let kn = declare_constant f (DefinitionEntry ce,IsDefinition kind) in
- ConstRef kn
+ let gr = ConstRef kn in
+ if Impargs.is_implicit_args () || imps <> [] then
+ declare_manual_implicits false gr (Impargs.is_implicit_args ()) imps;
+ gr
+
let prepare_recursive_declaration fixnames fixtypes fixdefs =
let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
@@ -766,6 +779,7 @@ let compute_possible_guardness_evidences (n,_) fixl fixtype =
but doing it properly involves delta-reduction, and it finally
doesn't seem to worth the effort (except for huge mutual
fixpoints ?) *)
+ (* FIXME, local_binders_length does not give the size of the final product if typeclasses are used *)
let m = local_binders_length fixl.fix_binders in
let ctx = fst (Sign.decompose_prod_n_assum m fixtype) in
list_map_i (fun i _ -> i) 0 ctx
@@ -778,6 +792,11 @@ let interp_recursive fixkind l boxed =
(* Interp arities allowing for unresolved types *)
let evdref = ref (Evd.create_evar_defs Evd.empty) in
+ let fiximps =
+ List.map
+ (fun x -> Implicit_quantifiers.implicits_of_binders x.fix_binders)
+ fixl
+ in
let fixctxs = List.map (interp_fix_context evdref env) fixl in
let fixccls = List.map2 (interp_fix_ccl evdref) fixctxs fixl in
let fixtypes = List.map2 build_fix_type fixctxs fixccls in
@@ -816,7 +835,7 @@ let interp_recursive fixkind l boxed =
in
(* Declare the recursive definitions *)
- ignore (list_map3 (declare_fix boxed kind) fixnames fixdecls fixtypes);
+ ignore (list_map4 (declare_fix boxed kind) fixnames fixdecls fixtypes fiximps);
if_verbose ppnl (recursive_message kind indexes fixnames);
(* Declare notations *)
diff --git a/toplevel/command.mli b/toplevel/command.mli
index 3f7e50285..6b15479d7 100644
--- a/toplevel/command.mli
+++ b/toplevel/command.mli
@@ -32,7 +32,7 @@ open Redexpr
val declare_definition : identifier -> definition_kind ->
local_binder list -> red_expr option -> constr_expr ->
- constr_expr option -> declaration_hook -> unit
+ constr_expr option -> declaration_hook -> unit
val syntax_definition : identifier -> constr_expr -> bool -> bool -> unit
@@ -43,6 +43,16 @@ val declare_assumption : identifier located list ->
coercion_flag -> assumption_kind -> local_binder list -> constr_expr -> bool
->unit
+val compute_interning_datas : Environ.env ->
+ 'a list ->
+ 'b list ->
+ Term.types list ->
+ 'a list *
+ ('b *
+ (Names.identifier list * Impargs.implicits_list *
+ Topconstr.scope_name option list))
+ list
+
val build_mutual : (inductive_expr * decl_notation) list -> bool -> unit
val declare_mutual_with_eliminations :
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index 8d9eabd7e..95ffeb44a 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -71,6 +71,7 @@ let hm2 s =
(* The list of all theories in the standard library /!\ order does matter *)
let theories_dirs_map = [
"theories/Unicode", "Unicode" ;
+ "theories/Classes", "Classes" ;
"theories/Program", "Program" ;
"theories/FSets", "FSets" ;
"theories/IntMap", "IntMap" ;
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index e381fe753..128091f44 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -21,6 +21,7 @@ open Sign
open Environ
open Pretype_errors
open Type_errors
+open Typeclasses_errors
open Indrec
open Reduction
open Cases
@@ -441,6 +442,38 @@ let explain_pretype_error env err =
| NoOccurrenceFound c -> explain_no_occurrence_found env c
| CannotUnifyBindingType (m,n) -> explain_cannot_unify_binding_type env m n
+
+(* Typeclass errors *)
+
+let explain_unbound_class env (_,id) =
+ str "Unbound class name " ++ Nameops.pr_id id
+
+let explain_unbound_method env cid id =
+ str "Unbound method name " ++ Nameops.pr_id (snd id) ++ spc () ++ str"of class" ++ spc () ++
+ Nameops.pr_id cid
+
+let pr_constr_exprs exprs =
+ hv 0 (List.fold_right
+ (fun d pps -> ws 2 ++ Ppconstr.pr_constr_expr d ++ pps)
+ exprs (mt ()))
+
+let explain_no_instance env (_,id) l =
+ str "No instance found for class " ++ Nameops.pr_id id ++ spc () ++
+ str "applied to arguments" ++ spc () ++
+ prlist_with_sep pr_spc (pr_lconstr_env env) l
+
+let explain_mismatched_contexts env c i j =
+ str"Mismatched contexts while declaring instance: " ++ brk (1,1) ++
+ hov 1 (str"Expected:" ++ brk (1, 1) ++ pr_named_context env j) ++ fnl () ++ brk (1,1) ++
+ hov 1 (str"Found:" ++ brk (1, 1) ++ pr_constr_exprs i)
+
+let explain_typeclass_error env err =
+ match err with
+ | UnboundClass id -> explain_unbound_class env id
+ | UnboundMethod (cid, id) -> explain_unbound_method env cid id
+ | NoInstance (id, l) -> explain_no_instance env id l
+ | MismatchedContextInstance (c, i, j) -> explain_mismatched_contexts env c i j
+
(* Refiner errors *)
let explain_refiner_bad_type arg ty conclty =
diff --git a/toplevel/himsg.mli b/toplevel/himsg.mli
index 05a52b8b7..d7a72bede 100644
--- a/toplevel/himsg.mli
+++ b/toplevel/himsg.mli
@@ -15,6 +15,7 @@ open Indtypes
open Environ
open Type_errors
open Pretype_errors
+open Typeclasses_errors
open Indrec
open Cases
open Logic
@@ -28,6 +29,8 @@ val explain_pretype_error : env -> pretype_error -> std_ppcmds
val explain_inductive_error : inductive_error -> std_ppcmds
+val explain_typeclass_error : env -> typeclass_error -> Pp.std_ppcmds
+
val explain_recursion_scheme_error : recursion_scheme_error -> std_ppcmds
val explain_refiner_error : refiner_error -> std_ppcmds
diff --git a/toplevel/record.ml b/toplevel/record.ml
index e80001d25..73e4e42f7 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -228,4 +228,5 @@ let definition_structure ((is_coe,(_,idstruc)),ps,cfs,idbuild,s) =
let kinds,sp_projs = declare_projections rsp coers fields in
let build = ConstructRef (rsp,1) in (* This is construct path of idbuild *)
if is_coe then Class.try_add_new_coercion build Global;
- Recordops.declare_structure(rsp,idbuild,List.rev kinds,List.rev sp_projs)
+ Recordops.declare_structure(rsp,idbuild,List.rev kinds,List.rev sp_projs);
+ kn
diff --git a/toplevel/record.mli b/toplevel/record.mli
index 5074719bc..e322dcfd0 100644
--- a/toplevel/record.mli
+++ b/toplevel/record.mli
@@ -25,4 +25,4 @@ val declare_projections :
val definition_structure :
lident with_coercion * local_binder list *
- (local_decl_expr with_coercion) list * identifier * sorts -> unit
+ (local_decl_expr with_coercion) list * identifier * sorts -> kernel_name
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index b71f9767e..0aea24d47 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -486,7 +486,7 @@ let vernac_record struc binders sort nameopt cfs =
| Sort s -> s
| _ -> user_err_loc
(constr_loc sort,"definition_structure", str "Sort expected") in
- Record.definition_structure (struc,binders,cfs,const,s)
+ ignore(Record.definition_structure (struc,binders,cfs,const,s))
(* Sections *)
@@ -522,6 +522,22 @@ let vernac_identity_coercion stre id qids qidt =
let source = cl_of_qualid qids in
Class.try_add_new_identity_coercion id stre source target
+(* Type classes *)
+let vernac_class id par ar sup props =
+ Classes.new_class id par ar sup props
+
+let vernac_instance sup inst props =
+ Classes.new_instance sup inst props
+
+let vernac_context l =
+ Classes.context l
+
+let vernac_declare_instance id =
+ Classes.declare_instance id
+
+(* Default tactics for solving evars management. *)
+let vernac_set_instantiation_tac tac =
+ Classes.set_instantiation_tactic tac
(***********)
(* Solving *)
@@ -656,9 +672,10 @@ let vernac_hints = Auto.add_hints
let vernac_syntactic_definition = Command.syntax_definition
-let vernac_declare_implicits local r = function
+let vernac_declare_implicits local r e = function
| Some imps ->
- Impargs.declare_manual_implicits local (global_with_alias r) imps
+ Impargs.declare_manual_implicits local (global_with_alias r) e
+ (List.map (fun (ex,b,f) -> ex, (b,f)) imps)
| None ->
Impargs.declare_implicits local (global_with_alias r)
@@ -712,6 +729,14 @@ let _ =
optread = Impargs.is_contextual_implicit_args;
optwrite = Impargs.make_contextual_implicit_args }
+(* let _ = *)
+(* declare_bool_option *)
+(* { optsync = true; *)
+(* optname = "forceable implicit arguments"; *)
+(* optkey = (SecondaryTable ("Forceable","Implicit")); *)
+(* optread = Impargs.is_forceable_implicit_args; *)
+(* optwrite = Impargs.make_forceable_implicit_args } *)
+
let _ =
declare_bool_option
{ optsync = true;
@@ -945,6 +970,8 @@ let vernac_print = function
| PrintOpaqueName qid -> msg (print_opaque_name qid)
| PrintGraph -> ppnl (Prettyp.print_graph())
| PrintClasses -> ppnl (Prettyp.print_classes())
+ | PrintTypeClasses -> ppnl (Prettyp.print_typeclasses())
+ | PrintInstances c -> ppnl (Prettyp.print_instances c)
| PrintLtac qid -> ppnl (Tacinterp.print_ltac (snd (qualid_of_reference qid)))
| PrintCoercions -> ppnl (Prettyp.print_coercions())
| PrintCoercionPaths (cls,clt) ->
@@ -1188,6 +1215,15 @@ let interp c = match c with
| VernacCoercion (str,r,s,t) -> vernac_coercion str r s t
| VernacIdentityCoercion (str,(_,id),s,t) -> vernac_identity_coercion str id s t
+ (* Type classes *)
+ | VernacClass (id, par, ar, sup, props) -> vernac_class id par ar sup props
+
+ | VernacInstance (sup, inst, props) -> vernac_instance sup inst props
+ | VernacContext sup -> vernac_context sup
+ | VernacDeclareInstance id -> vernac_declare_instance id
+
+ | VernacSetInstantiationTactic (tac) -> vernac_set_instantiation_tac tac
+
(* Solving *)
| VernacSolve (n,tac,b) -> vernac_solve n tac b
| VernacSolveExistential (n,c) -> vernac_solve_existential n c
@@ -1222,7 +1258,7 @@ let interp c = match c with
| VernacDeclareTacticDefinition (x,l) -> vernac_declare_tactic_definition x l
| VernacHints (local,dbnames,hints) -> vernac_hints local dbnames hints
| VernacSyntacticDefinition (id,c,l,b) ->vernac_syntactic_definition id c l b
- | VernacDeclareImplicits (local,qid,l) ->vernac_declare_implicits local qid l
+ | VernacDeclareImplicits (local,qid,e,l) ->vernac_declare_implicits local qid e l
| VernacReserve (idl,c) -> vernac_reserve idl c
| VernacSetOpacity (opaq, qidl) -> List.iter (vernac_set_opacity opaq) qidl
| VernacSetOption (key,v) -> vernac_set_option key v
diff --git a/toplevel/vernacexpr.ml b/toplevel/vernacexpr.ml
index c89393301..a5c04a561 100644
--- a/toplevel/vernacexpr.ml
+++ b/toplevel/vernacexpr.ml
@@ -48,6 +48,8 @@ type printable =
| PrintOpaqueName of reference
| PrintGraph
| PrintClasses
+ | PrintTypeClasses
+ | PrintInstances of reference
| PrintLtac of reference
| PrintCoercions
| PrintCoercionPaths of class_rawexpr * class_rawexpr
@@ -145,10 +147,12 @@ type sort_expr = Rawterm.rawsort
type decl_notation = (string * constr_expr * scope_name option) option
type simple_binder = lident list * constr_expr
+type class_binder = lident * constr_expr list
type 'a with_coercion = coercion_flag * 'a
type constructor_expr = (lident * constr_expr) with_coercion
type inductive_expr =
lident * local_binder list * constr_expr * constructor_expr list
+
type definition_expr =
| ProveBody of local_binder list * constr_expr
| DefineBody of local_binder list * raw_red_expr option * constr_expr
@@ -220,6 +224,26 @@ type vernac_expr =
| VernacIdentityCoercion of strength * lident *
class_rawexpr * class_rawexpr
+ (* Type classes *)
+ | VernacClass of
+ lident * (* name *)
+ (lident * constr_expr) list * (* params *)
+ sort_expr located * (* arity *)
+ typeclass_context * (* super *)
+ (lident * constr_expr) list (* props *)
+
+ | VernacInstance of
+ typeclass_context * (* super *)
+ typeclass_constraint * (* instance name, class name, params *)
+ (lident * lident list * constr_expr) list (* props *)
+
+ | VernacContext of typeclass_context
+
+ | VernacDeclareInstance of
+ lident (* instance name *)
+
+ | VernacSetInstantiationTactic of raw_tactic_expr
+
(* Modules and Module Types *)
| VernacDeclareModule of bool option * lident *
module_binder list * (module_type_ast * bool)
@@ -260,12 +284,12 @@ type vernac_expr =
(* Commands *)
| VernacDeclareTacticDefinition of
- rec_flag * (lident * raw_tactic_expr) list
+ rec_flag * (lident * bool * raw_tactic_expr) list
| VernacHints of locality_flag * lstring list * hints
| VernacSyntacticDefinition of identifier * constr_expr * locality_flag *
onlyparsing_flag
- | VernacDeclareImplicits of locality_flag * lreference *
- (explicitation * bool) list option
+ | VernacDeclareImplicits of locality_flag * lreference * bool *
+ (explicitation * bool * bool) list option
| VernacReserve of lident list * constr_expr
| VernacSetOpacity of opacity_flag * lreference list
| VernacUnsetOption of Goptions.option_name
diff --git a/toplevel/whelp.ml4 b/toplevel/whelp.ml4
index c7a9774e0..d76c281be 100644
--- a/toplevel/whelp.ml4
+++ b/toplevel/whelp.ml4
@@ -151,12 +151,12 @@ let rec uri_of_constr c =
let inst,rest = merge (section_parameters f) args in
uri_of_constr f; url_char ' '; uri_params uri_of_constr inst;
url_list_with_sep " " uri_of_constr rest
- | RLambda (_,na,ty,c) ->
+ | RLambda (_,na,k,ty,c) ->
url_string "\\lambda "; url_of_name na; url_string ":";
uri_of_constr ty; url_string "."; uri_of_constr c
- | RProd (_,Anonymous,ty,c) ->
+ | RProd (_,Anonymous,k,ty,c) ->
uri_of_constr ty; url_string "\\to "; uri_of_constr c
- | RProd (_,Name id,ty,c) ->
+ | RProd (_,Name id,k,ty,c) ->
url_string "\\forall "; url_id id; url_string ":";
uri_of_constr ty; url_string "."; uri_of_constr c
| RLetIn (_,na,b,c) ->