summaryrefslogtreecommitdiff
path: root/contrib/dp
diff options
context:
space:
mode:
Diffstat (limited to 'contrib/dp')
-rw-r--r--contrib/dp/TODO28
-rw-r--r--contrib/dp/dp.ml760
-rw-r--r--contrib/dp/dp.mli12
-rw-r--r--contrib/dp/dp_cvcl.ml112
-rw-r--r--contrib/dp/dp_cvcl.mli4
-rw-r--r--contrib/dp/dp_simplify.ml117
-rw-r--r--contrib/dp/dp_simplify.mli4
-rw-r--r--contrib/dp/dp_sorts.ml51
-rw-r--r--contrib/dp/dp_sorts.mli4
-rw-r--r--contrib/dp/dp_why.ml139
-rw-r--r--contrib/dp/dp_zenon.ml103
-rw-r--r--contrib/dp/dp_zenon.mli4
-rw-r--r--contrib/dp/fol.mli48
-rw-r--r--contrib/dp/g_dp.ml438
-rw-r--r--contrib/dp/test2.v78
-rw-r--r--contrib/dp/tests.v220
16 files changed, 1722 insertions, 0 deletions
diff --git a/contrib/dp/TODO b/contrib/dp/TODO
new file mode 100644
index 00000000..387cacdf
--- /dev/null
+++ b/contrib/dp/TODO
@@ -0,0 +1,28 @@
+
+TODO
+----
+
+- axiomes pour les prédicats récursifs comme
+
+ Fixpoint even (n:nat) : Prop :=
+ match n with
+ O => True
+ | S O => False
+ | S (S p) => even p
+ end.
+
+ ou encore In sur les listes du module Coq List.
+
+- discriminate
+
+- inversion (Set et Prop)
+
+
+BUGS
+----
+
+- value = Some : forall A:Set, A -> option A
+
+ -> eta_expanse échoue sur assert false (ligne 147)
+
+
diff --git a/contrib/dp/dp.ml b/contrib/dp/dp.ml
new file mode 100644
index 00000000..af684e6e
--- /dev/null
+++ b/contrib/dp/dp.ml
@@ -0,0 +1,760 @@
+(* Authors: Nicolas Ayache and Jean-Christophe Filliâtre *)
+(* Tactics to call decision procedures *)
+
+(* Works in two steps:
+
+ - first the Coq context and the current goal are translated in
+ Polymorphic First-Order Logic (see fol.mli in this directory)
+
+ - then the resulting query is passed to the Why tool that translates
+ it to the syntax of the selected prover (Simplify, CVC Lite, haRVey,
+ Zenon)
+*)
+
+open Util
+open Pp
+open Term
+open Tacmach
+open Tactics
+open Tacticals
+open Fol
+open Names
+open Nameops
+open Termops
+open Coqlib
+open Hipattern
+open Libnames
+open Declarations
+
+let debug = ref false
+
+let logic_dir = ["Coq";"Logic";"Decidable"]
+let coq_modules =
+ init_modules @ [logic_dir] @ arith_modules @ zarith_base_modules
+ @ [["Coq"; "omega"; "OmegaLemmas"]]
+
+let constant = gen_constant_in_modules "dp" coq_modules
+
+let coq_Z = lazy (constant "Z")
+let coq_Zplus = lazy (constant "Zplus")
+let coq_Zmult = lazy (constant "Zmult")
+let coq_Zopp = lazy (constant "Zopp")
+let coq_Zminus = lazy (constant "Zminus")
+let coq_Zdiv = lazy (constant "Zdiv")
+let coq_Zs = lazy (constant "Zs")
+let coq_Zgt = lazy (constant "Zgt")
+let coq_Zle = lazy (constant "Zle")
+let coq_Zge = lazy (constant "Zge")
+let coq_Zlt = lazy (constant "Zlt")
+let coq_Z0 = lazy (constant "Z0")
+let coq_Zpos = lazy (constant "Zpos")
+let coq_Zneg = lazy (constant "Zneg")
+let coq_xH = lazy (constant "xH")
+let coq_xI = lazy (constant "xI")
+let coq_xO = lazy (constant "xO")
+
+(* not Prop typed expressions *)
+exception NotProp
+
+(* not first-order expressions *)
+exception NotFO
+
+(* Renaming of Coq globals *)
+
+let global_names = Hashtbl.create 97
+let used_names = Hashtbl.create 97
+
+let rename_global r =
+ try
+ Hashtbl.find global_names r
+ with Not_found ->
+ let rec loop id =
+ if Hashtbl.mem used_names id then
+ loop (lift_ident id)
+ else begin
+ Hashtbl.add used_names id ();
+ let s = string_of_id id in
+ Hashtbl.add global_names r s;
+ s
+ end
+ in
+ loop (Nametab.id_of_global r)
+
+let foralls =
+ List.fold_right
+ (fun (x,t) p -> Forall (x, t, p))
+
+let fresh_var = function
+ | Anonymous -> rename_global (VarRef (id_of_string "x"))
+ | Name x -> rename_global (VarRef x)
+
+(* coq_rename_vars env [(x1,t1);...;(xn,tn)] renames the xi outside of
+ env names, and returns the new variables together with the new
+ environment *)
+let coq_rename_vars env vars =
+ let avoid = ref (ids_of_named_context (Environ.named_context env)) in
+ List.fold_right
+ (fun (na,t) (newvars, newenv) ->
+ let id = next_name_away na !avoid in
+ avoid := id :: !avoid;
+ id :: newvars, Environ.push_named (id, None, t) newenv)
+ vars ([],env)
+
+(* extract the prenex type quantifications i.e.
+ type_quantifiers env (A1:Set)...(Ak:Set)t = A1...An, (env+Ai), t *)
+let decomp_type_quantifiers env t =
+ let rec loop vars t = match kind_of_term t with
+ | Prod (n, a, t) when is_Set a ->
+ loop ((n,a) :: vars) t
+ | _ ->
+ let vars, env = coq_rename_vars env vars in
+ let t = substl (List.map mkVar vars) t in
+ List.rev vars, env, t
+ in
+ loop [] t
+
+(* same thing with lambda binders (for axiomatize body) *)
+let decomp_type_lambdas env t =
+ let rec loop vars t = match kind_of_term t with
+ | Lambda (n, a, t) when is_Set a ->
+ loop ((n,a) :: vars) t
+ | _ ->
+ let vars, env = coq_rename_vars env vars in
+ let t = substl (List.map mkVar vars) t in
+ List.rev vars, env, t
+ in
+ loop [] t
+
+let decompose_arrows =
+ let rec arrows_rec l c = match kind_of_term c with
+ | Prod (_,t,c) when not (dependent (mkRel 1) c) -> arrows_rec (t :: l) c
+ | Cast (c,_,_) -> arrows_rec l c
+ | _ -> List.rev l, c
+ in
+ arrows_rec []
+
+let rec eta_expanse t vars env i =
+ assert (i >= 0);
+ if i = 0 then
+ t, vars, env
+ else
+ match kind_of_term (Typing.type_of env Evd.empty t) with
+ | Prod (n, a, b) when not (dependent (mkRel 1) b) ->
+ let avoid = ids_of_named_context (Environ.named_context env) in
+ let id = next_name_away n avoid in
+ let env' = Environ.push_named (id, None, a) env in
+ let t' = mkApp (t, [| mkVar id |]) in
+ eta_expanse t' (id :: vars) env' (pred i)
+ | _ ->
+ assert false
+
+let rec skip_k_args k cl = match k, cl with
+ | 0, _ -> cl
+ | _, _ :: cl -> skip_k_args (k-1) cl
+ | _, [] -> raise NotFO
+
+(* Coq global references *)
+
+type global = Gnot_fo | Gfo of Fol.decl
+
+let globals = ref Refmap.empty
+let globals_stack = ref []
+
+(* synchronization *)
+let () =
+ Summary.declare_summary "Dp globals"
+ { Summary.freeze_function = (fun () -> !globals, !globals_stack);
+ Summary.unfreeze_function =
+ (fun (g,s) -> globals := g; globals_stack := s);
+ Summary.init_function = (fun () -> ());
+ Summary.survive_module = false;
+ Summary.survive_section = false }
+
+let add_global r d = globals := Refmap.add r d !globals
+let mem_global r = Refmap.mem r !globals
+let lookup_global r = match Refmap.find r !globals with
+ | Gnot_fo -> raise NotFO
+ | Gfo d -> d
+
+let locals = Hashtbl.create 97
+
+let lookup_local r = match Hashtbl.find locals r with
+ | Gnot_fo -> raise NotFO
+ | Gfo d -> d
+
+let iter_all_constructors i f =
+ let _, oib = Global.lookup_inductive i in
+ Array.iteri
+ (fun j tj -> f j (mkConstruct (i, j+1)))
+ oib.mind_nf_lc
+
+
+(* injection c [t1,...,tn] adds the injection axiom
+ forall x1:t1,...,xn:tn,y1:t1,...,yn:tn.
+ c(x1,...,xn)=c(y1,...,yn) -> x1=y1 /\ ... /\ xn=yn *)
+
+let injection c l =
+ let i = ref 0 in
+ let var s = incr i; id_of_string (s ^ string_of_int !i) in
+ let xl = List.map (fun t -> rename_global (VarRef (var "x")), t) l in
+ i := 0;
+ let yl = List.map (fun t -> rename_global (VarRef (var "y")), t) l in
+ let f =
+ List.fold_right2
+ (fun (x,_) (y,_) p -> And (Fatom (Eq (App (x,[]),App (y,[]))), p))
+ xl yl True
+ in
+ let vars = List.map (fun (x,_) -> App(x,[])) in
+ let f = Imp (Fatom (Eq (App (c, vars xl), App (c, vars yl))), f) in
+ let foralls = List.fold_right (fun (x,t) p -> Forall (x, t, p)) in
+ let f = foralls xl (foralls yl f) in
+ let ax = Axiom ("injection_" ^ c, f) in
+ globals_stack := ax :: !globals_stack
+
+(* rec_names_for c [|n1;...;nk|] builds the list of constant names for
+ identifiers n1...nk with the same path as c, if they exist; otherwise
+ raises Not_found *)
+let rec_names_for c =
+ let mp,dp,_ = Names.repr_con c in
+ array_map_to_list
+ (function
+ | Name id ->
+ let c' = Names.make_con mp dp (label_of_id id) in
+ ignore (Global.lookup_constant c');
+ msgnl (Printer.pr_constr (mkConst c'));
+ c'
+ | Anonymous ->
+ raise Not_found)
+
+(* abstraction tables *)
+
+let term_abstractions = Hashtbl.create 97
+
+let new_abstraction =
+ let r = ref 0 in fun () -> incr r; "abstraction_" ^ string_of_int !r
+
+(* Arithmetic constants *)
+
+exception NotArithConstant
+
+(* translates a closed Coq term p:positive into a FOL term of type int *)
+let rec tr_positive p = match kind_of_term p with
+ | Term.Construct _ when p = Lazy.force coq_xH ->
+ Cst 1
+ | Term.App (f, [|a|]) when f = Lazy.force coq_xI ->
+ Plus (Mult (Cst 2, tr_positive a), Cst 1)
+ | Term.App (f, [|a|]) when f = Lazy.force coq_xO ->
+ Mult (Cst 2, tr_positive a)
+ | Term.Cast (p, _, _) ->
+ tr_positive p
+ | _ ->
+ raise NotArithConstant
+
+(* translates a closed Coq term t:Z into a FOL term of type int *)
+let rec tr_arith_constant t = match kind_of_term t with
+ | Term.Construct _ when t = Lazy.force coq_Z0 ->
+ Cst 0
+ | Term.App (f, [|a|]) when f = Lazy.force coq_Zpos ->
+ tr_positive a
+ | Term.App (f, [|a|]) when f = Lazy.force coq_Zneg ->
+ Moins (Cst 0, tr_positive a)
+ | Term.Cast (t, _, _) ->
+ tr_arith_constant t
+ | _ ->
+ raise NotArithConstant
+
+(* translate a Coq term t:Set into a FOL type expression;
+ tv = list of type variables *)
+and tr_type tv env t =
+ let t = Reductionops.nf_betadeltaiota env Evd.empty t in
+ if t = Lazy.force coq_Z then
+ Tid ("int", [])
+ else match kind_of_term t with
+ | Var x when List.mem x tv ->
+ Tvar (string_of_id x)
+ | _ ->
+ let f, cl = decompose_app t in
+ begin try
+ let r = global_of_constr f in
+ match tr_global env r with
+ | DeclType (id, k) ->
+ assert (k = List.length cl); (* since t:Set *)
+ Tid (id, List.map (tr_type tv env) cl)
+ | _ ->
+ raise NotFO
+ with
+ | Not_found ->
+ raise NotFO
+ | NotFO ->
+ (* we need to abstract some part of (f cl) *)
+ (*TODO*)
+ raise NotFO
+ end
+
+and make_term_abstraction tv env c =
+ let ty = Typing.type_of env Evd.empty c in
+ let id = new_abstraction () in
+ match tr_decl env id ty with
+ | DeclFun (id,_,_,_) as d ->
+ begin try
+ Hashtbl.find term_abstractions c
+ with Not_found ->
+ Hashtbl.add term_abstractions c id;
+ globals_stack := d :: !globals_stack;
+ id
+ end
+ | _ ->
+ raise NotFO
+
+(* translate a Coq declaration id:ty in a FOL declaration, that is either
+ - a type declaration : DeclType (id, n) where n:int is the type arity
+ - a function declaration : DeclFun (id, tl, t) ; that includes constants
+ - a predicate declaration : DeclPred (id, tl)
+ - an axiom : Axiom (id, p)
+ *)
+and tr_decl env id ty =
+ let tv, env, t = decomp_type_quantifiers env ty in
+ if is_Set t then
+ DeclType (id, List.length tv)
+ else if is_Prop t then
+ DeclPred (id, List.length tv, [])
+ else
+ let s = Typing.type_of env Evd.empty t in
+ if is_Prop s then
+ Axiom (id, tr_formula tv [] env t)
+ else
+ let l, t = decompose_arrows t in
+ let l = List.map (tr_type tv env) l in
+ if is_Prop t then
+ DeclPred(id, List.length tv, l)
+ else
+ let s = Typing.type_of env Evd.empty t in
+ if is_Set s then
+ DeclFun(id, List.length tv, l, tr_type tv env t)
+ else
+ raise NotFO
+
+(* tr_global(r) = tr_decl(id(r),typeof(r)) + a cache mechanism *)
+and tr_global env r = match r with
+ | VarRef id ->
+ lookup_local id
+ | r ->
+ try
+ lookup_global r
+ with Not_found ->
+ try
+ let ty = Global.type_of_global r in
+ let id = rename_global r in
+ let d = tr_decl env id ty in
+ (* r can be already declared if it is a constructor *)
+ if not (mem_global r) then begin
+ add_global r (Gfo d);
+ globals_stack := d :: !globals_stack
+ end;
+ begin try axiomatize_body env r id d with NotFO -> () end;
+ d
+ with NotFO ->
+ add_global r Gnot_fo;
+ raise NotFO
+
+and axiomatize_body env r id d = match r with
+ | VarRef _ ->
+ assert false
+ | ConstRef c ->
+ begin match (Global.lookup_constant c).const_body with
+ | Some b ->
+ let b = force b in
+ let tv, env, b = decomp_type_lambdas env b in
+ let axioms =
+ (match d with
+ | DeclPred (id, _, []) ->
+ let value = tr_formula tv [] env b in
+ [id, Iff (Fatom (Pred (id, [])), value)]
+ | DeclFun (id, _, [], _) ->
+ let value = tr_term tv [] env b in
+ [id, Fatom (Eq (Fol.App (id, []), value))]
+ | DeclFun (id, _, l, _) | DeclPred (id, _, l) ->
+ Format.eprintf "axiomatize_body %S@." id;
+ let b = match kind_of_term b with
+ (* a single recursive function *)
+ | Fix (_, (_,_,[|b|])) ->
+ subst1 (mkConst c) b
+ (* mutually recursive functions *)
+ | Fix ((_,i), (names,_,bodies)) ->
+ (* we only deal with named functions *)
+ begin try
+ let l = rec_names_for c names in
+ substl (List.rev_map mkConst l) bodies.(i)
+ with Not_found ->
+ b
+ end
+ | _ ->
+ b
+ in
+ let vars, t = decompose_lam b in
+ let n = List.length l in
+ let k = List.length vars in
+ assert (k <= n);
+ let vars, env = coq_rename_vars env vars in
+ let t = substl (List.map mkVar vars) t in
+ let t, vars, env = eta_expanse t vars env (n-k) in
+ let vars = List.rev vars in
+ let bv = vars in
+ let vars = List.map (fun x -> string_of_id x) vars in
+ let fol_var x =
+ Fol.App (x, []) in
+ let fol_vars = List.map fol_var vars in
+ let vars = List.combine vars l in
+ begin match d with
+ | DeclFun _ ->
+ begin match kind_of_term t with
+ | Case (ci, _, e, br) ->
+ equations_for_case env id vars tv bv ci e br
+ | _ ->
+ let p =
+ Fatom (Eq (App (id, fol_vars),
+ tr_term tv bv env t))
+ in
+ [id, foralls vars p]
+ end
+ | DeclPred _ ->
+ let value = tr_formula tv bv env t in
+ let p = Iff (Fatom (Pred (id, fol_vars)), value) in
+ [id, foralls vars p]
+ | _ ->
+ assert false
+ end
+ | DeclType _ ->
+ raise NotFO
+ | Axiom _ -> assert false)
+ in
+ let axioms = List.map (fun (id,ax) -> Axiom (id, ax)) axioms in
+ globals_stack := axioms @ !globals_stack
+ | None ->
+ () (* Coq axiom *)
+ end
+ | IndRef i ->
+ iter_all_constructors i
+ (fun _ c ->
+ let rc = reference_of_constr c in
+ try
+ begin match tr_global env rc with
+ | DeclFun (_, _, [], _) -> ()
+ | DeclFun (idc, _, al, _) -> injection idc al
+ | _ -> ()
+ end
+ with NotFO ->
+ ())
+ | _ -> ()
+
+and equations_for_case env id vars tv bv ci e br = match kind_of_term e with
+ | Var x when List.exists (fun (y, _) -> string_of_id x = y) vars ->
+ let eqs = ref [] in
+ iter_all_constructors ci.ci_ind
+ (fun j cj ->
+ try
+ let cjr = reference_of_constr cj in
+ begin match tr_global env cjr with
+ | DeclFun (idc, _, l, _) ->
+ let b = br.(j) in
+ let rec_vars, b = decompose_lam b in
+ let rec_vars, env = coq_rename_vars env rec_vars in
+ let b = substl (List.map mkVar rec_vars) b in
+ let rec_vars = List.rev rec_vars in
+ let bv = bv @ rec_vars in
+ let rec_vars = List.map string_of_id rec_vars in
+ let fol_var x =
+ Fol.App (x, []) in
+ let fol_rec_vars = List.map fol_var rec_vars in
+ let fol_rec_term = App (idc, fol_rec_vars) in
+ let rec_vars = List.combine rec_vars l in
+ let fol_vars = List.map fst vars in
+ let fol_vars = List.map fol_var fol_vars in
+ let fol_vars = List.map (fun y -> match y with
+ | App (id, _) ->
+ if id = string_of_id x
+ then fol_rec_term
+ else y
+ | _ -> y)
+ fol_vars in
+ let vars = vars @ rec_vars in
+ let rec remove l e = match l with
+ | [] -> []
+ | (y, t)::l' -> if y = string_of_id e then l'
+ else (y, t)::(remove l' e) in
+ let vars = remove vars x in
+ let p =
+ Fatom (Eq (App (id, fol_vars),
+ tr_term tv bv env b))
+ in
+ eqs := (id ^ "_" ^ idc, foralls vars p) :: !eqs
+ | _ ->
+ assert false end
+ with NotFO ->
+ ());
+ !eqs
+ | _ ->
+ raise NotFO
+
+(* assumption: t:T:Set *)
+and tr_term tv bv env t = match kind_of_term t with
+ | Term.App (f, [|a;b|]) when f = Lazy.force coq_Zplus ->
+ Plus (tr_term tv bv env a, tr_term tv bv env b)
+ | Term.App (f, [|a;b|]) when f = Lazy.force coq_Zminus ->
+ Moins (tr_term tv bv env a, tr_term tv bv env b)
+ | Term.App (f, [|a;b|]) when f = Lazy.force coq_Zmult ->
+ Mult (tr_term tv bv env a, tr_term tv bv env b)
+ | Term.App (f, [|a;b|]) when f = Lazy.force coq_Zdiv ->
+ Div (tr_term tv bv env a, tr_term tv bv env b)
+ | Term.Var id when List.mem id bv ->
+ App (string_of_id id, [])
+ | _ ->
+ try
+ tr_arith_constant t
+ with NotArithConstant ->
+ let f, cl = decompose_app t in
+ begin try
+ let r = global_of_constr f in
+ match tr_global env r with
+ | DeclFun (s, k, _, _) ->
+ let cl = skip_k_args k cl in
+ Fol.App (s, List.map (tr_term tv bv env) cl)
+ | _ ->
+ raise NotFO
+ with
+ | Not_found ->
+ raise NotFO
+ | NotFO -> (* we need to abstract some part of (f cl) *)
+ let rec abstract app = function
+ | [] ->
+ Fol.App (make_term_abstraction tv env app, [])
+ | x :: l as args ->
+ begin try
+ let s = make_term_abstraction tv env app in
+ Fol.App (s, List.map (tr_term tv bv env) args)
+ with NotFO ->
+ abstract (applist (app, [x])) l
+ end
+ in
+ let app,l = match cl with
+ | x :: l -> applist (f, [x]), l | [] -> raise NotFO
+ in
+ abstract app l
+ end
+
+and quantifiers n a b tv bv env =
+ let vars, env = coq_rename_vars env [n,a] in
+ let id = match vars with [x] -> x | _ -> assert false in
+ let b = subst1 (mkVar id) b in
+ let t = tr_type tv env a in
+ let bv = id :: bv in
+ id, t, bv, env, b
+
+(* assumption: f is of type Prop *)
+and tr_formula tv bv env f =
+ let c, args = decompose_app f in
+ match kind_of_term c, args with
+ | Var id, [] ->
+ Fatom (Pred (rename_global (VarRef id), []))
+ | _, [t;a;b] when c = build_coq_eq () ->
+ let ty = Typing.type_of env Evd.empty t in
+ if is_Set ty then
+ let _ = tr_type tv env t in
+ Fatom (Eq (tr_term tv bv env a, tr_term tv bv env b))
+ else
+ raise NotFO
+ | _, [a;b] when c = Lazy.force coq_Zle ->
+ Fatom (Le (tr_term tv bv env a, tr_term tv bv env b))
+ | _, [a;b] when c = Lazy.force coq_Zlt ->
+ Fatom (Lt (tr_term tv bv env a, tr_term tv bv env b))
+ | _, [a;b] when c = Lazy.force coq_Zge ->
+ Fatom (Ge (tr_term tv bv env a, tr_term tv bv env b))
+ | _, [a;b] when c = Lazy.force coq_Zgt ->
+ Fatom (Gt (tr_term tv bv env a, tr_term tv bv env b))
+ | _, [] when c = build_coq_False () ->
+ False
+ | _, [] when c = build_coq_True () ->
+ True
+ | _, [a] when c = build_coq_not () ->
+ Not (tr_formula tv bv env a)
+ | _, [a;b] when c = build_coq_and () ->
+ And (tr_formula tv bv env a, tr_formula tv bv env b)
+ | _, [a;b] when c = build_coq_or () ->
+ Or (tr_formula tv bv env a, tr_formula tv bv env b)
+ | Prod (n, a, b), _ ->
+ if is_imp_term f then
+ Imp (tr_formula tv bv env a, tr_formula tv bv env b)
+ else
+ let id, t, bv, env, b = quantifiers n a b tv bv env in
+ Forall (string_of_id id, t, tr_formula tv bv env b)
+ | _, [_; a] when c = build_coq_ex () ->
+ begin match kind_of_term a with
+ | Lambda(n, a, b) ->
+ let id, t, bv, env, b = quantifiers n a b tv bv env in
+ Exists (string_of_id id, t, tr_formula tv bv env b)
+ | _ ->
+ (* unusual case of the shape (ex p) *)
+ raise NotFO (* TODO: we could eta-expanse *)
+ end
+ | _ ->
+ begin try
+ let r = global_of_constr c in
+ match tr_global env r with
+ | DeclPred (s, k, _) ->
+ let args = skip_k_args k args in
+ Fatom (Pred (s, List.map (tr_term tv bv env) args))
+ | _ ->
+ raise NotFO
+ with Not_found ->
+ raise NotFO
+ end
+
+
+let tr_goal gl =
+ Hashtbl.clear locals;
+ let tr_one_hyp (id, ty) =
+ try
+ let s = rename_global (VarRef id) in
+ let d = tr_decl (pf_env gl) s ty in
+ Hashtbl.add locals id (Gfo d);
+ d
+ with NotFO ->
+ Hashtbl.add locals id Gnot_fo;
+ raise NotFO
+ in
+ let hyps =
+ List.fold_right
+ (fun h acc -> try tr_one_hyp h :: acc with NotFO -> acc)
+ (pf_hyps_types gl) []
+ in
+ let c = tr_formula [] [] (pf_env gl) (pf_concl gl) in
+ let hyps = List.rev_append !globals_stack (List.rev hyps) in
+ hyps, c
+
+
+type prover = Simplify | CVCLite | Harvey | Zenon
+
+let remove_files = List.iter (fun f -> try Sys.remove f with _ -> ())
+
+let sprintf = Format.sprintf
+
+let call_simplify fwhy =
+ if Sys.command (sprintf "why --simplify %s" fwhy) <> 0 then
+ anomaly ("call to why --simplify " ^ fwhy ^ " failed; please report");
+ let fsx = Filename.chop_suffix fwhy ".why" ^ "_why.sx" in
+ let cmd =
+ sprintf "timeout 10 Simplify %s > out 2>&1 && grep -q -w Valid out" fsx
+ in
+ let out = Sys.command cmd in
+ let r = if out = 0 then Valid else if out = 1 then Invalid else Timeout in
+ if not !debug then remove_files [fwhy; fsx];
+ r
+
+let call_zenon fwhy =
+ let cmd = sprintf "why --no-prelude --no-zenon-prelude --zenon %s" fwhy in
+ if Sys.command cmd <> 0 then
+ anomaly ("call to " ^ cmd ^ " failed; please report");
+ let fznn = Filename.chop_suffix fwhy ".why" ^ "_why.znn" in
+ let cmd =
+ sprintf "timeout 10 zenon %s > out 2>&1 && grep -q PROOF-FOUND out" fznn
+ in
+ let out = Sys.command cmd in
+ let r =
+ if out = 0 then Valid
+ else if out = 1 then Invalid
+ else if out = 137 then Timeout
+ else anomaly ("malformed Zenon input file " ^ fznn)
+ in
+ if not !debug then remove_files [fwhy; fznn];
+ r
+
+let call_cvcl fwhy =
+ if Sys.command (sprintf "why --cvcl %s" fwhy) <> 0 then
+ anomaly ("call to why --cvcl " ^ fwhy ^ " failed; please report");
+ let fcvc = Filename.chop_suffix fwhy ".why" ^ "_why.cvc" in
+ let cmd =
+ sprintf "timeout 10 cvcl < %s > out 2>&1 && grep -q -w Valid out" fcvc
+ in
+ let out = Sys.command cmd in
+ let r = if out = 0 then Valid else if out = 1 then Invalid else Timeout in
+ if not !debug then remove_files [fwhy; fcvc];
+ r
+
+let call_harvey fwhy =
+ if Sys.command (sprintf "why --harvey %s" fwhy) <> 0 then
+ anomaly ("call to why --harvey " ^ fwhy ^ " failed; please report");
+ let frv = Filename.chop_suffix fwhy ".why" ^ "_why.rv" in
+ let out = Sys.command (sprintf "rvc -e -t %s > /dev/null 2>&1" frv) in
+ if out <> 0 then anomaly ("call to rvc -e -t " ^ frv ^ " failed");
+ let f = Filename.chop_suffix frv ".rv" ^ "-0.baf" in
+ let outf = Filename.temp_file "rv" ".out" in
+ let out =
+ Sys.command (sprintf "timeout 10 rv -e\"-T 2000\" %s > %s 2>&1" f outf)
+ in
+ let r =
+ if out <> 0 then
+ Timeout
+ else
+ let cmd =
+ sprintf "grep \"Proof obligation in\" %s | grep -q \"is valid\"" outf
+ in
+ if Sys.command cmd = 0 then Valid else Invalid
+ in
+ if not !debug then remove_files [fwhy; frv; outf];
+ r
+
+let call_prover prover q =
+ let fwhy = Filename.temp_file "coq_dp" ".why" in
+ Dp_why.output_file fwhy q;
+ if !debug then ignore (Sys.command (sprintf "cat %s" fwhy));
+ match prover with
+ | Simplify -> call_simplify fwhy
+ | Zenon -> call_zenon fwhy
+ | CVCLite -> call_cvcl fwhy
+ | Harvey -> call_harvey fwhy
+
+let dp prover gl =
+ let concl_type = pf_type_of gl (pf_concl gl) in
+ if not (is_Prop concl_type) then error "Conclusion is not a Prop";
+ try
+ let q = tr_goal gl in
+ begin match call_prover prover q with
+ | Valid -> Tactics.admit_as_an_axiom gl
+ | Invalid -> error "Invalid"
+ | DontKnow -> error "Don't know"
+ | Timeout -> error "Timeout"
+ end
+ with NotFO ->
+ error "Not a first order goal"
+
+
+let simplify = tclTHEN intros (dp Simplify)
+let cvc_lite = tclTHEN intros (dp CVCLite)
+let harvey = dp Harvey
+let zenon = tclTHEN intros (dp Zenon)
+
+let dp_hint l =
+ let env = Global.env () in
+ let one_hint (qid,r) =
+ if not (mem_global r) then begin
+ let ty = Global.type_of_global r in
+ let s = Typing.type_of env Evd.empty ty in
+ if is_Prop s then
+ try
+ let id = rename_global r in
+ let d = Axiom (id, tr_formula [] [] env ty) in
+ add_global r (Gfo d);
+ globals_stack := d :: !globals_stack
+ with NotFO ->
+ add_global r Gnot_fo;
+ msg_warning
+ (pr_reference qid ++
+ str " ignored (not a first order proposition)")
+ else begin
+ add_global r Gnot_fo;
+ msg_warning
+ (pr_reference qid ++ str " ignored (not a proposition)")
+ end
+ end
+ in
+ List.iter one_hint (List.map (fun qid -> qid, Nametab.global qid) l)
diff --git a/contrib/dp/dp.mli b/contrib/dp/dp.mli
new file mode 100644
index 00000000..3dad469c
--- /dev/null
+++ b/contrib/dp/dp.mli
@@ -0,0 +1,12 @@
+
+open Libnames
+open Proof_type
+
+val simplify : tactic
+val cvc_lite : tactic
+val harvey : tactic
+val zenon : tactic
+
+val dp_hint : reference list -> unit
+
+
diff --git a/contrib/dp/dp_cvcl.ml b/contrib/dp/dp_cvcl.ml
new file mode 100644
index 00000000..05d43081
--- /dev/null
+++ b/contrib/dp/dp_cvcl.ml
@@ -0,0 +1,112 @@
+
+open Format
+open Fol
+
+let rec print_list sep print fmt = function
+ | [] -> ()
+ | [x] -> print fmt x
+ | x :: r -> print fmt x; sep fmt (); print_list sep print fmt r
+
+let space fmt () = fprintf fmt "@ "
+let comma fmt () = fprintf fmt ",@ "
+
+let rec print_term fmt = function
+ | Cst n ->
+ fprintf fmt "%d" n
+ | Plus (a, b) ->
+ fprintf fmt "@[(%a@ +@ %a)@]" print_term a print_term b
+ | Moins (a, b) ->
+ fprintf fmt "@[(%a@ -@ %a)@]" print_term a print_term b
+ | Mult (a, b) ->
+ fprintf fmt "@[(%a@ *@ %a)@]" print_term a print_term b
+ | Div (a, b) ->
+ fprintf fmt "@[(%a@ /@ %a)@]" print_term a print_term b
+ | App (id, []) ->
+ fprintf fmt "@[%s@]" id
+ | App (id, tl) ->
+ fprintf fmt "@[%s(%a)@]" id print_terms tl
+
+and print_terms fmt tl =
+ print_list comma print_term fmt tl
+
+let rec print_predicate fmt p =
+ let pp = print_predicate in
+ match p with
+ | True ->
+ fprintf fmt "TRUE"
+ | False ->
+ fprintf fmt "FALSE"
+ | Fatom (Eq (a, b)) ->
+ fprintf fmt "@[(%a = %a)@]" print_term a print_term b
+ | Fatom (Le (a, b)) ->
+ fprintf fmt "@[(%a@ <= %a)@]" print_term a print_term b
+ | Fatom (Lt (a, b))->
+ fprintf fmt "@[(%a@ < %a)@]" print_term a print_term b
+ | Fatom (Ge (a, b)) ->
+ fprintf fmt "@[(%a@ >= %a)@]" print_term a print_term b
+ | Fatom (Gt (a, b)) ->
+ fprintf fmt "@[(%a@ > %a)@]" print_term a print_term b
+ | Fatom (Pred (id, [])) ->
+ fprintf fmt "@[%s@]" id
+ | Fatom (Pred (id, tl)) ->
+ fprintf fmt "@[%s(%a)@]" id print_terms tl
+ | Imp (a, b) ->
+ fprintf fmt "@[(%a@ => %a)@]" pp a pp b
+ | And (a, b) ->
+ fprintf fmt "@[(%a@ AND@ %a)@]" pp a pp b
+ | Or (a, b) ->
+ fprintf fmt "@[(%a@ OR@ %a)@]" pp a pp b
+ | Not a ->
+ fprintf fmt "@[(NOT@ %a)@]" pp a
+ | Forall (id, t, p) ->
+ fprintf fmt "@[(FORALL (%s:%s): %a)@]" id t pp p
+ | Exists (id, t, p) ->
+ fprintf fmt "@[(EXISTS (%s:%s): %a)@]" id t pp p
+
+let rec string_of_type_list = function
+ | [] -> assert false
+ | [e] -> e
+ | e :: l' -> e ^ ", " ^ (string_of_type_list l')
+
+let print_query fmt (decls,concl) =
+ let print_decl = function
+ | DeclVar (id, [], t) ->
+ fprintf fmt "@[%s: %s;@]@\n" id t
+ | DeclVar (id, [e], t) ->
+ fprintf fmt "@[%s: [%s -> %s];@]@\n"
+ id e t
+ | DeclVar (id, l, t) ->
+ fprintf fmt "@[%s: [[%s] -> %s];@]@\n"
+ id (string_of_type_list l) t
+ | DeclPred (id, []) ->
+ fprintf fmt "@[%s: BOOLEAN;@]@\n" id
+ | DeclPred (id, [e]) ->
+ fprintf fmt "@[%s: [%s -> BOOLEAN];@]@\n"
+ id e
+ | DeclPred (id, l) ->
+ fprintf fmt "@[%s: [[%s] -> BOOLEAN];@]@\n"
+ id (string_of_type_list l)
+ | DeclType id ->
+ fprintf fmt "@[%s: TYPE;@]@\n" id
+ | Assert (id, f) ->
+ fprintf fmt "@[ASSERT %% %s@\n %a;@]@\n" id print_predicate f
+ in
+ List.iter print_decl decls;
+ fprintf fmt "QUERY %a;" print_predicate concl
+
+let call q =
+ let f = Filename.temp_file "coq_dp" ".cvc" in
+ let c = open_out f in
+ let fmt = formatter_of_out_channel c in
+ fprintf fmt "@[%a@]@." print_query q;
+ close_out c;
+ ignore (Sys.command (sprintf "cat %s" f));
+ let cmd =
+ sprintf "timeout 10 cvcl < %s > out 2>&1 && grep -q -w Valid out" f
+ in
+ prerr_endline cmd; flush stderr;
+ let out = Sys.command cmd in
+ if out = 0 then Valid else if out = 1 then Invalid else Timeout
+ (* TODO: effacer le fichier f et le fichier out *)
+
+
diff --git a/contrib/dp/dp_cvcl.mli b/contrib/dp/dp_cvcl.mli
new file mode 100644
index 00000000..03b6d347
--- /dev/null
+++ b/contrib/dp/dp_cvcl.mli
@@ -0,0 +1,4 @@
+
+open Fol
+
+val call : query -> prover_answer
diff --git a/contrib/dp/dp_simplify.ml b/contrib/dp/dp_simplify.ml
new file mode 100644
index 00000000..d5376b8d
--- /dev/null
+++ b/contrib/dp/dp_simplify.ml
@@ -0,0 +1,117 @@
+
+open Format
+open Fol
+
+let is_simplify_ident s =
+ let is_simplify_char = function
+ | 'a'..'z' | 'A'..'Z' | '0'..'9' -> true
+ | _ -> false
+ in
+ try
+ String.iter (fun c -> if not (is_simplify_char c) then raise Exit) s; true
+ with Exit ->
+ false
+
+let ident fmt s =
+ if is_simplify_ident s then fprintf fmt "%s" s else fprintf fmt "|%s|" s
+
+let rec print_list sep print fmt = function
+ | [] -> ()
+ | [x] -> print fmt x
+ | x :: r -> print fmt x; sep fmt (); print_list sep print fmt r
+
+let space fmt () = fprintf fmt "@ "
+let comma fmt () = fprintf fmt ",@ "
+
+let rec print_term fmt = function
+ | Cst n ->
+ fprintf fmt "%d" n
+ | Plus (a, b) ->
+ fprintf fmt "@[(+@ %a@ %a)@]" print_term a print_term b
+ | Moins (a, b) ->
+ fprintf fmt "@[(-@ %a@ %a)@]" print_term a print_term b
+ | Mult (a, b) ->
+ fprintf fmt "@[(*@ %a@ %a)@]" print_term a print_term b
+ | Div (a, b) ->
+ fprintf fmt "@[(/@ %a@ %a)@]" print_term a print_term b
+ | App (id, []) ->
+ fprintf fmt "%a" ident id
+ | App (id, tl) ->
+ fprintf fmt "@[(%a@ %a)@]" ident id print_terms tl
+
+and print_terms fmt tl =
+ print_list space print_term fmt tl
+
+let rec print_predicate fmt p =
+ let pp = print_predicate in
+ match p with
+ | True ->
+ fprintf fmt "TRUE"
+ | False ->
+ fprintf fmt "FALSE"
+ | Fatom (Eq (a, b)) ->
+ fprintf fmt "@[(EQ %a@ %a)@]" print_term a print_term b
+ | Fatom (Le (a, b)) ->
+ fprintf fmt "@[(<= %a@ %a)@]" print_term a print_term b
+ | Fatom (Lt (a, b))->
+ fprintf fmt "@[(< %a@ %a)@]" print_term a print_term b
+ | Fatom (Ge (a, b)) ->
+ fprintf fmt "@[(>= %a@ %a)@]" print_term a print_term b
+ | Fatom (Gt (a, b)) ->
+ fprintf fmt "@[(> %a@ %a)@]" print_term a print_term b
+ | Fatom (Pred (id, tl)) ->
+ fprintf fmt "@[(EQ (%a@ %a) |@@true|)@]" ident id print_terms tl
+ | Imp (a, b) ->
+ fprintf fmt "@[(IMPLIES@ %a@ %a)@]" pp a pp b
+ | And (a, b) ->
+ fprintf fmt "@[(AND@ %a@ %a)@]" pp a pp b
+ | Or (a, b) ->
+ fprintf fmt "@[(OR@ %a@ %a)@]" pp a pp b
+ | Not a ->
+ fprintf fmt "@[(NOT@ %a)@]" pp a
+ | Forall (id, _, p) ->
+ fprintf fmt "@[(FORALL (%a)@ %a)@]" ident id pp p
+ | Exists (id, _, p) ->
+ fprintf fmt "@[(EXISTS (%a)@ %a)@]" ident id pp p
+
+(**
+let rec string_list l = match l with
+ [] -> ""
+ | [e] -> e
+ | e::l' -> e ^ " " ^ (string_list l')
+**)
+
+let print_query fmt (decls,concl) =
+ let print_decl = function
+ | DeclVar (id, [], t) ->
+ fprintf fmt "@[;; %s : %s@]@\n" id t
+ | DeclVar (id, l, t) ->
+ fprintf fmt "@[;; %s : %a -> %s@]@\n"
+ id (print_list comma pp_print_string) l t
+ | DeclPred (id, []) ->
+ fprintf fmt "@[;; %s : BOOLEAN @]@\n" id
+ | DeclPred (id, l) ->
+ fprintf fmt "@[;; %s : %a -> BOOLEAN@]@\n"
+ id (print_list comma pp_print_string) l
+ | DeclType id ->
+ fprintf fmt "@[;; %s : TYPE@]@\n" id
+ | Assert (id, f) ->
+ fprintf fmt "@[(BG_PUSH ;; %s@\n %a)@]@\n" id print_predicate f
+ in
+ List.iter print_decl decls;
+ fprintf fmt "%a@." print_predicate concl
+
+let call q =
+ let f = Filename.temp_file "coq_dp" ".sx" in
+ let c = open_out f in
+ let fmt = formatter_of_out_channel c in
+ fprintf fmt "@[%a@]@." print_query q;
+ close_out c;
+ ignore (Sys.command (sprintf "cat %s" f));
+ let cmd =
+ sprintf "timeout 10 Simplify %s > out 2>&1 && grep -q -w Valid out" f
+ in
+ prerr_endline cmd; flush stderr;
+ let out = Sys.command cmd in
+ if out = 0 then Valid else if out = 1 then Invalid else Timeout
+ (* TODO: effacer le fichier f et le fichier out *)
diff --git a/contrib/dp/dp_simplify.mli b/contrib/dp/dp_simplify.mli
new file mode 100644
index 00000000..03b6d347
--- /dev/null
+++ b/contrib/dp/dp_simplify.mli
@@ -0,0 +1,4 @@
+
+open Fol
+
+val call : query -> prover_answer
diff --git a/contrib/dp/dp_sorts.ml b/contrib/dp/dp_sorts.ml
new file mode 100644
index 00000000..7dbdfa56
--- /dev/null
+++ b/contrib/dp/dp_sorts.ml
@@ -0,0 +1,51 @@
+
+open Fol
+
+let term_has_sort x s = Fatom (Pred ("%sort_" ^ s, [x]))
+
+let has_sort x s = term_has_sort (App (x, [])) s
+
+let rec form = function
+ | True | False | Fatom _ as f -> f
+ | Imp (f1, f2) -> Imp (form f1, form f2)
+ | And (f1, f2) -> And (form f1, form f2)
+ | Or (f1, f2) -> Or (form f1, form f2)
+ | Not f -> Not (form f)
+ | Forall (x, ("INT" as t), f) -> Forall (x, t, form f)
+ | Forall (x, t, f) -> Forall (x, t, Imp (has_sort x t, form f))
+ | Exists (x, ("INT" as t), f) -> Exists (x, t, form f)
+ | Exists (x, t, f) -> Exists (x, t, Imp (has_sort x t, form f))
+
+let sort_ax = let r = ref 0 in fun () -> incr r; "sort_ax_" ^ string_of_int !r
+
+let hyp acc = function
+ | Assert (id, f) ->
+ (Assert (id, form f)) :: acc
+ | DeclVar (id, _, "INT") as d ->
+ d :: acc
+ | DeclVar (id, [], t) as d ->
+ (Assert (sort_ax (), has_sort id t)) :: d :: acc
+ | DeclVar (id, l, t) as d ->
+ let n = ref 0 in
+ let xi =
+ List.fold_left
+ (fun l t -> incr n; ("x" ^ string_of_int !n, t) :: l) [] l
+ in
+ let f =
+ List.fold_left
+ (fun f (x,t) -> if t = "INT" then f else Imp (has_sort x t, f))
+ (term_has_sort
+ (App (id, List.rev_map (fun (x,_) -> App (x,[])) xi)) t)
+ xi
+ in
+ let f = List.fold_left (fun f (x,t) -> Forall (x, t, f)) f xi in
+ (Assert (sort_ax (), f)) :: d :: acc
+ | DeclPred _ as d ->
+ d :: acc
+ | DeclType t as d ->
+ (DeclPred ("%sort_" ^ t, [t])) :: d :: acc
+
+let query (hyps, f) =
+ let hyps' = List.fold_left hyp [] hyps in
+ List.rev hyps', form f
+
diff --git a/contrib/dp/dp_sorts.mli b/contrib/dp/dp_sorts.mli
new file mode 100644
index 00000000..9e74f997
--- /dev/null
+++ b/contrib/dp/dp_sorts.mli
@@ -0,0 +1,4 @@
+
+open Fol
+
+val query : query -> query
diff --git a/contrib/dp/dp_why.ml b/contrib/dp/dp_why.ml
new file mode 100644
index 00000000..e1ddb039
--- /dev/null
+++ b/contrib/dp/dp_why.ml
@@ -0,0 +1,139 @@
+
+(* Pretty-print PFOL (see fol.mli) in Why syntax *)
+
+open Format
+open Fol
+
+let rec print_list sep print fmt = function
+ | [] -> ()
+ | [x] -> print fmt x
+ | x :: r -> print fmt x; sep fmt (); print_list sep print fmt r
+
+let space fmt () = fprintf fmt "@ "
+let comma fmt () = fprintf fmt ",@ "
+
+let is_why_keyword =
+ let h = Hashtbl.create 17 in
+ List.iter
+ (fun s -> Hashtbl.add h s ())
+ ["absurd"; "and"; "array"; "as"; "assert"; "axiom"; "begin";
+ "bool"; "do"; "done"; "else"; "end"; "exception"; "exists";
+ "external"; "false"; "for"; "forall"; "fun"; "function"; "goal";
+ "if"; "in"; "int"; "invariant"; "label"; "let"; "logic"; "not";
+ "of"; "or"; "parameter"; "predicate"; "prop"; "raise"; "raises";
+ "reads"; "real"; "rec"; "ref"; "returns"; "then"; "true"; "try";
+ "type"; "unit"; "variant"; "void"; "while"; "with"; "writes" ];
+ Hashtbl.mem h
+
+let ident fmt s =
+ if is_why_keyword s then fprintf fmt "coq__%s" s else fprintf fmt "%s" s
+
+let rec print_typ fmt = function
+ | Tvar x -> fprintf fmt "'%a" ident x
+ | Tid ("int", []) -> fprintf fmt "int"
+ | Tid (x, []) -> fprintf fmt "%a" ident x
+ | Tid (x, [t]) -> fprintf fmt "%a %a" print_typ t ident x
+ | Tid (x,tl) -> fprintf fmt "(%a) %a" (print_list comma print_typ) tl ident x
+
+let rec print_term fmt = function
+ | Cst n ->
+ fprintf fmt "%d" n
+ | Plus (a, b) ->
+ fprintf fmt "@[(%a +@ %a)@]" print_term a print_term b
+ | Moins (a, b) ->
+ fprintf fmt "@[(%a -@ %a)@]" print_term a print_term b
+ | Mult (a, b) ->
+ fprintf fmt "@[(%a *@ %a)@]" print_term a print_term b
+ | Div (a, b) ->
+ fprintf fmt "@[(%a /@ %a)@]" print_term a print_term b
+ | App (id, []) ->
+ fprintf fmt "%a" ident id
+ | App (id, tl) ->
+ fprintf fmt "@[%a(%a)@]" ident id print_terms tl
+
+and print_terms fmt tl =
+ print_list comma print_term fmt tl
+
+let rec print_predicate fmt p =
+ let pp = print_predicate in
+ match p with
+ | True ->
+ fprintf fmt "true"
+ | False ->
+ fprintf fmt "false"
+ | Fatom (Eq (a, b)) ->
+ fprintf fmt "@[(%a =@ %a)@]" print_term a print_term b
+ | Fatom (Le (a, b)) ->
+ fprintf fmt "@[(%a <=@ %a)@]" print_term a print_term b
+ | Fatom (Lt (a, b))->
+ fprintf fmt "@[(%a <@ %a)@]" print_term a print_term b
+ | Fatom (Ge (a, b)) ->
+ fprintf fmt "@[(%a >=@ %a)@]" print_term a print_term b
+ | Fatom (Gt (a, b)) ->
+ fprintf fmt "@[(%a >@ %a)@]" print_term a print_term b
+ | Fatom (Pred (id, [])) ->
+ fprintf fmt "%a" ident id
+ | Fatom (Pred (id, tl)) ->
+ fprintf fmt "@[%a(%a)@]" ident id print_terms tl
+ | Imp (a, b) ->
+ fprintf fmt "@[(%a ->@ %a)@]" pp a pp b
+ | Iff (a, b) ->
+ fprintf fmt "@[(%a <->@ %a)@]" pp a pp b
+ | And (a, b) ->
+ fprintf fmt "@[(%a and@ %a)@]" pp a pp b
+ | Or (a, b) ->
+ fprintf fmt "@[(%a or@ %a)@]" pp a pp b
+ | Not a ->
+ fprintf fmt "@[(not@ %a)@]" pp a
+ | Forall (id, t, p) ->
+ fprintf fmt "@[(forall %a:%a.@ %a)@]" ident id print_typ t pp p
+ | Exists (id, t, p) ->
+ fprintf fmt "@[(exists %a:%a.@ %a)@]" ident id print_typ t pp p
+
+let print_query fmt (decls,concl) =
+ let print_dtype = function
+ | DeclType (id, 0) ->
+ fprintf fmt "@[type %a@]@\n@\n" ident id
+ | DeclType (id, 1) ->
+ fprintf fmt "@[type 'a %a@]@\n@\n" ident id
+ | DeclType (id, n) ->
+ fprintf fmt "@[type (";
+ for i = 1 to n do
+ fprintf fmt "'a%d" i; if i < n then fprintf fmt ", "
+ done;
+ fprintf fmt ") %a@]@\n@\n" ident id
+ | DeclFun _ | DeclPred _ | Axiom _ ->
+ ()
+ in
+ let print_dvar_dpred = function
+ | DeclFun (id, _, [], t) ->
+ fprintf fmt "@[logic %a : -> %a@]@\n@\n" ident id print_typ t
+ | DeclFun (id, _, l, t) ->
+ fprintf fmt "@[logic %a : %a -> %a@]@\n@\n"
+ ident id (print_list comma print_typ) l print_typ t
+ | DeclPred (id, _, []) ->
+ fprintf fmt "@[logic %a : -> prop @]@\n@\n" ident id
+ | DeclPred (id, _, l) ->
+ fprintf fmt "@[logic %a : %a -> prop@]@\n@\n"
+ ident id (print_list comma print_typ) l
+ | DeclType _ | Axiom _ ->
+ ()
+ in
+ let print_assert = function
+ | Axiom (id, f) ->
+ fprintf fmt "@[<hov 2>axiom %a:@ %a@]@\n@\n" ident id print_predicate f
+ | DeclType _ | DeclFun _ | DeclPred _ ->
+ ()
+ in
+ List.iter print_dtype decls;
+ List.iter print_dvar_dpred decls;
+ List.iter print_assert decls;
+ fprintf fmt "@[<hov 2>goal coq___goal: %a@]" print_predicate concl
+
+let output_file f q =
+ let c = open_out f in
+ let fmt = formatter_of_out_channel c in
+ fprintf fmt "@[%a@]@." print_query q;
+ close_out c
+
+
diff --git a/contrib/dp/dp_zenon.ml b/contrib/dp/dp_zenon.ml
new file mode 100644
index 00000000..57b0a44f
--- /dev/null
+++ b/contrib/dp/dp_zenon.ml
@@ -0,0 +1,103 @@
+
+open Format
+open Fol
+
+let rec print_list sep print fmt = function
+ | [] -> ()
+ | [x] -> print fmt x
+ | x :: r -> print fmt x; sep fmt (); print_list sep print fmt r
+
+let space fmt () = fprintf fmt "@ "
+
+let rec print_term fmt = function
+ | Cst n ->
+ fprintf fmt "%d" n
+ | Plus (a, b) ->
+ fprintf fmt "@[(+@ %a@ %a)@]" print_term a print_term b
+ | Moins (a, b) ->
+ fprintf fmt "@[(-@ %a@ %a)@]" print_term a print_term b
+ | Mult (a, b) ->
+ fprintf fmt "@[(*@ %a@ %a)@]" print_term a print_term b
+ | Div (a, b) ->
+ fprintf fmt "@[(/@ %a@ %a)@]" print_term a print_term b
+ | App (id, []) ->
+ fprintf fmt "%s" id
+ | App (id, tl) ->
+ fprintf fmt "@[(%s@ %a)@]" id print_terms tl
+
+and print_terms fmt tl =
+ print_list space print_term fmt tl
+
+let rec print_predicate fmt p =
+ let pp = print_predicate in
+ match p with
+ | True ->
+ fprintf fmt "True"
+ | False ->
+ fprintf fmt "False"
+ | Fatom (Eq (a, b)) ->
+ fprintf fmt "@[(= %a@ %a)@]" print_term a print_term b
+ | Fatom (Le (a, b)) ->
+ fprintf fmt "@[(<= %a@ %a)@]" print_term a print_term b
+ | Fatom (Lt (a, b))->
+ fprintf fmt "@[(< %a@ %a)@]" print_term a print_term b
+ | Fatom (Ge (a, b)) ->
+ fprintf fmt "@[(>= %a@ %a)@]" print_term a print_term b
+ | Fatom (Gt (a, b)) ->
+ fprintf fmt "@[(> %a@ %a)@]" print_term a print_term b
+ | Fatom (Pred (id, tl)) ->
+ fprintf fmt "@[(%s@ %a)@]" id print_terms tl
+ | Imp (a, b) ->
+ fprintf fmt "@[(=>@ %a@ %a)@]" pp a pp b
+ | And (a, b) ->
+ fprintf fmt "@[(/\\@ %a@ %a)@]" pp a pp b
+ | Or (a, b) ->
+ fprintf fmt "@[(\\/@ %a@ %a)@]" pp a pp b
+ | Not a ->
+ fprintf fmt "@[(-.@ %a)@]" pp a
+ | Forall (id, t, p) ->
+ fprintf fmt "@[(A. ((%s \"%s\")@ %a))@]" id t pp p
+ | Exists (id, t, p) ->
+ fprintf fmt "@[(E. ((%s \"%s\")@ %a))@]" id t pp p
+
+let rec string_of_type_list = function
+ | [] -> ""
+ | e :: l' -> e ^ " -> " ^ (string_of_type_list l')
+
+let print_query fmt (decls,concl) =
+ let print_decl = function
+ | DeclVar (id, [], t) ->
+ fprintf fmt "@[;; %s: %s@]@\n" id t
+ | DeclVar (id, l, t) ->
+ fprintf fmt "@[;; %s: %s%s@]@\n"
+ id (string_of_type_list l) t
+ | DeclPred (id, l) ->
+ fprintf fmt "@[;; %s: %sBOOLEAN@]@\n"
+ id (string_of_type_list l)
+ | DeclType id ->
+ fprintf fmt "@[;; %s: TYPE@]@\n" id
+ | Assert (id, f) ->
+ fprintf fmt "@[\"%s\" %a@]@\n" id print_predicate f
+ in
+ List.iter print_decl decls;
+ fprintf fmt "$goal %a@." print_predicate concl
+
+let call q =
+ let f = Filename.temp_file "coq_dp" ".znn" in
+ let c = open_out f in
+ let fmt = formatter_of_out_channel c in
+ fprintf fmt "@[%a@]@." print_query q;
+ close_out c;
+ ignore (Sys.command (sprintf "cat %s" f));
+ let cmd =
+ sprintf "timeout 10 zenon %s > out 2>&1 && grep -q PROOF-FOUND out" f
+ in
+ prerr_endline cmd; flush stderr;
+ let out = Sys.command cmd in
+ if out = 0 then Valid
+ else if out = 1 then Invalid
+ else if out = 137 then Timeout
+ else Util.anomaly "malformed Zenon input file"
+ (* TODO: effacer le fichier f et le fichier out *)
+
+
diff --git a/contrib/dp/dp_zenon.mli b/contrib/dp/dp_zenon.mli
new file mode 100644
index 00000000..03b6d347
--- /dev/null
+++ b/contrib/dp/dp_zenon.mli
@@ -0,0 +1,4 @@
+
+open Fol
+
+val call : query -> prover_answer
diff --git a/contrib/dp/fol.mli b/contrib/dp/fol.mli
new file mode 100644
index 00000000..a85469cc
--- /dev/null
+++ b/contrib/dp/fol.mli
@@ -0,0 +1,48 @@
+
+(* Polymorphic First-Order Logic (that is Why's input logic) *)
+
+type typ =
+ | Tvar of string
+ | Tid of string * typ list
+
+type term =
+ | Cst of int
+ | Plus of term * term
+ | Moins of term * term
+ | Mult of term * term
+ | Div of term * term
+ | App of string * term list
+
+and atom =
+ | Eq of term * term
+ | Le of term * term
+ | Lt of term * term
+ | Ge of term * term
+ | Gt of term * term
+ | Pred of string * term list
+
+and form =
+ | Fatom of atom
+ | Imp of form * form
+ | Iff of form * form
+ | And of form * form
+ | Or of form * form
+ | Not of form
+ | Forall of string * typ * form
+ | Exists of string * typ * form
+ | True
+ | False
+
+(* the integer indicates the number of type variables *)
+type decl =
+ | DeclType of string * int
+ | DeclFun of string * int * typ list * typ
+ | DeclPred of string * int * typ list
+ | Axiom of string * form
+
+type query = decl list * form
+
+
+(* prover result *)
+
+type prover_answer = Valid | Invalid | DontKnow | Timeout
diff --git a/contrib/dp/g_dp.ml4 b/contrib/dp/g_dp.ml4
new file mode 100644
index 00000000..eb7fb73b
--- /dev/null
+++ b/contrib/dp/g_dp.ml4
@@ -0,0 +1,38 @@
+(************************************************************************)
+(* 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: g_dp.ml4 7165 2005-06-24 12:56:46Z coq $ *)
+
+open Dp
+
+TACTIC EXTEND Simplify
+ [ "simplify" ] -> [ simplify ]
+END
+
+TACTIC EXTEND CVCLite
+ [ "cvcl" ] -> [ cvc_lite ]
+END
+
+TACTIC EXTEND Harvey
+ [ "harvey" ] -> [ harvey ]
+END
+
+TACTIC EXTEND Zenon
+ [ "zenon" ] -> [ zenon ]
+END
+
+(* should be part of basic tactics syntax *)
+TACTIC EXTEND admit
+ [ "admit" ] -> [ Tactics.admit_as_an_axiom ]
+END
+
+VERNAC COMMAND EXTEND Dp_hint
+ [ "Dp_hint" ne_global_list(l) ] -> [ dp_hint l ]
+END
diff --git a/contrib/dp/test2.v b/contrib/dp/test2.v
new file mode 100644
index 00000000..4e933a3c
--- /dev/null
+++ b/contrib/dp/test2.v
@@ -0,0 +1,78 @@
+Require Import ZArith.
+Require Import Classical.
+Require Import List.
+
+Open Scope list_scope.
+Open Scope Z_scope.
+
+Definition neg (z:Z) : Z := match z with
+ | Z0 => Z0
+ | Zpos p => Zneg p
+ | Zneg p => Zpos p
+ end.
+
+Goal forall z, neg (neg z) = z.
+ Admitted.
+
+Open Scope nat_scope.
+Print plus.
+
+Goal forall x, x+0=x.
+ induction x.
+ zenon.
+ zenon.
+ (* simplify resoud le premier, pas le second *)
+ Admitted.
+
+Goal 1::2::3::nil = 1::2::(1+2)::nil.
+ zenon.
+ Admitted.
+
+Definition T := nat.
+Parameter fct : T -> nat.
+Goal fct O = O.
+ Admitted.
+
+Fixpoint even (n:nat) : Prop :=
+ match n with
+ O => True
+ | S O => False
+ | S (S p) => even p
+ end.
+
+Goal even 4%nat.
+ try zenon.
+ Admitted.
+
+Definition p (A B:Set) (a:A) (b:B) : list (A*B) := cons (a,b) nil.
+
+Definition head :=
+fun (A : Set) (l : list A) =>
+match l with
+| nil => None (A:=A)
+| x :: _ => Some x
+end.
+
+Goal forall x, head _ (p _ _ 1 2) = Some x -> fst x = 1.
+
+Admitted.
+
+(*
+BUG avec head prédéfini : manque eta-expansion sur A:Set
+
+Goal forall x, head _ (p _ _ 1 2) = Some x -> fst x = 1.
+
+Print value.
+Print Some.
+
+zenon.
+*)
+
+Inductive IN (A:Set) : A -> list A -> Prop :=
+ | IN1 : forall x l, IN A x (x::l)
+ | IN2: forall x l, IN A x l -> forall y, IN A x (y::l).
+Implicit Arguments IN [A].
+
+Goal forall x, forall (l:list nat), IN x l -> IN x (1%nat::l).
+ zenon.
+Print In.
diff --git a/contrib/dp/tests.v b/contrib/dp/tests.v
new file mode 100644
index 00000000..52a57a0c
--- /dev/null
+++ b/contrib/dp/tests.v
@@ -0,0 +1,220 @@
+
+Require Import ZArith.
+Require Import Classical.
+
+(* First example with the 0 and the equality translated *)
+
+Goal 0 = 0.
+zenon.
+Qed.
+
+
+(* Examples in the Propositional Calculus
+ and theory of equality *)
+
+Parameter A C : Prop.
+
+Goal A -> A.
+zenon.
+Qed.
+
+
+Goal A -> (A \/ C).
+
+zenon.
+Qed.
+
+
+Parameter x y z : Z.
+
+Goal x = y -> y = z -> x = z.
+
+zenon.
+Qed.
+
+
+Goal ((((A -> C) -> A) -> A) -> C) -> C.
+
+zenon.
+Qed.
+
+
+(* Arithmetic *)
+Open Scope Z_scope.
+
+Goal 1 + 1 = 2.
+simplify.
+Qed.
+
+
+Goal 2*x + 10 = 18 -> x = 4.
+
+simplify.
+Qed.
+
+
+(* Universal quantifier *)
+
+Goal (forall (x y : Z), x = y) -> 0=1.
+try zenon.
+simplify.
+Qed.
+
+Goal forall (x: nat), (x + 0 = x)%nat.
+
+induction x0.
+zenon.
+zenon.
+Qed.
+
+
+(* No decision procedure can solve this problem
+ Goal forall (x a b : Z), a * x + b = 0 -> x = - b/a.
+*)
+
+
+(* Functions definitions *)
+
+Definition fst (x y : Z) : Z := x.
+
+Goal forall (g : Z -> Z) (x y : Z), g (fst x y) = g x.
+
+simplify.
+Qed.
+
+
+(* Eta-expansion example *)
+
+Definition snd_of_3 (x y z : Z) : Z := y.
+
+Definition f : Z -> Z -> Z := snd_of_3 0.
+
+Goal forall (x y z z1 : Z), snd_of_3 x y z = f y z1.
+
+simplify.
+Qed.
+
+
+(* Inductive types definitions - call to incontrib/dp/jection function *)
+
+Inductive even : Z -> Prop :=
+| even_0 : even 0
+| even_plus2 : forall z : Z, even z -> even (z + 2).
+
+
+(* Simplify and Zenon can't prove this goal before the timeout
+ unlike CVC Lite *)
+
+Goal even 4.
+cvcl.
+Qed.
+
+
+Definition skip_z (z : Z) (n : nat) := n.
+
+Definition skip_z1 := skip_z.
+
+Goal forall (z : Z) (n : nat), skip_z z n = skip_z1 z n.
+
+zenon.
+Qed.
+
+
+(* Axioms definitions and dp_hint *)
+
+Parameter add : nat -> nat -> nat.
+Axiom add_0 : forall (n : nat), add 0%nat n = n.
+Axiom add_S : forall (n1 n2 : nat), add (S n1) n2 = S (add n1 n2).
+
+Dp_hint add_0.
+Dp_hint add_S.
+
+(* Simplify can't prove this goal before the timeout
+ unlike zenon *)
+
+Goal forall n : nat, add n 0 = n.
+
+induction n ; zenon.
+Qed.
+
+
+Definition pred (n : nat) : nat := match n with
+ | 0%nat => 0%nat
+ | S n' => n'
+end.
+
+Goal forall n : nat, n <> 0%nat -> pred (S n) <> 0%nat.
+
+zenon.
+Qed.
+
+
+Fixpoint plus (n m : nat) {struct n} : nat :=
+ match n with
+ | 0%nat => m
+ | S n' => S (plus n' m)
+end.
+
+Goal forall n : nat, plus n 0%nat = n.
+
+induction n; zenon.
+Qed.
+
+
+(* Mutually recursive functions *)
+
+Fixpoint even_b (n : nat) : bool := match n with
+ | O => true
+ | S m => odd_b m
+end
+with odd_b (n : nat) : bool := match n with
+ | O => false
+ | S m => even_b m
+end.
+
+Goal even_b (S (S O)) = true.
+
+zenon.
+Qed.
+
+
+(* sorts issues *)
+
+Parameter foo : Set.
+Parameter ff : nat -> foo -> foo -> nat.
+Parameter g : foo -> foo.
+Goal (forall x:foo, ff 0 x x = O) -> forall y, ff 0 (g y) (g y) = O.
+zenon.
+Qed.
+
+
+
+(* abstractions *)
+
+Parameter poly_f : forall A:Set, A->A.
+
+Goal forall x:nat, poly_f nat x = poly_f nat x.
+zenon.
+Qed.
+
+
+
+(* Anonymous mutually recursive functions : no equations are produced
+
+Definition mrf :=
+ fix even2 (n : nat) : bool := match n with
+ | O => true
+ | S m => odd2 m
+ end
+ with odd2 (n : nat) : bool := match n with
+ | O => false
+ | S m => even2 m
+ end for even.
+
+ Thus this goal is unsolvable
+
+Goal mrf (S (S O)) = true.
+
+zenon.
+
+*)