ssreflect-1.5/0000755000175000017500000000000012175455021012337 5ustar garesgaresssreflect-1.5/INSTALL0000644000175000017500000000570612175455021013400 0ustar garesgaresINSTALLATION PROCEDURE FOR SSREFLECT ------------------------------------ Summary: - Requirements - Building Ssreflect - Customization of the Proof General Emacs interface - REQUIREMENTS ============ - Coq version 8.3pl4 or 8.4 BUILDING SSREFLECT ================================================== - We suppose that the Coq system has been compiled from sources and installed on your system using a standard installation process. - Your COQBIN environment variable must be point to directory where Coq's binaries are. - Your PATH environment variable value must make Coq binaries (coqtop, coq_makefile,...) accessible. E.g. export PATH="$COQBIN:$PATH" - Download and unpack the archive of the ssreflect distribution in a directory whose name does not contain spaces (like " Documents and Settings" ...). - Go to the root of the ssreflect directory creating by the previous unpack and type `make && make install`. - Note for Mac OSX users: If you encounter a "stack overflow" error message, then set the following environment variable and try again: OCAMLOPT=ocamlopt.opt - Note for Mac OSX users and Windows users: If you encounter a linking error it may be caused by an old version of OCaml that does not support dynamic loding of plugins on you platform. Edit the Make file and uncomment the lines marked for static linking. In this way a statically linked binary bin/ssrcoq will be produced. It is up to you to copy it inside $COQBIN. CUSTOMIZATION OF THE PROOF GENERAL EMACS INTERFACE ================================================== The ssreflect distribution comes with a small configuration file pg-ssr.el to extend ProofGeneral (PG), a generic interface for proof assistants based on the customizable text editor Emacs, to the syntax of ssreflect. The >= 3.7 versions of ProofGeneral support this extension. - Following the installation instructions, unpack the sources of PG in a directory, for instance /ProofGeneral-3.7, and add the following line to your .emacs file: - under Unix/MacOS: (load-file "/ProofGeneral-3.7/generic/proof-site.el" ) - under Windows+Cygwin: (load-file "C:\\\\ProofGeneral-3.7\\generic\\proof-site.el") where is the location of your own ProofGeneral directory. -Immediately after the previous line added to your .emacs, add this one: (load-file "/pg-ssr.el") respectively (load-file "\\pg-ssr.el") for Windows+Cygwin users, where is the location of your pg-ssr.el file. Coq sources have a .v extension. Opening any .v file should automatically launch ProofGeneral. Try this on a foo.v file. In the menu 'ProofGeneral', choose the item: 'Advanced/Customize/Coq/Coq Prog Name' Change the value of the variable to /bin/ssrcoq or \\bin\\ssrcoq for Windows+Cygwin users, where is the location of your coq directory. ssreflect-1.5/AUTHORS0000644000175000017500000000172112175455021013410 0ustar garesgaresAndrea Asperti University of Bologna - Microsoft Inria Joint Centre Jeremy Avigad Carnegie Mellon University - Microsoft Inria Joint Centre Yves Bertot Inria Sophia Antipolis - Microsoft Inria Joint Centre Cyril Cohen LIX École Polytechnique - Microsoft Inria Joint Centre François Garillot Microsoft Inria Joint Centre Georges Gonthier Microsoft Research Cambridge - Microsoft Inria Joint Centre Stéphane Le Roux Microsoft Inria Joint Centre Assia Mahboubi Inria Saclay - Microsoft Inria Joint Centre Sidi Ould Biha Inria Sophia Antipolis - Microsoft Inria Joint Centre Ioana Pasca Inria Sophia Antipolis - Microsoft Inria Joint Centre Laurence Rideau Inria Sophia Antipolis - Microsoft Inria Joint Centre Enrico Tassi Inria Saclay - Microsoft Inria Joint Centre Laurent Théry Inria Sophia Antipolis - Microsoft Inria Joint Centre Russell O'Connor Mc Master University - Microsoft Inria Joint Centre ssreflect-1.5/ANNOUNCE0000644000175000017500000000625612175455021013501 0ustar garesgaresAnnouncing Ssreflect version 1.4 We are pleased to announce the new release of the Ssreflect extension for the Coq proof assistant. This release includes: - an update of the tactic language which complies with the new version of Coq; - an extended update of the libraries distributed in the previous release of ssreflect. The name Ssreflect stands for "small scale reflection", a style of proof that evolved from the computer-checked proof of the Four Colour Theorem and which leverages the higher-order nature of Coq's underlying logic to provide effective automation for many small, clerical proof steps. This is often accomplished by restating ("reflecting") problems in a more concrete form, hence the name. For example, in the Ssreflect library arithmetic comparison is not an abstract predicate, but a function computing a boolean. Along with documentation, also available at [1], the Ssreflect distribution comprises two parts: - A new tactic language, which promotes more structured, concise and robust proof scripts, and is in fact independent from the "reflection" proof style. It is implemented as a linkable extension to the Coq system. - A set of Coq libraries that provide core "reflection-oriented" theories for + basic combinatorics: arithmetic, lists, graphs, and finite sets. + abstract algebra: an algebraic hierarchy from additive groups to closed fields, polynomials, matrix, general linear algebra, basic and advanced finite group theory, infrastructure for finite summations,...) Some features of the tactic language: - It provides tacticals for most structural steps (e.g., moving assumptions), so that script steps mostly match logical steps. - It provides tactics and tatical to support structured layout, including reordering subgoals and supporting "without loss of generality" arguments. - It provides a powerful rewriting tactic that supports chained rules, automatic unfolding of definitions and conditional rewriting, as well as precise control over where and how much rewriting occurs. - It can be used in combination with the classic Coq tactic language. Some features of the library: - Exploits advanced features of Coq such as coercions and canonical projections to build generic theories (e.g., for decidable equality). - Uses rewrite rules and dependent predicate families to state lemmas that can be applied deeply and bidirectionally. This means fewer structural steps and a smaller library, respectively. - Uses boolean functions to represent sets (i.e., comprehensions), rather than an ad hoc set algebra. The Ssreflect 1.4 distribution is available at [2]. It is distributed under the CeCILL-B licence (the French equivalent of the BSD license). Comments and bug reports are of course most welcome, and can be directed at ssreflect(at-sign)msr-inria.inria.fr. To subscribe, either send an email to sympa@msr-inria.inria.fr, whose title contains the word ssreflect, or use the following web interface at [3]. Enjoy! The Mathematical Components Team, at the Microsoft Research-Inria Joint Center [1] : http://hal.inria.fr/inria-00258384 [2] : http://www.msr-inria.inria.fr/Projects/math-components [3] : https://www.msr-inria.inria.fr/sympa ssreflect-1.5/pg-ssr.el0000644000175000017500000000336212175455021014100 0ustar garesgares;; Customization of PG for ssreflect syntax ;; Assia Mahboubi 2007 (defcustom coq-user-tactics-db '(("nat_congr" "ncongr" "nat_congr" t "nat_congr") ("nat_norm" "nnorm" "nat_norm" t "nat_norm") ("bool_congr" "bcongr" "bool_congr" t "bool_congr") ("prop_congr" "prcongr" "prop_congr" t "prop_congr") ("move" "m" "move" t "move") ("set" "set" "set # := #" t "set") ("have" "hv" "have # : #" t "have") ("congr" "con" "congr #" t "congr") ("wlog" "wlog" "wlog : / #" t "wlog") ("without loss" "wilog" "without loss #" t "without loss") ("unlock" "unlock" "unlock #" t "unlock") ("suffices" "suffices" "suffices # : #" t "suffices") ("suff" "suff" "suff # : #" t "suff") ) "Extended list of tactics, includings ssr and user defined ones") (defcustom coq-user-commands-db '(("Prenex Implicits" "pi" "Prenex Implicits #" t "Prenex\\s-+Implicits") ("Hint View for" "hv" "Hint View for #" t "Hint\\s-+View\\s-+for") ("inside" "ins" nil f "inside") ("outside" "outs" nil f "outside") ("Canonical " nil "Canonical #." t "Canonical") ) "Extended list of commands, includings ssr and user defined ones") (defcustom coq-user-tacticals-db '(("last" "lst" nil t "last")) "Extended list of tacticals, includings ssr and user defined ones") (defcustom coq-user-reserved-db '("is" "nosimpl" "of") "Extended list of keywords, includings ssr and user defined ones") (defcustom coq-user-solve-tactics-db '(("done" nil "done" nil "done") ) "Extended list of closing tactic(al)s, includings ssr and user defined ones") ;; This works only with the cvs version (> 3.7) of Proof General. (defcustom coq-variable-highlight-enable nil "Activates partial bound variable highlighting" ) ssreflect-1.5/src/0000755000175000017500000000000012175455021013126 5ustar garesgaresssreflect-1.5/src/ssrmatching.ml40000644000175000017500000013727212175455021016102 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) (* Defining grammar rules with "xx" in it automatically declares keywords too, * we thus save the lexer to restore it at the end of the file *) let frozen_lexer = Lexer.freeze () ;; (*i camlp4use: "pa_extend.cmo" i*) (*i camlp4deps: "grammar/grammar.cma" i*) open Names open Pp open Pcoq open Genarg open Term open Topconstr open Libnames open Tactics open Tacticals open Termops open Namegen open Recordops open Tacmach open Coqlib open Glob_term open Util open Evd open Extend open Goptions open Tacexpr open Tacinterp open Pretyping open Constr open Tactic open Extraargs open Ppconstr open Printer type loc = Util.loc let errorstrm = errorlabstrm "ssreflect" let loc_error loc msg = user_err_loc (loc, msg, str msg) (* 0 cost pp function. Active only if env variable SSRDEBUG is set *) (* or if SsrDebug is Set *) let pp_ref = ref (fun _ -> ()) let ssr_pp s = pperrnl (str"SSR: "++Lazy.force s) let _ = try ignore(Sys.getenv "SSRDEBUG"); pp_ref := ssr_pp with Not_found -> () let debug b = if b then pp_ref := ssr_pp else pp_ref := fun _ -> () let _ = Goptions.declare_bool_option { Goptions.optsync = false; Goptions.optname = "ssrmatching debugging"; Goptions.optkey = ["SsrMatchingDebug"]; Goptions.optdepr = false; Goptions.optread = (fun _ -> !pp_ref == ssr_pp); Goptions.optwrite = debug } let pp s = !pp_ref s (** Utils {{{ *****************************************************************) let env_size env = List.length (Environ.named_context env) let safeDestApp c = match kind_of_term c with App (f, a) -> f, a | _ -> c, [| |] let get_index = function ArgArg i -> i | _ -> anomaly "Uninterpreted index" (* Toplevel constr must be globalized twice ! *) let glob_constr ist gsigma genv = function | _, Some ce -> let ltacvars = List.map fst ist.lfun, [] in Constrintern.intern_gen false ~ltacvars:ltacvars gsigma genv ce | rc, None -> rc (* Term printing utilities functions for deciding bracketing. *) let pr_paren prx x = hov 1 (str "(" ++ prx x ++ str ")") (* String lexing utilities *) let skip_wschars s = let rec loop i = match s.[i] with '\n'..' ' -> loop (i + 1) | _ -> i in loop (* We also guard characters that might interfere with the ssreflect *) (* tactic syntax. *) let guard_term ch1 s i = match s.[i] with | '(' -> false | '{' | '/' | '=' -> true | _ -> ch1 = '(' (* The call 'guard s i' should return true if the contents of s *) (* starting at i need bracketing to avoid ambiguities. *) let pr_guarded guard prc c = msg_with Format.str_formatter (prc c); let s = Format.flush_str_formatter () ^ "$" in if guard s (skip_wschars s 0) then pr_paren prc c else prc c (* More sensible names for constr printers *) let pr_constr = pr_constr let prl_glob_constr c = pr_lglob_constr_env (Global.env ()) c let pr_glob_constr c = pr_glob_constr_env (Global.env ()) c let prl_constr_expr = pr_lconstr_expr let pr_constr_expr = pr_constr_expr let prl_glob_constr_and_expr = function | _, Some c -> prl_constr_expr c | c, None -> prl_glob_constr c let pr_glob_constr_and_expr = function | _, Some c -> pr_constr_expr c | c, None -> pr_glob_constr c let pr_term (k, c) = pr_guarded (guard_term k) pr_glob_constr_and_expr c let prl_term (k, c) = pr_guarded (guard_term k) prl_glob_constr_and_expr c (** Adding a new uninterpreted generic argument type *) let add_genarg tag pr = let wit, globwit, rawwit as wits = create_arg None tag in let glob _ rarg = in_gen globwit (out_gen rawwit rarg) in let interp _ gl garg = Tacmach.project gl,in_gen wit (out_gen globwit garg) in let subst _ garg = garg in add_interp_genarg tag (glob, interp, subst); let gen_pr _ _ _ = pr in Pptactic.declare_extra_genarg_pprule (rawwit, gen_pr) (globwit, gen_pr) (wit, gen_pr); wits (** Constructors for cast type *) let dC t = CastConv (DEFAULTcast, t) (** Constructors for constr_expr *) let isCVar = function CRef (Ident _) -> true | _ -> false let destCVar = function CRef (Ident (_, id)) -> id | _ -> anomaly "not a CRef" let mkCHole loc = CHole (loc, None) let mkCLambda loc name ty t = CLambdaN (loc, [[loc, name], Default Explicit, ty], t) let mkCLetIn loc name bo t = CLetIn (loc, (loc, name), bo, t) let mkCCast loc t ty = CCast (loc,t, dC ty) (** Constructors for rawconstr *) let mkRHole = GHole (dummy_loc, InternalHole) let mkRApp f args = if args = [] then f else GApp (dummy_loc, f, args) let mkRCast rc rt = GCast (dummy_loc, rc, dC rt) let mkRLambda n s t = GLambda (dummy_loc, n, Explicit, s, t) (* ssrterm conbinators *) let combineCG t1 t2 f g = match t1, t2 with | (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None) | (x, (_, Some t1)), (_, (_, Some t2)) -> x, (mkRHole, Some (f t1 t2)) | _, (_, (_, None)) -> anomaly "have: mixed C-G constr" | _ -> anomaly "have: mixed G-C constr" let loc_ofCG = function | (_, (s, None)) -> loc_of_glob_constr s | (_, (_, Some s)) -> constr_loc s let mk_term k c = k, (mkRHole, Some c) let mk_lterm = mk_term ' ' (* }}} *) (** Profiling {{{ *************************************************************) type profiler = { profile : 'a 'b. ('a -> 'b) -> 'a -> 'b; reset : unit -> unit; print : unit -> unit } let profile_now = ref false let something_profiled = ref false let profilers = ref [] let add_profiler f = profilers := f :: !profilers;; let profile b = profile_now := b; if b then List.iter (fun f -> f.reset ()) !profilers; if not b then List.iter (fun f -> f.print ()) !profilers ;; let _ = Goptions.declare_bool_option { Goptions.optsync = false; Goptions.optname = "ssrmatching profiling"; Goptions.optkey = ["SsrMatchingProfiling"]; Goptions.optread = (fun _ -> !profile_now); Goptions.optdepr = false; Goptions.optwrite = profile } let () = let prof_total = let init = ref 0.0 in { profile = (fun f x -> assert false); reset = (fun () -> init := Unix.gettimeofday ()); print = (fun () -> if !something_profiled then prerr_endline (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" "total" 0 (Unix.gettimeofday() -. !init) 0.0 0.0)) } in let prof_legenda = { profile = (fun f x -> assert false); reset = (fun () -> ()); print = (fun () -> if !something_profiled then begin prerr_endline (Printf.sprintf "!! %39s ---------- --------- --------- ---------" (String.make 39 '-')); prerr_endline (Printf.sprintf "!! %-39s %10s %9s %9s %9s" "function" "#calls" "total" "max" "average") end) } in add_profiler prof_legenda; add_profiler prof_total ;; let mk_profiler s = let total, calls, max = ref 0.0, ref 0, ref 0.0 in let reset () = total := 0.0; calls := 0; max := 0.0 in let profile f x = if not !profile_now then f x else let before = Unix.gettimeofday () in try incr calls; let res = f x in let after = Unix.gettimeofday () in let delta = after -. before in total := !total +. delta; if delta > !max then max := delta; res with exc -> let after = Unix.gettimeofday () in let delta = after -. before in total := !total +. delta; if delta > !max then max := delta; raise exc in let print () = if !calls <> 0 then begin something_profiled := true; prerr_endline (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" s !calls !total !max (!total /. (float_of_int !calls))) end in let prof = { profile = profile; reset = reset; print = print } in add_profiler prof; prof ;; (* }}} *) exception NoProgress (** Unification procedures. *) (* To enforce the rigidity of the rooted match we always split *) (* top applications, so the unification procedures operate on *) (* arrays of patterns and terms. *) (* We perform three kinds of unification: *) (* EQ: exact conversion check *) (* FO: first-order unification of evars, without conversion *) (* HO: higher-order unification with conversion *) (* The subterm unification strategy is to find the first FO *) (* match, if possible, and the first HO match otherwise, then *) (* compute all the occurrences that are EQ matches for the *) (* relevant subterm. *) (* Additional twists: *) (* - If FO/HO fails then we attempt to fill evars using *) (* typeclasses before raising an outright error. We also *) (* fill typeclasses even after a successful match, since *) (* beta-reduction and canonical instances may leave *) (* undefined evars. *) (* - We do postchecks to rule out matches that are not *) (* closed or that assign to a global evar; these can be *) (* disabled for rewrite or dependent family matches. *) (* - We do a full FO scan before turning to HO, as the FO *) (* comparison can be much faster than the HO one. *) let unif_EQ env sigma p c = let evars = existential_opt_value sigma in try let _ = Reduction.conv env p ~evars c in true with _ -> false let unif_EQ_args env sigma pa a = let n = Array.length pa in let rec loop i = (i = n) || unif_EQ env sigma pa.(i) a.(i) && loop (i + 1) in loop 0 let prof_unif_eq_args = mk_profiler "unif_EQ_args";; let unif_EQ_args env sigma pa a = prof_unif_eq_args.profile (unif_EQ_args env sigma pa) a ;; let unif_HO env ise p c = Evarconv.the_conv_x env p c ise let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env p c ise let unif_HO_args env ise0 pa i ca = let n = Array.length pa in let rec loop ise j = if j = n then ise else loop (unif_HO env ise pa.(j) ca.(i + j)) (j + 1) in loop ise0 0 (* FO unification should boil down to calling w_unify with no_delta, but *) (* alas things are not so simple: w_unify does partial type-checking, *) (* which breaks down when the no-delta flag is on (as the Coq type system *) (* requires full convertibility. The workaround here is to convert all *) (* evars into metas, since 8.2 does not TC metas. This means some lossage *) (* for HO evars, though hopefully Miller patterns can pick up some of *) (* those cases, and HO matching will mop up the rest. *) let flags_FO = {Unification.default_no_delta_unify_flags with Unification.modulo_conv_on_closed_terms = None; Unification.modulo_delta_types = full_transparent_state; Unification.allow_K_in_toplevel_higher_order_unification=false} let unif_FO env ise p c = Unification.w_unify env ise Reduction.CONV ~flags:flags_FO p c (* Perform evar substitution in main term and prune substitution. *) let nf_open_term sigma0 ise c = let s = ise and s' = ref sigma0 in let rec nf c' = match kind_of_term c' with | Evar ex -> begin try nf (existential_value s ex) with _ -> let k, a = ex in let a' = Array.map nf a in if not (Evd.mem !s' k) then s' := Evd.add !s' k (Evarutil.nf_evar_info s (Evd.find s k)); mkEvar (k, a') end | _ -> map_constr nf c' in let copy_def k evi () = if evar_body evi != Evd.Evar_empty then () else match Evd.evar_body (Evd.find s k) with | Evar_defined c' -> s' := Evd.define k (nf c') !s' | _ -> () in let c' = nf c in let _ = Evd.fold copy_def sigma0 () in !s', c' let unif_end env sigma0 ise0 pt ok = let ise = Evarconv.consider_remaining_unif_problems env ise0 in let s, t = nf_open_term sigma0 ise pt in let ise1 = create_evar_defs s in let ise2 = Typeclasses.resolve_typeclasses ~fail:true env ise1 in if not (ok ise) then raise NoProgress else if ise2 == ise1 then (s, t) else nf_open_term sigma0 ise2 t let pf_unif_HO gl sigma pt p c = let env = pf_env gl in let ise = unif_HO env (create_evar_defs sigma) p c in unif_end env (project gl) ise pt (fun _ -> true) let unify_HO env sigma0 t1 t2 = let sigma = unif_HO env sigma0 t1 t2 in let sigma, _ = unif_end env sigma0 sigma t2 (fun _ -> true) in sigma let pf_unify_HO gl t1 t2 = let env, sigma0, si = pf_env gl, project gl, sig_it gl in let sigma = unify_HO env sigma0 t1 t2 in re_sig si sigma (* This is what the definition of iter_constr should be... *) let iter_constr_LR f c = match kind_of_term c with | Evar (k, a) -> Array.iter f a | Cast (cc, _, t) -> f cc; f t | Prod (_, t, b) | Lambda (_, t, b) -> f t; f b | LetIn (_, v, t, b) -> f v; f t; f b | App (cf, a) -> f cf; Array.iter f a | Case (_, p, v, b) -> f v; f p; Array.iter f b | Fix (_, (_, t, b)) | CoFix (_, (_, t, b)) -> for i = 0 to Array.length t - 1 do f t.(i); f b.(i) done | _ -> () (* The comparison used to determine which subterms matches is KEYED *) (* CONVERSION. This looks for convertible terms that either have the same *) (* same head constant as pat if pat is an application (after beta-iota), *) (* or start with the same constr constructor (esp. for LetIn); this is *) (* disregarded if the head term is let x := ... in x, and casts are always *) (* ignored and removed). *) (* Record projections get special treatment: in addition to the projection *) (* constant itself, ssreflect also recognizes head constants of canonical *) (* projections. *) exception NoMatch type ssrdir = L2R | R2L let pr_dir_side = function L2R -> str "LHS" | R2L -> str "RHS" let inv_dir = function L2R -> R2L | R2L -> L2R type pattern_class = | KpatFixed | KpatConst | KpatEvar of existential_key | KpatLet | KpatLam | KpatRigid | KpatFlex | KpatProj of constant type tpattern = { up_k : pattern_class; up_FO : constr; up_f : constr; up_a : constr array; up_t : constr; (* equation proof term or matched term *) up_dir : ssrdir; (* direction of the rule *) up_ok : constr -> evar_map -> bool; (* progess test for rewrite *) } let all_ok _ _ = true let proj_nparams c = try 1 + Recordops.find_projection_nparams (ConstRef c) with _ -> 0 let isFixed c = match kind_of_term c with | Var _ | Ind _ | Construct _ | Const _ -> true | _ -> false let isRigid c = match kind_of_term c with | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true | _ -> false exception UndefPat let hole_var = mkVar (id_of_string "_") let pr_constr_pat c0 = let rec wipe_evar c = if isEvar c then hole_var else map_constr wipe_evar c in pr_constr (wipe_evar c0) (* Turn (new) evars into metas *) let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 = let ise = ref ise0 in let sigma = ref ise0 in let nenv = env_size env + if hack then 1 else 0 in let rec put c = match kind_of_term c with | Evar (k, a as ex) -> begin try put (existential_value !sigma ex) with NotInstantiatedEvar -> if Evd.mem sigma0 k then map_constr put c else let evi = Evd.find !sigma k in let dc = list_firstn (max 0 (Array.length a - nenv)) (evar_filtered_context evi) in let abs_dc (d, c) = function | x, Some b, t -> d, mkNamedLetIn x (put b) (put t) c | x, None, t -> mkVar x :: d, mkNamedProd x (put t) c in let a, t = Sign.fold_named_context_reverse abs_dc ~init:([], (put evi.evar_concl)) dc in let m = Evarutil.new_meta () in ise := meta_declare m t !ise; sigma := Evd.define k (applist (mkMeta m, a)) !sigma; put (existential_value !sigma ex) end | _ -> map_constr put c in let c1 = put c0 in !ise, c1 (* Compile a match pattern from a term; t is the term to fill. *) (* p_origin can be passed to obtain a better error message *) let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p = let k, f, a = let f, a = Reductionops.whd_betaiota_stack ise p in match kind_of_term f with | Const p -> let np = proj_nparams p in if np = 0 || np > List.length a then KpatConst, f, a else let a1, a2 = list_chop np a in KpatProj p, applist(f, a1), a2 | Var _ | Ind _ | Construct _ -> KpatFixed, f, a | Evar (k, _) -> if Evd.mem sigma0 k then KpatEvar k, f, a else if a <> [] then KpatFlex, f, a else (match p_origin with None -> error "indeterminate pattern" | Some (dir, rule) -> errorstrm (str "indeterminate " ++ pr_dir_side dir ++ str " in " ++ pr_constr_pat rule)) | LetIn (_, v, _, b) -> if b <> mkRel 1 then KpatLet, f, a else KpatFlex, v, a | Lambda _ -> KpatLam, f, a | _ -> KpatRigid, f, a in let aa = Array.of_list a in let ise', p' = evars_for_FO ~hack env sigma0 ise (mkApp (f, aa)) in ise', { up_k = k; up_FO = p'; up_f = f; up_a = aa; up_ok = ok; up_dir = dir; up_t = t} (* Specialize a pattern after a successful match: assign a precise head *) (* kind and arity for Proj and Flex patterns. *) let ungen_upat lhs (sigma, t) u = let f, a = safeDestApp lhs in let k = match kind_of_term f with | Var _ | Ind _ | Construct _ -> KpatFixed | Const _ -> KpatConst | Evar (k, _) -> if is_defined sigma k then raise NoMatch else KpatEvar k | LetIn _ -> KpatLet | Lambda _ -> KpatLam | _ -> KpatRigid in sigma, {u with up_k = k; up_FO = lhs; up_f = f; up_a = a; up_t = t} let nb_cs_proj_args pc f u = let na k = List.length (lookup_canonical_conversion (ConstRef pc, k)).o_TCOMPS in try match kind_of_term f with | Prod _ -> na Prod_cs | Sort s -> na (Sort_cs (family_of_sort s)) | Const c' when c' = pc -> Array.length (snd (destApp u.up_f)) | Var _ | Ind _ | Construct _ | Const _ -> na (Const_cs (global_of_constr f)) | _ -> -1 with Not_found -> -1 let isEvar_k k f = match kind_of_term f with Evar (k', _) -> k = k' | _ -> false let nb_args c = match kind_of_term c with App (_, a) -> Array.length a | _ -> 0 let mkSubArg i a = if i = Array.length a then a else Array.sub a 0 i let mkSubApp f i a = if i = 0 then f else mkApp (f, mkSubArg i a) let splay_app ise = let rec loop c a = match kind_of_term c with | App (f, a') -> loop f (Array.append a' a) | Cast (c', _, _) -> loop c' a | Evar ex -> (try loop (existential_value ise ex) a with _ -> c, a) | _ -> c, a in fun c -> match kind_of_term c with | App (f, a) -> loop f a | Cast _ | Evar _ -> loop c [| |] | _ -> c, [| |] let filter_upat i0 f n u fpats = let na = Array.length u.up_a in if n < na then fpats else let np = match u.up_k with | KpatConst when eq_constr u.up_f f -> na | KpatFixed when u.up_f = f -> na | KpatEvar k when isEvar_k k f -> na | KpatLet when isLetIn f -> na | KpatLam when isLambda f -> na | KpatRigid when isRigid f -> na | KpatFlex -> na | KpatProj pc -> let np = na + nb_cs_proj_args pc f u in if n < np then -1 else np | _ -> -1 in if np < na then fpats else let () = if !i0 < np then i0 := n in (u, np) :: fpats let filter_upat_FO i0 f n u fpats = let np = nb_args u.up_FO in if n < np then fpats else let ok = match u.up_k with | KpatConst -> eq_constr u.up_f f | KpatFixed -> u.up_f = f | KpatEvar k -> isEvar_k k f | KpatLet -> isLetIn f | KpatLam -> isLambda f | KpatRigid -> isRigid f | KpatProj pc -> f = mkConst pc | KpatFlex -> i0 := n; true in if ok then begin if !i0 < np then i0 := np; (u, np) :: fpats end else fpats exception FoundUnif of (evar_map * tpattern) (* Note: we don't update env as we descend into the term, as the primitive *) (* unification procedure always rejects subterms with bound variables. *) (* We are forced to duplicate code between the FO/HO matching because we *) (* have to work around several kludges in unify.ml: *) (* - w_unify drops into second-order unification when the pattern is an *) (* application whose head is a meta. *) (* - w_unify tries to unify types without subsumption when the pattern *) (* head is an evar or meta (e.g., it fails on ?1 = nat when ?1 : Type). *) (* - w_unify expands let-in (zeta conversion) eagerly, whereas we want to *) (* match a head let rigidly. *) let match_upats_FO upats env sigma0 ise = let rec loop c = let f, a = splay_app ise c in let i0 = ref (-1) in let fpats = List.fold_right (filter_upat_FO i0 f (Array.length a)) upats [] in while !i0 >= 0 do let i = !i0 in i0 := -1; let c' = mkSubApp f i a in let one_match (u, np) = let skip = if i <= np then i < np else if u.up_k == KpatFlex then begin i0 := i - 1; false end else begin if !i0 < np then i0 := np; true end in if skip || not (closed0 c') then () else try let _ = match u.up_k with | KpatFlex -> let kludge v = mkLambda (Anonymous, mkProp, v) in unif_FO env ise (kludge u.up_FO) (kludge c') | KpatLet -> let kludge vla = let vl, a = safeDestApp vla in let x, v, t, b = destLetIn vl in mkApp (mkLambda (x, t, b), array_cons v a) in unif_FO env ise (kludge u.up_FO) (kludge c') | _ -> unif_FO env ise u.up_FO c' in let ise' = (* Unify again using HO to assign evars *) let p = mkApp (u.up_f, u.up_a) in try unif_HO env ise p c' with _ -> raise NoMatch in let lhs = mkSubApp f i a in let pt' = unif_end env sigma0 ise' u.up_t (u.up_ok lhs) in raise (FoundUnif (ungen_upat lhs pt' u)) with FoundUnif _ as sigma_u -> raise sigma_u | Not_found -> anomaly "incomplete ise in match_upats_FO" | _ -> () in List.iter one_match fpats done; iter_constr_LR loop f; Array.iter loop a in fun c -> try loop c with Invalid_argument _ -> anomaly "IN FO" let prof_FO = mk_profiler "match_upats_FO";; let match_upats_FO upats env sigma0 ise c = prof_FO.profile (match_upats_FO upats env sigma0) ise c ;; let match_upats_HO upats env sigma0 ise c = let it_did_match = ref false in let rec aux upats env sigma0 ise c = let f, a = splay_app ise c in let i0 = ref (-1) in let fpats = List.fold_right (filter_upat i0 f (Array.length a)) upats [] in while !i0 >= 0 do let i = !i0 in i0 := -1; let one_match (u, np) = let skip = if i <= np then i < np else if u.up_k == KpatFlex then begin i0 := i - 1; false end else begin if !i0 < np then i0 := np; true end in if skip then () else try let ise' = match u.up_k with | KpatFixed | KpatConst -> ise | KpatEvar _ -> let _, pka = destEvar u.up_f and _, ka = destEvar f in unif_HO_args env ise pka 0 ka | KpatLet -> let x, v, t, b = destLetIn f in let _, pv, _, pb = destLetIn u.up_f in let ise' = unif_HO env ise pv v in unif_HO (Environ.push_rel (x, None, t) env) ise' pb b | KpatFlex | KpatProj _ -> unif_HO env ise u.up_f (mkSubApp f (i - Array.length u.up_a) a) | _ -> unif_HO env ise u.up_f f in let ise'' = unif_HO_args env ise' u.up_a (i - Array.length u.up_a) a in let lhs = mkSubApp f i a in let pt' = unif_end env sigma0 ise'' u.up_t (u.up_ok lhs) in raise (FoundUnif (ungen_upat lhs pt' u)) with FoundUnif _ as sigma_u -> raise sigma_u | NoProgress -> it_did_match := true | _ -> () in List.iter one_match fpats done; iter_constr_LR (aux upats env sigma0 ise) f; Array.iter (aux upats env sigma0 ise) a in aux upats env sigma0 ise c; if !it_did_match then raise NoProgress let prof_HO = mk_profiler "match_upats_HO";; let match_upats_HO upats env sigma0 ise c = prof_HO.profile (match_upats_HO upats env sigma0) ise c ;; let fixed_upat = function | {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false | {up_t = t} -> not (occur_existential t) let do_once r f = match !r with Some _ -> () | None -> r := Some (f ()) let assert_done r = match !r with Some x -> x | None -> anomaly "do_once never called" type subst = Environ.env -> Term.constr -> int -> Term.constr type find_P = Environ.env -> Term.constr -> int -> k:subst -> Term.constr type conclude = unit -> Term.constr * ssrdir * (Evd.evar_map * Term.constr) (* upats_origin makes a better error message only *) let mk_tpattern_matcher ?(raise_NoMatch=false) ?upats_origin sigma0 occ (ise, upats) = let nocc = ref 0 and skip_occ = ref false in let use_occ, occ_list = match occ with | Some (true, ol) -> ol = [], ol | Some (false, ol) -> ol <> [], ol | None -> false, [] in let max_occ = List.fold_right max occ_list 0 in let subst_occ = let occ_set = Array.make max_occ (not use_occ) in let _ = List.iter (fun i -> occ_set.(i - 1) <- use_occ) occ_list in let _ = if max_occ = 0 then skip_occ := use_occ in fun () -> incr nocc; if !nocc = max_occ then skip_occ := use_occ; if !nocc <= max_occ then occ_set.(!nocc - 1) else not use_occ in let upat_that_matched = ref None in let match_EQ env sigma u = match u.up_k with | KpatLet -> let x, pv, t, pb = destLetIn u.up_f in let env' = Environ.push_rel (x, None, t) env in let match_let f = match kind_of_term f with | LetIn (_, v, _, b) -> unif_EQ env sigma pv v && unif_EQ env' sigma pb b | _ -> false in match_let | KpatFixed -> (=) u.up_f | KpatConst -> eq_constr u.up_f | KpatLam -> fun c -> (match kind_of_term c with | Lambda _ -> unif_EQ env sigma u.up_f c | _ -> false) | _ -> unif_EQ env sigma u.up_f in let p2t p = mkApp(p.up_f,p.up_a) in let source () = match upats_origin, upats with | None, [p] -> (if fixed_upat p then str"term " else str"partial term ") ++ pr_constr_pat (p2t p) ++ spc() | Some (dir,rule), [p] -> str"The " ++ pr_dir_side dir ++ str" of " ++ pr_constr_pat rule ++ fnl() ++ ws 4 ++ pr_constr_pat (p2t p) ++ fnl() | Some (dir,rule), _ -> str"The " ++ pr_dir_side dir ++ str" of " ++ pr_constr_pat rule ++ spc() | _, [] | None, _::_::_ -> anomaly "mk_tpattern_matcher with no upats_origin" in ((fun env c h ~k -> do_once upat_that_matched (fun () -> try match_upats_FO upats env sigma0 ise c; match_upats_HO upats env sigma0 ise c; raise NoMatch with FoundUnif sigma_u -> sigma_u | NoMatch when (not raise_NoMatch) -> errorstrm (source () ++ str "does not match any subterm of the goal") | NoProgress when (not raise_NoMatch) -> let dir = match upats_origin with Some (d,_) -> d | _ -> anomaly "mk_tpattern_matcher with no upats_origin" in errorstrm (str"all matches of "++source()++ str"are equal to the " ++ pr_dir_side (inv_dir dir)) | NoProgress -> raise NoMatch); let sigma, ({up_f = pf; up_a = pa} as u) = assert_done upat_that_matched in if !skip_occ then (ignore(k env u.up_t 0); c) else let match_EQ = match_EQ env sigma u in let pn = Array.length pa in let rec subst_loop (env,h as acc) c' = if !skip_occ then c' else let f, a = splay_app sigma c' in if Array.length a >= pn && match_EQ f && unif_EQ_args env sigma pa a then let a1, a2 = array_chop (Array.length pa) a in let fa1 = mkApp (f, a1) in let f' = if subst_occ () then k env u.up_t h else fa1 in mkApp (f', array_map_left (subst_loop acc) a2) else (* TASSI: clear letin values to avoid unfolding *) let inc_h (n,_,ty) (env,h') = Environ.push_rel (n,None,ty) env, h' + 1 in let f' = map_constr_with_binders_left_to_right inc_h subst_loop acc f in mkApp (f', array_map_left (subst_loop acc) a) in subst_loop (env,h) c) : find_P), ((fun () -> let sigma, ({up_f = pf; up_a = pa} as u) = match !upat_that_matched with | Some x -> x | None when raise_NoMatch -> raise NoMatch | None -> anomaly "companion function never called" in let p' = mkApp (pf, pa) in if max_occ <= !nocc then p', u.up_dir, (sigma, u.up_t) else errorstrm (str"Only " ++ int !nocc ++ str" < " ++ int max_occ ++ str(plural !nocc " occurence") ++ match upats_origin with | None -> str" of" ++ spc() ++ pr_constr_pat p' | Some (dir,rule) -> str" of the " ++ pr_dir_side dir ++ fnl() ++ ws 4 ++ pr_constr_pat p' ++ fnl () ++ str"of " ++ pr_constr_pat rule)) : conclude) type ('ident, 'term) ssrpattern = | T of 'term | In_T of 'term | X_In_T of 'ident * 'term | In_X_In_T of 'ident * 'term | E_In_X_In_T of 'term * 'ident * 'term | E_As_X_In_T of 'term * 'ident * 'term let pr_pattern = function | T t -> prl_term t | In_T t -> str "in " ++ prl_term t | X_In_T (x,t) -> prl_term x ++ str " in " ++ prl_term t | In_X_In_T (x,t) -> str "in " ++ prl_term x ++ str " in " ++ prl_term t | E_In_X_In_T (e,x,t) -> prl_term e ++ str " in " ++ prl_term x ++ str " in " ++ prl_term t | E_As_X_In_T (e,x,t) -> prl_term e ++ str " as " ++ prl_term x ++ str " in " ++ prl_term t let pr_pattern_w_ids = function | T t -> prl_term t | In_T t -> str "in " ++ prl_term t | X_In_T (x,t) -> pr_id x ++ str " in " ++ prl_term t | In_X_In_T (x,t) -> str "in " ++ pr_id x ++ str " in " ++ prl_term t | E_In_X_In_T (e,x,t) -> prl_term e ++ str " in " ++ pr_id x ++ str " in " ++ prl_term t | E_As_X_In_T (e,x,t) -> prl_term e ++ str " as " ++ pr_id x ++ str " in " ++ prl_term t let pr_pattern_aux pr_constr = function | T t -> pr_constr t | In_T t -> str "in " ++ pr_constr t | X_In_T (x,t) -> pr_constr x ++ str " in " ++ pr_constr t | In_X_In_T (x,t) -> str "in " ++ pr_constr x ++ str " in " ++ pr_constr t | E_In_X_In_T (e,x,t) -> pr_constr e ++ str " in " ++ pr_constr x ++ str " in " ++ pr_constr t | E_As_X_In_T (e,x,t) -> pr_constr e ++ str " as " ++ pr_constr x ++ str " in " ++ pr_constr t let pp_pattern (sigma, p) = pr_pattern_aux (fun t -> pr_constr (snd (nf_open_term sigma sigma t))) p let pr_cpattern = pr_term let pr_rpattern _ _ _ = pr_pattern let pr_option f = function None -> mt() | Some x -> f x let pr_ssrpattern _ _ _ = pr_option pr_pattern let pr_pattern_squarep = pr_option (fun r -> str "[" ++ pr_pattern r ++ str "]") let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep let pr_pattern_roundp = pr_option (fun r -> str "(" ++ pr_pattern r ++ str ")") let pr_ssrpattern_roundp _ _ _ = pr_pattern_roundp let wit_rpatternty, globwit_rpatternty, rawwit_rpatternty = add_genarg "rpatternty" pr_pattern ARGUMENT EXTEND rpattern TYPED AS rpatternty PRINTED BY pr_rpattern | [ lconstr(c) ] -> [ T (mk_lterm c) ] | [ "in" lconstr(c) ] -> [ In_T (mk_lterm c) ] | [ lconstr(x) "in" lconstr(c) ] -> [ X_In_T (mk_lterm x, mk_lterm c) ] | [ "in" lconstr(x) "in" lconstr(c) ] -> [ In_X_In_T (mk_lterm x, mk_lterm c) ] | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] -> [ E_In_X_In_T (mk_lterm e, mk_lterm x, mk_lterm c) ] | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] -> [ E_As_X_In_T (mk_lterm e, mk_lterm x, mk_lterm c) ] END type cpattern = char * glob_constr_and_expr let tag_of_cpattern = fst let loc_of_cpattern = loc_ofCG let cpattern_of_term t = t type occ = (bool * int list) option type rpattern = (cpattern, cpattern) ssrpattern let pr_rpattern = pr_pattern type pattern = Evd.evar_map * (Term.constr, Term.constr) ssrpattern let id_of_cpattern = function | _,(_,Some (CRef (Ident (_, x)))) -> Some x | _,(_,Some (CAppExpl (_, (_, Ident (_, x)), []))) -> Some x | _,(GRef (_, VarRef x) ,None) -> Some x | _ -> None let id_of_Cterm t = match id_of_cpattern t with | Some x -> x | None -> loc_error (loc_of_cpattern t) "Only identifiers are allowed here" let interp_wit globwit wit ist gl x = let globarg = in_gen globwit x in let sigma, arg = interp_genarg ist gl globarg in sigma, out_gen wit arg let interp_constr = interp_wit globwit_constr wit_constr let interp_open_constr ist gl gc = interp_wit globwit_open_constr wit_open_constr ist gl ((), gc) let pf_intern_term ist gl (_, c) = glob_constr ist (project gl) (pf_env gl) c let interp_term ist gl (_, c) = snd (interp_open_constr ist gl c) let glob_ssrterm gs = function | k, (_, Some c) -> k, Tacinterp.intern_constr gs c | ct -> ct let subst_ssrterm s (k, c) = k, Tacinterp.subst_glob_constr_and_expr s c let pr_ssrterm _ _ _ = pr_term let input_ssrtermkind strm = match Compat.get_tok (stream_nth 0 strm) with | Tok.KEYWORD "(" -> '(' | Tok.KEYWORD "@" -> '@' | _ -> ' ' let ssrtermkind = Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind (* This piece of code asserts the following notations are reserved *) (* Reserved Notation "( a 'in' b )" (at level 0). *) (* Reserved Notation "( a 'as' b )" (at level 0). *) (* Reserved Notation "( a 'in' b 'in' c )" (at level 0). *) (* Reserved Notation "( a 'as' b 'in' c )" (at level 0). *) let glob_cpattern gs p = pp(lazy(str"globbing pattern: " ++ pr_term p)); let glob x = snd (glob_ssrterm gs (mk_lterm x)) in let encode k s l = let name = Name (id_of_string ("_ssrpat_" ^ s)) in k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None) in let bind_in t1 t2 = let d = dummy_loc in let n = Name (destCVar t1) in fst (glob (mkCCast d (mkCHole d) (mkCLambda d n (mkCHole d) t2))) in let check_var t2 = if not (isCVar t2) then loc_error (constr_loc t2) "Only identifiers are allowed here" in match p with | _, (_, None) as x -> x | k, (v, Some t) as orig -> if k = 'x' then glob_ssrterm gs ('(', (v, Some t)) else match t with | CNotation(_, "( _ in _ )", ([t1; t2], [], [])) -> (try match glob t1, glob t2 with | (r1, None), (r2, None) -> encode k "In" [r1;r2] | (r1, Some _), (r2, Some _) when isCVar t1 -> encode k "In" [r1; r2; bind_in t1 t2] | (r1, Some _), (r2, Some _) -> encode k "In" [r1; r2] | _ -> anomaly "where are we?" with _ when isCVar t1 -> encode k "In" [bind_in t1 t2]) | CNotation(_, "( _ in _ in _ )", ([t1; t2; t3], [], [])) -> check_var t2; encode k "In" [fst (glob t1); bind_in t2 t3] | CNotation(_, "( _ as _ )", ([t1; t2], [], [])) -> encode k "As" [fst (glob t1); fst (glob t2)] | CNotation(_, "( _ as _ in _ )", ([t1; t2; t3], [], [])) -> check_var t2; encode k "As" [fst (glob t1); bind_in t2 t3] | _ -> glob_ssrterm gs orig ;; let interp_ssrterm _ gl t = Tacmach.project gl, t ARGUMENT EXTEND cpattern PRINTED BY pr_ssrterm INTERPRETED BY interp_ssrterm GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm RAW_TYPED AS cpattern RAW_PRINTED BY pr_ssrterm GLOB_TYPED AS cpattern GLOB_PRINTED BY pr_ssrterm | [ "Qed" constr(c) ] -> [ mk_lterm c ] END GEXTEND Gram GLOBAL: cpattern; cpattern: [[ k = ssrtermkind; c = constr -> let pattern = mk_term k c in if loc_ofCG pattern <> loc && k = '(' then mk_term 'x' c else pattern ]]; END ARGUMENT EXTEND lcpattern PRINTED BY pr_ssrterm INTERPRETED BY interp_ssrterm GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm RAW_TYPED AS cpattern RAW_PRINTED BY pr_ssrterm GLOB_TYPED AS cpattern GLOB_PRINTED BY pr_ssrterm | [ "Qed" lconstr(c) ] -> [ mk_lterm c ] END GEXTEND Gram GLOBAL: lcpattern; lcpattern: [[ k = ssrtermkind; c = lconstr -> let pattern = mk_term k c in if loc_ofCG pattern <> loc && k = '(' then mk_term 'x' c else pattern ]]; END let interp_pattern ist gl red redty = pp(lazy(str"interpreting: " ++ pr_pattern red)); let xInT x y = X_In_T(x,y) and inXInT x y = In_X_In_T(x,y) in let inT x = In_T x and eInXInT e x t = E_In_X_In_T(e,x,t) in let eAsXInT e x t = E_As_X_In_T(e,x,t) in let mkG ?(k=' ') x = k,(x,None) in let decode t f g = try match (pf_intern_term ist gl t) with | GCast(_,GHole _,CastConv(_,GLambda(_,Name x,_,_,c))) -> f x (' ',(c,None)) | it -> g t with _ -> g t in let decodeG t f g = decode (mkG t) f g in let bad_enc id _ = anomaly ("bad encoding for pattern " ^ id) in let cleanup_XinE h x rp sigma = let h_k = match kind_of_term h with Evar (k,_) -> k | _ -> assert false in let to_clean, update = (* handle rename if x is already used *) let ctx = pf_hyps gl in let len = Sign.named_context_length ctx in let name = ref None in try ignore(Sign.lookup_named x ctx); (name, fun k -> if !name = None then let nctx = Evd.evar_context (Evd.find sigma k) in let nlen = Sign.named_context_length nctx in if nlen > len then begin name := Some (pi1 (List.nth nctx (nlen - len - 1))) end) with Not_found -> ref (Some x), fun _ -> () in let sigma0 = project gl in let new_evars = let rec aux acc t = match kind_of_term t with | Evar (k,_) -> if k = h_k || List.mem k acc || Evd.mem sigma0 k then acc else (update k; k::acc) | _ -> fold_constr aux acc t in aux [] (Evarutil.nf_evar sigma rp) in let sigma = List.fold_left (fun sigma e -> if Evd.is_defined sigma e then sigma else (* clear may be recursiv *) let name = Option.get !to_clean in let g = Goal.build e in pp(lazy(pr_id name)); try snd(Logic.prim_refiner (Proof_type.Thin [name]) sigma g) with Evarutil.ClearDependencyError _ -> sigma) sigma new_evars in sigma in let red = match red with | T(k,(GCast (_,GHole _,(CastConv(_,GLambda (_,Name id,_,_,t)))),None)) when let id = string_of_id id in let len = String.length id in (len > 8 && String.sub id 0 8 = "_ssrpat_") -> let id = string_of_id id in let len = String.length id in (match String.sub id 8 (len - 8), t with | "In", GApp(_, _, [t]) -> decodeG t xInT (fun x -> T x) | "In", GApp(_, _, [e; t]) -> decodeG t (eInXInT (mkG e)) (bad_enc id) | "In", GApp(_, _, [e; t; e_in_t]) -> decodeG t (eInXInT (mkG e)) (fun _ -> decodeG e_in_t xInT (fun _ -> assert false)) | "As", GApp(_, _, [e; t]) -> decodeG t (eAsXInT (mkG e)) (bad_enc id) | _ -> bad_enc id ()) | T t -> decode t xInT (fun x -> T x) | In_T t -> decode t inXInT inT | X_In_T (e,t) -> decode t (eInXInT e) (fun x -> xInT (id_of_Cterm e) x) | In_X_In_T (e,t) -> inXInT (id_of_Cterm e) t | E_In_X_In_T (e,x,rp) -> eInXInT e (id_of_Cterm x) rp | E_As_X_In_T (e,x,rp) -> eAsXInT e (id_of_Cterm x) rp in pp(lazy(str"decoded as: " ++ pr_pattern_w_ids red)); let red = match redty with None -> red | Some ty -> let ty = ' ', ty in match red with | T t -> T (combineCG t ty (mkCCast (loc_ofCG t)) mkRCast) | X_In_T (x,t) -> let ty = pf_intern_term ist gl ty in E_As_X_In_T (mkG (mkRCast mkRHole ty), x, t) | E_In_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term ist gl ty) in E_In_X_In_T (combineCG e ty (mkCCast (loc_ofCG t)) mkRCast, x, t) | E_As_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term ist gl ty) in E_As_X_In_T (combineCG e ty (mkCCast (loc_ofCG t)) mkRCast, x, t) | red -> red in pp(lazy(str"typed as: " ++ pr_pattern_w_ids red)); let mkXLetIn loc x (a,(g,c)) = match c with | Some b -> a,(g,Some (mkCLetIn loc x (CHole(loc,None)) b)) | None -> a,(GLetIn (loc,x,(GHole (loc, BinderType x)), g), None) in match red with | T t -> let sigma, t = interp_term ist gl t in sigma, T t | In_T t -> let sigma, t = interp_term ist gl t in sigma, In_T t | X_In_T (x, rp) | In_X_In_T (x, rp) -> let mk x p = match red with X_In_T _ -> X_In_T(x,p) | _ -> In_X_In_T(x,p) in let rp = mkXLetIn dummy_loc (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in let rp = subst1 h (Evarutil.nf_evar sigma rp) in sigma, mk h rp | E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) -> let mk e x p = match red with E_In_X_In_T _ ->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p) in let rp = mkXLetIn dummy_loc (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in let rp = subst1 h (Evarutil.nf_evar sigma rp) in let sigma, e = interp_term ist (re_sig (sig_it gl) sigma) e in sigma, mk e h rp ;; let interp_cpattern ist gl red redty = interp_pattern ist gl (T red) redty;; let interp_rpattern ist gl red = interp_pattern ist gl red None;; (* The full occurrence set *) let noindex = Some(false,[]) (* calls do_subst on every sub-term identified by (pattern,occ) *) let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let fs sigma x = Reductionops.nf_evar sigma x in let pop_evar sigma e p = let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in let e_body = match e_body with Evar_defined c -> c | _ -> errorstrm (str "Matching the pattern " ++ pr_constr p ++ str " did not instantiate ?" ++ int e ++ spc () ++ str "Does the variable bound by the \"in\" construct occur "++ str "in the pattern?") in let sigma = Evd.add (Evd.remove sigma e) e {e_def with Evd.evar_body = Evar_empty} in sigma, e_body in let ex_value hole = match kind_of_term hole with Evar (e,_) -> e | _ -> assert false in let mk_upat_for ?hack env sigma0 (sigma, t) ?(p=t) ok = let sigma,pat= mk_tpattern ?hack env sigma0 (sigma,p) ok L2R (fs sigma t) in sigma, [pat] in match pattern with | None -> do_subst env0 concl0 1 | Some (sigma, (T rp | In_T rp)) -> let rp = fs sigma rp in let ise = create_evar_defs sigma in let occ = match pattern with Some (_, T _) -> occ | _ -> noindex in let rp = mk_upat_for env0 sigma0 (ise, rp) all_ok in let find_T, end_T = mk_tpattern_matcher ?raise_NoMatch sigma0 occ rp in let concl = find_T env0 concl0 1 do_subst in let _ = end_T () in concl | Some (sigma, (X_In_T (hole, p) | In_X_In_T (hole, p))) -> let p = fs sigma p in let occ = match pattern with Some (_, X_In_T _) -> occ | _ -> noindex in let ex = ex_value hole in let rp = mk_upat_for ~hack:true env0 sigma0 (sigma, p) all_ok in let find_T, end_T = mk_tpattern_matcher sigma0 noindex rp in (* we start from sigma, so hole is considered a rigid head *) let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher ?raise_NoMatch sigma occ holep in let concl = find_T env0 concl0 1 (fun env c h -> let p_sigma = unify_HO env (create_evar_defs sigma) c p in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h (fun env _ -> do_subst env e_body))) in let _ = end_X () in let _ = end_T () in concl | Some (sigma, E_In_X_In_T (e, hole, p)) -> let p, e = fs sigma p, fs sigma e in let ex = ex_value hole in let rp = mk_upat_for ~hack:true env0 sigma0 (sigma, p) all_ok in let find_T, end_T = mk_tpattern_matcher sigma0 noindex rp in let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher sigma noindex holep in let re = mk_upat_for env0 sigma0 (sigma, e) all_ok in let find_E, end_E = mk_tpattern_matcher ?raise_NoMatch sigma0 occ re in let concl = find_T env0 concl0 1 (fun env c h -> let p_sigma = unify_HO env (create_evar_defs sigma) c p in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h (fun env c h -> find_E env e_body h do_subst))) in let _ = end_E () in let _ = end_X () in let _ = end_T () in concl | Some (sigma, E_As_X_In_T (e, hole, p)) -> let p, e = fs sigma p, fs sigma e in let ex = ex_value hole in let rp = let e_sigma = unify_HO env0 sigma hole e in e_sigma, fs e_sigma p in let rp = mk_upat_for ~hack:true env0 sigma0 rp all_ok in let find_TE, end_TE = mk_tpattern_matcher sigma0 noindex rp in let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher sigma occ holep in let concl = find_TE env0 concl0 1 (fun env c h -> let p_sigma = unify_HO env (create_evar_defs sigma) c p in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h (fun env c h -> let e_sigma = unify_HO env sigma e_body e in let e_body = fs e_sigma e in do_subst env e_body h))) in let _ = end_X () in let _ = end_TE () in concl ;; let redex_of_pattern (sigma, p) = let e = match p with | In_T _ | In_X_In_T _ -> anomaly "pattern without redex" | T e | X_In_T (e, _) | E_As_X_In_T (e, _, _) | E_In_X_In_T (e, _, _) -> e in Reductionops.nf_evar sigma e let fill_occ_pattern ?raise_NoMatch env sigma cl pat occ h = let find_R, conclude = let r = ref None in (fun env c h' -> do_once r (fun () -> c); mkRel (h'+h-1)), (fun _ -> if !r = None then redex_of_pattern pat else assert_done r) in let cl = eval_pattern ?raise_NoMatch env sigma cl (Some pat) occ find_R in let e = conclude cl in e, cl ;; (* clenup interface for external use *) let mk_tpattern ?p_origin env sigma0 sigma_t f dir c = mk_tpattern ?p_origin env sigma0 sigma_t f dir c ;; let pf_fill_occ env concl occ sigma0 p (sigma, t) ok h = let ise = create_evar_defs sigma in let ise, u = mk_tpattern env sigma0 (ise,t) ok L2R p in let find_U, end_U = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in let concl = find_U env concl h (fun _ _ -> mkRel) in let rdx, _, (sigma, p) = end_U () in sigma, p, concl, rdx let fill_occ_term env cl occ sigma0 (sigma, t) = try let sigma',t',cl,_ = pf_fill_occ env cl occ sigma0 t (sigma, t) all_ok 1 in if sigma' != sigma0 then error "matching impacts evars" else cl, (sigma',t') with NoMatch -> try let sigma', t' = unif_end env sigma0 (create_evar_defs sigma) t (fun _ -> true) in if sigma' != sigma0 then raise NoMatch else cl, (sigma', t') with _ -> errorstrm (str "partial term " ++ pr_constr_pat t ++ str " does not match any subterm of the goal") let pf_fill_occ_term gl occ t = let sigma0 = project gl and env = pf_env gl and concl = pf_concl gl in let cl,(_,t) = fill_occ_term env concl occ sigma0 t in cl, t let cpattern_of_id id = ' ', (GRef (dummy_loc, VarRef id), None) let is_wildcard = function | _,(_,Some (CHole _)|GHole _,None) -> true | _ -> false type ltacctx = int let pr_ltacctx _ _ _ _ = mt () let ltacctxs = ref (1, []) let interp_ltacctx ist gl n0 = Tacmach.project gl, if n0 = 0 then 0 else let n, s = !ltacctxs in let n' = if n >= max_int then 1 else n + 1 in ltacctxs := (n', (n, ist) :: s); n let noltacctx = 0 let rawltacctx = 1 ARGUMENT EXTEND ltacctx TYPED AS int PRINTED BY pr_ltacctx INTERPRETED BY interp_ltacctx | [ ] -> [ rawltacctx ] END let get_ltacctx i = match !ltacctxs with | _ when i = noltacctx -> anomaly "Missing Ltac context" | n, (i', ist) :: s when i' = i -> ltacctxs := (n, s); ist | _ -> anomaly "Bad scope in SSR tactical" (* "ssrpattern" *) let pr_ssrpatternarg _ _ _ cpat = pr_rpattern cpat ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY pr_ssrpatternarg | [ "[" rpattern(pat) "]" ] -> [ pat ] END let ssrpatterntac ctx arg gl = let ist = get_ltacctx ctx in let pat = interp_rpattern ist gl arg in let sigma0 = project gl in let concl0 = pf_concl gl in let t, concl_x = fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in let tty = pf_type_of gl t in let concl = mkLetIn (Name (id_of_string "toto"), t, tty, concl_x) in convert_concl concl DEFAULTcast gl TACTIC EXTEND ssrat | [ "ssrpattern" ssrpatternarg(arg) ltacctx(ctx) ] -> [ ssrpatterntac ctx arg ] END (* We wipe out all the keywords generated by the grammar rules we defined. *) (* The user is supposed to Require Import ssreflect or Require ssreflect *) (* and Import ssreflectSsrSyntax to obtain these keywords and as a *) (* consequence the extended ssreflect grammar. *) let () = Lexer.unfreeze frozen_lexer ;; (* vim: set filetype=ocaml foldmethod=marker: *) ssreflect-1.5/src/ssrmatching.mli0000644000175000017500000002477712175455021016174 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) open Genarg open Environ open Evd open Proof_type open Term (** ******** Small Scale Reflection pattern matching facilities ************* *) (** Pattern parsing *) (** The type of context patterns, the patterns of the [set] tactic and [:] tactical. These are patterns that identify a precise subterm. *) type cpattern val pr_cpattern : cpattern -> Pp.std_ppcmds (** CS cpattern: (f _), (X in t), (t in X in t), (t as X in t) *) val cpattern : cpattern Pcoq.Gram.entry val globwit_cpattern : (cpattern, glevel) abstract_argument_type val rawwit_cpattern : (cpattern, rlevel) abstract_argument_type val wit_cpattern : (cpattern, tlevel) abstract_argument_type (** OS cpattern: f _, (X in t), (t in X in t), (t as X in t) *) val lcpattern : cpattern Pcoq.Gram.entry val globwit_lcpattern : (cpattern, glevel) abstract_argument_type val rawwit_lcpattern : (cpattern, rlevel) abstract_argument_type val wit_lcpattern : (cpattern, tlevel) abstract_argument_type (** The type of rewrite patterns, the patterns of the [rewrite] tactic. These patterns also include patterns that identify all the subterms of a context (i.e. "in" prefix) *) type rpattern val pr_rpattern : rpattern -> Pp.std_ppcmds (** OS rpattern: f _, in t, X in t, in X in t, t in X in t, t as X in t *) val rpattern : rpattern Pcoq.Gram.entry val globwit_rpattern : (rpattern, glevel) abstract_argument_type val rawwit_rpattern : (rpattern, rlevel) abstract_argument_type val wit_rpattern : (rpattern, tlevel) abstract_argument_type (** Pattern interpretation and matching *) exception NoMatch exception NoProgress (** AST for [rpattern] (and consequently [cpattern]) *) type ('ident, 'term) ssrpattern = | T of 'term | In_T of 'term | X_In_T of 'ident * 'term | In_X_In_T of 'ident * 'term | E_In_X_In_T of 'term * 'ident * 'term | E_As_X_In_T of 'term * 'ident * 'term type pattern = evar_map * (constr, constr) ssrpattern val pp_pattern : pattern -> Pp.std_ppcmds (** Extracts the redex and applies to it the substitution part of the pattern. @raise Anomaly if called on [In_T] or [In_X_In_T] *) val redex_of_pattern : pattern -> constr (** [interp_rpattern ise gl rpat] "internalizes" and "interprets" [rpat] in the current [Ltac] interpretation signature [ise] and tactic input [gl]*) val interp_rpattern : Tacinterp.interp_sign -> goal Tacmach.sigma -> rpattern -> pattern (** [interp_cpattern ise gl cpat ty] "internalizes" and "interprets" [cpat] in the current [Ltac] interpretation signature [ise] and tactic input [gl]. [ty] is an optional type for the redex of [cpat] *) val interp_cpattern : Tacinterp.interp_sign -> goal Tacmach.sigma -> cpattern -> glob_constr_and_expr option -> pattern (** The set of occurrences to be matched. The boolean is set to true * to signal the complement of this set (i.e. {-1 3}) *) type occ = (bool * int list) option (** Substitution function. The [int] argument is the number of binders traversed so far *) type subst = env -> constr -> int -> constr (** [eval_pattern b env sigma t pat occ subst] maps [t] calling [subst] on every [occ] occurrence of [pat]. The [int] argument is the number of binders traversed. If [pat] is [None] then then subst is called on [t]. [t] must live in [env] and [sigma], [pat] must have been interpreted in (an extension of) [sigma]. @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) @return [t] where all [occ] occurrences of [pat] have been mapped using [subst] *) val eval_pattern : ?raise_NoMatch:bool -> env -> evar_map -> constr -> pattern option -> occ -> subst -> constr (** [fill_occ_pattern b env sigma t pat occ h] is a simplified version of [eval_pattern]. It replaces all [occ] occurrences of [pat] in [t] with Rel [h]. [t] must live in [env] and [sigma], [pat] must have been interpreted in (an extension of) [sigma]. @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) @return the instance of the redex of [pat] that was matched and [t] transformed as described above. *) val fill_occ_pattern : ?raise_NoMatch:bool -> env -> evar_map -> constr -> pattern -> occ -> int -> constr * constr (** Grammar entry to grab [Ltac] context, needed for the functions above. It must appear after the last tactic argument and only once. For example {[ TACTIC EXTEND ssrclear | [ "clear" natural(n) ltacctx(ctx) ] -> [poptac ~ist:(get_ltacctx ctx) n] END ]} *) type ltacctx val get_ltacctx : ltacctx -> Tacinterp.interp_sign val ltacctx : ltacctx Pcoq.Gram.entry val globwit_ltacctx : (ltacctx, glevel) abstract_argument_type val rawwit_ltacctx : (ltacctx, rlevel) abstract_argument_type val wit_ltacctx : (ltacctx, tlevel) abstract_argument_type (** *************************** Low level APIs ****************************** *) (* The primitive matching facility. It matches of a term with holes, like the T pattern above, and calls a continuation on its occurrences. *) type ssrdir = L2R | R2L val pr_dir_side : ssrdir -> Pp.std_ppcmds (** a pattern for a term with wildcards *) type tpattern (** [mk_tpattern env sigma0 sigma_t ok p_origin dir p] compiles a term [t] living in [env] [sigma] (an extension of [sigma0]) intro a [tpattern]. The [tpattern] can hold a (proof) term [p] and a diction [dir]. The [ok] callback is used to filter occurrences. @return the compiled [tpattern] and its [evar_map] @raise UserEerror is the pattern is a wildcard *) val mk_tpattern : ?p_origin:ssrdir * constr -> env -> evar_map -> evar_map * constr -> (constr -> evar_map -> bool) -> ssrdir -> constr -> evar_map * tpattern (** [findP env t i k] is a stateful function that finds the next occurrence of a tpattern and calls the callback [k] to map the subterm matched. The [int] argument passed to [k] is the number of binders traversed so far plus the initial value [i]. @return [t] where the subterms identified by the selected occurrences of the patter have been mapped using [k] @raise NoMatch if the raise_NoMatch flag given to [mk_tpattern_matcher] is [true] and if the pattern did not match @raise UserEerror if the raise_NoMatch flag given to [mk_tpattern_matcher] is [false] and if the pattern did not match *) type find_P = env -> constr -> int -> k:subst -> constr (** [conclude ()] asserts that all mentioned ocurrences have been visited. @return the instance of the pattern, the evarmap after the pattern instantiation, the proof term and the ssrdit stored in the tpattern @raise UserEerror if too many occurrences were specified *) type conclude = unit -> constr * ssrdir * (evar_map * constr) (** [mk_tpattern_matcher b o sigma0 occ sigma_tplist] creates a pair a function [find_P] and [conclude] with the behaviour explained above. The flag [b] (default [false]) changes the error reporting behaviour of [find_P] if none of the [tpattern] matches. The argument [o] can be passed to tune the [UserError] eventually raised (useful if the pattern is coming from the LHS/RHS of an equation) *) val mk_tpattern_matcher : ?raise_NoMatch:bool -> ?upats_origin:ssrdir * constr -> evar_map -> occ -> evar_map * tpattern list -> find_P * conclude (** Example of [mk_tpattern_matcher] to implement [rewrite \{occ\}\[in t\]rules]. It first matches "in t" (called [pat]), then in all matched subterms it matches the LHS of the rules using [find_R]. [concl0] is the initial goal, [concl] will be the goal where some terms are replaced by a De Bruijn index. The [rw_progress] extra check selects only occurrences that are not rewritten to themselves (e.g. an occurrence "x + x" rewritten with the commutativity law of addition is skipped) {[ let find_R, conclude = match pat with | Some (_, In_T _) -> let aux (sigma, pats) (d, r, lhs, rhs) = let sigma, pat = mk_tpattern env0 sigma0 (sigma, r) (rw_progress rhs) d lhs in sigma, pats @ [pat] in let rpats = List.fold_left aux (r_sigma, []) rules in let find_R, end_R = mk_tpattern_matcher sigma0 occ rpats in find_R ~k:(fun _ _ h -> mkRel h), fun cl -> let rdx, d, r = end_R () in (d,r),rdx | _ -> ... in let concl = eval_pattern env0 sigma0 concl0 pat occ find_R in let (d, r), rdx = conclude concl in ]} *) (* convenience shortcut: [pf_fill_occ_term gl occ (sigma,t)] returns * the conclusion of [gl] where [occ] occurrences of [t] have been replaced * by [Rel 1] and the instance of [t] *) val pf_fill_occ_term : goal Tacmach.sigma -> occ -> evar_map * constr -> constr * constr (* It may be handy to inject a simple term into the first form of cpattern *) val cpattern_of_term : char * glob_constr_and_expr -> cpattern (** Helpers to make stateful closures. Example: a [find_P] function may be called many times, but the pattern instantiation phase is performed only the first time. The corresponding [conclude] has to return the instantiated pattern redex. Since it is up to [find_P] to raise [NoMatch] if the pattern has no instance, [conclude] considers it an anomaly if the pattern did not match *) (** [do_once r f] calls [f] and updates the ref only once *) val do_once : 'a option ref -> (unit -> 'a) -> unit (** [assert_done r] return the content of r. @raise Anomaly is r is [None] *) val assert_done : 'a option ref -> 'a (** Very low level APIs. these are calls to evarconv's [the_conv_x] followed by [consider_remaining_unif_problems] and [resolve_typeclasses]. In case of failure they raise [NoMatch] *) val unify_HO : env -> evar_map -> constr -> constr -> evar_map val pf_unify_HO : goal Tacmach.sigma -> constr -> constr -> goal Tacmach.sigma (** Some more low level functions needed to implement the full SSR language on top of the former APIs *) val tag_of_cpattern : cpattern -> char val loc_of_cpattern : cpattern -> Util.loc val id_of_cpattern : cpattern -> Names.variable option val is_wildcard : cpattern -> bool val cpattern_of_id : Names.variable -> cpattern val rawltacctx : ltacctx val cpattern_of_id : Names.variable -> cpattern val rawltacctx : ltacctx val pr_constr_pat : constr -> Pp.std_ppcmds (* One can also "Set SsrMatchingDebug" from a .v *) val debug : bool -> unit (* One should delimit a snippet with "Set SsrMatchingProfiling" and * "Unset SsrMatchingProfiling" to get timings *) val profile : bool -> unit (* eof *) ssreflect-1.5/src/ssreflect.ml40000644000175000017500000064644012175455021015554 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) (* This line is read by the Makefile's dist target: do not remove. *) let ssrversion = "1.5";; let ssrAstVersion = 1;; let () = Mltop.add_known_plugin (fun () -> if Flags.is_verbose () && not !Flags.batch_mode then begin Printf.printf "\nSmall Scale Reflection version %s loaded.\n" ssrversion; Printf.printf "Copyright 2005-2012 Microsoft Corporation and INRIA.\n"; Printf.printf "Distributed under the terms of the CeCILL-B license.\n\n" end; (* Disable any semantics associated with bullets *) Goptions.set_string_option_value_gen (Some false) ["Bullet";"Behavior"] "None") "ssreflect" ;; (* Defining grammar rules with "xx" in it automatically declares keywords too *) let frozen_lexer = Lexer.freeze () ;; (*i camlp4use: "pa_extend.cmo" i*) (*i camlp4deps: "parsing/grammar.cma" i*) open Names open Pp open Pcoq open Genarg open Term open Topconstr open Libnames open Tactics open Tacticals open Termops open Namegen open Recordops open Tacmach open Coqlib open Glob_term open Util open Evd open Extend open Goptions open Tacexpr open Tacinterp open Pretyping.Default open Constr open Tactic open Extraargs open Ppconstr open Printer open Ssrmatching (** look up a name in the ssreflect internals module *) let ssrdirpath = make_dirpath [id_of_string "ssreflect"] let ssrqid name = make_qualid ssrdirpath (id_of_string name) let ssrtopqid name = make_short_qualid (id_of_string name) let locate_reference qid = Smartlocate.global_of_extended_global (Nametab.locate_extended qid) let mkSsrRef name = try locate_reference (ssrqid name) with Not_found -> try locate_reference (ssrtopqid name) with Not_found -> error "Small scale reflection library not loaded" let mkSsrRRef name = GRef (dummy_loc, mkSsrRef name) let mkSsrConst name = constr_of_reference (mkSsrRef name) let loc_error loc msg = user_err_loc (loc, msg, str msg) let errorstrm = errorlabstrm "ssreflect" (** Ssreflect load check. *) (* To allow ssrcoq to be fully compatible with the "plain" Coq, we only *) (* turn on its incompatible features (the new rewrite syntax, and the *) (* reserved identifiers) when the theory library (ssreflect.v) has *) (* has actually been required, or is being defined. Because this check *) (* needs to be done often (for each identifier lookup), we implement *) (* some caching, repeating the test only when the environment changes. *) (* We check for protect_term because it is the first constant loaded; *) (* ssr_have would ultimately be a better choice. *) let ssr_loaded = let cache = ref (Global.safe_env (), false) in fun () -> Lexer.is_keyword "is" && let new_lbl = Global.safe_env () in match !cache with | lbl, loaded when lbl == new_lbl -> loaded | _ -> let loaded = (try ignore (mkSsrRef "protect_term"); true with _ -> false) in cache := new_lbl, loaded; loaded (* 0 cost pp function. Active only if env variable SSRDEBUG is set *) (* or if SsrDebug is Set *) let pp_ref = ref (fun _ -> ()) let ssr_pp s = pperrnl (str"SSR: "++Lazy.force s) let _ = try ignore(Sys.getenv "SSRDEBUG"); pp_ref := ssr_pp with Not_found -> () let _ = Goptions.declare_bool_option { Goptions.optsync = false; Goptions.optname = "ssreflect debugging"; Goptions.optkey = ["SsrDebug"]; Goptions.optdepr = false; Goptions.optread = (fun _ -> !pp_ref == ssr_pp); Goptions.optwrite = (fun b -> Ssrmatching.debug b; if b then pp_ref := ssr_pp else pp_ref := fun _ -> ()) } let pp s = !pp_ref s (** Utils {{{ *****************************************************************) let env_size env = List.length (Environ.named_context env) let safeDestApp c = match kind_of_term c with App (f, a) -> f, a | _ -> c, [| |] let get_index = function ArgArg i -> i | _ -> anomaly "Uninterpreted index" (* Toplevel constr must be globalized twice ! *) let glob_constr ist gsigma genv = function | _, Some ce -> let ltacvars = List.map fst ist.lfun, [] in Constrintern.intern_gen false ~ltacvars:ltacvars gsigma genv ce | rc, None -> rc (* Term printing utilities functions for deciding bracketing. *) let pr_paren prx x = hov 1 (str "(" ++ prx x ++ str ")") (* String lexing utilities *) let skip_wschars s = let rec loop i = match s.[i] with '\n'..' ' -> loop (i + 1) | _ -> i in loop let skip_numchars s = let rec loop i = match s.[i] with '0'..'9' -> loop (i + 1) | _ -> i in loop (* We also guard characters that might interfere with the ssreflect *) (* tactic syntax. *) let guard_term ch1 s i = match s.[i] with | '(' -> false | '{' | '/' | '=' -> true | _ -> ch1 = '(' (* The call 'guard s i' should return true if the contents of s *) (* starting at i need bracketing to avoid ambiguities. *) let pr_guarded guard prc c = msg_with Format.str_formatter (prc c); let s = Format.flush_str_formatter () ^ "$" in if guard s (skip_wschars s 0) then pr_paren prc c else prc c (* More sensible names for constr printers *) let prl_constr = pr_lconstr let pr_constr = pr_constr let prl_glob_constr c = pr_lglob_constr_env (Global.env ()) c let pr_glob_constr c = pr_glob_constr_env (Global.env ()) c let prl_constr_expr = pr_lconstr_expr let pr_constr_expr = pr_constr_expr let prl_glob_constr_and_expr = function | _, Some c -> prl_constr_expr c | c, None -> prl_glob_constr c let pr_glob_constr_and_expr = function | _, Some c -> pr_constr_expr c | c, None -> pr_glob_constr c let pr_term (k, c) = pr_guarded (guard_term k) pr_glob_constr_and_expr c let prl_term (k, c) = pr_guarded (guard_term k) prl_glob_constr_and_expr c (** Adding a new uninterpreted generic argument type *) let add_genarg tag pr = let wit, globwit, rawwit as wits = create_arg None tag in let glob _ rarg = in_gen globwit (out_gen rawwit rarg) in let interp _ gl garg = Tacmach.project gl,in_gen wit (out_gen globwit garg) in let subst _ garg = garg in add_interp_genarg tag (glob, interp, subst); let gen_pr _ _ _ = pr in Pptactic.declare_extra_genarg_pprule (rawwit, gen_pr) (globwit, gen_pr) (wit, gen_pr); wits (** Constructors for cast type *) let dC t = CastConv (DEFAULTcast,t) (** Constructors for constr_expr *) let mkCProp loc = CSort (loc, GProp Null) let mkCType loc = CSort (loc, GType None) let mkCVar loc id = CRef (Ident (loc, id)) let isCVar = function CRef (Ident _) -> true | _ -> false let destCVar = function CRef (Ident (_, id)) -> id | _ -> anomaly "not a CRef" let rec mkCHoles loc n = if n <= 0 then [] else CHole (loc, None) :: mkCHoles loc (n - 1) let mkCHole loc = CHole (loc, None) let rec isCHoles = function CHole _ :: cl -> isCHoles cl | cl -> cl = [] let mkCExplVar loc id n = CAppExpl (loc, (None, Ident (loc, id)), mkCHoles loc n) let mkCLambda loc name ty t = CLambdaN (loc, [[loc, name], Default Explicit, ty], t) let mkCLetIn loc name bo t = CLetIn (loc, (loc, name), bo, t) let mkCArrow loc ty t = CProdN (loc, [[dummy_loc,Anonymous], Default Explicit, ty], t) let mkCCast loc t ty = CCast (loc,t, dC ty) (** Constructors for rawconstr *) let mkRHole = GHole (dummy_loc, InternalHole) let rec mkRHoles n = if n > 0 then mkRHole :: mkRHoles (n - 1) else [] let rec isRHoles = function GHole _ :: cl -> isRHoles cl | cl -> cl = [] let mkRApp f args = if args = [] then f else GApp (dummy_loc, f, args) let mkRVar id = GRef (dummy_loc, VarRef id) let mkRltacVar id = GVar (dummy_loc, id) let mkRCast rc rt = GCast (dummy_loc, rc, dC rt) let mkRType = GSort (dummy_loc, GType None) let mkRProp = GSort (dummy_loc, GProp Null) let mkRArrow rt1 rt2 = GProd (dummy_loc, Anonymous, Explicit, rt1, rt2) let mkRConstruct c = GRef (dummy_loc, ConstructRef c) let mkRInd mind = GRef (dummy_loc, IndRef mind) let mkRLambda n s t = GLambda (dummy_loc, n, Explicit, s, t) (** Constructors for constr *) let mkAppRed f c = match kind_of_term f with | Lambda (_, _, b) -> subst1 c b | _ -> mkApp (f, [|c|]) let mkProt t c = mkApp (mkSsrConst "protect_term", [|t; c|]) let mkRefl t c = mkApp ((build_coq_eq_data()).refl, [|t; c|]) (* Application to a sequence of n rels (for building eta-expansions). *) (* The rel indices decrease down to imin (inclusive), unless n < 0, *) (* in which case they're incresing (from imin). *) let mkEtaApp c n imin = if n = 0 then c else let nargs, mkarg = if n < 0 then -n, (fun i -> mkRel (imin + i)) else let imax = imin + n - 1 in n, (fun i -> mkRel (imax - i)) in mkApp (c, Array.init nargs mkarg) (* Same, but optimizing head beta redexes *) let rec whdEtaApp c n = if n = 0 then c else match kind_of_term c with | Lambda (_, _, c') -> whdEtaApp c' (n - 1) | _ -> mkEtaApp (lift n c) n 1 (* ssrterm conbinators *) let combineCG t1 t2 f g = match t1, t2 with | (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None) | (x, (_, Some t1)), (_, (_, Some t2)) -> x, (mkRHole, Some (f t1 t2)) | _, (_, (_, None)) -> anomaly "have: mixed C-G constr" | _ -> anomaly "have: mixed G-C constr" let loc_ofCG = function | (_, (s, None)) -> loc_of_glob_constr s | (_, (_, Some s)) -> constr_loc s let mk_term k c = k, (mkRHole, Some c) let mk_lterm = mk_term ' ' (* }}} *) (** Profiling {{{ *************************************************************) type profiler = { profile : 'a 'b. ('a -> 'b) -> 'a -> 'b; reset : unit -> unit; print : unit -> unit } let profile_now = ref false let something_profiled = ref false let profilers = ref [] let add_profiler f = profilers := f :: !profilers;; let _ = Goptions.declare_bool_option { Goptions.optsync = false; Goptions.optname = "ssreflect profiling"; Goptions.optkey = ["SsrProfiling"]; Goptions.optread = (fun _ -> !profile_now); Goptions.optdepr = false; Goptions.optwrite = (fun b -> Ssrmatching.profile b; profile_now := b; if b then List.iter (fun f -> f.reset ()) !profilers; if not b then List.iter (fun f -> f.print ()) !profilers) } let () = let prof_total = let init = ref 0.0 in { profile = (fun f x -> assert false); reset = (fun () -> init := Unix.gettimeofday ()); print = (fun () -> if !something_profiled then prerr_endline (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" "total" 0 (Unix.gettimeofday() -. !init) 0.0 0.0)) } in let prof_legenda = { profile = (fun f x -> assert false); reset = (fun () -> ()); print = (fun () -> if !something_profiled then begin prerr_endline (Printf.sprintf "!! %39s ---------- --------- --------- ---------" (String.make 39 '-')); prerr_endline (Printf.sprintf "!! %-39s %10s %9s %9s %9s" "function" "#calls" "total" "max" "average") end) } in add_profiler prof_legenda; add_profiler prof_total ;; let mk_profiler s = let total, calls, max = ref 0.0, ref 0, ref 0.0 in let reset () = total := 0.0; calls := 0; max := 0.0 in let profile f x = if not !profile_now then f x else let before = Unix.gettimeofday () in try incr calls; let res = f x in let after = Unix.gettimeofday () in let delta = after -. before in total := !total +. delta; if delta > !max then max := delta; res with exc -> let after = Unix.gettimeofday () in let delta = after -. before in total := !total +. delta; if delta > !max then max := delta; raise exc in let print () = if !calls <> 0 then begin something_profiled := true; prerr_endline (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" s !calls !total !max (!total /. (float_of_int !calls))) end in let prof = { profile = profile; reset = reset; print = print } in add_profiler prof; prof ;; (* }}} *) let inVersion = Libobject.declare_object { (Libobject.default_object "SSRASTVERSION") with Libobject.load_function = (fun _ (_,v) -> if v <> ssrAstVersion then error "Please recompile your .vo files"); Libobject.classify_function = (fun v -> Libobject.Keep v); } let _ = Goptions.declare_bool_option { Goptions.optsync = false; Goptions.optname = "ssreflect version"; Goptions.optkey = ["SsrAstVersion"]; Goptions.optread = (fun _ -> true); Goptions.optdepr = false; Goptions.optwrite = (fun _ -> Lib.add_anonymous_leaf (inVersion ssrAstVersion)) } let tactic_expr = Tactic.tactic_expr let gallina_ext = Vernac_.gallina_ext let sprintf = Printf.sprintf let tactic_mode = G_vernac.tactic_mode (** 1. Utilities *) let ssroldreworder = ref true let _ = Goptions.declare_bool_option { Goptions.optsync = false; Goptions.optname = "ssreflect 1.3 compatibility flag"; Goptions.optkey = ["SsrOldRewriteGoalsOrder"]; Goptions.optread = (fun _ -> !ssroldreworder); Goptions.optdepr = false; Goptions.optwrite = (fun b -> ssroldreworder := b) } (** Primitive parsing to avoid syntax conflicts with basic tactics. *) let accept_before_syms syms strm = match Compat.get_tok (stream_nth 1 strm) with | Tok.KEYWORD sym when List.mem sym syms -> () | _ -> raise Stream.Failure let accept_before_syms_or_any_id syms strm = match Compat.get_tok (stream_nth 1 strm) with | Tok.KEYWORD sym when List.mem sym syms -> () | Tok.IDENT _ -> () | _ -> raise Stream.Failure let accept_before_syms_or_ids syms ids strm = match Compat.get_tok (stream_nth 1 strm) with | Tok.KEYWORD sym when List.mem sym syms -> () | Tok.IDENT id when List.mem id ids -> () | _ -> raise Stream.Failure (** Pretty-printing utilities *) let pr_id = Ppconstr.pr_id let pr_name = function Name id -> pr_id id | Anonymous -> str "_" let pr_spc () = str " " let pr_bar () = Pp.cut() ++ str "|" let pr_list = prlist_with_sep let tacltop = (5,Ppextend.E) (** Tactic-level diagnosis *) let pf_pr_constr gl = pr_constr_env (pf_env gl) let pf_pr_glob_constr gl = pr_glob_constr_env (pf_env gl) (* debug *) let pf_msg gl = let ppgl = pr_lconstr_env (pf_env gl) (pf_concl gl) in msgnl (str "goal is " ++ ppgl) let msgtac gl = pf_msg gl; tclIDTAC gl (** Tactic utilities *) let last_goal gls = let sigma, gll = Refiner.unpackage gls in Refiner.repackage sigma (List.nth gll (List.length gll - 1)) let pf_type_id gl t = id_of_string (hdchar (pf_env gl) t) let not_section_id id = not (is_section_variable id) let is_pf_var c = isVar c && not_section_id (destVar c) let pf_ids_of_proof_hyps gl = let add_hyp (id, _, _) ids = if not_section_id id then id :: ids else ids in Sign.fold_named_context add_hyp (pf_hyps gl) ~init:[] (* Basic tactics *) let convert_concl_no_check t = convert_concl_no_check t DEFAULTcast let convert_concl t = convert_concl t DEFAULTcast let reduct_in_concl t = reduct_in_concl (t, DEFAULTcast) let havetac id = pose_proof (Name id) let settac id c = letin_tac None (Name id) c None let posetac id cl = settac id cl nowhere let basecuttac name c = apply (mkApp (mkSsrConst name, [|c|])) (* we reduce head beta redexes *) let betared env = Closure.create_clos_infos (Closure.RedFlags.mkflags [Closure.RedFlags.fBETA]) env ;; let introid name = tclTHEN (fun gl -> let g, env = pf_concl gl, pf_env gl in match kind_of_term g with | App (hd, _) when isLambda hd -> let g = Closure.whd_val (betared env) (Closure.inject g) in convert_concl_no_check g gl | _ -> tclIDTAC gl) (intro_mustbe_force name) ;; (** Name generation {{{ *******************************************************) (* Since Coq now does repeated internal checks of its external lexical *) (* rules, we now need to carve ssreflect reserved identifiers out of *) (* out of the user namespace. We use identifiers of the form _id_ for *) (* this purpose, e.g., we "anonymize" an identifier id as _id_, adding *) (* an extra leading _ if this might clash with an internal identifier. *) (* We check for ssreflect identifiers in the ident grammar rule; *) (* when the ssreflect Module is present this is normally an error, *) (* but we provide a compatibility flag to reduce this to a warning. *) let ssr_reserved_ids = ref true let _ = Goptions.declare_bool_option { Goptions.optsync = true; Goptions.optname = "ssreflect identifiers"; Goptions.optkey = ["SsrIdents"]; Goptions.optdepr = false; Goptions.optread = (fun _ -> !ssr_reserved_ids); Goptions.optwrite = (fun b -> ssr_reserved_ids := b) } let is_ssr_reserved s = let n = String.length s in n > 2 && s.[0] = '_' && s.[n - 1] = '_' let internal_names = ref [] let add_internal_name pt = internal_names := pt :: !internal_names let is_internal_name s = List.exists (fun p -> p s) !internal_names let ssr_id_of_string loc s = if is_ssr_reserved s && ssr_loaded () then begin if !ssr_reserved_ids then loc_error loc ("The identifier " ^ s ^ " is reserved.") else if is_internal_name s then warning ("Conflict between " ^ s ^ " and ssreflect internal names.") else warning ( "The name " ^ s ^ " fits the _xxx_ format used for anonymous variables.\n" ^ "Scripts with explicit references to anonymous variables are fragile.") end; id_of_string s let ssr_null_entry = Gram.Entry.of_parser "ssr_null" (fun _ -> ()) GEXTEND Gram GLOBAL: Prim.ident; Prim.ident: [[ s = IDENT; ssr_null_entry -> ssr_id_of_string loc s ]]; END let mk_internal_id s = let s' = sprintf "_%s_" s in for i = 1 to String.length s do if s'.[i] = ' ' then s'.[i] <- '_' done; add_internal_name ((=) s'); id_of_string s' let same_prefix s t n = let rec loop i = i = n || s.[i] = t.[i] && loop (i + 1) in loop 0 let skip_digits s = let n = String.length s in let rec loop i = if i < n && is_digit s.[i] then loop (i + 1) else i in loop let mk_tagged_id t i = id_of_string (sprintf "%s%d_" t i) let is_tagged t s = let n = String.length s - 1 and m = String.length t in m < n && s.[n] = '_' && same_prefix s t m && skip_digits s m = n let perm_tag = "_perm_Hyp_" let _ = add_internal_name (is_tagged perm_tag) let mk_perm_id = let salt = ref 1 in fun () -> salt := !salt mod 10000 + 1; mk_tagged_id perm_tag !salt let evar_tag = "_evar_" let _ = add_internal_name (is_tagged evar_tag) let mk_evar_name n = Name (mk_tagged_id evar_tag n) let nb_evar_deps = function | Name id -> let s = string_of_id id in if not (is_tagged evar_tag s) then 0 else let m = String.length evar_tag in (try int_of_string (String.sub s m (String.length s - 1 - m)) with _ -> 0) | _ -> 0 let discharged_tag = "_discharged_" let mk_discharged_id id = id_of_string (sprintf "%s%s_" discharged_tag (string_of_id id)) let has_discharged_tag s = let m = String.length discharged_tag and n = String.length s - 1 in m < n && s.[n] = '_' && same_prefix s discharged_tag m let _ = add_internal_name has_discharged_tag let is_discharged_id id = has_discharged_tag (string_of_id id) let wildcard_tag = "_the_" let wildcard_post = "_wildcard_" let mk_wildcard_id i = id_of_string (sprintf "%s%s%s" wildcard_tag (ordinal i) wildcard_post) let has_wildcard_tag s = let n = String.length s in let m = String.length wildcard_tag in let m' = String.length wildcard_post in n < m + m' + 2 && same_prefix s wildcard_tag m && String.sub s (n - m') m' = wildcard_post && skip_digits s m = n - m' - 2 let _ = add_internal_name has_wildcard_tag let max_suffix m (t, j0 as tj0) id = let s = string_of_id id in let n = String.length s - 1 in let dn = String.length t - 1 - n in let i0 = j0 - dn in if not (i0 >= m && s.[n] = '_' && same_prefix s t m) then tj0 else let rec loop i = if i < i0 && s.[i] = '0' then loop (i + 1) else if (if i < i0 then skip_digits s i = n else le_s_t i) then s, i else tj0 and le_s_t i = let ds = s.[i] and dt = t.[i + dn] in if ds = dt then i = n || le_s_t (i + 1) else dt < ds && skip_digits s i = n in loop m let mk_anon_id t gl = let m, si0, id0 = let s = ref (sprintf "_%s_" t) in if is_internal_name !s then s := "_" ^ !s; let n = String.length !s - 1 in let rec loop i j = let d = !s.[i] in if not (is_digit d) then i + 1, j else loop (i - 1) (if d = '0' then j else i) in let m, j = loop (n - 1) n in m, (!s, j), id_of_string !s in let gl_ids = pf_ids_of_hyps gl in if not (List.mem id0 gl_ids) then id0 else let s, i = List.fold_left (max_suffix m) si0 gl_ids in let n = String.length s - 1 in let rec loop i = if s.[i] = '9' then (s.[i] <- '0'; loop (i - 1)) else if i < m then (s.[n] <- '0'; s.[m] <- '1'; s ^ "_") else (s.[i] <- Char.chr (Char.code s.[i] + 1); s) in id_of_string (loop (n - 1)) (* We must not anonymize context names discharged by the "in" tactical. *) let ssr_anon_hyp = "Hyp" let anontac (x, _, _) gl = let id = match x with | Name id -> if is_discharged_id id then id else mk_anon_id (string_of_id id) gl | _ -> mk_anon_id ssr_anon_hyp gl in introid id gl let rec constr_name c = match kind_of_term c with | Var id -> Name id | Cast (c', _, _) -> constr_name c' | Const cn -> Name (id_of_label (con_label cn)) | App (c', _) -> constr_name c' | _ -> Anonymous (* }}} *) (** Open term to lambda-term coercion {{{ ************************************) (* This operation takes a goal gl and an open term (sigma, t), and *) (* returns a term t' where all the new evars in sigma are abstracted *) (* with the mkAbs argument, i.e., for mkAbs = mkLambda then there is *) (* some duplicate-free array args of evars of sigma such that the *) (* term mkApp (t', args) is convertible to t. *) (* This makes a useful shorthand for local definitions in proofs, *) (* i.e., pose succ := _ + 1 means pose succ := fun n : nat => n + 1, *) (* and, in context of the the 4CT library, pose mid := maps id means *) (* pose mid := fun d : detaSet => @maps d d (@id (datum d)) *) (* Note that this facility does not extend to set, which tries *) (* instead to fill holes by matching a goal subterm. *) (* The argument to "have" et al. uses product abstraction, e.g. *) (* have Hmid: forall s, (maps id s) = s. *) (* stands for *) (* have Hmid: forall (d : dataSet) (s : seq d), (maps id s) = s. *) (* We also use this feature for rewrite rules, so that, e.g., *) (* rewrite: (plus_assoc _ 3). *) (* will execute as *) (* rewrite (fun n => plus_assoc n 3) *) (* i.e., it will rewrite some subterm .. + (3 + ..) to .. + 3 + ... *) (* The convention is also used for the argument of the congr tactic, *) (* e.g., congr (x + _ * 1). *) (* Replace new evars with lambda variables, retaining local dependencies *) (* but stripping global ones. We use the variable names to encode the *) (* the number of dependencies, so that the transformation is reversible. *) let pf_abs_evars gl (sigma, c0) = let sigma0 = project gl in let nenv = env_size (pf_env gl) in let abs_evar n k = let evi = Evd.find sigma k in let dc = list_firstn n (evar_filtered_context evi) in let abs_dc c = function | x, Some b, t -> mkNamedLetIn x b t (mkArrow t c) | x, None, t -> mkNamedProd x t c in let t = Sign.fold_named_context_reverse abs_dc ~init:evi.evar_concl dc in Evarutil.nf_evar sigma t in let rec put evlist c = match kind_of_term c with | Evar (k, a) -> if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else let n = max 0 (Array.length a - nenv) in let t = abs_evar n k in (k, (n, t)) :: put evlist t | _ -> fold_constr put evlist c in let evlist = put [] c0 in if evlist = [] then 0, c0 else let rec lookup k i = function | [] -> 0, 0 | (k', (n, _)) :: evl -> if k = k' then i, n else lookup k (i + 1) evl in let rec get i c = match kind_of_term c with | Evar (ev, a) -> let j, n = lookup ev i evlist in if j = 0 then map_constr (get i) c else if n = 0 then mkRel j else mkApp (mkRel j, Array.init n (fun k -> get i a.(n - 1 - k))) | _ -> map_constr_with_binders ((+) 1) get i c in let rec loop c i = function | (_, (n, t)) :: evl -> loop (mkLambda (mk_evar_name n, get (i - 1) t, c)) (i - 1) evl | [] -> c in List.length evlist, loop (get 1 c0) 1 evlist (* As before but if (?i : T(?j)) and (?j : P : Prop), then the lambda for i * looks like (fun evar_i : (forall pi : P. T(pi))) thanks to "loopP" and all * occurrences of evar_i are replaced by (evar_i evar_j) thanks to "app". * * If P can be solved by ssrautoprop (that defaults to trivial), then * the corresponding lambda looks like (fun evar_i : T(c)) where c is * the solution found by ssrautoprop. *) let ssrautoprop_tac = ref (fun gl -> assert false) (* Thanks to Arnaud Spiwack for this snippet *) let call_on_evar tac e s = let { it = gs ; sigma = s } = tac { it = Goal.build e ; sigma = s } in gs, s let pf_abs_evars_pirrel gl (sigma, c0) = pp(lazy(str"==PF_ABS_EVARS_PIRREL==")); pp(lazy(str"c0= " ++ pr_constr c0)); let sigma0 = project gl in let nenv = env_size (pf_env gl) in let abs_evar n k = let evi = Evd.find sigma k in let dc = list_firstn n (evar_filtered_context evi) in let abs_dc c = function | x, Some b, t -> mkNamedLetIn x b t (mkArrow t c) | x, None, t -> mkNamedProd x t c in let t = Sign.fold_named_context_reverse abs_dc ~init:evi.evar_concl dc in Evarutil.nf_evar sigma t in let rec put evlist c = match kind_of_term c with | Evar (k, a) -> if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else let n = max 0 (Array.length a - nenv) in let k_ty = Retyping.get_sort_family_of (pf_env gl) sigma (Evd.evar_concl (Evd.find sigma k)) in let is_prop = k_ty = InProp in let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t | _ -> fold_constr put evlist c in let evlist = put [] c0 in if evlist = [] then 0, c0 else let pr_constr t = pr_constr (Reductionops.nf_beta (project gl) t) in let evplist = let depev = List.fold_left (fun evs (_,(_,t,_)) -> Intset.union evs (Evarutil.evars_of_term t)) Intset.empty evlist in List.filter (fun i,(_,_,b) -> b && Intset.mem i depev) evlist in let evlist, evplist, sigma = if evplist = [] then evlist, [], sigma else List.fold_left (fun (ev, evp, sigma) (i, (_,t,_) as p) -> try let ng, sigma = call_on_evar !ssrautoprop_tac i sigma in if (ng <> []) then errorstrm (str "Should we tell the user?"); List.filter (fun (j,_) -> j <> i) ev, evp, sigma with _ -> ev, p::evp, sigma) (evlist, [], sigma) (List.rev evplist) in let c0 = Evarutil.nf_evar sigma c0 in let evlist = List.map (fun (x,(y,t,z)) -> x,(y,Evarutil.nf_evar sigma t,z)) evlist in let evplist = List.map (fun (x,(y,t,z)) -> x,(y,Evarutil.nf_evar sigma t,z)) evplist in pp(lazy(str"c0= " ++ pr_constr c0)); let rec lookup k i = function | [] -> 0, 0 | (k', (n,_,_)) :: evl -> if k = k' then i,n else lookup k (i + 1) evl in let rec get evlist i c = match kind_of_term c with | Evar (ev, a) -> let j, n = lookup ev i evlist in if j = 0 then map_constr (get evlist i) c else if n = 0 then mkRel j else mkApp (mkRel j, Array.init n (fun k -> get evlist i a.(n - 1 - k))) | _ -> map_constr_with_binders ((+) 1) (get evlist) i c in let rec app extra_args i c = match decompose_app c with | hd, args when isRel hd && destRel hd = i -> let j = destRel hd in mkApp (mkRel j, Array.of_list (List.map (lift (i-1)) extra_args @ args)) | _ -> map_constr_with_binders ((+) 1) (app extra_args) i c in let rec loopP evlist c i = function | (_, (n, t, _)) :: evl -> let t = get evlist (i - 1) t in let n = Name (id_of_string (ssr_anon_hyp ^ string_of_int n)) in loopP evlist (mkProd (n, t, c)) (i - 1) evl | [] -> c in let rec loop c i = function | (_, (n, t, _)) :: evl -> let evs = Evarutil.evars_of_term t in let t_evplist = List.filter (fun (k,_) -> Intset.mem k evs) evplist in let t = loopP t_evplist (get t_evplist 1 t) 1 t_evplist in let t = get evlist (i - 1) t in let extra_args = List.map (fun (k,_) -> mkRel (fst (lookup k i evlist))) (List.rev t_evplist) in let c = if extra_args = [] then c else app extra_args 1 c in loop (mkLambda (mk_evar_name n, t, c)) (i - 1) evl | [] -> c in let res = loop (get evlist 1 c0) 1 evlist in pp(lazy(str"res= " ++ pr_constr res)); List.length evlist, res (* Strip all non-essential dependencies from an abstracted term, generating *) (* standard names for the abstracted holes. *) let pf_abs_cterm gl n c0 = if n <= 0 then c0 else let noargs = [|0|] in let eva = Array.make n noargs in let rec strip i c = match kind_of_term c with | App (f, a) when isRel f -> let j = i - destRel f in if j >= n || eva.(j) = noargs then mkApp (f, Array.map (strip i) a) else let dp = eva.(j) in let nd = Array.length dp - 1 in let mkarg k = strip i a.(if k < nd then dp.(k + 1) - j else k + dp.(0)) in mkApp (f, Array.init (Array.length a - dp.(0)) mkarg) | _ -> map_constr_with_binders ((+) 1) strip i c in let rec strip_ndeps j i c = match kind_of_term c with | Prod (x, t, c1) when i < j -> let dl, c2 = strip_ndeps j (i + 1) c1 in if noccurn 1 c2 then dl, lift (-1) c2 else i :: dl, mkProd (x, strip i t, c2) | LetIn (x, b, t, c1) when i < j -> let _, _, c1' = destProd c1 in let dl, c2 = strip_ndeps j (i + 1) c1' in if noccurn 1 c2 then dl, lift (-1) c2 else i :: dl, mkLetIn (x, strip i b, strip i t, c2) | _ -> [], strip i c in let rec strip_evars i c = match kind_of_term c with | Lambda (x, t1, c1) when i < n -> let na = nb_evar_deps x in let dl, t2 = strip_ndeps (i + na) i t1 in let na' = List.length dl in eva.(i) <- Array.of_list (na - na' :: dl); let x' = if na' = 0 then Name (pf_type_id gl t2) else mk_evar_name na' in mkLambda (x', t2, strip_evars (i + 1) c1) (* if noccurn 1 c2 then lift (-1) c2 else mkLambda (Name (pf_type_id gl t2), t2, c2) *) | _ -> strip i c in strip_evars 0 c0 (* Undo the evar abstractions. Also works for non-evar variables. *) let pf_unabs_evars gl ise n c0 = if n = 0 then c0 else let evv = Array.make n mkProp in let nev = ref 0 in let env0 = pf_env gl in let nenv0 = env_size env0 in let rec unabs i c = match kind_of_term c with | Rel j when i - j < !nev -> evv.(i - j) | App (f, a0) when isRel f -> let a = Array.map (unabs i) a0 in let j = i - destRel f in if j >= !nev then mkApp (f, a) else let ev, eva = destEvar evv.(j) in let nd = Array.length eva - nenv0 in if nd = 0 then mkApp (evv.(j), a) else let evarg k = if k < nd then a.(nd - 1 - k) else eva.(k) in let c' = mkEvar (ev, Array.init (nd + nenv0) evarg) in let na = Array.length a - nd in if na = 0 then c' else mkApp (c', Array.sub a nd na) | _ -> map_constr_with_binders ((+) 1) unabs i c in let push_rel = Environ.push_rel in let rec mk_evar j env i c = match kind_of_term c with | Prod (x, t, c1) when i < j -> mk_evar j (push_rel (x, None, unabs i t) env) (i + 1) c1 | LetIn (x, b, t, c1) when i < j -> let _, _, c2 = destProd c1 in mk_evar j (push_rel (x, Some (unabs i b), unabs i t) env) (i + 1) c2 | _ -> Evarutil.e_new_evar ise env (unabs i c) in let rec unabs_evars c = if !nev = n then unabs n c else match kind_of_term c with | Lambda (x, t, c1) when !nev < n -> let i = !nev in evv.(i) <- mk_evar (i + nb_evar_deps x) env0 i t; incr nev; unabs_evars c1 | _ -> unabs !nev c in unabs_evars c0 (* }}} *) (** Tactical extensions. {{{ **************************************************) (* The TACTIC EXTEND facility can't be used for defining new user *) (* tacticals, because: *) (* - the concrete syntax must start with a fixed string *) (* - the lexical Ltac environment is NOT used to interpret tactic *) (* arguments *) (* The second limitation means that the extended tacticals will *) (* exhibit run-time scope errors if used inside Ltac functions or *) (* pattern-matching constructs. *) (* We use the following workaround: *) (* - We use the (unparsable) "Qed" token for tacticals that *) (* don't start with a token, then redefine the grammar and *) (* printer using GEXTEND and set_pr_ssrtac, respectively. *) (* - We use a global stack and side effects to pass the lexical *) (* Ltac evaluation context to the extended tactical. The context *) (* is grabbed by interpreting an (empty) ltacctx argument, *) (* which should appear last in the grammar rules; the *) (* get_ltacctx function pops the stack and returns the context. *) (* For additional safety, the push returns an integer key that *) (* is checked by the pop; since arguments are interpreted *) (* left-to-right, this checks that only one tactic argument *) (* pushes a context. *) (* - To avoid a spurrious option type, we don't push the context *) (* for a null tag. *) type ssrargfmt = ArgSsr of string | ArgCoq of argument_type | ArgSep of string let set_pr_ssrtac name prec afmt = let fmt = List.map (function ArgSep s -> Some s | _ -> None) afmt in let rec mk_akey = function | ArgSsr s :: afmt' -> ExtraArgType ("ssr" ^ s) :: mk_akey afmt' | ArgCoq a :: afmt' -> a :: mk_akey afmt' | ArgSep _ :: afmt' -> mk_akey afmt' | [] -> [] in let tacname = "ssr" ^ name in Pptactic.declare_extra_tactic_pprule (tacname, mk_akey afmt, (prec, fmt)) let ssrtac_atom loc name args = TacExtend (loc, "ssr" ^ name, args) let ssrtac_expr loc name args = TacAtom (loc, ssrtac_atom loc name args) let ssrevaltac ist gtac = interp_tac_gen ist.lfun [] ist.debug (globTacticIn (fun _ -> gtac)) (* fun gl -> let lfun = [tacarg_id, val_interp ist gl gtac] in interp_tac_gen lfun [] ist.debug tacarg_expr gl *) (** Generic argument-based globbing/typing utilities *) let interp_wit globwit wit ist gl x = let globarg = in_gen globwit x in let sigma, arg = interp_genarg ist gl globarg in sigma, out_gen wit arg let interp_intro_pattern = interp_wit globwit_intro_pattern wit_intro_pattern let interp_constr = interp_wit globwit_constr wit_constr let interp_open_constr ist gl gc = interp_wit globwit_open_constr wit_open_constr ist gl ((), gc) let interp_refine ist gl rc = let roc = (), (rc, None) in interp_wit globwit_casted_open_constr wit_casted_open_constr ist gl roc let pf_match = pf_apply (fun e s c t -> understand_tcc s e ~expected_type:t c) (* Estimate a bound on the number of arguments of a raw constr. *) (* This is not perfect, because the unifier may fail to *) (* typecheck the partial application, so we use a minimum of 5. *) (* Also, we don't handle delayed or iterated coercions to *) (* FUNCLASS, which is probably just as well since these can *) (* lead to infinite arities. *) let splay_open_constr gl (sigma, c) = let env = pf_env gl in let t = Retyping.get_type_of env sigma c in Reductionops.splay_prod env sigma t let nbargs_open_constr gl oc = let pl, _ = splay_open_constr gl oc in List.length pl let interp_nbargs ist gl rc = try let rc6 = mkRApp rc (mkRHoles 6) in let sigma, t = interp_open_constr ist gl (rc6, None) in let si = sig_it gl in let gl = re_sig si sigma in 6 + nbargs_open_constr gl t with _ -> 5 let pf_nbargs gl c = nbargs_open_constr gl (project gl, c) let isAppInd gl c = try ignore (pf_reduce_to_atomic_ind gl c); true with _ -> false let interp_view_nbimps ist gl rc = try let sigma, t = interp_open_constr ist gl (rc, None) in let si = sig_it gl in let gl = re_sig si sigma in let pl, c = splay_open_constr gl t in if isAppInd gl c then List.length pl else ~-(List.length pl) with _ -> 0 (* }}} *) (** Vernacular commands: Prenex Implicits and Search {{{ **********************) (* This should really be implemented as an extension to the implicit *) (* arguments feature, but unfortuately that API is sealed. The current *) (* workaround uses a combination of notations that works reasonably, *) (* with the following caveats: *) (* - The pretty-printing always elides prenex implicits, even when *) (* they are obviously needed. *) (* - Prenex Implicits are NEVER exported from a module, because this *) (* would lead to faulty pretty-printing and scoping errors. *) (* - The command "Import Prenex Implicits" can be used to reassert *) (* Prenex Implicits for all the visible constants that had been *) (* declared as Prenex Implicits. *) let declare_one_prenex_implicit locality f = let fref = try Smartlocate.global_with_alias f with _ -> errorstrm (pr_reference f ++ str " is not declared") in let rec loop = function | a :: args' when Impargs.is_status_implicit a -> (ExplByName (Impargs.name_of_implicit a), (true, true, true)) :: loop args' | args' when List.exists Impargs.is_status_implicit args' -> errorstrm (str "Expected prenex implicits for " ++ pr_reference f) | _ -> [] in let impls = match Impargs.implicits_of_global fref with | [cond,impls] -> impls | [] -> errorstrm (str "Expected some implicits for " ++ pr_reference f) | _ -> errorstrm (str "Multiple implicits not supported") in match loop impls with | [] -> errorstrm (str "Expected some implicits for " ++ pr_reference f) | impls -> Impargs.declare_manual_implicits locality fref ~enriching:false [impls] VERNAC COMMAND EXTEND Ssrpreneximplicits | [ "Prenex" "Implicits" ne_global_list(fl) ] -> [ let locality = Vernacexpr.use_section_locality () in List.iter (declare_one_prenex_implicit locality) fl ] END (* Vernac grammar visibility patch *) GEXTEND Gram GLOBAL: gallina_ext; gallina_ext: [ [ IDENT "Import"; IDENT "Prenex"; IDENT "Implicits" -> Vernacexpr.VernacUnsetOption (None, ["Printing"; "Implicit"; "Defensive"]) ] ] ; END (* Remove the silly restriction that forces coercion classes to be precise *) (* aliases, e.g., allowing notations that specify some class parameters. *) let qualify_ref clref = let loc, qid = qualid_of_reference clref in try match Nametab.locate_extended qid with | TrueGlobal _ -> clref | SynDef kn -> let rec head_of = function | ARef gref -> Qualid (loc, Nametab.shortest_qualid_of_global Idset.empty gref) | AApp (rc, _) -> head_of rc | ACast (rc, _) -> head_of rc | ALetIn (_, _, rc) -> head_of rc | rc -> user_err_loc (loc, "qualify_ref", str "The definition of " ++ Ppconstr.pr_qualid qid ++ str " does not have a head constant") in head_of (snd (Syntax_def.search_syntactic_definition kn)) with _ -> clref let class_rawexpr = G_vernac.class_rawexpr in GEXTEND Gram GLOBAL: class_rawexpr; ssrqref: [[ gref = global -> qualify_ref gref ]]; class_rawexpr: [[ class_ref = ssrqref -> Vernacexpr.RefClass (Genarg.AN class_ref) ]]; END (** Extend Search to subsume SearchAbout, also adding hidden Type coercions. *) (* Main prefilter *) let pr_search_item = function | Search.GlobSearchString s -> str s | Search.GlobSearchSubPattern p -> pr_constr_pattern p let wit_ssr_searchitem, globwit_ssr_searchitem, rawwit_ssr_searchitem = add_genarg "ssr_searchitem" pr_search_item let interp_search_notation loc s opt_scope = try let interp = Notation.interp_notation_as_global_reference loc in let ref = interp (fun _ -> true) s opt_scope in Search.GlobSearchSubPattern (Pattern.PRef ref) with _ -> let diagnosis = try let ntns = Notation.locate_notation pr_glob_constr s opt_scope in let ambig = "This string refers to a complex or ambiguous notation." in str ambig ++ str "\nTry searching with one of\n" ++ ntns with _ -> str "This string is not part of an identifier or notation." in user_err_loc (loc, "interp_search_notation", diagnosis) let pr_ssr_search_item _ _ _ = pr_search_item (* Workaround the notation API that can only print notations *) let is_ident s = try Lexer.check_ident s; true with _ -> false let is_ident_part s = is_ident ("H" ^ s) let interp_search_notation loc tag okey = let err msg = user_err_loc (loc, "interp_search_notation", msg) in let mk_pntn s for_key = let n = String.length s in let s' = String.make (n + 2) ' ' in let rec loop i i' = if i >= n then s', i' - 2 else if s.[i] = ' ' then loop (i + 1) i' else let j = try String.index_from s (i + 1) ' ' with _ -> n in let m = j - i in if s.[i] = '\'' && i < j - 2 && s.[j - 1] = '\'' then (String.blit s (i + 1) s' i' (m - 2); loop (j + 1) (i' + m - 1)) else if for_key && is_ident (String.sub s i m) then (s'.[i'] <- '_'; loop (j + 1) (i' + 2)) else (String.blit s i s' i' m; loop (j + 1) (i' + m + 1)) in loop 0 1 in let trim_ntn (pntn, m) = String.sub pntn 1 (max 0 m) in let pr_ntn ntn = str "(" ++ str ntn ++ str ")" in let pr_and_list pr = function | [x] -> pr x | x :: lx -> pr_list pr_comma pr lx ++ pr_comma () ++ str "and " ++ pr x | [] -> mt () in let pr_sc sc = str (if sc = "" then "independently" else sc) in let pr_scs = function | [""] -> pr_sc "" | scs -> str "in " ++ pr_and_list pr_sc scs in let generator, pr_tag_sc = let ign _ = mt () in match okey with | Some key -> let sc = Notation.find_delimiters_scope loc key in let pr_sc s_in = str s_in ++ spc() ++ str sc ++ pr_comma() in Notation.pr_scope ign sc, pr_sc | None -> Notation.pr_scopes ign, ign in let qtag s_in = pr_tag_sc s_in ++ qstring tag ++ spc()in let ptag, ttag = let ptag, m = mk_pntn tag false in if m <= 0 then err (str "empty notation fragment"); ptag, trim_ntn (ptag, m) in let last = ref "" and last_sc = ref "" in let scs = ref [] and ntns = ref [] in let push_sc sc = match !scs with | "" :: scs' -> scs := "" :: sc :: scs' | scs' -> scs := sc :: scs' in let get s _ _ = match !last with | "Scope " -> last_sc := s; last := "" | "Lonely notation" -> last_sc := ""; last := "" | "\"" -> let pntn, m = mk_pntn s true in if string_string_contains pntn ptag then begin let ntn = trim_ntn (pntn, m) in match !ntns with | [] -> ntns := [ntn]; scs := [!last_sc] | ntn' :: _ when ntn' = ntn -> push_sc !last_sc | _ when ntn = ttag -> ntns := ntn :: !ntns; scs := [!last_sc] | _ :: ntns' when List.mem ntn ntns' -> () | ntn' :: ntns' -> ntns := ntn' :: ntn :: ntns' end; last := "" | _ -> last := s in pp_with (Format.make_formatter get (fun _ -> ())) generator; let ntn = match !ntns with | [] -> err (hov 0 (qtag "in" ++ str "does not occur in any notation")) | ntn :: ntns' when ntn = ttag -> if ntns' <> [] then begin let pr_ntns' = pr_and_list pr_ntn ntns' in msg_warning (hov 4 (qtag "In" ++ str "also occurs in " ++ pr_ntns')) end; ntn | [ntn] -> msgnl (hov 4 (qtag "In" ++ str "is part of notation " ++ pr_ntn ntn)); ntn | ntns' -> let e = str "occurs in" ++ spc() ++ pr_and_list pr_ntn ntns' in err (hov 4 (str "ambiguous: " ++ qtag "in" ++ e)) in let (nvars, body), ((_, pat), osc) = match !scs with | [sc] -> Notation.interp_notation loc ntn (None, [sc]) | scs' -> try Notation.interp_notation loc ntn (None, []) with _ -> let e = pr_ntn ntn ++ spc() ++ str "is defined " ++ pr_scs scs' in err (hov 4 (str "ambiguous: " ++ pr_tag_sc "in" ++ e)) in let sc = Option.default "" osc in let _ = let m_sc = if osc <> None then str "In " ++ str sc ++ pr_comma() else mt() in let ntn_pat = trim_ntn (mk_pntn pat false) in let rbody = glob_constr_of_aconstr loc body in let m_body = hov 0 (Constrextern.without_symbols prl_glob_constr rbody) in let m = m_sc ++ pr_ntn ntn_pat ++ spc () ++ str "denotes " ++ m_body in msgnl (hov 0 m) in if List.length !scs > 1 then let scs' = list_remove sc !scs in let w = pr_ntn ntn ++ str " is also defined " ++ pr_scs scs' in msg_warning (hov 4 w) else if string_string_contains ntn " .. " then err (pr_ntn ntn ++ str " is an n-ary notation"); let nvars = List.filter (fun (_,(_,typ)) -> typ = NtnTypeConstr) nvars in let rec sub () = function | AVar x when List.mem_assoc x nvars -> GPatVar (loc, (false, x)) | c -> glob_constr_of_aconstr_with_binders loc (fun _ x -> (), x) sub () c in let _, npat = Pattern.pattern_of_glob_constr (sub () body) in Search.GlobSearchSubPattern npat ARGUMENT EXTEND ssr_search_item TYPED AS ssr_searchitem PRINTED BY pr_ssr_search_item | [ string(s) ] -> [ if is_ident_part s then Search.GlobSearchString s else interp_search_notation loc s None ] | [ string(s) "%" preident(key) ] -> [ interp_search_notation loc s (Some key) ] | [ constr_pattern(p) ] -> [ try let intern = Constrintern.intern_constr_pattern Evd.empty in Search.GlobSearchSubPattern (snd (intern (Global.env()) p)) with e -> raise (Cerrors.process_vernac_interp_error e) ] END let pr_ssr_search_arg _ _ _ = let pr_item (b, p) = str (if b then "-" else "") ++ pr_search_item p in pr_list spc pr_item ARGUMENT EXTEND ssr_search_arg TYPED AS (bool * ssr_searchitem) list PRINTED BY pr_ssr_search_arg | [ "-" ssr_search_item(p) ssr_search_arg(a) ] -> [ (false, p) :: a ] | [ ssr_search_item(p) ssr_search_arg(a) ] -> [ (true, p) :: a ] | [ ] -> [ [] ] END (* Main type conclusion pattern filter *) let rec splay_search_pattern na = function | Pattern.PApp (fp, args) -> splay_search_pattern (na + Array.length args) fp | Pattern.PLetIn (_, _, bp) -> splay_search_pattern na bp | Pattern.PRef hr -> hr, na | _ -> error "no head constant in head search pattern" let coerce_search_pattern_to_sort hpat = let env = Global.env () and sigma = Evd.empty in let mkPApp fp n_imps args = let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in Pattern.PApp (fp, args') in let hr, na = splay_search_pattern 0 hpat in let dc, ht = Reductionops.splay_prod env sigma (Global.type_of_global hr) in let np = List.length dc in if np < na then error "too many arguments in head search pattern" else let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in let warn () = msg_warning (str "Listing only lemmas with conclusion matching " ++ pr_constr_pattern hpat') in if isSort ht then begin warn (); true, hpat' end else let filter_head, coe_path = try let _, cp = Classops.lookup_path_to_sort_from (push_rels_assum dc env) sigma ht in warn (); true, cp with _ -> false, [] in let coerce hp coe_index = let coe = Classops.get_coercion_value coe_index in try let coe_ref = reference_of_constr coe in let n_imps = Option.get (Classops.hide_coercion coe_ref) in mkPApp (Pattern.PRef coe_ref) n_imps [|hp|] with _ -> errorstrm (str "need explicit coercion " ++ pr_constr coe ++ spc () ++ str "to interpret head search pattern as type") in filter_head, List.fold_left coerce hpat' coe_path let rec interp_head_pat hpat = let filter_head, p = coerce_search_pattern_to_sort hpat in let rec loop c = match kind_of_term c with | Cast (c', _, _) -> loop c' | Prod (_, _, c') -> loop c' | LetIn (_, _, _, c') -> loop c' | _ -> Matching.is_matching p c in filter_head, loop let all_true _ = true let interp_search_arg a = let hpat, a1 = match a with | (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a' | (true, Search.GlobSearchSubPattern p) :: a' -> let filter_head, p = interp_head_pat p in if filter_head then p, a' else all_true, a | _ -> all_true, a in let is_string = function (_, Search.GlobSearchString _) -> true | _ -> false in let a2, a3 = List.partition is_string a1 in hpat, a2 @ a3 (* Module path postfilter *) let pr_modloc (b, m) = if b then str "-" ++ pr_reference m else pr_reference m let wit_ssrmodloc, globwit_ssrmodloc, rawwit_ssrmodloc = add_genarg "ssrmodloc" pr_modloc let pr_ssr_modlocs _ _ _ ml = if ml = [] then str "" else spc () ++ str "in " ++ pr_list spc pr_modloc ml ARGUMENT EXTEND ssr_modlocs TYPED AS ssrmodloc list PRINTED BY pr_ssr_modlocs | [ ] -> [ [] ] END GEXTEND Gram GLOBAL: ssr_modlocs; modloc: [[ "-"; m = global -> true, m | m = global -> false, m]]; ssr_modlocs: [[ "in"; ml = LIST1 modloc -> ml ]]; END let interp_modloc mr = let interp_mod (_, mr) = let (loc, qid) = qualid_of_reference mr in try Nametab.full_name_module qid with Not_found -> user_err_loc (loc, "interp_modloc", str "No Module " ++ pr_qualid qid) in let mr_out, mr_in = List.partition fst mr in let interp_bmod b rmods = if rmods = [] then fun _ _ _ -> true else Search.filter_by_module_from_list (List.map interp_mod rmods, b) in let is_in = interp_bmod false mr_in and is_out = interp_bmod true mr_out in fun gr env -> is_in gr env () && is_out gr env () (* The unified, extended vernacular "Search" command *) let ssrdisplaysearch gr env t = let pr_res = pr_global gr ++ spc () ++ str " " ++ pr_lconstr_env env t in msg (hov 2 pr_res ++ fnl ()) VERNAC COMMAND EXTEND SsrSearchPattern | [ "Search" ssr_search_arg(a) ssr_modlocs(mr) ] -> [ let hpat, a' = interp_search_arg a in let in_mod = interp_modloc mr in let post_filter gr env typ = in_mod gr env && hpat typ in Search.raw_search_about post_filter ssrdisplaysearch a' ] END (* }}} *) (** Alternative notations for "match" and anonymous arguments. {{{ ************) (* Syntax: *) (* if is then ... else ... *) (* if is [in ..] return ... then ... else ... *) (* let: := in ... *) (* let: [in ...] := return ... in ... *) (* The scope of a top-level 'as' in the pattern extends over the *) (* 'return' type (dependent if/let). *) (* Note that the optional "in ..." appears next to the *) (* rather than the in then "let:" syntax. The alternative *) (* would lead to ambiguities in, e.g., *) (* let: p1 := (*v---INNER LET:---v *) *) (* let: p2 := let: p3 := e3 in k return t in k2 in k1 return t' *) (* in b (*^--ALTERNATIVE INNER LET--------^ *) *) (* Caveat : There is no pretty-printing support, since this would *) (* require a modification to the Coq kernel (adding a new match *) (* display style -- why aren't these strings?); also, the v8.1 *) (* pretty-printer only allows extension hooks for printing *) (* integer or string literals. *) (* Also note that in the v8 grammar "is" needs to be a keyword; *) (* as this can't be done from an ML extension file, the new *) (* syntax will only work when ssreflect.v is imported. *) let no_ct = None, None and no_rt = None in let aliasvar = function | [_, [CPatAlias (loc, _, id)]] -> Some (loc,Name id) | _ -> None in let mk_cnotype mp = aliasvar mp, None in let mk_ctype mp t = aliasvar mp, Some t in let mk_rtype t = Some t in let mk_dthen loc (mp, ct, rt) c = (loc, mp, c), ct, rt in let mk_let loc rt ct mp c1 = CCases (loc, LetPatternStyle, rt, ct, [loc, mp, c1]) in GEXTEND Gram GLOBAL: binder_constr; ssr_rtype: [[ "return"; t = operconstr LEVEL "100" -> mk_rtype t ]]; ssr_mpat: [[ p = pattern -> [loc, [p]] ]]; ssr_dpat: [ [ mp = ssr_mpat; "in"; t = lconstr; rt = ssr_rtype -> mp, mk_ctype mp t, rt | mp = ssr_mpat; rt = ssr_rtype -> mp, mk_cnotype mp, rt | mp = ssr_mpat -> mp, no_ct, no_rt ] ]; ssr_dthen: [[ dp = ssr_dpat; "then"; c = lconstr -> mk_dthen loc dp c ]]; ssr_elsepat: [[ "else" -> [loc, [CPatAtom (loc, None)]] ]]; ssr_else: [[ mp = ssr_elsepat; c = lconstr -> loc, mp, c ]]; binder_constr: [ [ "if"; c = operconstr LEVEL "200"; "is"; db1 = ssr_dthen; b2 = ssr_else -> let b1, ct, rt = db1 in CCases (loc, MatchStyle, rt, [c, ct], [b1; b2]) | "if"; c = operconstr LEVEL "200";"isn't";db1 = ssr_dthen; b2 = ssr_else -> let b1, ct, rt = db1 in let b1, b2 = let (l1, p1, r1), (l2, p2, r2) = b1, b2 in (l1, p1, r2), (l2, p2, r1) in CCases (loc, MatchStyle, rt, [c, ct], [b1; b2]) | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; "in"; c1 = lconstr -> mk_let loc no_rt [c, no_ct] mp c1 | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; rt = ssr_rtype; "in"; c1 = lconstr -> mk_let loc rt [c, mk_cnotype mp] mp c1 | "let"; ":"; mp = ssr_mpat; "in"; t = lconstr; ":="; c = lconstr; rt = ssr_rtype; "in"; c1 = lconstr -> mk_let loc rt [c, mk_ctype mp t] mp c1 ] ]; END GEXTEND Gram GLOBAL: closed_binder; closed_binder: [ [ ["of" | "&"]; c = operconstr LEVEL "99" -> [LocalRawAssum ([loc, Anonymous], Default Explicit, c)] ] ]; END (* }}} *) (** Tacticals (+, -, *, done, by, do, =>, first, and last). {{{ ***************) (** Bracketing tactical *) (* The tactic pretty-printer doesn't know that some extended tactics *) (* are actually tacticals. To prevent it from improperly removing *) (* parentheses we override the parsing rule for bracketed tactic *) (* expressions so that the pretty-print always reflects the input. *) (* (Removing user-specified parentheses is dubious anyway). *) GEXTEND Gram GLOBAL: tactic_expr; ssrparentacarg: [[ "("; tac = tactic_expr; ")" -> loc, Tacexp tac ]]; tactic_expr: LEVEL "0" [[ arg = ssrparentacarg -> TacArg arg ]]; END (** The internal "done" and "ssrautoprop" tactics. *) (* For additional flexibility, "done" and "ssrautoprop" are *) (* defined in Ltac. *) (* Although we provide a default definition in ssreflect, *) (* we look up the definition dynamically at each call point, *) (* to allow for user extensions. "ssrautoprop" defaults to *) (* trivial. *) let donetac gl = let tacname = try Nametab.locate_tactic (qualid_of_ident (id_of_string "done")) with Not_found -> try Nametab.locate_tactic (ssrqid "done") with Not_found -> Util.error "The ssreflect library was not loaded" in let tacexpr = dummy_loc, Tacexpr.Reference (ArgArg (dummy_loc, tacname)) in eval_tactic (Tacexpr.TacArg tacexpr) gl let prof_donetac = mk_profiler "donetac";; let donetac gl = prof_donetac.profile donetac gl;; let ssrautoprop gl = try let tacname = try Nametab.locate_tactic (qualid_of_ident (id_of_string "ssrautoprop")) with Not_found -> Nametab.locate_tactic (ssrqid "ssrautoprop") in let tacexpr = dummy_loc, Tacexpr.Reference (ArgArg (dummy_loc, tacname)) in eval_tactic (Tacexpr.TacArg tacexpr) gl with Not_found -> Auto.full_trivial [] gl let () = ssrautoprop_tac := ssrautoprop let tclBY tac = tclTHEN tac donetac (** Tactical arguments. *) (* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) (* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) (* and subgoal reordering tacticals (; first & ; last), respectively. *) (* Force use of the tactic_expr parsing entry, to rule out tick marks. *) let pr_ssrtacarg _ _ prt = prt tacltop ARGUMENT EXTEND ssrtacarg TYPED AS tactic PRINTED BY pr_ssrtacarg | [ "Qed" ] -> [ Util.anomaly "Grammar placeholder match" ] END GEXTEND Gram GLOBAL: ssrtacarg; ssrtacarg: [[ tac = tactic_expr LEVEL "5" -> tac ]]; END (* Lexically closed tactic for tacticals. *) let pr_ssrtclarg _ _ prt (tac, _) = prt tacltop tac ARGUMENT EXTEND ssrtclarg TYPED AS ssrtacarg * ltacctx PRINTED BY pr_ssrtclarg | [ ssrtacarg(tac) ] -> [ tac, rawltacctx ] END let eval_tclarg (tac, ctx) = ssrevaltac (get_ltacctx ctx) tac let pr_ortacs prt = let rec pr_rec = function | [None] -> spc() ++ str "|" ++ spc() | None :: tacs -> spc() ++ str "|" ++ pr_rec tacs | Some tac :: tacs -> spc() ++ str "| " ++ prt tacltop tac ++ pr_rec tacs | [] -> mt() in function | [None] -> spc() | None :: tacs -> pr_rec tacs | Some tac :: tacs -> prt tacltop tac ++ pr_rec tacs | [] -> mt() let pr_ssrortacs _ _ = pr_ortacs ARGUMENT EXTEND ssrortacs TYPED AS tactic option list PRINTED BY pr_ssrortacs | [ ssrtacarg(tac) "|" ssrortacs(tacs) ] -> [ Some tac :: tacs ] | [ ssrtacarg(tac) "|" ] -> [ [Some tac; None] ] | [ ssrtacarg(tac) ] -> [ [Some tac] ] | [ "|" ssrortacs(tacs) ] -> [ None :: tacs ] | [ "|" ] -> [ [None; None] ] END let pr_hintarg prt = function | true, tacs -> hv 0 (str "[ " ++ pr_ortacs prt tacs ++ str " ]") | false, [Some tac] -> prt tacltop tac | _, _ -> mt() let pr_ssrhintarg _ _ = pr_hintarg let mk_hint tac = false, [Some tac] let mk_orhint tacs = true, tacs let nullhint = true, [] let nohint = false, [] ARGUMENT EXTEND ssrhintarg TYPED AS bool * ssrortacs PRINTED BY pr_ssrhintarg | [ "[" "]" ] -> [ nullhint ] | [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ] | [ ssrtacarg(arg) ] -> [ mk_hint arg ] END ARGUMENT EXTEND ssrortacarg TYPED AS ssrhintarg PRINTED BY pr_ssrhintarg | [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ] END let hinttac ist is_by (is_or, atacs) = let dtac = if is_by then donetac else tclIDTAC in let mktac = function | Some atac -> tclTHEN (ssrevaltac ist atac) dtac | _ -> dtac in match List.map mktac atacs with | [] -> if is_or then dtac else tclIDTAC | [tac] -> tac | tacs -> tclFIRST tacs (** The "-"/"+"/"*" tacticals. *) (* These are just visual cues to flag the beginning of the script for *) (* new subgoals, when indentation is not appropriate (typically after *) (* tactics that generate more than two subgoals). *) TACTIC EXTEND ssrtclplus | [ "Qed" "+" ssrtclarg(arg) ] -> [ eval_tclarg arg ] END set_pr_ssrtac "tclplus" 5 [ArgSep "+ "; ArgSsr "tclarg"] TACTIC EXTEND ssrtclminus | [ "Qed" "-" ssrtclarg(arg) ] -> [ eval_tclarg arg ] END set_pr_ssrtac "tclminus" 5 [ArgSep "- "; ArgSsr "tclarg"] TACTIC EXTEND ssrtclstar | [ "Qed" "*" ssrtclarg(arg) ] -> [ eval_tclarg arg ] END set_pr_ssrtac "tclstar" 5 [ArgSep "- "; ArgSsr "tclarg"] let gen_tclarg = in_gen rawwit_ssrtclarg GEXTEND Gram GLOBAL: tactic tactic_mode; tactic: [ [ "+"; tac = ssrtclarg -> ssrtac_expr loc "tclplus" [gen_tclarg tac] | "-"; tac = ssrtclarg -> ssrtac_expr loc "tclminus" [gen_tclarg tac] | "*"; tac = ssrtclarg -> ssrtac_expr loc "tclstar" [gen_tclarg tac] ] ]; tactic_mode: [ [ "+"; tac = G_vernac.subgoal_command -> tac None | "-"; tac = G_vernac.subgoal_command -> tac None | "*"; tac = G_vernac.subgoal_command -> tac None ] ]; END (** The "by" tactical. *) let pr_hint prt arg = if arg = nohint then mt() else str "by " ++ pr_hintarg prt arg let pr_ssrhint _ _ = pr_hint ARGUMENT EXTEND ssrhint TYPED AS ssrhintarg PRINTED BY pr_ssrhint | [ ] -> [ nohint ] END TACTIC EXTEND ssrtclby | [ "Qed" ssrhint(tac) ltacctx(ctx)] -> [ hinttac (get_ltacctx ctx) true tac ] END set_pr_ssrtac "tclby" 0 [ArgSsr "hint"; ArgSsr "ltacctx"] (* We can't parse "by" in ARGUMENT EXTEND because it will only be made *) (* into a keyword in ssreflect.v; so we anticipate this in GEXTEND. *) GEXTEND Gram GLOBAL: ssrhint simple_tactic; ssrhint: [[ "by"; arg = ssrhintarg -> arg ]]; simple_tactic: [ [ "by"; arg = ssrhintarg -> let garg = in_gen rawwit_ssrhint arg in let gctx = in_gen rawwit_ltacctx rawltacctx in ssrtac_atom loc "tclby" [garg; gctx] ] ]; END (* }}} *) (** Bound assumption argument *) (* The Ltac API does have a type for assumptions but it is level-dependent *) (* and therefore impratical to use for complex arguments, so we substitute *) (* our own to have a uniform representation. Also, we refuse to intern *) (* idents that match global/section constants, since this would lead to *) (* fragile Ltac scripts. *) type ssrhyp = SsrHyp of loc * identifier let hyp_id (SsrHyp (_, id)) = id let pr_hyp (SsrHyp (_, id)) = pr_id id let pr_ssrhyp _ _ _ = pr_hyp let wit_ssrhyprep, globwit_ssrhyprep, rawwit_ssrhyprep = add_genarg "ssrhyprep" pr_hyp let hyp_err loc msg id = user_err_loc (loc, "ssrhyp", str msg ++ pr_id id) let intern_hyp ist (SsrHyp (loc, id) as hyp) = let _ = intern_genarg ist (in_gen rawwit_var (loc, id)) in if not_section_id id then hyp else hyp_err loc "Can't clear section hypothesis " id let interp_hyp ist gl (SsrHyp (loc, id)) = let s, id' = interp_wit globwit_var wit_var ist gl (loc, id) in if not_section_id id' then s, SsrHyp (loc, id') else hyp_err loc "Can't clear section hypothesis " id' ARGUMENT EXTEND ssrhyp TYPED AS ssrhyprep PRINTED BY pr_ssrhyp INTERPRETED BY interp_hyp GLOBALIZED BY intern_hyp | [ ident(id) ] -> [ SsrHyp (loc, id) ] END type ssrhyps = ssrhyp list let pr_hyps = pr_list pr_spc pr_hyp let pr_ssrhyps _ _ _ = pr_hyps let hyps_ids = List.map hyp_id let rec check_hyps_uniq ids = function | SsrHyp (loc, id) :: _ when List.mem id ids -> hyp_err loc "Duplicate assumption " id | SsrHyp (_, id) :: hyps -> check_hyps_uniq (id :: ids) hyps | [] -> () let check_hyp_exists hyps (SsrHyp(_, id)) = try ignore(Sign.lookup_named id hyps) with Not_found -> errorstrm (str"No assumption is named " ++ pr_id id) let interp_hyps ist gl ghyps = let hyps = List.map snd (List.map (interp_hyp ist gl) ghyps) in check_hyps_uniq [] hyps; Tacmach.project gl, hyps ARGUMENT EXTEND ssrhyps TYPED AS ssrhyp list PRINTED BY pr_ssrhyps INTERPRETED BY interp_hyps | [ ssrhyp_list(hyps) ] -> [ check_hyps_uniq [] hyps; hyps ] END (** Terms parsing. {{{ ********************************************************) (* Because we allow wildcards in term references, we need to stage the *) (* interpretation of terms so that it occurs at the right time during *) (* the execution of the tactic (e.g., so that we don't report an error *) (* for a term that isn't actually used in the execution). *) (* The term representation tracks whether the concrete initial term *) (* started with an opening paren, which might avoid a conflict between *) (* the ssrreflect term syntax and Gallina notation. *) (* kinds of terms *) type ssrtermkind = char (* print flag *) let input_ssrtermkind strm = match Compat.get_tok (stream_nth 0 strm) with | Tok.KEYWORD "(" -> '(' | Tok.KEYWORD "@" -> '@' | _ -> ' ' let ssrtermkind = Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind let id_of_Cterm t = match id_of_cpattern t with | Some x -> x | None -> loc_error (loc_of_cpattern t) "Only identifiers are allowed here" let ssrhyp_of_ssrterm = function | k, (_, Some c) as o -> SsrHyp (constr_loc c, id_of_Cterm (cpattern_of_term o)), String.make 1 k | _, (_, None) -> assert false (* terms *) let pr_ssrterm _ _ _ = pr_term let pf_intern_term ist gl (_, c) = glob_constr ist (project gl) (pf_env gl) c let intern_term ist sigma env (_, c) = glob_constr ist sigma env c let interp_term ist gl (_, c) = snd (interp_open_constr ist gl c) let force_term ist gl (_, c) = interp_constr ist gl c let glob_ssrterm gs = function | k, (_, Some c) -> k, Tacinterp.intern_constr gs c | ct -> ct let subst_ssrterm s (k, c) = k, Tacinterp.subst_glob_constr_and_expr s c let interp_ssrterm _ gl t = Tacmach.project gl, t ARGUMENT EXTEND ssrterm PRINTED BY pr_ssrterm INTERPRETED BY interp_ssrterm GLOBALIZED BY glob_ssrterm SUBSTITUTED BY subst_ssrterm RAW_TYPED AS cpattern RAW_PRINTED BY pr_ssrterm GLOB_TYPED AS cpattern GLOB_PRINTED BY pr_ssrterm | [ "Qed" constr(c) ] -> [ mk_lterm c ] END GEXTEND Gram GLOBAL: ssrterm; ssrterm: [[ k = ssrtermkind; c = constr -> mk_term k c ]]; END (* }}} *) (** The "in" pseudo-tactical {{{ **********************************************) (* We can't make "in" into a general tactical because this would create a *) (* crippling conflict with the ltac let .. in construct. Hence, we add *) (* explicitly an "in" suffix to all the extended tactics for which it is *) (* relevant (including move, case, elim) and to the extended do tactical *) (* below, which yields a general-purpose "in" of the form do [...] in ... *) (* This tactical needs to come before the intro tactics because the latter *) (* must take precautions in order not to interfere with the discharged *) (* assumptions. This is especially difficult for discharged "let"s, which *) (* the default simpl and unfold tactics would erase blindly. *) (** Clear switch *) type ssrclear = ssrhyps let pr_clear_ne clr = str "{" ++ pr_hyps clr ++ str "}" let pr_clear sep clr = if clr = [] then mt () else sep () ++ pr_clear_ne clr let pr_ssrclear _ _ _ = pr_clear mt ARGUMENT EXTEND ssrclear_ne TYPED AS ssrhyps PRINTED BY pr_ssrclear | [ "{" ne_ssrhyp_list(clr) "}" ] -> [ check_hyps_uniq [] clr; clr ] END ARGUMENT EXTEND ssrclear TYPED AS ssrclear_ne PRINTED BY pr_ssrclear | [ ssrclear_ne(clr) ] -> [ clr ] | [ ] -> [ [] ] END let cleartac clr = check_hyps_uniq [] clr; clear (hyps_ids clr) (* type ssrwgen = ssrclear * ssrhyp * string *) let pr_wgen = function | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hyp id | (clr, Some((id,k),Some p)) -> spc() ++ pr_clear mt clr ++ str"(" ++ str k ++ pr_hyp id ++ str ":=" ++ pr_cpattern p ++ str ")" | (clr, None) -> spc () ++ pr_clear mt clr let pr_ssrwgen _ _ _ = pr_wgen (* no globwith for char *) ARGUMENT EXTEND ssrwgen TYPED AS ssrclear * ((ssrhyp * string) * cpattern option) option PRINTED BY pr_ssrwgen | [ ssrclear_ne(clr) ] -> [ clr, None ] | [ ssrhyp(id) ] -> [ [], Some((id, " "), None) ] | [ "@" ssrhyp(id) ] -> [ [], Some((id, "@"), None) ] | [ "(" ssrhyp(id) ":=" lcpattern(p) ")" ] -> [ [], Some ((id," "),Some p) ] | [ "(" ssrhyp(id) ")" ] -> [ [], Some ((id,"("), None) ] | [ "(@" ssrhyp(id) ":=" lcpattern(p) ")" ] -> [ [], Some ((id,"@"),Some p) ] | [ "(" "@" ssrhyp(id) ":=" lcpattern(p) ")" ] -> [ [], Some ((id,"@"),Some p) ] END type ssrclseq = InGoal | InHyps | InHypsGoal | InHypsSeqGoal | InSeqGoal | InHypsSeq | InAll | InAllHyps let pr_clseq = function | InGoal | InHyps -> mt () | InSeqGoal -> str "|- *" | InHypsSeqGoal -> str " |- *" | InHypsGoal -> str " *" | InAll -> str "*" | InHypsSeq -> str " |-" | InAllHyps -> str "* |-" let wit_ssrclseq, globwit_ssrclseq, rawwit_ssrclseq = add_genarg "ssrclseq" pr_clseq let pr_clausehyps = pr_list pr_spc pr_wgen let pr_ssrclausehyps _ _ _ = pr_clausehyps ARGUMENT EXTEND ssrclausehyps TYPED AS ssrwgen list PRINTED BY pr_ssrclausehyps | [ ssrwgen(hyp) "," ssrclausehyps(hyps) ] -> [ hyp :: hyps ] | [ ssrwgen(hyp) ssrclausehyps(hyps) ] -> [ hyp :: hyps ] | [ ssrwgen(hyp) ] -> [ [hyp] ] END (* type ssrclauses = ssrahyps * ssrclseq *) let pr_clauses (hyps, clseq) = if clseq = InGoal then mt () else str "in " ++ pr_clausehyps hyps ++ pr_clseq clseq let pr_ssrclauses _ _ _ = pr_clauses ARGUMENT EXTEND ssrclauses TYPED AS ssrwgen list * ssrclseq PRINTED BY pr_ssrclauses | [ "in" ssrclausehyps(hyps) "|-" "*" ] -> [ hyps, InHypsSeqGoal ] | [ "in" ssrclausehyps(hyps) "|-" ] -> [ hyps, InHypsSeq ] | [ "in" ssrclausehyps(hyps) "*" ] -> [ hyps, InHypsGoal ] | [ "in" ssrclausehyps(hyps) ] -> [ hyps, InHyps ] | [ "in" "|-" "*" ] -> [ [], InSeqGoal ] | [ "in" "*" ] -> [ [], InAll ] | [ "in" "*" "|-" ] -> [ [], InAllHyps ] | [ ] -> [ [], InGoal ] END let nohide = mkRel 0 let hidden_goal_tag = "the_hidden_goal" (* Reduction that preserves the Prod/Let spine of the "in" tactical. *) let inc_safe n = if n = 0 then n else n + 1 let rec safe_depth c = match kind_of_term c with | LetIn (Name x, _, _, c') when is_discharged_id x -> safe_depth c' + 1 | LetIn (_, _, _, c') | Prod (_, _, c') -> inc_safe (safe_depth c') | _ -> 0 let red_safe r e s c0 = let rec red_to e c n = match kind_of_term c with | Prod (x, t, c') when n > 0 -> let t' = r e s t in let e' = Environ.push_rel (x, None, t') e in mkProd (x, t', red_to e' c' (n - 1)) | LetIn (x, b, t, c') when n > 0 -> let t' = r e s t in let e' = Environ.push_rel (x, None, t') e in mkLetIn (x, r e s b, t', red_to e' c' (n - 1)) | _ -> r e s c in red_to e c0 (safe_depth c0) let check_wgen_uniq gens = let clears = List.flatten (List.map fst gens) in check_hyps_uniq [] clears; let ids = Util.list_map_filter (function (_,Some ((SsrHyp (_,id),_),_)) -> Some id | _ -> None) gens in let rec check ids = function | id :: _ when List.mem id ids -> errorstrm (str"Duplicate generalization " ++ pr_id id) | id :: hyps -> check (id :: ids) hyps | [] -> () in check [] ids let pf_clauseids gl gens clseq = let keep_clears = List.map (fun x, _ -> x, None) in if gens <> [] then (check_wgen_uniq gens; gens) else if clseq <> InAll && clseq <> InAllHyps then keep_clears gens else Util.error "assumptions should be named explicitly" let hidden_clseq = function InHyps | InHypsSeq | InAllHyps -> true | _ -> false let hidetacs clseq idhide cl0 = if not (hidden_clseq clseq) then [] else [posetac idhide cl0; convert_concl_no_check (mkVar idhide)] let discharge_hyp (id', (id, mode)) gl = let cl' = subst_var id (pf_concl gl) in match pf_get_hyp gl id, mode with | (_, None, t), _ | (_, Some _, t), "(" -> apply_type (mkProd (Name id', t, cl')) [mkVar id] gl | (_, Some v, t), _ -> convert_concl (mkLetIn (Name id', v, t, cl')) gl let endclausestac id_map clseq gl_id cl0 gl = let not_hyp' id = not (List.mem_assoc id id_map) in let orig_id id = try List.assoc id id_map with _ -> id in let dc, c = Term.decompose_prod_assum (pf_concl gl) in let hide_goal = hidden_clseq clseq in let c_hidden = hide_goal && c = mkVar gl_id in let rec fits forced = function | (id, _) :: ids, (Name id', _, _) :: dc' when id' = id -> fits true (ids, dc') | ids, dc' -> forced && ids = [] && (not hide_goal || dc' = [] && c_hidden) in let rec unmark c = match kind_of_term c with | Var id when hidden_clseq clseq && id = gl_id -> cl0 | Prod (Name id, t, c') when List.mem_assoc id id_map -> mkProd (Name (orig_id id), unmark t, unmark c') | LetIn (Name id, v, t, c') when List.mem_assoc id id_map -> mkLetIn (Name (orig_id id), unmark v, unmark t, unmark c') | _ -> map_constr unmark c in let utac hyp = convert_hyp_no_check (map_named_declaration unmark hyp) in let utacs = List.map utac (pf_hyps gl) in let ugtac gl' = convert_concl_no_check (unmark (pf_concl gl')) gl' in let ctacs = if hide_goal then [clear [gl_id]] else [] in let mktac itacs = tclTHENLIST (itacs @ utacs @ ugtac :: ctacs) in let itac (_, id) = introduction id in if fits false (id_map, List.rev dc) then mktac (List.map itac id_map) gl else let all_ids = ids_of_rel_context dc @ pf_ids_of_hyps gl in if List.for_all not_hyp' all_ids && not c_hidden then mktac [] gl else Util.error "tampering with discharged assumptions of \"in\" tactical" let is_id_constr c = match kind_of_term c with | Lambda(_,_,c) when isRel c -> 1 = destRel c | _ -> false let red_product_skip_id env sigma c = match kind_of_term c with | App(hd,args) when Array.length args = 1 && is_id_constr hd -> args.(0) | _ -> try Tacred.red_product env sigma c with _ -> c let abs_wgen keep_let ist gl f gen (args,c) = let sigma, env = project gl, pf_env gl in let evar_closed t p = if occur_existential t then Util.user_err_loc (loc_of_cpattern p,"ssreflect", pr_constr_pat t ++ str" contains holes and matches no subterm of the goal") in match gen with | _, Some ((SsrHyp(_,x), mode), None) when mode = "@" || (mode = " " && keep_let) -> let _, bo, ty = pf_get_hyp gl x in (if bo <> None then args else mkVar x :: args), mkProd_or_LetIn (Name (f x),bo,ty) (subst_var x c) | _, Some ((SsrHyp(_,x), _), None) -> mkVar x :: args, mkProd (Name (f x), pf_get_hyp_typ gl x, subst_var x c) | _, Some ((SsrHyp(_,x), "@"), Some p) -> let cp = interp_cpattern ist gl p None in let t, c = try fill_occ_pattern ~raise_NoMatch:true env sigma c cp None 1 with NoMatch -> redex_of_pattern cp, c in evar_closed t p; let ut = red_product_skip_id env sigma t in args, mkLetIn(Name (f x), ut, pf_type_of gl t, c) | _, Some ((SsrHyp(_,x), _), Some p) -> let cp = interp_cpattern ist gl p None in let t, c = try fill_occ_pattern ~raise_NoMatch:true env sigma c cp None 1 with NoMatch -> redex_of_pattern cp, c in evar_closed t p; t :: args, mkProd(Name (f x), pf_type_of gl t, c) | _ -> args, c let clr_of_wgen gen clrs = match gen with | clr, Some ((x, _), None) -> cleartac clr :: cleartac [x] :: clrs | clr, _ -> cleartac clr :: clrs let tclCLAUSES ist tac (gens, clseq) gl = if clseq = InGoal || clseq = InSeqGoal then tac gl else let clr_gens = pf_clauseids gl gens clseq in let clear = tclTHENLIST (List.rev(List.fold_right clr_of_wgen clr_gens [])) in let gl_id = mk_anon_id hidden_goal_tag gl in let cl0 = pf_concl gl in let dtac gl = let c = pf_concl gl in let args, c = List.fold_right (abs_wgen true ist gl mk_discharged_id) gens ([], c) in apply_type c args gl in let endtac = let id_map = Util.list_map_filter (function | _, Some ((SsrHyp(_,id),_),_) -> Some (mk_discharged_id id, id) | _, None -> None) gens in endclausestac id_map clseq gl_id cl0 in tclTHENLIST (hidetacs clseq gl_id cl0 @ [dtac; clear; tac; endtac]) gl (* }}} *) (** Simpl switch *) type ssrsimpl = Simpl | Cut | SimplCut | Nop let pr_simpl = function | Simpl -> str "/=" | Cut -> str "//" | SimplCut -> str "//=" | Nop -> mt () let pr_ssrsimpl _ _ _ = pr_simpl let wit_ssrsimplrep, globwit_ssrsimplrep, rawwit_ssrsimplrep = add_genarg "ssrsimplrep" pr_simpl ARGUMENT EXTEND ssrsimpl_ne TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl | [ "/=" ] -> [ Simpl ] | [ "//" ] -> [ Cut ] | [ "//=" ] -> [ SimplCut ] END ARGUMENT EXTEND ssrsimpl TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl | [ ssrsimpl_ne(sim) ] -> [ sim ] | [ ] -> [ Nop ] END (* We must avoid zeta-converting any "let"s created by the "in" tactical. *) let safe_simpltac gl = let cl' = red_safe Tacred.simpl (pf_env gl) (project gl) (pf_concl gl) in convert_concl_no_check cl' gl let simpltac = function | Simpl -> safe_simpltac | Cut -> tclTRY donetac | SimplCut -> tclTHEN safe_simpltac (tclTRY donetac) | Nop -> tclIDTAC (** Rewriting direction *) let pr_dir = function L2R -> str "->" | R2L -> str "<-" let pr_rwdir = function L2R -> mt() | R2L -> str "-" let rewritetac dir c = (* Due to the new optional arg ?tac, application shouldn't be too partial *) Equality.general_rewrite (dir = L2R) all_occurrences true false c let wit_ssrdir, globwit_ssrdir, rawwit_ssrdir = add_genarg "ssrdir" pr_dir let dir_org = function L2R -> 1 | R2L -> 2 (** Indexes *) (* Since SSR indexes are always positive numbers, we use the 0 value *) (* to encode an omitted index. We reuse the in or_var type, but we *) (* supply our own interpretation function, which checks for non *) (* positive values, and allows the use of constr numerals, so that *) (* e.g., "let n := eval compute in (1 + 3) in (do n!clear)" works. *) type ssrindex = int or_var let pr_index = function | ArgVar (_, id) -> pr_id id | ArgArg n when n > 0 -> pr_int n | _ -> mt () let pr_ssrindex _ _ _ = pr_index let noindex = ArgArg 0 let allocc = Some(false,[]) let check_index loc i = if i > 0 then i else loc_error loc "Index not positive" let mk_index loc = function ArgArg i -> ArgArg (check_index loc i) | iv -> iv let interp_index ist gl idx = Tacmach.project gl, match idx with | ArgArg _ -> idx | ArgVar (loc, id) -> let i = try match List.assoc id ist.lfun with | VInteger i -> i | VConstr ([],c) -> let rc = Detyping.detype false [] [] c in begin match Notation.uninterp_prim_token rc with | _, Numeral bigi -> int_of_string (Bigint.to_string bigi) | _ -> raise Not_found end | _ -> raise Not_found with _ -> loc_error loc "Index not a number" in ArgArg (check_index loc i) ARGUMENT EXTEND ssrindex TYPED AS int_or_var PRINTED BY pr_ssrindex INTERPRETED BY interp_index | [ int_or_var(i) ] -> [ mk_index loc i ] END (** Occurrence switch *) (* The standard syntax of complemented occurrence lists involves a single *) (* initial "-", e.g., {-1 3 5}. An initial *) (* "+" may be used to indicate positive occurrences (the default). The *) (* "+" is optional, except if the list of occurrences starts with a *) (* variable or is empty (to avoid confusion with a clear switch). The *) (* empty positive switch "{+}" selects no occurrences, while the empty *) (* negative switch "{-}" selects all occurrences explicitly; this is the *) (* default, but "{-}" prevents the implicit clear, and can be used to *) (* force dependent elimination -- see ndefectelimtac below. *) type ssrocc = occ let pr_occ = function | Some (true, occ) -> str "{-" ++ pr_list pr_spc int occ ++ str "}" | Some (false, occ) -> str "{+" ++ pr_list pr_spc int occ ++ str "}" | None -> str "{}" let pr_ssrocc _ _ _ = pr_occ ARGUMENT EXTEND ssrocc TYPED AS (bool * int list) option PRINTED BY pr_ssrocc | [ natural(n) natural_list(occ) ] -> [ Some (false, List.map (check_index loc) (n::occ)) ] | [ "-" natural_list(occ) ] -> [ Some (true, occ) ] | [ "+" natural_list(occ) ] -> [ Some (false, occ) ] END let pf_mkprod gl c ?(name=constr_name c) cl = let t = pf_type_of gl c in if name <> Anonymous || noccurn 1 cl then mkProd (name, t, cl) else mkProd (Name (pf_type_id gl t), t, cl) let pf_abs_prod name gl c cl = pf_mkprod gl c ~name (subst_term c cl) (** Discharge occ switch (combined occurrence / clear switch *) type ssrdocc = ssrclear option * ssrocc option let mkocc occ = None, occ let noclr = mkocc None let mkclr clr = Some clr, None let nodocc = mkclr [] let pr_docc = function | None, occ -> pr_occ occ | Some clr, _ -> pr_clear mt clr let pr_ssrdocc _ _ _ = pr_docc ARGUMENT EXTEND ssrdocc TYPED AS ssrclear option * ssrocc PRINTED BY pr_ssrdocc | [ "{" ne_ssrhyp_list(clr) "}" ] -> [ mkclr clr ] | [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ] END (** View hint database and View application. {{{ ******************************) (* There are three databases of lemmas used to mediate the application *) (* of reflection lemmas: one for forward chaining, one for backward *) (* chaining, and one for secondary backward chaining. *) (* View hints *) let rec isCxHoles = function (CHole _, None) :: ch -> isCxHoles ch | _ -> false let pr_raw_ssrhintref prc _ _ = function | CAppExpl (_, (None, r), args) when isCHoles args -> prc (CRef r) ++ str "|" ++ int (List.length args) | CApp (_, (_, CRef _), _) as c -> prc c | CApp (_, (_, c), args) when isCxHoles args -> prc c ++ str "|" ++ int (List.length args) | c -> prc c let pr_rawhintref = function | GApp (_, f, args) when isRHoles args -> pr_glob_constr f ++ str "|" ++ int (List.length args) | c -> pr_glob_constr c let pr_glob_ssrhintref _ _ _ (c, _) = pr_rawhintref c let pr_ssrhintref prc _ _ = prc let mkhintref loc c n = match c with | CRef r -> CAppExpl (loc, (None, r), mkCHoles loc n) | _ -> mkAppC (c, mkCHoles loc n) ARGUMENT EXTEND ssrhintref PRINTED BY pr_ssrhintref RAW_TYPED AS constr RAW_PRINTED BY pr_raw_ssrhintref GLOB_TYPED AS constr GLOB_PRINTED BY pr_glob_ssrhintref | [ constr(c) ] -> [ c ] | [ constr(c) "|" natural(n) ] -> [ mkhintref loc c n ] END (* View purpose *) let pr_viewpos = function | 0 -> str " for move/" | 1 -> str " for apply/" | 2 -> str " for apply//" | _ -> mt () let pr_ssrviewpos _ _ _ = pr_viewpos ARGUMENT EXTEND ssrviewpos TYPED AS int PRINTED BY pr_ssrviewpos | [ "for" "move" "/" ] -> [ 0 ] | [ "for" "apply" "/" ] -> [ 1 ] | [ "for" "apply" "/" "/" ] -> [ 2 ] | [ "for" "apply" "//" ] -> [ 2 ] | [ ] -> [ 3 ] END let pr_ssrviewposspc _ _ _ i = pr_viewpos i ++ spc () ARGUMENT EXTEND ssrviewposspc TYPED AS ssrviewpos PRINTED BY pr_ssrviewposspc | [ ssrviewpos(i) ] -> [ i ] END (* The table and its display command *) let viewtab : glob_constr list array = Array.make 3 [] let _ = let init () = Array.fill viewtab 0 3 [] in let freeze () = Array.copy viewtab in let unfreeze vt = Array.blit vt 0 viewtab 0 3 in Summary.declare_summary "ssrview" { Summary.freeze_function = freeze; Summary.unfreeze_function = unfreeze; Summary.init_function = init } let mapviewpos f n k = if n < 3 then f n else for i = 0 to k - 1 do f i done let print_view_hints i = let pp_viewname = str "Hint View" ++ pr_viewpos i ++ str " " in let pp_hints = pr_list spc pr_rawhintref viewtab.(i) in ppnl (pp_viewname ++ hov 0 pp_hints ++ Pp.cut ()) VERNAC COMMAND EXTEND PrintView | [ "Print" "Hint" "View" ssrviewpos(i) ] -> [ mapviewpos print_view_hints i 3 ] END (* Populating the table *) let cache_viewhint (_, (i, lvh)) = let mem_raw h = List.exists (Topconstr.eq_glob_constr h) in let add_hint h hdb = if mem_raw h hdb then hdb else h :: hdb in viewtab.(i) <- List.fold_right add_hint lvh viewtab.(i) let subst_viewhint ( subst, (i, lvh as ilvh)) = let lvh' = list_smartmap (Detyping.subst_glob_constr subst) lvh in if lvh' == lvh then ilvh else i, lvh' let classify_viewhint x = Libobject.Substitute x let in_viewhint = Libobject.declare_object {(Libobject.default_object "VIEW_HINTS") with Libobject.open_function = (fun i o -> if i = 1 then cache_viewhint o); Libobject.cache_function = cache_viewhint; Libobject.subst_function = subst_viewhint; Libobject.classify_function = classify_viewhint } let glob_view_hints lvh = List.map (Constrintern.intern_constr Evd.empty (Global.env ())) lvh let add_view_hints lvh i = Lib.add_anonymous_leaf (in_viewhint (i, lvh)) VERNAC COMMAND EXTEND HintView | [ "Hint" "View" ssrviewposspc(n) ne_ssrhintref_list(lvh) ] -> [ mapviewpos (add_view_hints (glob_view_hints lvh)) n 2 ] END (** Views *) (* Views for the "move" and "case" commands are actually open *) (* terms, but this is handled by interp_view, which is called *) (* by interp_casearg. We use lists, to support the *) (* "double-view" feature of the apply command. *) (* type ssrview = ssrterm list *) let pr_view = pr_list mt (fun c -> str "/" ++ pr_term c) let pr_ssrview _ _ _ = pr_view ARGUMENT EXTEND ssrview TYPED AS ssrterm list PRINTED BY pr_ssrview | [ "/" constr(c) ] -> [ [mk_term ' ' c] ] | [ "/" constr(c) ssrview(w) ] -> [ (mk_term ' ' c) :: w ] END (* There are two ways of "applying" a view to term: *) (* 1- using a view hint if the view is an instance of some *) (* (reflection) inductive predicate. *) (* 2- applying the view if it coerces to a function, adding *) (* implicit arguments. *) (* They require guessing the view hints and the number of *) (* implicits, respectively, which we do by brute force. *) let view_error s gv = errorstrm (str ("Cannot " ^ s ^ " view ") ++ pr_term gv) let interp_view ist si env sigma gv rid = match intern_term ist sigma env gv with | GApp (loc, GHole _, rargs) -> let rv = GApp (loc, rid, rargs) in snd (interp_open_constr ist (re_sig si sigma) (rv, None)) | rv -> let interp rc rargs = interp_open_constr ist (re_sig si sigma) (mkRApp rc rargs, None) in let rec simple_view rargs n = if n < 0 then view_error "use" gv else try interp rv rargs with _ -> simple_view (mkRHole :: rargs) (n - 1) in let view_nbimps = interp_view_nbimps ist (re_sig si sigma) rv in let view_args = [mkRApp rv (mkRHoles view_nbimps); rid] in let rec view_with = function | [] -> simple_view [rid] (interp_nbargs ist (re_sig si sigma) rv) | hint :: hints -> try interp hint view_args with _ -> view_with hints in snd (view_with (if view_nbimps < 0 then [] else viewtab.(0))) let top_id = mk_internal_id "top assumption" let with_view ist si env gl0 c name cl prune = let c2r ist x = { ist with lfun = (top_id, VConstr ([], x)) :: ist.lfun } in let rec loop (sigma, c') = function | f :: view -> let rid, ist = match kind_of_term c' with | Var id -> mkRVar id, ist | _ -> mkRltacVar top_id, c2r ist c' in loop (interp_view ist si env sigma f rid) view | [] -> let n, c' = pf_abs_evars gl0 (sigma, c') in let c' = if not prune then c' else pf_abs_cterm gl0 n c' in pf_abs_prod name gl0 c' (prod_applist cl [c]), c' in loop let pf_with_view ist gl (prune, view) cl c = let env, sigma, si = pf_env gl, project gl, sig_it gl in with_view ist si env gl c (constr_name c) cl prune (sigma, c) view (* }}} *) (** Extended intro patterns {{{ ***********************************************) type ssrtermrep = char * glob_constr_and_expr type ssripat = | IpatSimpl of ssrclear * ssrsimpl | IpatId of identifier | IpatWild | IpatCase of ssripats list | IpatRw of ssrocc * ssrdir | IpatAll | IpatAnon | IpatView of ssrtermrep list | IpatNoop and ssripats = ssripat list let remove_loc = snd let rec ipat_of_intro_pattern = function | IntroIdentifier id -> IpatId id | IntroWildcard -> IpatWild | IntroOrAndPattern iorpat -> IpatCase (List.map (List.map ipat_of_intro_pattern) (List.map (List.map remove_loc) iorpat)) | IntroAnonymous -> IpatAnon | IntroRewrite b -> IpatRw (allocc, if b then L2R else R2L) | IntroFresh id -> IpatAnon | IntroForthcoming _ -> (* Unable to determine which kind of ipat interp_introid could return [HH] *) assert false let rec pr_ipat = function | IpatId id -> pr_id id | IpatSimpl (clr, sim) -> pr_clear mt clr ++ pr_simpl sim | IpatCase iorpat -> hov 1 (str "[" ++ pr_iorpat iorpat ++ str "]") | IpatRw (occ, dir) -> pr_occ occ ++ pr_dir dir | IpatAll -> str "*" | IpatWild -> str "_" | IpatAnon -> str "?" | IpatView v -> pr_view v | IpatNoop -> str "-" and pr_iorpat iorpat = pr_list pr_bar pr_ipats iorpat and pr_ipats ipats = pr_list spc pr_ipat ipats let wit_ssripatrep, globwit_ssripatrep, rawwit_ssripatrep = add_genarg "ssripatrep" pr_ipat let pr_ssripat _ _ _ = pr_ipat let pr_ssripats _ _ _ = pr_ipats let pr_ssriorpat _ _ _ = pr_iorpat let intern_ipat ist ipat = let rec check_pat = function | IpatSimpl (clr, _) -> ignore (List.map (intern_hyp ist) clr) | IpatCase iorpat -> List.iter (List.iter check_pat) iorpat | _ -> () in check_pat ipat; ipat let intern_ipats ist = List.map (intern_ipat ist) let interp_introid ist gl id = try IntroIdentifier (hyp_id (snd (interp_hyp ist gl (SsrHyp (dummy_loc, id))))) with _ -> snd(snd (interp_intro_pattern ist gl (dummy_loc,IntroIdentifier id))) let rec add_intro_pattern_hyps (loc, ipat) hyps = match ipat with | IntroIdentifier id -> if not_section_id id then SsrHyp (loc, id) :: hyps else hyp_err loc "Can't delete section hypothesis " id | IntroWildcard -> hyps | IntroOrAndPattern iorpat -> List.fold_right (List.fold_right add_intro_pattern_hyps) iorpat hyps | IntroAnonymous -> [] | IntroFresh _ -> [] | IntroRewrite _ -> hyps | IntroForthcoming _ -> (* As in ipat_of_intro_pattern, was unable to determine which kind of ipat interp_introid could return [HH] *) assert false let rec interp_ipat ist gl = let ltacvar id = List.mem_assoc id ist.lfun in let rec interp = function | IpatId id when ltacvar id -> ipat_of_intro_pattern (interp_introid ist gl id) | IpatSimpl (clr, sim) -> let add_hyps (SsrHyp (loc, id) as hyp) hyps = if not (ltacvar id) then hyp :: hyps else add_intro_pattern_hyps (loc, (interp_introid ist gl id)) hyps in let clr' = List.fold_right add_hyps clr [] in check_hyps_uniq [] clr'; IpatSimpl (clr', sim) | IpatCase iorpat -> IpatCase (List.map (List.map interp) iorpat) | ipat -> ipat in interp let interp_ipats ist gl l = project gl, List.map (interp_ipat ist gl) l let pushIpatRw = function | pats :: orpat -> (IpatRw (allocc, L2R) :: pats) :: orpat | [] -> [] let pushIpatNoop = function | pats :: orpat -> (IpatNoop :: pats) :: orpat | [] -> [] ARGUMENT EXTEND ssripat TYPED AS ssripatrep list PRINTED BY pr_ssripats INTERPRETED BY interp_ipats GLOBALIZED BY intern_ipats | [ "_" ] -> [ [IpatWild] ] | [ "*" ] -> [ [IpatAll] ] | [ ident(id) ] -> [ [IpatId id] ] | [ "?" ] -> [ [IpatAnon] ] | [ ssrsimpl_ne(sim) ] -> [ [IpatSimpl ([], sim)] ] | [ ssrdocc(occ) "->" ] -> [ match occ with | None, occ -> [IpatRw (occ, L2R)] | Some clr, _ -> [IpatSimpl (clr, Nop); IpatRw (allocc, L2R)]] | [ ssrdocc(occ) "<-" ] -> [ match occ with | None, occ -> [IpatRw (occ, R2L)] | Some clr, _ -> [IpatSimpl (clr, Nop); IpatRw (allocc, R2L)]] | [ ssrdocc(occ) ] -> [ match occ with | Some cl, _ -> check_hyps_uniq [] cl; [IpatSimpl (cl, Nop)] | _ -> loc_error loc "Only identifiers are allowed here"] | [ "->" ] -> [ [IpatRw (allocc, L2R)] ] | [ "<-" ] -> [ [IpatRw (allocc, R2L)] ] | [ "-" ] -> [ [IpatNoop] ] | [ "-/" "=" ] -> [ [IpatNoop;IpatSimpl([],Simpl)] ] | [ "-/=" ] -> [ [IpatNoop;IpatSimpl([],Simpl)] ] | [ "-/" "/" ] -> [ [IpatNoop;IpatSimpl([],Cut)] ] | [ "-//" ] -> [ [IpatNoop;IpatSimpl([],Cut)] ] | [ "-/" "/=" ] -> [ [IpatNoop;IpatSimpl([],SimplCut)] ] | [ "-//" "=" ] -> [ [IpatNoop;IpatSimpl([],SimplCut)] ] | [ "-//=" ] -> [ [IpatNoop;IpatSimpl([],SimplCut)] ] | [ ssrview(v) ] -> [ [IpatView v] ] END ARGUMENT EXTEND ssripats TYPED AS ssripat PRINTED BY pr_ssripats | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ] | [ ] -> [ [] ] END ARGUMENT EXTEND ssriorpat TYPED AS ssripat list PRINTED BY pr_ssriorpat | [ ssripats(pats) "|" ssriorpat(orpat) ] -> [ pats :: orpat ] | [ ssripats(pats) "|-" ">" ssriorpat(orpat) ] -> [ pats :: pushIpatRw orpat ] | [ ssripats(pats) "|-" ssriorpat(orpat) ] -> [ pats :: pushIpatNoop orpat ] | [ ssripats(pats) "|->" ssriorpat(orpat) ] -> [ pats :: pushIpatRw orpat ] | [ ssripats(pats) "||" ssriorpat(orpat) ] -> [ pats :: [] :: orpat ] | [ ssripats(pats) "|||" ssriorpat(orpat) ] -> [ pats :: [] :: [] :: orpat ] | [ ssripats(pats) "||||" ssriorpat(orpat) ] -> [ [pats; []; []; []] @ orpat ] | [ ssripats(pats) ] -> [ [pats] ] END ARGUMENT EXTEND ssrcpat TYPED AS ssripatrep PRINTED BY pr_ssripat | [ "[" ssriorpat(iorpat) "]" ] -> [ IpatCase iorpat ] END GEXTEND Gram GLOBAL: ssripat; ssripat: [[ pat = ssrcpat -> [pat] ]]; END ARGUMENT EXTEND ssripats_ne TYPED AS ssripat PRINTED BY pr_ssripats | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ] END (* subsets of patterns *) let check_ssrhpats loc w_binders ipats = let err_loc s = user_err_loc (loc, "ssreflect", s) in let clr, ipats = let rec aux clr = function | IpatSimpl (cl, Nop) :: tl -> aux (clr @ cl) tl | IpatSimpl (cl, sim) :: tl -> clr @ cl, IpatSimpl ([], sim) :: tl | tl -> clr, tl in aux [] ipats in let simpl, ipats = match List.rev ipats with | IpatSimpl ([],_) as s :: tl -> [s], List.rev tl | _ -> [], ipats in if simpl <> [] && not w_binders then err_loc (str "No s-item allowed here: " ++ pr_ipats simpl); let ipat, binders = let rec loop ipat = function | [] -> ipat, [] | ( IpatId _| IpatAnon| IpatCase _| IpatRw _ as i) :: tl -> if w_binders then if simpl <> [] && tl <> [] then err_loc(str"binders XOR s-item allowed here: "++pr_ipats(tl@simpl)) else if not (List.for_all (function IpatId _ -> true | _ -> false) tl) then err_loc (str "Only binders allowed here: " ++ pr_ipats tl) else ipat @ [i], tl else if tl = [] then ipat @ [i], [] else err_loc (str "No binder or s-item allowed here: " ++ pr_ipats tl) | hd :: tl -> loop (ipat @ [hd]) tl in loop [] ipats in ((clr, ipat), binders), simpl let single loc = function [x] -> x | _ -> loc_error loc "Only one intro pattern is allowed" let pr_hpats (((clr, ipat), binders), simpl) = pr_clear mt clr ++ pr_ipats ipat ++ pr_ipats binders ++ pr_ipats simpl let pr_ssrhpats _ _ _ = pr_hpats ARGUMENT EXTEND ssrhpats TYPED AS ((ssrclear * ssripat) * ssripat) * ssripat PRINTED BY pr_ssrhpats | [ ssripats(i) ] -> [ check_ssrhpats loc true i ] END ARGUMENT EXTEND ssrhpats_nobs TYPED AS ((ssrclear * ssripat) * ssripat) * ssripat PRINTED BY pr_ssrhpats | [ ssripats(i) ] -> [ check_ssrhpats loc false i ] END ARGUMENT EXTEND ssrrpat TYPED AS ssripatrep PRINTED BY pr_ssripat | [ "->" ] -> [ IpatRw (allocc, L2R) ] | [ "<-" ] -> [ IpatRw (allocc, R2L) ] END type ssrintros = ssripats * ltacctx let pr_intros sep (intrs, _) = if intrs = [] then mt() else sep () ++ str "=> " ++ pr_ipats intrs let pr_ssrintros _ _ _ = pr_intros mt ARGUMENT EXTEND ssrintros_ne TYPED AS ssripat * ltacctx PRINTED BY pr_ssrintros | [ "=>" ssripats_ne(pats) ] -> [ pats, rawltacctx ] END ARGUMENT EXTEND ssrintros TYPED AS ssrintros_ne PRINTED BY pr_ssrintros | [ ssrintros_ne(intrs) ] -> [ intrs ] | [ ] -> [ [], rawltacctx ] END let injecteq_id = mk_internal_id "injection equation" let pf_nb_prod gl = nb_prod (pf_concl gl) let rev_id = mk_internal_id "rev concl" let revtoptac n0 gl = let n = pf_nb_prod gl - n0 in let dc, cl = decompose_prod_n n (pf_concl gl) in let dc' = dc @ [Name rev_id, compose_prod (List.rev dc) cl] in let f = compose_lam dc' (mkEtaApp (mkRel (n + 1)) (-n) 1) in refine (mkApp (f, [|Evarutil.mk_new_meta ()|])) gl let equality_inj l b id c gl = let msg = ref "" in try Equality.inj l b c gl with Compat.Loc.Exc_located(_,UserError (_,s)) | UserError (_,s) when msg := Pp.string_of_ppcmds s; !msg = "Not a projectable equality but a discriminable one." || !msg = "Nothing to inject." -> msg_warning (str !msg); discharge_hyp (id, (id, "")) gl let injectidl2rtac id c gl = tclTHEN (equality_inj [] true id c) (revtoptac (pf_nb_prod gl)) gl let injectl2rtac c = match kind_of_term c with | Var id -> injectidl2rtac id (mkVar id, NoBindings) | _ -> let id = injecteq_id in tclTHENLIST [havetac id c; injectidl2rtac id (mkVar id, NoBindings); clear [id]] let is_injection_case c gl = let mind, _ = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in mkInd mind = build_coq_eq () let perform_injection c gl = let mind, t = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in let dc, eqt = decompose_prod t in if dc = [] then injectl2rtac c gl else if not (closed0 eqt) then error "can't decompose a quantified equality" else let cl = pf_concl gl in let n = List.length dc in let c_eq = mkEtaApp c n 2 in let cl1 = mkLambda (Anonymous, mkArrow eqt cl, mkApp (mkRel 1, [|c_eq|])) in let id = injecteq_id in let id_with_ebind = (mkVar id, NoBindings) in let injtac = tclTHEN (introid id) (injectidl2rtac id id_with_ebind) in tclTHENLAST (apply (compose_lam dc cl1)) injtac gl let simplest_newcase_ref = ref (fun t gl -> assert false) let simplest_newcase x gl = !simplest_newcase_ref x gl let ssrscasetac c gl = if is_injection_case c gl then perform_injection c gl else simplest_newcase c gl let intro_all gl = let dc, _ = Term.decompose_prod_assum (pf_concl gl) in tclTHENLIST (List.map anontac (List.rev dc)) gl let rec intro_anon gl = try anontac (List.hd (fst (Term.decompose_prod_n_assum 1 (pf_concl gl)))) gl with err0 -> try tclTHEN red_in_concl intro_anon gl with _ -> raise err0 (* with _ -> error "No product even after reduction" *) let with_top tac = tclTHENLIST [introid top_id; tac (mkVar top_id); clear [top_id]] let rec mapLR f = function [] -> [] | x :: s -> let y = f x in y :: mapLR f s let wild_ids = ref [] let new_wild_id () = let i = 1 + List.length !wild_ids in let id = mk_wildcard_id i in wild_ids := id :: !wild_ids; id let clear_wilds wilds gl = clear (List.filter (fun id -> List.mem id wilds) (pf_ids_of_hyps gl)) gl let clear_with_wilds wilds clr0 gl = let extend_clr clr (id, _, _ as nd) = if List.mem id clr || not (List.mem id wilds) then clr else let vars = global_vars_set_of_decl (pf_env gl) nd in let occurs id' = Idset.mem id' vars in if List.exists occurs clr then id :: clr else clr in clear (Sign.fold_named_context_reverse extend_clr ~init:clr0 (pf_hyps gl)) gl let tclTHENS_nonstrict tac tacl taclname gl = let tacres = tac gl in let n_gls = List.length (sig_it tacres) in let n_tac = List.length tacl in if n_gls = n_tac then tclTHENS (fun _ -> tacres) tacl gl else if n_gls = 0 then tacres else let pr_only n1 n2 = if n1 < n2 then str "only " else mt () in let pr_nb n1 n2 name = pr_only n1 n2 ++ int n1 ++ str (" " ^ plural n1 name) in errorstrm (pr_nb n_tac n_gls taclname ++ spc () ++ str "for " ++ pr_nb n_gls n_tac "subgoal") (* Forward reference to extended rewrite *) let ipat_rewritetac = ref (fun _ -> rewritetac) let rec is_name_in_ipats name = function | IpatSimpl(clr,_) :: tl -> List.exists (function SsrHyp(_,id) -> id = name) clr || is_name_in_ipats name tl | IpatId id :: tl -> id = name || is_name_in_ipats name tl | IpatCase l :: tl -> is_name_in_ipats name (List.flatten l @ tl) | _ :: tl -> is_name_in_ipats name tl | [] -> false let move_top_with_view = ref (fun _ -> assert false) (* introstac: for "move" and "clear", tclEQINTROS: for "case" and "elim" *) (* This block hides the spaghetti-code needed to implement the only two *) (* tactics that should be used to process intro patters. *) (* The difficulty is that we don't want to always rename, but we can *) (* compute needeed renamings only at runtime, so we theread a tree like *) (* imperativestructure so that outer renamigs are inherited by inner *) (* ipts and that the cler performed at the end of ipatstac clears hyps *) (* eventually renamed at runtime. *) (* TODO: hide wild_ids in this block too *) let introstac, tclEQINTROS = let rec map_acc_k f k = function | [] -> (* tricky: we save wilds now, we get to_cler (aka k) later *) let clear_ww = clear_with_wilds !wild_ids in [fun gl -> clear_ww (hyps_ids (List.flatten (List.map (!) k))) gl] | x :: xs -> let k, x = f k xs x in x :: map_acc_k f k xs in let rename force to_clr rest clr gl = let hyps = pf_hyps gl in pp(lazy(str"rename " ++ pr_clear spc clr)); let () = if not force then List.iter (check_hyp_exists hyps) clr in if List.exists (fun x -> force || is_name_in_ipats (hyp_id x) rest) clr then let ren_clr, ren = List.split (List.map (fun x -> let x = hyp_id x in let x' = mk_anon_id (string_of_id x) gl in SsrHyp (dummy_loc, x'), (x, x')) clr) in let () = to_clr := ren_clr in rename_hyp ren gl else let () = to_clr := clr in tclIDTAC gl in let rec ipattac ?ist k rest = function | IpatWild -> k, introid (new_wild_id ()) | IpatCase iorpat -> k, tclIORPAT ?ist k (with_top ssrscasetac) iorpat | IpatRw (occ, dir) -> k, with_top (!ipat_rewritetac occ dir) | IpatId id -> k, introid id | IpatSimpl (clr, sim) -> let to_clr = ref [] in to_clr :: k, tclTHEN (rename false to_clr rest clr) (simpltac sim) | IpatAll -> k, intro_all | IpatAnon -> k, intro_anon | IpatNoop -> k, tclIDTAC | IpatView v -> match ist with | None -> anomaly "ipattac with no ist but view" | Some ist -> match rest with | (IpatCase _ | IpatRw _)::_ -> let to_clr = ref [] in let top_id = ref top_id in to_clr :: k, tclTHEN (!move_top_with_view false top_id (false,v) ist) (fun gl -> rename true to_clr rest [SsrHyp (dummy_loc, !top_id)]gl) | _ -> k, !move_top_with_view true (ref top_id) (true,v) ist and tclIORPAT ?ist k tac = function | [[]] -> tac | orp -> tclTHENS_nonstrict tac (mapLR (ipatstac ?ist k) orp) "intro pattern" and ipatstac ?ist k ipats = tclTHENLIST (map_acc_k (ipattac ?ist) k ipats) in let introstac ?ist ipats = wild_ids := []; let tac = ipatstac ?ist [] ipats in tclTHENLIST [tac; clear_wilds !wild_ids] in let tclEQINTROS ?ist tac eqtac ipats = wild_ids := []; let rec split_itacs to_clr tac' = function | (IpatSimpl _ as spat) :: ipats' -> let to_clr, tac = ipattac ?ist to_clr ipats' spat in split_itacs to_clr (tclTHEN tac' tac) ipats' | IpatCase iorpat :: ipats' -> to_clr, tclIORPAT ?ist to_clr tac' iorpat, ipats' | ipats' -> to_clr, tac', ipats' in let to_clr, tac1, ipats' = split_itacs [] tac ipats in let tac2 = ipatstac ?ist to_clr ipats' in tclTHENLIST [tac1; eqtac; tac2; clear_wilds !wild_ids] in introstac, tclEQINTROS ;; let rec eqmoveipats eqpat = function | (IpatSimpl _ as ipat) :: ipats -> ipat :: eqmoveipats eqpat ipats | (IpatAll :: _ | []) as ipats -> IpatAnon :: eqpat :: ipats | ipat :: ipats -> ipat :: eqpat :: ipats (* General case *) let tclINTROS tac (ipats, ctx) = let ist = get_ltacctx ctx in tclEQINTROS ~ist (tac ist) tclIDTAC ipats (** The "=>" tactical *) let ssrintros_sep = let atom_sep = function | TacSplit (_,_, [NoBindings]) -> mt | TacLeft (_, NoBindings) -> mt | TacRight (_, NoBindings) -> mt (* | TacExtend (_, "ssrapply", []) -> mt *) | _ -> spc in function | TacId [] -> mt | TacArg (_, Tacexp _) -> mt | TacArg (_, Reference _) -> mt | TacAtom (_, atom) -> atom_sep atom | _ -> spc let pr_ssrintrosarg _ _ prt (tac, ipats) = prt tacltop tac ++ pr_intros (ssrintros_sep tac) ipats ARGUMENT EXTEND ssrintrosarg TYPED AS tactic * ssrintros PRINTED BY pr_ssrintrosarg | [ "Qed" ssrtacarg(arg) ssrintros_ne(ipats) ] -> [ arg, ipats ] END TACTIC EXTEND ssrtclintros | [ "Qed" ssrintrosarg(arg) ] -> [ let tac, intros = arg in tclINTROS (fun ist -> ssrevaltac ist tac) intros ] END set_pr_ssrtac "tclintros" 0 [ArgSsr "introsarg"] let tclintros_expr loc tac ipats = let args = [in_gen rawwit_ssrintrosarg (tac, ipats)] in ssrtac_expr loc "tclintros" args GEXTEND Gram GLOBAL: tactic_expr; tactic_expr: LEVEL "1" [ RIGHTA [ tac = tactic_expr; intros = ssrintros_ne -> tclintros_expr loc tac intros ] ]; END (* }}} *) (** Multipliers {{{ ***********************************************************) (* modality *) type ssrmmod = May | Must | Once let pr_mmod = function May -> str "?" | Must -> str "!" | Once -> mt () let wit_ssrmmod, globwit_ssrmmod, rawwit_ssrmmod = add_genarg "ssrmmod" pr_mmod let ssrmmod = Gram.Entry.create "ssrmmod" GEXTEND Gram GLOBAL: ssrmmod; ssrmmod: [[ "!" -> Must | LEFTQMARK -> May | "?" -> May]]; END (* tactical *) let tclID tac = tac let tclDOTRY n tac = if n <= 0 then tclIDTAC else let rec loop i gl = if i = n then tclTRY tac gl else tclTRY (tclTHEN tac (loop (i + 1))) gl in loop 1 let tclDO n tac = let prefix i = str"At iteration " ++ int i ++ str": " in let tac_err_at i gl = try tac gl with | UserError (l, s) -> raise (UserError (l, prefix i ++ s)) | Compat.Loc.Exc_located(loc, UserError (l, s)) -> raise (Compat.Loc.Exc_located(loc, UserError (l, prefix i ++ s))) in let rec loop i gl = if i = n then tac_err_at i gl else (tclTHEN (tac_err_at i) (loop (i + 1))) gl in loop 1 let tclMULT = function | 0, May -> tclREPEAT | 1, May -> tclTRY | n, May -> tclDOTRY n | 0, Must -> tclAT_LEAST_ONCE | n, Must when n > 1 -> tclDO n | _ -> tclID (** The "do" tactical. ********************************************************) (* type ssrdoarg = ((ssrindex * ssrmmod) * (ssrhint * ltacctx)) * ssrclauses *) let pr_ssrdoarg prc _ prt (((n, m), (tac, _)), clauses) = pr_index n ++ pr_mmod m ++ pr_hintarg prt tac ++ pr_clauses clauses ARGUMENT EXTEND ssrdoarg TYPED AS ((ssrindex * ssrmmod) * (ssrhintarg * ltacctx)) * ssrclauses PRINTED BY pr_ssrdoarg | [ "Qed" ] -> [ anomaly "Grammar placeholder match" ] END let ssrdotac (((n, m), (tac, ctx)), clauses) = let mul = get_index n, m in let ist = get_ltacctx ctx in tclCLAUSES ist (tclMULT mul (hinttac ist false tac)) clauses TACTIC EXTEND ssrtcldo | [ "Qed" "do" ssrdoarg(arg) ] -> [ ssrdotac arg ] END set_pr_ssrtac "tcldo" 3 [ArgSep "do "; ArgSsr "doarg"] let ssrdotac_expr loc n m tac clauses = let arg = ((n, m), (tac, rawltacctx)), clauses in ssrtac_expr loc "tcldo" [in_gen rawwit_ssrdoarg arg] GEXTEND Gram GLOBAL: tactic_expr; ssrdotac: [ [ tac = tactic_expr LEVEL "3" -> mk_hint tac | tacs = ssrortacarg -> tacs ] ]; tactic_expr: LEVEL "3" [ RIGHTA [ IDENT "do"; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses -> ssrdotac_expr loc noindex m tac clauses | IDENT "do"; tac = ssrortacarg; clauses = ssrclauses -> ssrdotac_expr loc noindex Once tac clauses | IDENT "do"; n = int_or_var; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses -> ssrdotac_expr loc (mk_index loc n) m tac clauses ] ]; END (* }}} *) (** The "first" and "last" tacticals. {{{ *************************************) (* type ssrseqarg = ssrindex * (ssrtacarg * ssrtac option) *) let pr_seqtacarg prt = function | (is_first, []), _ -> str (if is_first then "first" else "last") | tac, Some dtac -> hv 0 (pr_hintarg prt tac ++ spc() ++ str "|| " ++ prt tacltop dtac) | tac, _ -> pr_hintarg prt tac let pr_ssrseqarg _ _ prt = function | ArgArg 0, tac -> pr_seqtacarg prt tac | i, tac -> pr_index i ++ str " " ++ pr_seqtacarg prt tac (* We must parse the index separately to resolve the conflict with *) (* an unindexed tactic. *) ARGUMENT EXTEND ssrseqarg TYPED AS ssrindex * (ssrhintarg * tactic option) PRINTED BY pr_ssrseqarg | [ "Qed" ] -> [ anomaly "Grammar placeholder match" ] END let sq_brace_tacnames = ["first"; "solve"; "do"; "rewrite"; "have"; "suffices"; "wlog"] (* "by" is a keyword *) let accept_ssrseqvar strm = match Compat.get_tok (stream_nth 0 strm) with | Tok.IDENT id when not (List.mem id sq_brace_tacnames) -> accept_before_syms_or_ids ["["] ["first";"last"] strm | _ -> raise Stream.Failure let test_ssrseqvar = Gram.Entry.of_parser "test_ssrseqvar" accept_ssrseqvar let swaptacarg (loc, b) = (b, []), Some (TacAtom (loc, TacRevert [])) let check_seqtacarg dir arg = match snd arg, dir with | ((true, []), Some (TacAtom (loc, _))), L2R -> loc_error loc "expected \"last\"" | ((false, []), Some (TacAtom (loc, _))), R2L -> loc_error loc "expected \"first\"" | _, _ -> arg let ssrorelse = Gram.Entry.create "ssrorelse" GEXTEND Gram GLOBAL: ssrorelse ssrseqarg; ssrseqidx: [ [ test_ssrseqvar; id = Prim.ident -> ArgVar (loc, id) | n = Prim.natural -> ArgArg (check_index loc n) ] ]; ssrswap: [[ IDENT "first" -> loc, true | IDENT "last" -> loc, false ]]; ssrorelse: [[ "||"; tac = tactic_expr LEVEL "2" -> tac ]]; ssrseqarg: [ [ arg = ssrswap -> noindex, swaptacarg arg | i = ssrseqidx; tac = ssrortacarg; def = OPT ssrorelse -> i, (tac, def) | i = ssrseqidx; arg = ssrswap -> i, swaptacarg arg | tac = tactic_expr LEVEL "3" -> noindex, (mk_hint tac, None) ] ]; END let tclPERM perm tac gls = let subgls = tac gls in let sigma, subgll = Refiner.unpackage subgls in let subgll' = perm subgll in Refiner.repackage sigma subgll' (* let tclPERM perm tac gls = let mkpft n g r = {Proof_type.open_subgoals = n; Proof_type.goal = g; Proof_type.ref = r} in let mkleaf g = mkpft 0 g None in let mkprpft n g pr a = mkpft n g (Some (Proof_type.Prim pr, a)) in let mkrpft n g c = mkprpft n g (Proof_type.Refine c) in let mkipft n g = let mki pft (id, _, _ as d) = let g' = {g with evar_concl = mkNamedProd_or_LetIn d g.evar_concl} in mkprpft n g' (Proof_type.Intro id) [pft] in List.fold_left mki in let gl = Refiner.sig_it gls in let mkhyp subgl = let rec chop_section = function | (x, _, _ as d) :: e when not_section_id x -> d :: chop_section e | _ -> [] in let lhyps = Environ.named_context_of_val subgl.evar_hyps in mk_perm_id (), subgl, chop_section lhyps in let mkpfvar (hyp, subgl, lhyps) = let mkarg args (lhyp, body, _) = if body = None then mkVar lhyp :: args else args in mkrpft 0 subgl (applist (mkVar hyp, List.fold_left mkarg [] lhyps)) [] in let mkpfleaf (_, subgl, lhyps) = mkipft 1 gl (mkleaf subgl) lhyps in let mkmeta _ = Evarutil.mk_new_meta () in let mkhypdecl (hyp, subgl, lhyps) = hyp, None, it_mkNamedProd_or_LetIn subgl.evar_concl lhyps in let subgls, v as res0 = tac gls in let sigma, subgll = Refiner.unpackage subgls in let n = List.length subgll in if n = 0 then res0 else let hyps = List.map mkhyp subgll in let hyp_decls = List.map mkhypdecl (List.rev (perm hyps)) in let c = applist (mkmeta (), List.map mkmeta subgll) in let pft0 = mkipft 0 gl (v (List.map mkpfvar hyps)) hyp_decls in let pft1 = mkrpft n gl c (pft0 :: List.map mkpfleaf (perm hyps)) in let subgll', v' = Refiner.frontier pft1 in Refiner.repackage sigma subgll', v' *) let tclREV tac gl = tclPERM List.rev tac gl let rot_hyps dir i hyps = let n = List.length hyps in if i = 0 then List.rev hyps else if i > n then error "Not enough subgoals" else let rec rot i l_hyps = function | hyp :: hyps' when i > 0 -> rot (i - 1) (hyp :: l_hyps) hyps' | hyps' -> hyps' @ (List.rev l_hyps) in rot (match dir with L2R -> i | R2L -> n - i) [] hyps let tclSEQAT (atac1, ctx) dir (ivar, ((_, atacs2), atac3)) = let i = get_index ivar in let evtac = ssrevaltac (get_ltacctx ctx) in let tac1 = evtac atac1 in if atacs2 = [] && atac3 <> None then tclPERM (rot_hyps dir i) tac1 else let evotac = function Some atac -> evtac atac | _ -> tclIDTAC in let tac3 = evotac atac3 in let rec mk_pad n = if n > 0 then tac3 :: mk_pad (n - 1) else [] in match dir, mk_pad (i - 1), List.map evotac atacs2 with | L2R, [], [tac2] when atac3 = None -> tclTHENFIRST tac1 tac2 | L2R, [], [tac2] when atac3 = None -> tclTHENLAST tac1 tac2 | L2R, pad, tacs2 -> tclTHENSFIRSTn tac1 (Array.of_list (pad @ tacs2)) tac3 | R2L, pad, tacs2 -> tclTHENSLASTn tac1 tac3 (Array.of_list (tacs2 @ pad)) (* We can't actually parse the direction separately because this *) (* would introduce conflicts with the basic ltac syntax. *) let pr_ssrseqdir _ _ _ = function | L2R -> str ";" ++ spc () ++ str "first " | R2L -> str ";" ++ spc () ++ str "last " ARGUMENT EXTEND ssrseqdir TYPED AS ssrdir PRINTED BY pr_ssrseqdir | [ "Qed" ] -> [ anomaly "Grammar placeholder match" ] END TACTIC EXTEND ssrtclseq | [ "Qed" ssrtclarg(tac) ssrseqdir(dir) ssrseqarg(arg) ] -> [ tclSEQAT tac dir arg ] END set_pr_ssrtac "tclseq" 5 [ArgSsr "tclarg"; ArgSsr "seqdir"; ArgSsr "seqarg"] let tclseq_expr loc tac dir arg = let arg1 = in_gen rawwit_ssrtclarg (tac, rawltacctx) in let arg2 = in_gen rawwit_ssrseqdir dir in let arg3 = in_gen rawwit_ssrseqarg (check_seqtacarg dir arg) in ssrtac_expr loc "tclseq" [arg1; arg2; arg3] GEXTEND Gram GLOBAL: tactic_expr; ssr_first: [ [ tac = ssr_first; ipats = ssrintros_ne -> tclintros_expr loc tac ipats | "["; tacl = LIST0 tactic_expr SEP "|"; "]" -> TacFirst tacl ] ]; ssr_first_else: [ [ tac1 = ssr_first; tac2 = ssrorelse -> TacOrelse (tac1, tac2) | tac = ssr_first -> tac ]]; tactic_expr: LEVEL "4" [ LEFTA [ tac1 = tactic_expr; ";"; IDENT "first"; tac2 = ssr_first_else -> TacThen (tac1,[||], tac2,[||]) | tac = tactic_expr; ";"; IDENT "first"; arg = ssrseqarg -> tclseq_expr loc tac L2R arg | tac = tactic_expr; ";"; IDENT "last"; arg = ssrseqarg -> tclseq_expr loc tac R2L arg ] ]; END (* }}} *) (** 5. Bookkeeping tactics (clear, move, case, elim) *) (* post-interpretation of terms *) let all_ok _ _ = true let pf_abs_ssrterm ist gl t = let n, c = pf_abs_evars gl (interp_term ist gl t) in pf_abs_cterm gl n c let pf_interp_ty ist gl ty = let n_binders = ref 0 in let ty = match ty with | a, (t, None) -> let rec force_type = function | GProd (l, x, k, s, t) -> incr n_binders; GProd (l, x, k, s, force_type t) | GLetIn (l, x, v, t) -> incr n_binders; GLetIn (l, x, v, force_type t) | ty -> mkRCast ty mkRType in a, (force_type t, None) | _, (_, Some ty) -> let rec force_type = function | CProdN (l, abs, t) -> n_binders := !n_binders + List.length (List.flatten (List.map pi1 abs)); CProdN (l, abs, force_type t) | CLetIn (l, n, v, t) -> incr n_binders; CLetIn (l, n, v, force_type t) | ty -> mkCCast dummy_loc ty (mkCType dummy_loc) in mk_term ' ' (force_type ty) in let strip_cast (sigma, t) = let rec aux t = match kind_of_type t with | CastType (t, ty) when !n_binders = 0 && isSort ty -> t | ProdType(n,s,t) -> decr n_binders; mkProd (n, s, aux t) | LetInType(n,v,ty,t) -> decr n_binders; mkLetIn (n, v, ty, aux t) | _ -> anomaly "pf_interp_ty: ssr Type cast deleted by typecheck" in sigma, aux t in let ty = strip_cast (interp_term ist gl ty) in let n, c = pf_abs_evars gl ty in let lam_c = pf_abs_cterm gl n c in let ctx, c = decompose_lam_n n lam_c in n, compose_prod ctx c, lam_c ;; let whd_app f args = Reductionops.whd_betaiota Evd.empty (mkApp (f, args)) let pr_cargs a = str "[" ++ pr_list pr_spc pr_constr (Array.to_list a) ++ str "]" let pp_term gl t = let t = Reductionops.nf_evar (project gl) t in pr_constr t let pp_concat hd ?(sep=str", ") = function [] -> hd | x :: xs -> hd ++ List.fold_left (fun acc x -> acc ++ sep ++ x) x xs let fake_pmatcher_end () = mkProp, L2R, (Evd.empty, mkProp) (* TASSI: given (c : ty), generates (c ??? : ty[???/...]) with m evars *) exception NotEnoughProducts let prof_saturate_whd = mk_profiler "saturate.whd";; let saturate ?(beta=false) env sigma c ?(ty=Retyping.get_type_of env sigma c) m = let rec loop ty args sigma n = if n = 0 then let args = List.rev args in (if beta then Reductionops.whd_beta sigma else fun x -> x) (mkApp (c, Array.of_list (List.map snd args))), ty, args, sigma else match kind_of_type ty with | ProdType (_, src, tgt) -> let sigma, x = Evarutil.new_evar (create_evar_defs sigma) env src in loop (subst1 x tgt) ((m - n,x) :: args) sigma (n-1) | CastType (t, _) -> loop t args sigma n | LetInType (_, v, _, t) -> loop (subst1 v t) args sigma n | SortType _ -> assert false | AtomicType _ -> let ty = prof_saturate_whd.profile (Reductionops.whd_betadeltaiota env sigma) ty in match kind_of_type ty with | ProdType _ -> loop ty args sigma n | _ -> anomaly "saturate did not find enough products" in loop ty [] sigma m let pf_saturate ?beta gl c ?ty m = let env, sigma, si = pf_env gl, project gl, sig_it gl in let t, ty, args, sigma = saturate ?beta env sigma c ?ty m in t, ty, args, re_sig si sigma (** Rewrite redex switch *) (** Generalization (discharge) item *) (* An item is a switch + term pair. *) (* type ssrgen = ssrdocc * ssrterm *) let pr_gen (docc, dt) = pr_docc docc ++ pr_cpattern dt let pr_ssrgen _ _ _ = pr_gen ARGUMENT EXTEND ssrgen TYPED AS ssrdocc * cpattern PRINTED BY pr_ssrgen | [ ssrdocc(docc) cpattern(dt) ] -> [ docc, dt ] | [ cpattern(dt) ] -> [ nodocc, dt ] END let has_occ ((_, occ), _) = occ <> None let hyp_of_var v = SsrHyp (dummy_loc, destVar v) let interp_clr = function | Some clr, (k, c) when (k = ' ' || k = '@') && is_pf_var c -> hyp_of_var c :: clr | Some clr, _ -> clr | None, _ -> [] (* XXX the k of the redex should percolate out *) let pf_interp_gen_aux ist gl to_ind ((oclr, occ), t) = let pat = interp_cpattern ist gl t None in (* UGLY API *) let c = redex_of_pattern pat in let cl, env, sigma = pf_concl gl, pf_env gl, project gl in let c, cl = try fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1 with NoMatch -> c, cl in let clr = interp_clr (oclr, (tag_of_cpattern t, c)) in if not(occur_existential c) then if tag_of_cpattern t = '@' then if not (isVar c) then errorstrm (str "@ can be used with variables only") else match pf_get_hyp gl (destVar c) with | _, None, _ -> errorstrm (str "@ can be used with let-ins only") | name, Some bo, ty -> true, pat, mkLetIn (Name name, bo, ty, cl), c, clr else false, pat, pf_mkprod gl c cl, c, clr else if to_ind && occ = None then let nv, p = pf_abs_evars gl (fst pat, c) in if nv = 0 then anomaly "occur_existential but no evars" else false, pat, mkProd (constr_name c, pf_type_of gl p, pf_concl gl), p, clr else loc_error (loc_of_cpattern t) "generalized term didn't match" let genclrtac cl cs clr = tclTHEN (apply_type cl cs) (cleartac clr) let gentac ist gen gl = let conv, _, cl, c, clr = pf_interp_gen_aux ist gl false gen in if conv then tclTHEN (convert_concl cl) (cleartac clr) gl else genclrtac cl [c] clr gl let pf_interp_gen ist gl to_ind gen = let _, _, a, b ,c = pf_interp_gen_aux ist gl to_ind gen in a, b ,c (** Generalization (discharge) sequence *) (* A discharge sequence is represented as a list of up to two *) (* lists of d-items, plus an ident list set (the possibly empty *) (* final clear switch). The main list is empty iff the command *) (* is defective, and has length two if there is a sequence of *) (* dependent terms (and in that case it is the first of the two *) (* lists). Thus, the first of the two lists is never empty. *) (* type ssrgens = ssrgen list *) (* type ssrdgens = ssrgens list * ssrclear *) let gens_sep = function [], [] -> mt | _ -> spc let pr_dgens pr_gen (gensl, clr) = let prgens s gens = str s ++ pr_list spc pr_gen gens in let prdeps deps = prgens ": " deps ++ spc () ++ str "/" in match gensl with | [deps; []] -> prdeps deps ++ pr_clear pr_spc clr | [deps; gens] -> prdeps deps ++ prgens " " gens ++ pr_clear spc clr | [gens] -> prgens ": " gens ++ pr_clear spc clr | _ -> pr_clear pr_spc clr let pr_ssrdgens _ _ _ = pr_dgens pr_gen let cons_gen gen = function | gens :: gensl, clr -> (gen :: gens) :: gensl, clr | _ -> anomaly "missing gen list" let cons_dep (gensl, clr) = if List.length gensl = 1 then ([] :: gensl, clr) else error "multiple dependents switches '/'" ARGUMENT EXTEND ssrdgens_tl TYPED AS ssrgen list list * ssrclear PRINTED BY pr_ssrdgens | [ "{" ne_ssrhyp_list(clr) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> [ cons_gen (mkclr clr, dt) dgens ] | [ "{" ne_ssrhyp_list(clr) "}" ] -> [ [[]], clr ] | [ "{" ssrocc(occ) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> [ cons_gen (mkocc occ, dt) dgens ] | [ "/" ssrdgens_tl(dgens) ] -> [ cons_dep dgens ] | [ cpattern(dt) ssrdgens_tl(dgens) ] -> [ cons_gen (nodocc, dt) dgens ] | [ ] -> [ [[]], [] ] END ARGUMENT EXTEND ssrdgens TYPED AS ssrdgens_tl PRINTED BY pr_ssrdgens | [ ":" ssrgen(gen) ssrdgens_tl(dgens) ] -> [ cons_gen gen dgens ] END let genstac (gens, clr) ist = tclTHENLIST (cleartac clr :: List.rev_map (gentac ist) gens) (* Common code to handle generalization lists along with the defective case *) let with_defective maintac deps clr ist gl = let top_id = match kind_of_type (pf_concl gl) with | ProdType (Name id, _, _) when has_discharged_tag (string_of_id id) -> id | _ -> top_id in let top_gen = mkclr clr, cpattern_of_id top_id in tclTHEN (introid top_id) (maintac deps top_gen ist) gl let with_dgens (gensl, clr) maintac ist = match gensl with | [deps; []] -> with_defective maintac deps clr ist | [deps; gen :: gens] -> tclTHEN (genstac (gens, clr) ist) (maintac deps gen ist) | [gen :: gens] -> tclTHEN (genstac (gens, clr) ist) (maintac [] gen ist) | _ -> with_defective maintac [] clr ist let with_deps deps0 maintac cl0 cs0 clr0 ist gl0 = let rec loop gl cl cs clr args clrs = function | [] -> let n = List.length args in maintac (if n > 0 then applist (to_lambda n cl, args) else cl) clrs ist gl0 | dep :: deps -> let gl' = first_goal (genclrtac cl cs clr gl) in let cl', c', clr' = pf_interp_gen ist gl' false dep in loop gl' cl' [c'] clr' (c' :: args) (clr' :: clrs) deps in loop gl0 cl0 cs0 clr0 cs0 [clr0] (List.rev deps0) (** Equations *) (* argument *) type ssreqid = ssripat option let pr_eqid = function Some pat -> str " " ++ pr_ipat pat | None -> mt () let pr_ssreqid _ _ _ = pr_eqid (* We must use primitive parsing here to avoid conflicts with the *) (* basic move, case, and elim tactics. *) ARGUMENT EXTEND ssreqid TYPED AS ssripatrep option PRINTED BY pr_ssreqid | [ "Qed" ] -> [ Util.anomaly "Grammar placeholder match" ] END let accept_ssreqid strm = match Compat.get_tok (Util.stream_nth 0 strm) with | Tok.IDENT _ -> accept_before_syms [":"] strm | Tok.KEYWORD ":" -> () | Tok.KEYWORD pat when List.mem pat ["_"; "?"; "->"; "<-"] -> accept_before_syms [":"] strm | _ -> raise Stream.Failure let test_ssreqid = Gram.Entry.of_parser "test_ssreqid" accept_ssreqid GEXTEND Gram GLOBAL: ssreqid; ssreqpat: [ [ id = Prim.ident -> IpatId id | "_" -> IpatWild | "?" -> IpatAnon | occ = ssrdocc; "->" -> (match occ with | None, occ -> IpatRw (occ, L2R) | _ -> loc_error loc "Only occurrences are allowed here") | occ = ssrdocc; "<-" -> (match occ with | None, occ -> IpatRw (occ, R2L) | _ -> loc_error loc "Only occurrences are allowed here") | "->" -> IpatRw (allocc, L2R) | "<-" -> IpatRw (allocc, R2L) ]]; ssreqid: [ [ test_ssreqid; pat = ssreqpat -> Some pat | test_ssreqid -> None ]]; END (* creation *) let mkEq dir cl c t n = let eqargs = [|t; c; c|] in eqargs.(dir_org dir) <- mkRel n; mkArrow (mkApp (build_coq_eq(), eqargs)) (lift 1 cl), mkRefl t c let pushmoveeqtac cl c = let x, t, cl1 = destProd cl in let cl2, eqc = mkEq R2L cl1 c t 1 in apply_type (mkProd (x, t, cl2)) [c; eqc] let pushcaseeqtac cl gl = let cl1, args = destApplication cl in let n = Array.length args in let dc, cl2 = decompose_lam_n n cl1 in let _, t = List.nth dc (n - 1) in let cl3, eqc = mkEq R2L cl2 args.(0) t n in let cl4 = mkApp (compose_lam dc (mkProt (pf_type_of gl cl) cl3), args) in tclTHEN (apply_type cl4 [eqc]) (convert_concl cl4) gl let pushelimeqtac gl = let _, args = destApplication (pf_concl gl) in let x, t, _ = destLambda args.(1) in let cl1 = mkApp (args.(1), Array.sub args 2 (Array.length args - 2)) in let cl2, eqc = mkEq L2R cl1 args.(2) t 1 in tclTHEN (apply_type (mkProd (x, t, cl2)) [args.(2); eqc]) intro gl (** Bookkeeping (discharge-intro) argument *) (* Since all bookkeeping ssr commands have the same discharge-intro *) (* argument format we use a single grammar entry point to parse them. *) (* the entry point parses only non-empty arguments to avoid conflicts *) (* with the basic Coq tactics. *) (* type ssrarg = ssrview * (ssreqid * (ssrdgens * ssripats)) *) let pr_ssrarg _ _ _ (view, (eqid, (dgens, ipats))) = let pri = pr_intros (gens_sep dgens) in pr_view view ++ pr_eqid eqid ++ pr_dgens pr_gen dgens ++ pri ipats ARGUMENT EXTEND ssrarg TYPED AS ssrview * (ssreqid * (ssrdgens * ssrintros)) PRINTED BY pr_ssrarg | [ ssrview(view) ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> [ view, (eqid, (dgens, ipats)) ] | [ ssrview(view) ssrclear(clr) ssrintros(ipats) ] -> [ view, (None, (([], clr), ipats)) ] | [ ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> [ [], (eqid, (dgens, ipats)) ] | [ ssrclear_ne(clr) ssrintros(ipats) ] -> [ [], (None, (([], clr), ipats)) ] | [ ssrintros_ne(ipats) ] -> [ [], (None, (([], []), ipats)) ] END (** The "clear" tactic *) (* We just add a numeric version that clears the n top assumptions. *) let poptac ?ist n = introstac ?ist (list_tabulate (fun _ -> IpatWild) n) TACTIC EXTEND ssrclear | [ "clear" natural(n) ltacctx(ctx) ] -> [poptac ~ist:(get_ltacctx ctx) n] END (** The "move" tactic *) let rec improper_intros = function | IpatSimpl _ :: ipats -> improper_intros ipats | (IpatId _ | IpatAnon | IpatCase _ | IpatAll) :: _ -> false | _ -> true let check_movearg = function | view, (eqid, _) when view <> [] && eqid <> None -> Util.error "incompatible view and equation in move tactic" | view, (_, (([gen :: _], _), _)) when view <> [] && has_occ gen -> Util.error "incompatible view and occurrence switch in move tactic" | _, (_, ((dgens, _), _)) when List.length dgens > 1 -> Util.error "dependents switch `/' in move tactic" | _, (eqid, (_, (ipats, _))) when eqid <> None && improper_intros ipats -> Util.error "no proper intro pattern for equation in move tactic" | arg -> arg ARGUMENT EXTEND ssrmovearg TYPED AS ssrarg PRINTED BY pr_ssrarg | [ ssrarg(arg) ] -> [ check_movearg arg ] END let viewmovetac_aux clear name_ref (_, vl as v) _ gen ist gl = let cl, c, clr = pf_interp_gen ist gl false gen in let cl, c = if vl = [] then cl, c else pf_with_view ist gl v cl c in let clr = if clear then clr else [] in name_ref := (match id_of_cpattern (snd gen) with Some id -> id | _ -> top_id); genclrtac cl [c] clr gl let () = move_top_with_view := (fun c r v -> with_defective (viewmovetac_aux c r v) [] []) let viewmovetac v deps gen ist gl = viewmovetac_aux true (ref top_id) v deps gen ist gl let eqmovetac _ gen ist gl = let cl, c, _ = pf_interp_gen ist gl false gen in pushmoveeqtac cl c gl let movehnftac gl = match kind_of_term (pf_concl gl) with | Prod _ | LetIn _ -> tclIDTAC gl | _ -> hnf_in_concl gl let ssrmovetac ist = function | _::_ as view, (_, (dgens, (ipats, _))) -> let dgentac = with_dgens dgens (viewmovetac (true, view)) ist in tclTHEN dgentac (introstac ~ist ipats) | _, (Some pat, (dgens, (ipats, _))) -> let dgentac = with_dgens dgens eqmovetac ist in tclTHEN dgentac (introstac ~ist (eqmoveipats pat ipats)) | _, (_, (([gens], clr), (ipats, _))) -> let gentac = genstac (gens, clr) ist in tclTHEN gentac (introstac ~ist ipats) | _, (_, ((_, clr), (ipats, _))) -> tclTHENLIST [movehnftac; cleartac clr; introstac ~ist ipats] let ist_of_arg (_, (_, (_, (_, ctx)))) = get_ltacctx ctx TACTIC EXTEND ssrmove | [ "move" ssrmovearg(arg) ssrrpat(pat) ] -> [ let ist = ist_of_arg arg in tclTHEN (ssrmovetac ist arg) (introstac ~ist [pat]) ] | [ "move" ssrmovearg(arg) ssrclauses(clauses) ] -> [ let ist = ist_of_arg arg in tclCLAUSES ist (ssrmovetac ist arg) clauses ] | [ "move" ssrrpat(pat) ltacctx(ctx) ] -> [ introstac ~ist:(get_ltacctx ctx) [pat] ] | [ "move" ] -> [ movehnftac ] END (* TASSI: given the type of an elimination principle, it finds the higher order * argument (index), it computes it's arity and the arity of the eliminator and * checks if the eliminator is recursive or not *) let analyze_eliminator elimty env sigma = let rec loop ctx t = match kind_of_type t with | AtomicType (hd, args) when isRel hd -> ctx, destRel hd, not (noccurn 1 t), Array.length args | CastType (t, _) -> loop ctx t | ProdType (x, ty, t) -> loop ((x,None,ty) :: ctx) t | LetInType (x,b,ty,t) -> loop ((x,Some b,ty) :: ctx) (subst1 b t) | _ -> let env' = Environ.push_rel_context ctx env in let t' = Reductionops.whd_betadeltaiota env' sigma t in if not (eq_constr t t') then loop ctx t' else errorstrm (str"The eliminator has the wrong shape."++spc()++ str"A (applied) bound variable was expected as the conclusion of "++ str"the eliminator's"++Pp.cut()++str"type:"++spc()++pr_constr elimty) in let ctx, pred_id, elim_is_dep, n_pred_args = loop [] elimty in let n_elim_args = rel_context_nhyps ctx in let is_rec_elim = let count_occurn n term = let count = ref 0 in let rec occur_rec n c = match kind_of_term c with | Rel m -> if m = n then incr count | _ -> iter_constr_with_binders succ occur_rec n c in occur_rec n term; !count in let occurr2 n t = count_occurn n t > 1 in not (list_for_all_i (fun i (_,rd) -> pred_id <= i || not (occurr2 (pred_id - i) rd)) 1 (assums_of_rel_context ctx)) in n_elim_args - pred_id, n_elim_args, is_rec_elim, elim_is_dep, n_pred_args (* TASSI: This version of unprotects inlines the unfold tactic definition, * since we don't want to wipe out let-ins, and it seems there is no flag * to change that behaviour in the standard unfold code *) let unprotecttac gl = let prot = destConst (mkSsrConst "protect_term") in onClause (fun idopt -> let hyploc = Option.map (fun id -> id, InHyp) idopt in reduct_option (Reductionops.clos_norm_flags (Closure.RedFlags.mkflags [Closure.RedFlags.fBETA; Closure.RedFlags.fCONST prot; Closure.RedFlags.fIOTA]), DEFAULTcast) hyploc) allHypsAndConcl gl let dependent_apply_error = try Util.error "Could not fill dependent hole in \"apply\"" with err -> err (* TASSI: Sometimes Coq's apply fails. According to my experience it may be * related to goals that are products and with beta redexes. In that case it * guesses the wrong number of implicit arguments for your lemma. What follows * is just like apply, but with a user-provided number n of implicits. * * Refine.refine function that handles type classes and evars but fails to * handle "dependently typed higher order evars". * * Refiner.refiner that does not handle metas with a non ground type but works * with dependently typed higher order metas. *) let applyn ~with_evars n t gl = if with_evars then let t, sigma = if n = 0 then t, project gl else let t, _, _, gl = pf_saturate gl t n in (* saturate with evars *) t, project gl in pp(lazy(str"Refine.refine " ++ pr_constr t)); Refine.refine (sigma, t) gl else let t, gl = if n = 0 then t, gl else let sigma, si = project gl, sig_it gl in let rec loop sigma bo args = function (* saturate with metas *) | 0 -> mkApp (t, Array.of_list (List.rev args)), re_sig si sigma | n -> match kind_of_term bo with | Lambda (_, ty, bo) -> if not (closed0 ty) then raise dependent_apply_error; let m = Evarutil.new_meta () in loop (meta_declare m ty sigma) bo ((mkMeta m)::args) (n-1) | _ -> assert false in loop sigma t [] n in pp(lazy(str"Refiner.refiner " ++ pr_constr t)); Refiner.refiner (Proof_type.Prim (Proof_type.Refine t)) gl let refine_with ?(first_goes_last=false) ?(with_evars=true) oc gl = let rec mkRels = function 1 -> [] | n -> mkRel n :: mkRels (n-1) in let n, oc = pf_abs_evars_pirrel gl oc in let oc = if not first_goes_last || n <= 1 then oc else let l, c = decompose_lam oc in if not (list_for_all_i (fun i (_,t) -> closedn ~-i t) (1-n) l) then oc else compose_lam (let xs,y = list_chop (n-1) l in y @ xs) (mkApp (compose_lam l c, Array.of_list (mkRel 1 :: mkRels n))) in pp(lazy(str"after: " ++ pr_constr oc)); try applyn ~with_evars n oc gl with _ -> raise dependent_apply_error (** The "case" and "elim" tactic *) (* A case without explicit dependent terms but with both a view and an *) (* occurrence switch and/or an equation is treated as dependent, with the *) (* viewed term as the dependent term (the occurrence switch would be *) (* meaningless otherwise). When both a view and explicit dependents are *) (* present, it is forbidden to put a (meaningless) occurrence switch on *) (* the viewed term. *) (* This is both elim and case (defaulting to the former). If ~elim is omitted * the standard eliminator is chosen. The code is made of 4 parts: * 1. find the eliminator if not given as ~elim and analyze it * 2. build the patterns to be matched against the conclusion, looking at * (occ, c), deps and the pattern inferred from the type of the eliminator * 3. build the new predicate matching the patterns, and the tactic to * generalize the equality in case eqid is not None * 4. build the tactic handle intructions and clears as required in ipats and * by eqid *) let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = (* some sanity checks *) let oc, orig_clr, occ, c_gen = match what with | `EConstr(_,_,t) when isEvar t -> anomaly "elim called on a constr evar" | `EGen _ when ist = None -> anomaly "no ist and non simple elimination" | `EGen (_, g) when elim = None && is_wildcard g -> errorstrm(str"Indeterminate pattern and no eliminator") | `EGen ((Some clr,occ), g) when is_wildcard g -> None, clr, occ, None | `EGen ((None, occ), g) when is_wildcard g -> None,[],occ,None | `EGen ((_, occ), p as gen) -> let _, c, clr = pf_interp_gen (Option.get ist) gl true gen in Some c, clr, occ, Some p | `EConstr (clr, occ, c) -> Some c, clr, occ, None in let orig_gl, concl, env = gl, pf_concl gl, pf_env gl in pp(lazy(str(if is_case then "==CASE==" else "==ELIM=="))); (* Utils of local interest only *) let iD s ?t gl = let t = match t with None -> pf_concl gl | Some x -> x in pp(lazy(str s ++ pr_constr t)); tclIDTAC gl in let eq, protectC = build_coq_eq (), mkSsrConst "protect_term" in let fire_subst gl t = Reductionops.nf_evar (project gl) t in let fire_sigma sigma t = Reductionops.nf_evar sigma t in let is_undef_pat = function | sigma, T t -> (match kind_of_term (fire_sigma sigma t) with Evar _ -> true | _ -> false) | _ -> false in let match_pat env p occ h cl = let sigma0 = project orig_gl in pp(lazy(str"matching: " ++ pp_pattern p)); let c, cl = fill_occ_pattern ~raise_NoMatch:true env sigma0 cl p occ h in pp(lazy(str" got: " ++ pr_constr c)); c, cl in let mkTpat gl t = (* takes a term, refreshes it and makes a T pattern *) let n, t = pf_abs_evars orig_gl (project gl, fire_subst gl t) in let t, _, _, sigma = saturate ~beta:true env (project gl) t n in sigma, T t in let unif_redex gl (sigma, r as p) t = (* t is a hint for the redex of p *) let n, t = pf_abs_evars orig_gl (project gl, fire_subst gl t) in let t, _, _, sigma = saturate ~beta:true env sigma t n in match r with | X_In_T (e, p) -> sigma, E_As_X_In_T (t, e, p) | _ -> try unify_HO env sigma t (redex_of_pattern p), r with _ -> p in (* finds the eliminator applies it to evars and c saturated as needed *) (* obtaining "elim ??? (c ???)". pred is the higher order evar *) let cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl = match elim with | Some elim -> let elimty = pf_type_of gl elim in let pred_id, n_elim_args, is_rec, elim_is_dep, n_pred_args = analyze_eliminator elimty env (project gl) in let elim, elimty, elim_args, gl = pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in let pred = List.assoc pred_id elim_args in let elimty = Reductionops.whd_betadeltaiota env (project gl) elimty in None, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl | None -> let c = Option.get oc in let c_ty = pf_type_of gl c in let (kn, i) as ind, unfolded_c_ty = pf_reduce_to_quantified_ind gl c_ty in let sort = elimination_sort_of_goal gl in let elim = if not is_case then Indrec.lookup_eliminator ind sort else pf_apply Indrec.build_case_analysis_scheme gl ind true sort in let elimty = pf_type_of gl elim in let pred_id,n_elim_args,is_rec,elim_is_dep,n_pred_args = analyze_eliminator elimty env (project gl) in let rctx = fst (decompose_prod_assum unfolded_c_ty) in let n_c_args = rel_context_length rctx in let c, c_ty, t_args, gl = pf_saturate gl c ~ty:c_ty n_c_args in let elim, elimty, elim_args, gl = pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in let pred = List.assoc pred_id elim_args in let pc = match n_c_args, c_gen with | 0, Some p -> interp_cpattern (Option.get ist) orig_gl p None | _ -> mkTpat gl c in let cty = Some (c, c_ty, pc) in let elimty = Reductionops.whd_betadeltaiota env (project gl) elimty in cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl in pp(lazy(str"elim= "++ pr_constr_pat elim)); pp(lazy(str"elimty= "++ pr_constr elimty)); let inf_deps_r = match kind_of_type elimty with | AtomicType (_, args) -> List.rev (Array.to_list args) | _ -> assert false in let saturate_until gl c c_ty f = let rec loop n = try let c, c_ty, _, gl = pf_saturate gl c ~ty:c_ty n in let gl' = f c c_ty gl in Some (c, c_ty, gl, gl') with NotEnoughProducts -> None | _ -> loop (n+1) in loop 0 in let elim_is_dep, gl = match cty with | None -> true, gl | Some (c, c_ty, _) -> let res = if elim_is_dep then None else let arg = List.assoc (n_elim_args - 1) elim_args in let arg_ty = pf_type_of gl arg in match saturate_until gl c c_ty (fun c c_ty gl -> pf_unify_HO (pf_unify_HO gl c_ty arg_ty) arg c) with | Some (c, _, _, gl) -> Some (false, gl) | None -> None in match res with | Some x -> x | None -> let inf_arg = List.hd inf_deps_r in let inf_arg_ty = pf_type_of gl inf_arg in match saturate_until gl c c_ty (fun _ c_ty gl -> pf_unify_HO gl c_ty inf_arg_ty) with | Some (c, _, _,gl) -> true, gl | None -> errorstrm (str"Unable to apply the eliminator to the term"++ spc()++pr_constr c++spc()++str"or to unify it's type with"++ pr_constr inf_arg_ty) in pp(lazy(str"elim_is_dep= " ++ bool elim_is_dep)); let predty = pf_type_of gl pred in (* Patterns for the inductive types indexes to be bound in pred are computed * looking at the ones provided by the user and the inferred ones looking at * the type of the elimination principle *) let pp_pat (_,p,_,_) = pp_pattern p in let pp_inf_pat gl (_,_,t,_) = pr_constr_pat (fire_subst gl t) in let patterns, clr, gl = let rec loop patterns clr i = function | [],[] -> patterns, clr, gl | ((oclr, occ), t):: deps, inf_t :: inf_deps -> let ist = match ist with Some x -> x | None -> assert false in let p = interp_cpattern ist orig_gl t None in let clr_t = interp_clr (oclr,(tag_of_cpattern t,redex_of_pattern p))in (* if we are the index for the equation we do not clear *) let clr_t = if deps = [] && eqid <> None then [] else clr_t in let p = if is_undef_pat p then mkTpat gl inf_t else p in loop (patterns @ [i, p, inf_t, occ]) (clr_t @ clr) (i+1) (deps, inf_deps) | [], c :: inf_deps -> pp(lazy(str"adding inferred pattern " ++ pr_constr_pat c)); loop (patterns @ [i, mkTpat gl c, c, allocc]) clr (i+1) ([], inf_deps) | _::_, [] -> errorstrm (str "Too many dependent abstractions") in let deps, head_p, inf_deps_r = match what, elim_is_dep, cty with | `EConstr _, _, None -> anomaly "Simple welim with no term" | _, false, _ -> deps, [], inf_deps_r | `EGen gen, true, None -> deps @ [gen], [], inf_deps_r | _, true, Some (c, _, pc) -> let occ = if occ = None then allocc else occ in let inf_p, inf_deps_r = List.hd inf_deps_r, List.tl inf_deps_r in deps, [1, pc, inf_p, occ], inf_deps_r in let patterns, clr, gl = loop [] orig_clr (List.length head_p+1) (List.rev deps, inf_deps_r) in head_p @ patterns, Util.list_uniquize clr, gl in pp(lazy(pp_concat (str"patterns=") (List.map pp_pat patterns))); pp(lazy(pp_concat (str"inf. patterns=") (List.map (pp_inf_pat gl) patterns))); (* Predicate generation, and (if necessary) tactic to generalize the * equation asked by the user *) let elim_pred, gen_eq_tac, clr, gl = let error gl t inf_t = errorstrm (str"The given pattern matches the term"++ spc()++pp_term gl t++spc()++str"while the inferred pattern"++ spc()++pr_constr_pat (fire_subst gl inf_t)++spc()++ str"doesn't") in let match_or_postpone (cl, gl, post) (h, p, inf_t, occ) = let p = unif_redex gl p inf_t in if is_undef_pat p then let () = pp(lazy(str"postponing " ++ pp_pattern p)) in cl, gl, post @ [h, p, inf_t, occ] else try let c, cl = match_pat env p occ h cl in let gl = try pf_unify_HO gl inf_t c with _ -> error gl c inf_t in cl, gl, post with | NoMatch | NoProgress -> let e = redex_of_pattern p in let n, e = pf_abs_evars gl (fst p, e) in let e, _, _, gl = pf_saturate ~beta:true gl e n in let gl = try pf_unify_HO gl inf_t e with _ -> error gl e inf_t in cl, gl, post in let rec match_all concl gl patterns = let concl, gl, postponed = List.fold_left match_or_postpone (concl, gl, []) patterns in if postponed = [] then concl, gl else if List.length postponed = List.length patterns then errorstrm (str "Some patterns are undefined even after all"++spc()++ str"the defined ones matched") else match_all concl gl postponed in let concl, gl = match_all concl gl patterns in let pred_rctx, _ = decompose_prod_assum (fire_subst gl predty) in let concl, gen_eq_tac, clr = match eqid with | Some (IpatId _) when not is_rec -> let k = List.length deps in let c = fire_subst gl (List.assoc (n_elim_args - k - 1) elim_args) in let t = pf_type_of gl c in let gen_eq_tac = let refl = mkApp (eq, [|t; c; c|]) in let new_concl = mkArrow refl (lift 1 (pf_concl orig_gl)) in let new_concl = fire_subst gl new_concl in let erefl = fire_subst gl (mkRefl t c) in apply_type new_concl [erefl] in let rel = k + if elim_is_dep then 1 else 0 in let src = mkProt mkProp (mkApp (eq,[|t; c; mkRel rel|])) in let concl = mkArrow src (lift 1 concl) in let clr = if deps <> [] then clr else [] in concl, gen_eq_tac, clr | _ -> concl, tclIDTAC, clr in let mk_lam t r = mkLambda_or_LetIn r t in let concl = List.fold_left mk_lam concl pred_rctx in let concl = if eqid <> None && is_rec then mkProt (pf_type_of gl concl) concl else concl in concl, gen_eq_tac, clr, gl in pp(lazy(str"elim_pred=" ++ pp_term gl elim_pred)); let pty = Typing.type_of env (project gl) elim_pred in pp(lazy(str"elim_pred_ty=" ++ pp_term gl pty)); let gl = pf_unify_HO gl pred elim_pred in (* check that the patterns do not contain non instantiated dependent metas *) let () = let evars_of_term = Evarutil.evars_of_term in let patterns = List.map (fun (_,_,t,_) -> fire_subst gl t) patterns in let patterns_ev = List.map evars_of_term patterns in let ev = List.fold_left Intset.union Intset.empty patterns_ev in let ty_ev = Intset.fold (fun i e -> let ex = Evd.existential_of_int i in let i_ty = Evd.evar_concl (Evd.find (project gl) ex) in Intset.union e (evars_of_term i_ty)) ev Intset.empty in let inter = Intset.inter ev ty_ev in if not (Intset.is_empty inter) then begin let i = Intset.choose inter in let pat = List.find (fun t -> Intset.mem i (evars_of_term t)) patterns in errorstrm(str"Pattern"++spc()++pr_constr_pat pat++spc()++ str"was not completely instantiated and one of its variables"++spc()++ str"occurs in the type of another non instantieted pattern variable"); end in (* the elim tactic, with the eliminator and the predicated we computed *) let elim = project gl, fire_subst gl elim in let elim_tac gl = tclTHENLIST [refine_with ~with_evars:false elim; cleartac clr] gl in (* handling of following intro patterns and equation introduction if any *) let elim_intro_tac gl = let intro_eq = match eqid with | Some (IpatId ipat) when not is_rec -> let rec intro_eq gl = match kind_of_type (pf_concl gl) with | ProdType (_, src, tgt) -> (match kind_of_type src with | AtomicType (hd, _) when eq_constr hd protectC -> tclTHENLIST [unprotecttac; introid ipat] gl | _ -> tclTHENLIST [ iD "IA"; introstac [IpatAnon]; intro_eq] gl) |_ -> errorstrm (str "Too many names in intro pattern") in intro_eq | Some (IpatId ipat) -> let name gl = mk_anon_id "K" gl in let intro_lhs gl = let elim_name = match orig_clr, what with | [SsrHyp(_, x)], _ -> x | _, `EConstr(_,_,t) when isVar t -> destVar t | _ -> name gl in if is_name_in_ipats elim_name ipats then introid (name gl) gl else introid elim_name gl in let rec gen_eq_tac gl = let concl = pf_concl gl in let ctx, last = decompose_prod_assum concl in let args = match kind_of_type last with | AtomicType (hd, args) -> assert(eq_constr hd protectC); args | _ -> assert false in let case = args.(Array.length args-1) in if not(closed0 case) then tclTHEN (introstac [IpatAnon]) gen_eq_tac gl else let case_ty = pf_type_of gl case in let refl = mkApp (eq, [|lift 1 case_ty; mkRel 1; lift 1 case|]) in let new_concl = fire_subst gl (mkProd (Name (name gl), case_ty, mkArrow refl (lift 2 concl))) in let erefl = fire_subst gl (mkRefl case_ty case) in apply_type new_concl [case;erefl] gl in tclTHENLIST [gen_eq_tac; intro_lhs; introid ipat] | _ -> tclIDTAC in let unprot = if eqid <> None && is_rec then unprotecttac else tclIDTAC in tclEQINTROS ?ist elim_tac (tclTHENLIST [intro_eq; unprot]) ipats gl in tclTHENLIST [gen_eq_tac; elim_intro_tac] orig_gl ;; let simplest_newelim x= ssrelim ~is_case:false [] (`EConstr ([],None,x)) None [] let simplest_newcase x= ssrelim ~is_case:true [] (`EConstr ([],None,x)) None [] let _ = simplest_newcase_ref := simplest_newcase let check_casearg = function | view, (_, (([_; gen :: _], _), _)) when view <> [] && has_occ gen -> Util.error "incompatible view and occurrence switch in dependent case tactic" | arg -> arg ARGUMENT EXTEND ssrcasearg TYPED AS ssrarg PRINTED BY pr_ssrarg | [ ssrarg(arg) ] -> [ check_casearg arg ] END let ssrcasetac ist (view, (eqid, (dgens, (ipats, _)))) = let ndefectcasetac view eqid ipats deps ((_, occ), _ as gen) ist gl = let simple = (eqid = None && deps = [] && occ = None) in let cl, c, clr = pf_interp_gen ist gl true gen in let vc = if view = [] then c else snd(pf_with_view ist gl (false, view) cl c) in if simple && is_injection_case vc gl then tclTHENLIST [perform_injection vc; cleartac clr; introstac ~ist ipats] gl else (* macro for "case/v E: x" ---> "case E: x / (v x)" *) let deps, clr, occ = if view <> [] && eqid <> None && deps = [] then [gen], [], None else deps, clr, occ in ssrelim ~is_case:true ~ist deps (`EConstr (clr,occ, vc)) eqid ipats gl in with_dgens dgens (ndefectcasetac view eqid ipats) ist TACTIC EXTEND ssrcase | [ "case" ssrcasearg(arg) ssrclauses(clauses) ] -> [ let _, (_, (_, (_, ctx))) = arg in let ist = get_ltacctx ctx in tclCLAUSES ist (ssrcasetac ist arg) clauses ] | [ "case" ] -> [ with_top ssrscasetac ] END (** The "elim" tactic *) (* Elim views are elimination lemmas, so the eliminated term is not addded *) (* to the dependent terms as for "case", unless it actually occurs in the *) (* goal, the "all occurrences" {+} switch is used, or the equation switch *) (* is used and there are no dependents. *) let ssrelimtac ist (view, (eqid, (dgens, (ipats, _)))) = let ndefectelimtac view eqid ipats deps gen ist gl = let elim = match view with [v] -> Some (snd(force_term ist gl v)) | _ -> None in ssrelim ~ist deps (`EGen gen) ?elim eqid ipats gl in with_dgens dgens (ndefectelimtac view eqid ipats) ist TACTIC EXTEND ssrelim | [ "elim" ssrarg(arg) ssrclauses(clauses) ] -> [ let _, (_, (_, (_, ctx))) = arg in let ist = get_ltacctx ctx in tclCLAUSES ist (ssrelimtac ist arg) clauses ] | [ "elim" ] -> [ with_top simplest_newelim ] END (** 6. Backward chaining tactics: apply, exact, congr. *) (** The "apply" tactic *) let pr_agen (docc, dt) = pr_docc docc ++ pr_term dt let pr_ssragen _ _ _ = pr_agen let pr_ssragens _ _ _ = pr_dgens pr_agen ARGUMENT EXTEND ssragen TYPED AS ssrdocc * ssrterm PRINTED BY pr_ssragen | [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ] -> [ mkclr clr, dt ] | [ ssrterm(dt) ] -> [ nodocc, dt ] END ARGUMENT EXTEND ssragens TYPED AS ssragen list list * ssrclear PRINTED BY pr_ssragens | [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ssragens(agens) ] -> [ cons_gen (mkclr clr, dt) agens ] | [ "{" ne_ssrhyp_list(clr) "}" ] -> [ [[]], clr] | [ ssrterm(dt) ssragens(agens) ] -> [ cons_gen (nodocc, dt) agens ] | [ ] -> [ [[]], [] ] END let mk_applyarg views agens intros = views, (None, (agens, intros)) let pr_ssraarg _ _ _ (view, (eqid, (dgens, ipats))) = let pri = pr_intros (gens_sep dgens) in pr_view view ++ pr_eqid eqid ++ pr_dgens pr_agen dgens ++ pri ipats ARGUMENT EXTEND ssrapplyarg TYPED AS ssrview * (ssreqid * (ssragens * ssrintros)) PRINTED BY pr_ssraarg | [ ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> [ mk_applyarg [] (cons_gen gen dgens) intros ] | [ ssrclear_ne(clr) ssrintros(intros) ] -> [ mk_applyarg [] ([], clr) intros ] | [ ssrintros_ne(intros) ] -> [ mk_applyarg [] ([], []) intros ] | [ ssrview(view) ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> [ mk_applyarg view (cons_gen gen dgens) intros ] | [ ssrview(view) ssrclear(clr) ssrintros(intros) ] -> [ mk_applyarg view ([], clr) intros ] END let interp_agen ist gl ((goclr, _), (k, gc)) (clr, rcs) = let rc = glob_constr ist (project gl) (pf_env gl) gc in let rcs' = rc :: rcs in match goclr with | None -> clr, rcs' | Some ghyps -> let clr' = snd (interp_hyps ist gl ghyps) @ clr in if k <> ' ' then clr', rcs' else match rc with | GVar (loc, id) when not_section_id id -> SsrHyp (loc, id) :: clr', rcs' | GRef (loc, VarRef id) when not_section_id id -> SsrHyp (loc, id) :: clr', rcs' | _ -> clr', rcs' let interp_agens ist gl gagens = match List.fold_right (interp_agen ist gl) gagens ([], []) with | clr, rlemma :: args -> let n = interp_nbargs ist gl rlemma - List.length args in let rec loop i = if i > n then errorstrm (str "Cannot apply lemma " ++ pf_pr_glob_constr gl rlemma) else try interp_refine ist gl (mkRApp rlemma (mkRHoles i @ args)) with _ -> loop (i + 1) in clr, loop 0 | _ -> assert false let apply_rconstr ?ist t gl = let n = match ist, t with | None, (GVar (_, id) | GRef (_, VarRef id)) -> pf_nbargs gl (mkVar id) | Some ist, _ -> interp_nbargs ist gl t | _ -> anomaly "apply_rconstr without ist and not RVar" in let mkRlemma i = mkRApp t (mkRHoles i) in let cl = pf_concl gl in let rec loop i = if i > n then errorstrm (str"Cannot apply lemma "++pf_pr_glob_constr gl t) else try pf_match gl (mkRlemma i) cl with _ -> loop (i + 1) in refine_with (loop 0) gl let mkRAppView ist gl rv gv = let nb_view_imps = interp_view_nbimps ist gl rv in mkRApp rv (mkRHoles (abs nb_view_imps)) let prof_apply_interp_with = mk_profiler "ssrapplytac.interp_with";; let refine_interp_apply_view i ist gl gv = let pair i = List.map (fun x -> i, x) in let rv = pf_intern_term ist gl gv in let v = mkRAppView ist gl rv gv in let interp_with (i, hint) = interp_refine ist gl (mkRApp hint (v :: mkRHoles i)) in let interp_with x = prof_apply_interp_with.profile interp_with x in let rec loop = function | [] -> (try apply_rconstr ~ist rv gl with _ -> view_error "apply" gv) | h :: hs -> (try refine_with (snd (interp_with h)) gl with _ -> loop hs) in loop (pair i viewtab.(i) @ if i = 2 then pair 1 viewtab.(1) else []) let apply_top_tac gl = tclTHENLIST [introid top_id; apply_rconstr (mkRVar top_id); clear [top_id]] gl let inner_ssrapplytac gviews ggenl gclr ist gl = let _, clr = interp_hyps ist gl gclr in let vtac gv i gl' = refine_interp_apply_view i ist gl' gv in let ggenl, tclGENTAC = if gviews <> [] && ggenl <> [] then let ggenl= List.map (fun (x,g) -> x, cpattern_of_term g) (List.hd ggenl) in [], tclTHEN (genstac (ggenl,[]) ist) else ggenl, tclTHEN tclIDTAC in tclGENTAC (fun gl -> match gviews, ggenl with | v :: tl, [] -> let dbl = if List.length tl = 1 then 2 else 1 in tclTHEN (List.fold_left (fun acc v -> tclTHENLAST acc (vtac v dbl)) (vtac v 1) tl) (cleartac clr) gl | [], [agens] -> let clr', (_, lemma) = interp_agens ist gl agens in tclTHENLIST [cleartac clr; refine_with lemma; cleartac clr'] gl | _, _ -> tclTHEN apply_top_tac (cleartac clr) gl) gl let ssrapplytac (views, (_, ((gens, clr), intros))) = tclINTROS (inner_ssrapplytac views gens clr) intros let prof_ssrapplytac = mk_profiler "ssrapplytac";; let ssrapplytac arg gl = prof_ssrapplytac.profile (ssrapplytac arg) gl;; TACTIC EXTEND ssrapply | [ "apply" ssrapplyarg(arg) ] -> [ ssrapplytac arg ] | [ "apply" ] -> [ apply_top_tac ] END (** The "exact" tactic *) let mk_exactarg views dgens = mk_applyarg views dgens ([], rawltacctx) ARGUMENT EXTEND ssrexactarg TYPED AS ssrapplyarg PRINTED BY pr_ssraarg | [ ":" ssragen(gen) ssragens(dgens) ] -> [ mk_exactarg [] (cons_gen gen dgens) ] | [ ssrview(view) ssrclear(clr) ] -> [ mk_exactarg view ([], clr) ] | [ ssrclear_ne(clr) ] -> [ mk_exactarg [] ([], clr) ] END let vmexacttac pf gl = exact_no_check (mkCast (pf, VMcast, pf_concl gl)) gl TACTIC EXTEND ssrexact | [ "exact" ssrexactarg(arg) ] -> [ tclBY (ssrapplytac arg) ] | [ "exact" ] -> [ tclORELSE donetac (tclBY apply_top_tac) ] | [ "exact" "<:" lconstr(pf) ] -> [ vmexacttac pf ] END (** The "congr" tactic *) type ssrcongrarg = open_constr * (int * constr) let pr_ssrcongrarg _ _ _ (_, ((n, f), dgens)) = (if n <= 0 then mt () else str " " ++ pr_int n) ++ str " " ++ pr_term f ++ pr_dgens pr_gen dgens ARGUMENT EXTEND ssrcongrarg TYPED AS ltacctx * ((int * ssrterm) * ssrdgens) PRINTED BY pr_ssrcongrarg | [ natural(n) constr(c) ssrdgens(dgens) ] -> [ rawltacctx, ((n, mk_term ' ' c), dgens) ] | [ natural(n) constr(c) ] -> [ rawltacctx, ((n, mk_term ' ' c),([[]],[])) ] | [ constr(c) ssrdgens(dgens) ] -> [ rawltacctx, ((0, mk_term ' ' c), dgens) ] | [ constr(c) ] -> [ rawltacctx, ((0, mk_term ' ' c), ([[]],[])) ] END let rec mkRnat n = if n <= 0 then GRef (dummy_loc, glob_O) else mkRApp (GRef (dummy_loc, glob_S)) [mkRnat (n - 1)] let interp_congrarg_at ist gl n rf ty m = pp(lazy(str"===interp_congrarg_at===")); let congrn = mkSsrRRef "nary_congruence" in let args1 = mkRnat n :: mkRHoles n @ [ty] in let args2 = mkRHoles (3 * n) in let rec loop i = if i + n > m then None else try let rt = mkRApp congrn (args1 @ mkRApp rf (mkRHoles i) :: args2) in pp(lazy(str"rt=" ++ pr_glob_constr rt)); Some (interp_refine ist gl rt) with _ -> loop (i + 1) in loop 0 let pattern_id = mk_internal_id "pattern value" let congrtac ((n, t), ty) ist gl = pp(lazy(str"===congr===")); pp(lazy(str"concl=" ++ pr_constr (pf_concl gl))); let _, f = pf_abs_evars gl (interp_term ist gl t) in let ist' = {ist with lfun = [pattern_id, VConstr ([],f)]} in let rf = mkRltacVar pattern_id in let m = pf_nbargs gl f in let _, cf = if n > 0 then match interp_congrarg_at ist' gl n rf ty m with | Some cf -> cf | None -> errorstrm (str "No " ++ pr_int n ++ str "-congruence with " ++ pr_term t) else let rec loop i = if i > m then errorstrm (str "No congruence with " ++ pr_term t) else match interp_congrarg_at ist' gl i rf ty m with | Some cf -> cf | None -> loop (i + 1) in loop 1 in tclTHEN (refine_with cf) (tclTRY reflexivity) gl let newssrcongrtac arg ist gl = pp(lazy(str"===newcongr===")); pp(lazy(str"concl=" ++ pr_constr (pf_concl gl))); (* utils *) let fs gl t = Reductionops.nf_evar (project gl) t in let tclMATCH_GOAL (c, gl_c) proj t_ok t_fail gl = match try Some (pf_unify_HO gl_c (pf_concl gl) c) with _ -> None with | Some gl_c -> tclTHEN (convert_concl (fs gl_c c)) (t_ok (proj gl_c)) gl | None -> t_fail () gl in let mk_evar gl ty = let env, sigma, si = pf_env gl, project gl, sig_it gl in let sigma, x = Evarutil.new_evar (create_evar_defs sigma) env ty in x, re_sig si sigma in let ssr_congr lr = mkApp (mkSsrConst "ssr_congr_arrow",lr) in (* here thw two cases: simple equality or arrow *) let equality, _, eq_args, gl' = pf_saturate gl (build_coq_eq ()) 3 in tclMATCH_GOAL (equality, gl') (fun gl' -> fs gl' (List.assoc 0 eq_args)) (fun ty -> congrtac (arg, Detyping.detype false [] [] ty) ist) (fun () -> let lhs, gl' = mk_evar gl mkProp in let rhs, gl' = mk_evar gl' mkProp in let arrow = mkArrow lhs (lift 1 rhs) in tclMATCH_GOAL (arrow, gl') (fun gl' -> [|fs gl' lhs;fs gl' rhs|]) (fun lr -> tclTHEN (apply (ssr_congr lr)) (congrtac (arg, mkRType) ist)) (fun _ _ -> errorstrm (str"Conclusion is not an equality nor an arrow"))) gl ;; TACTIC EXTEND ssrcongr | [ "congr" ssrcongrarg(arg) ] -> [ let ctx, (arg, dgens) = arg in let ist = get_ltacctx ctx in match dgens with | [gens], clr -> tclTHEN (genstac (gens,clr) ist) (newssrcongrtac arg ist) | _ -> errorstrm (str"Dependent family abstractions not allowed in congr")] END (** 7. Rewriting tactics (rewrite, unlock) *) (** Coq rewrite compatibility flag *) let ssr_strict_match = ref false let _ = Goptions.declare_bool_option { Goptions.optsync = true; Goptions.optname = "strict redex matching"; Goptions.optkey = ["Match"; "Strict"]; Goptions.optread = (fun () -> !ssr_strict_match); Goptions.optdepr = false; Goptions.optwrite = (fun b -> ssr_strict_match := b) } (** Rewrite multiplier *) type ssrmult = int * ssrmmod let notimes = 0 let nomult = 1, Once let pr_mult (n, m) = if n > 0 && m <> Once then pr_int n ++ pr_mmod m else pr_mmod m let pr_ssrmult _ _ _ = pr_mult ARGUMENT EXTEND ssrmult_ne TYPED AS int * ssrmmod PRINTED BY pr_ssrmult | [ natural(n) ssrmmod(m) ] -> [ check_index loc n, m ] | [ ssrmmod(m) ] -> [ notimes, m ] END ARGUMENT EXTEND ssrmult TYPED AS ssrmult_ne PRINTED BY pr_ssrmult | [ ssrmult_ne(m) ] -> [ m ] | [ ] -> [ nomult ] END (** Rewrite clear/occ switches *) let pr_rwocc = function | None, None -> mt () | None, occ -> pr_occ occ | Some clr, _ -> pr_clear_ne clr let pr_ssrrwocc _ _ _ = pr_rwocc ARGUMENT EXTEND ssrrwocc TYPED AS ssrdocc PRINTED BY pr_ssrrwocc | [ "{" ssrhyp_list(clr) "}" ] -> [ mkclr clr ] | [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ] | [ ] -> [ noclr ] END (** Rewrite rules *) type ssrwkind = RWred of ssrsimpl | RWdef | RWeq (* type ssrrule = ssrwkind * ssrterm *) let pr_rwkind = function | RWred s -> pr_simpl s | RWdef -> str "/" | RWeq -> mt () let wit_ssrrwkind, globwit_ssrrwkind, rawwit_ssrrwkind = add_genarg "ssrrwkind" pr_rwkind let pr_rule = function | RWred s, _ -> pr_simpl s | RWdef, r-> str "/" ++ pr_term r | RWeq, r -> pr_term r let pr_ssrrule _ _ _ = pr_rule let noruleterm loc = mk_term ' ' (mkCProp loc) ARGUMENT EXTEND ssrrule_ne TYPED AS ssrrwkind * ssrterm PRINTED BY pr_ssrrule | [ ssrsimpl_ne(s) ] -> [ RWred s, noruleterm loc ] | [ "/" ssrterm(t) ] -> [ RWdef, t ] | [ ssrterm(t) ] -> [ RWeq, t ] END ARGUMENT EXTEND ssrrule TYPED AS ssrrule_ne PRINTED BY pr_ssrrule | [ ssrrule_ne(r) ] -> [ r ] | [ ] -> [ RWred Nop, noruleterm loc ] END (** Rewrite arguments *) (* type ssrrwarg = (ssrdir * ssrmult) * ((ssrdocc * ssrpattern) * ssrrule) *) let pr_option f = function None -> mt() | Some x -> f x let pr_pattern_squarep= pr_option (fun r -> str "[" ++ pr_rpattern r ++ str "]") let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep let pr_rwarg ((d, m), ((docc, rx), r)) = pr_rwdir d ++ pr_mult m ++ pr_rwocc docc ++ pr_pattern_squarep rx ++ pr_rule r let pr_ssrrwarg _ _ _ = pr_rwarg let mk_rwarg (d, (n, _ as m)) ((clr, occ as docc), rx) (rt, _ as r) = if rt <> RWeq then begin if rt = RWred Nop && not (m = nomult && occ = None && rx = None) && (clr = None || clr = Some []) then anomaly "Improper rewrite clear switch"; if d = R2L && rt <> RWdef then error "Right-to-left switch on simplification"; if n <> 1 && rt = RWred Cut then error "Bad or useless multiplier"; if occ <> None && rx = None && rt <> RWdef then error "Missing redex for simplification occurrence" end; (d, m), ((docc, rx), r) let norwmult = L2R, nomult let norwocc = noclr, None (* let pattern_ident = Prim.pattern_ident in GEXTEND Gram GLOBAL: pattern_ident; pattern_ident: [[c = pattern_ident -> (CRef (Ident (loc,c)), NoBindings)]]; END *) ARGUMENT EXTEND ssrpattern_squarep TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ] | [ ] -> [ None ] END ARGUMENT EXTEND ssrpattern_ne_squarep TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ] END ARGUMENT EXTEND ssrrwarg TYPED AS (ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule) PRINTED BY pr_ssrrwarg | [ "-" ssrmult(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> [ mk_rwarg (R2L, m) (docc, rx) r ] | [ "-/" ssrterm(t) ] -> (* just in case '-/' should become a token *) [ mk_rwarg (R2L, nomult) norwocc (RWdef, t) ] | [ ssrmult_ne(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> [ mk_rwarg (L2R, m) (docc, rx) r ] | [ "{" ne_ssrhyp_list(clr) "}" ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> [ mk_rwarg norwmult (mkclr clr, rx) r ] | [ "{" ne_ssrhyp_list(clr) "}" ssrrule(r) ] -> [ mk_rwarg norwmult (mkclr clr, None) r ] | [ "{" ssrocc(occ) "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> [ mk_rwarg norwmult (mkocc occ, rx) r ] | [ "{" "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> [ mk_rwarg norwmult (nodocc, rx) r ] | [ ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> [ mk_rwarg norwmult (noclr, rx) r ] | [ ssrrule_ne(r) ] -> [ mk_rwarg norwmult norwocc r ] END let simplintac occ rdx sim gl = let simptac gl = let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let simp env c _ = red_safe Tacred.simpl env sigma0 c in convert_concl_no_check (eval_pattern env0 sigma0 concl0 rdx occ simp) gl in match sim with | Simpl -> simptac gl | SimplCut -> tclTHEN simptac (tclTRY donetac) gl | _ -> simpltac sim gl let rec get_evalref c = match kind_of_term c with | Var id -> EvalVarRef id | Const k -> EvalConstRef k | App (c', _) -> get_evalref c' | Cast (c', _, _) -> get_evalref c' | _ -> errorstrm (str "The term " ++ pr_constr c ++ str " is not unfoldable") (* Strip a pattern generated by a prenex implicit to its constant. *) let strip_unfold_term ((sigma, t) as p) kt = match kind_of_term t with | App (f, a) when kt = ' ' && array_for_all isEvar a && isConst f -> (sigma, f), true | Const _ | Var _ -> p, true | _ -> p, false let unfoldintac occ rdx t (kt,_) gl = let fs sigma x = Reductionops.nf_evar sigma x in let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let (sigma, t), const = strip_unfold_term t kt in let body env t c = Tacred.unfoldn [(true, [1]), get_evalref t] env sigma0 c in let easy = occ = None && rdx = None in let red_flags = if easy then Closure.betaiotazeta else Closure.betaiota in let beta env = Reductionops.clos_norm_flags red_flags env sigma0 in let unfold, conclude = match rdx with | Some (_, (In_T _ | In_X_In_T _)) | None -> let ise = create_evar_defs sigma in let ise, u = mk_tpattern env0 sigma0 (ise,t) all_ok L2R t in let find_T, end_T = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in (fun env c h -> try find_T env c h (fun env c _ -> body env t c) with NoMatch when easy -> c | NoMatch | NoProgress -> errorstrm (str"No occurrence of " ++ pr_constr_pat t ++ spc() ++ str "in " ++ pr_constr c)), (fun () -> try end_T () with | NoMatch when easy -> fake_pmatcher_end () | NoMatch -> anomaly "unfoldintac") | _ -> (fun env (c as orig_c) h -> if const then let rec aux c = match kind_of_term c with | Const _ when eq_constr c t -> body env t t | App (f,a) when eq_constr f t -> mkApp (body env f f,a) | _ -> let c = Reductionops.whd_betaiotazeta sigma0 c in match kind_of_term c with | Const _ when eq_constr c t -> body env t t | App (f,a) when eq_constr f t -> mkApp (body env f f,a) | Const f -> aux (body env c c) | App (f, a) -> aux (mkApp (body env f f, a)) | _ -> errorstrm (str "The term "++pr_constr orig_c++ str" contains no " ++ pr_constr t ++ str" even after unfolding") in aux c else try body env t (fs (unify_HO env sigma c t) t) with _ -> errorstrm (str "The term " ++ pr_constr c ++spc()++ str "does not unify with " ++ pr_constr_pat t)), fake_pmatcher_end in let concl = try beta env0 (eval_pattern env0 sigma0 concl0 rdx occ unfold) with Option.IsNone -> errorstrm (str"Failed to unfold " ++ pr_constr_pat t) in let _ = conclude () in convert_concl concl gl ;; let foldtac occ rdx ft gl = let fs sigma x = Reductionops.nf_evar sigma x in let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let sigma, t = ft in let fold, conclude = match rdx with | Some (_, (In_T _ | In_X_In_T _)) | None -> let ise = create_evar_defs sigma in let ut = red_product_skip_id env0 sigma t in let ise, ut = mk_tpattern env0 sigma0 (ise,t) all_ok L2R ut in let find_T, end_T = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[ut]) in (fun env c h -> try find_T env c h (fun env t _ -> t) with NoMatch -> c), (fun () -> try end_T () with NoMatch -> fake_pmatcher_end ()) | _ -> (fun env c h -> try let sigma = unify_HO env sigma c t in fs sigma t with _ -> errorstrm (str "fold pattern " ++ pr_constr_pat t ++ spc () ++ str "does not match redex " ++ pr_constr_pat c)), fake_pmatcher_end in let concl = eval_pattern env0 sigma0 concl0 rdx occ fold in let _ = conclude () in convert_concl concl gl ;; let converse_dir = function L2R -> R2L | R2L -> L2R let rw_progress rhs lhs ise = not (eq_constr lhs (Evarutil.nf_evar ise rhs)) (* Coq has a more general form of "equation" (any type with a single *) (* constructor with no arguments with_rect_r elimination lemmas). *) (* However there is no clear way of determining the LHS and RHS of *) (* such a generic Leibnitz equation -- short of inspecting the type *) (* of the elimination lemmas. *) let rec strip_prod_assum c = match kind_of_term c with | Prod (_, _, c') -> strip_prod_assum c' | LetIn (_, v, _, c') -> strip_prod_assum (subst1 v c) | Cast (c', _, _) -> strip_prod_assum c' | _ -> c let rule_id = mk_internal_id "rewrite rule" exception PRtype_error exception PRindetermined_rhs of constr let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl = let env = pf_env gl in let beta = Reductionops.clos_norm_flags Closure.beta env sigma in let sigma, p = let sigma = create_evar_defs sigma in Evarutil.new_evar sigma env (beta (subst1 new_rdx pred)) in let pred = mkNamedLambda pattern_id rdx_ty pred in let elim = let (kn, i) as ind, unfolded_c_ty = pf_reduce_to_quantified_ind gl c_ty in let sort = elimination_sort_of_goal gl in let elim = Indrec.lookup_eliminator ind sort in if dir = R2L then elim else (* taken from Coq's rewrite *) let elim = destConst elim in let mp,dp,l = repr_con (constant_of_kn (canonical_con elim)) in let l' = label_of_id (Nameops.add_suffix (id_of_label l) "_r") in let c1' = Global.constant_of_delta_kn (canonical_con (make_con mp dp l')) in mkConst c1' in let proof = mkApp (elim, [| rdx_ty; new_rdx; pred; p; rdx; c |]) in (* We check the proof is well typed *) let proof_ty = try Typing.type_of env sigma proof with _ -> raise PRtype_error in pp(lazy(str"pirrel_rewrite proof term of type: " ++ pr_constr proof_ty)); try refine_with ~first_goes_last:(not !ssroldreworder) ~with_evars:false (sigma, proof) gl with _ -> (* we generate a msg like: "Unable to find an instance for the variable" *) let c = Reductionops.nf_evar sigma c in let hd_ty, miss = match kind_of_term c with | App (hd, args) -> let hd_ty = Retyping.get_type_of env sigma hd in let names = let rec aux t = function 0 -> [] | n -> let t = Reductionops.whd_betadeltaiota env sigma t in match kind_of_type t with | ProdType (name, _, t) -> name :: aux t (n-1) | _ -> assert false in aux hd_ty (Array.length args) in hd_ty, Util.list_map_filter (fun (t, name) -> let evs = Util.Intset.elements (Evarutil.evars_of_term t) in let open_evs = List.filter (fun k -> InProp <> Retyping.get_sort_family_of env sigma (Evd.evar_concl (Evd.find sigma k))) evs in if open_evs <> [] then Some name else None) (List.combine (Array.to_list args) names) | _ -> anomaly "rewrite rule not an application" in errorstrm (Himsg.explain_refiner_error (Logic.UnresolvedBindings miss) ++ (Pp.fnl()++str"Rule's type:" ++ spc() ++ pr_constr hd_ty)) ;; let rwcltac cl rdx dir sr gl = let n, r_n = pf_abs_evars gl sr in let r_n' = pf_abs_cterm gl n r_n in let r' = subst_var pattern_id r_n' in let rdxt = Retyping.get_type_of (pf_env gl) (fst sr) rdx in let cvtac, rwtac = if closed0 r' then let env, sigma, c, c_eq = pf_env gl, fst sr, snd sr, build_coq_eq () in let c_ty = Typing.type_of env sigma c in match kind_of_type (Reductionops.whd_betadeltaiota env sigma c_ty) with | AtomicType(e, a) when eq_constr e c_eq -> let new_rdx = if dir = L2R then a.(2) else a.(1) in pirrel_rewrite cl rdx rdxt new_rdx dir sr c_ty, tclIDTAC | _ -> let cl' = mkApp (mkNamedLambda pattern_id rdxt cl, [|rdx|]) in convert_concl cl', rewritetac dir r' else let dc, r2 = decompose_lam_n n r' in let r3, _, r3t = try destCast r2 with _ -> errorstrm (str "no cast from " ++ pr_constr_pat (snd sr) ++ str " to " ++ pr_constr r2) in let cl' = mkNamedProd rule_id (compose_prod dc r3t) (lift 1 cl) in let cl'' = mkNamedProd pattern_id rdxt cl' in let itacs = [introid pattern_id; introid rule_id] in let cltac = clear [pattern_id; rule_id] in let rwtacs = [rewritetac dir (mkVar rule_id); cltac] in (apply_type cl'' [rdx; compose_lam dc r3], tclTHENLIST (itacs @ rwtacs)) in let cvtac' _ = try cvtac gl with | PRtype_error -> if occur_existential (pf_concl gl) then errorstrm (str "Rewriting impacts evars") else errorstrm (str "Dependent type error in rewrite of " ++ pf_pr_constr gl (mkNamedLambda pattern_id rdxt cl)) | Util.UserError _ as e -> raise e | e -> anomaly ("cvtac's exception: " ^ Printexc.to_string e); in tclTHEN cvtac' rwtac gl let prof_rwcltac = mk_profiler "rwrxtac.rwcltac";; let rwcltac cl rdx dir sr gl = prof_rwcltac.profile (rwcltac cl rdx dir sr) gl ;; let lz_coq_prod = let prod = lazy (build_prod ()) in fun () -> Lazy.force prod let lz_setoid_relation = let sdir = ["Classes"; "RelationClasses"] in let last_srel = ref (Environ.empty_env, None) in fun env -> match !last_srel with | env', srel when env' == env -> srel | _ -> let srel = try Some (coq_constant "Class_setoid" sdir "RewriteRelation") with _ -> None in last_srel := (env, srel); srel let ssr_is_setoid env = match lz_setoid_relation env with | None -> fun _ _ _ -> false | Some srel -> fun sigma r args -> Rewrite.is_applied_rewrite_relation env sigma [] (mkApp (r, args)) <> None let prof_rwxrtac_find_rule = mk_profiler "rwrxtac.find_rule";; let closed0_check cl p gl = if closed0 cl then errorstrm (str"No occurrence of redex "++pf_pr_constr gl p) let rwrxtac occ rdx_pat dir rule gl = let env = pf_env gl in let coq_prod = lz_coq_prod () in let is_setoid = ssr_is_setoid env in let r_sigma, rules = let rec loop d sigma r t0 rs red = let t = if red = 1 then Tacred.hnf_constr env sigma t0 else Reductionops.whd_betaiotazeta sigma t0 in match kind_of_term t with | Prod (_, xt, at) -> let ise, x = Evarutil.new_evar (create_evar_defs sigma) env xt in loop d ise (mkApp (r, [|x|])) (subst1 x at) rs 0 | App (pr, a) when pr = coq_prod.Coqlib.typ -> let sr = match kind_of_term (Tacred.hnf_constr env sigma r) with | App (c, ra) when c = coq_prod.Coqlib.intro -> fun i -> ra.(i + 1) | _ -> let ra = Array.append a [|r|] in function 1 -> mkApp (coq_prod.Coqlib.proj1, ra) | _ -> mkApp (coq_prod.Coqlib.proj2, ra) in if a.(0) = build_coq_True () then loop (converse_dir d) sigma (sr 2) a.(1) rs 0 else let sigma2, rs2 = loop d sigma (sr 2) a.(1) rs 0 in loop d sigma2 (sr 1) a.(0) rs2 0 | App (r_eq, a) when Hipattern.match_with_equality_type t != None -> let ind = destInd r_eq and rhs = array_last a in let np, ndep = Inductiveops.inductive_nargs env ind in let ind_ct = Inductiveops.type_of_constructors env ind in let lhs0 = last_arg (strip_prod_assum ind_ct.(0)) in let rdesc = match kind_of_term lhs0 with | Rel i -> let lhs = a.(np - i) in let lhs, rhs = if d = L2R then lhs, rhs else rhs, lhs in (* msgnl (str "RW: " ++ pr_rwdir d ++ str " " ++ pr_constr_pat r ++ str " : " ++ pr_constr_pat lhs ++ str " ~> " ++ pr_constr_pat rhs); *) d, r, lhs, rhs (* let l_i, r_i = if d = L2R then i, 1 - ndep else 1 - ndep, i in let lhs = a.(np - l_i) and rhs = a.(np - r_i) in let a' = Array.copy a in let _ = a'.(np - l_i) <- mkVar pattern_id in let r' = mkCast (r, DEFAULTcast, mkApp (r_eq, a')) in (d, r', lhs, rhs) *) | _ -> let lhs = substl (array_list_of_tl (Array.sub a 0 np)) lhs0 in let lhs, rhs = if d = R2L then lhs, rhs else rhs, lhs in let d' = if Array.length a = 1 then d else converse_dir d in d', r, lhs, rhs in sigma, rdesc :: rs | App (s_eq, a) when is_setoid sigma s_eq a -> let np = Array.length a and i = 3 - dir_org d in let lhs = a.(np - i) and rhs = a.(np + i - 3) in let a' = Array.copy a in let _ = a'.(np - i) <- mkVar pattern_id in let r' = mkCast (r, DEFAULTcast, mkApp (s_eq, a')) in sigma, (d, r', lhs, rhs) :: rs | _ -> if red = 0 then loop d sigma r t rs 1 else errorstrm (str "not a rewritable relation: " ++ pr_constr_pat t ++ spc() ++ str "in rule " ++ pr_constr_pat (snd rule)) in let sigma, r = rule in let t = Retyping.get_type_of env sigma r in loop dir sigma r t [] 0 in let find_rule rdx = let rec rwtac = function | [] -> errorstrm (str "pattern " ++ pr_constr_pat rdx ++ str " does not match " ++ pr_dir_side dir ++ str " of " ++ pr_constr_pat (snd rule)) | (d, r, lhs, rhs) :: rs -> try let ise = unify_HO env (create_evar_defs r_sigma) lhs rdx in if not (rw_progress rhs rdx ise) then raise NoMatch else d, (ise, Reductionops.nf_evar ise r) with _ -> rwtac rs in rwtac rules in let find_rule rdx = prof_rwxrtac_find_rule.profile find_rule rdx in let sigma0, env0, concl0 = project gl, pf_env gl, pf_concl gl in let find_R, conclude = match rdx_pat with | Some (_, (In_T _ | In_X_In_T _)) | None -> let upats_origin = dir, snd rule in let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) = let sigma, pat = mk_tpattern env sigma0 (sigma,r) (rw_progress rhs) d lhs in sigma, pats @ [pat] in let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in let find_R, end_R = mk_tpattern_matcher sigma0 occ ~upats_origin rpats in find_R ~k:(fun _ _ h -> mkRel h), fun cl -> let rdx,d,r = end_R () in closed0_check cl rdx gl; (d,r),rdx | Some(_, (T e | X_In_T (_,e) | E_As_X_In_T (e,_,_) | E_In_X_In_T (e,_,_))) -> let r = ref None in (fun env c h -> do_once r (fun () -> find_rule c, c); mkRel h), (fun concl -> closed0_check concl e gl; assert_done r) in let concl = eval_pattern env0 sigma0 concl0 rdx_pat occ find_R in let (d, r), rdx = conclude concl in rwcltac concl rdx d r gl ;; let prof_rwxrtac = mk_profiler "rwrxtac";; let rwrxtac occ rdx_pat dir rule gl = prof_rwxrtac.profile (rwrxtac occ rdx_pat dir rule) gl ;; (* Resolve forward reference *) let _ = ipat_rewritetac := fun occ dir c gl -> rwrxtac occ None dir (project gl, c) gl let rwargtac ist ((dir, mult), (((oclr, occ), grx), (kind, gt))) gl = let fail = ref false in let interp_rpattern ist gl gc = try interp_rpattern ist gl gc with _ when snd mult = May -> fail := true; project gl, T mkProp in let interp gc gl = try interp_term ist gl gc with _ when snd mult = May -> fail := true; (project gl, mkProp) in let rwtac gl = let rx = Option.map (interp_rpattern ist gl) grx in let t = interp gt gl in (match kind with | RWred sim -> simplintac occ rx sim | RWdef -> if dir = R2L then foldtac occ rx t else unfoldintac occ rx t gt | RWeq -> rwrxtac occ rx dir t) gl in let ctac = cleartac (interp_clr (oclr, (fst gt, snd (interp gt gl)))) in if !fail then ctac gl else tclTHEN (tclMULT mult rwtac) ctac gl (** Rewrite argument sequence *) (* type ssrrwargs = ssrrwarg list * ltacctx *) let pr_ssrrwargs _ _ _ (rwargs, _) = pr_list spc pr_rwarg rwargs ARGUMENT EXTEND ssrrwargs TYPED AS ssrrwarg list * ltacctx PRINTED BY pr_ssrrwargs | [ "Qed" ] -> [ anomaly "Grammar placeholder match" ] END let ssr_rw_syntax = ref true let _ = Goptions.declare_bool_option { Goptions.optsync = true; Goptions.optname = "ssreflect rewrite"; Goptions.optkey = ["SsrRewrite"]; Goptions.optread = (fun _ -> !ssr_rw_syntax); Goptions.optdepr = false; Goptions.optwrite = (fun b -> ssr_rw_syntax := b) } let test_ssr_rw_syntax = let test strm = if not !ssr_rw_syntax then raise Stream.Failure else if ssr_loaded () then () else match Compat.get_tok (Util.stream_nth 0 strm) with | Tok.KEYWORD key when List.mem key.[0] ['{'; '['; '/'] -> () | _ -> raise Stream.Failure in Gram.Entry.of_parser "test_ssr_rw_syntax" test GEXTEND Gram GLOBAL: ssrrwargs; ssrrwargs: [[ test_ssr_rw_syntax; a = LIST1 ssrrwarg -> a, rawltacctx ]]; END (** The "rewrite" tactic *) let ssrrewritetac ist rwargs = tclTHENLIST (List.map (rwargtac ist) rwargs) TACTIC EXTEND ssrrewrite | [ "rewrite" ssrrwargs(args) ssrclauses(clauses) ] -> [ let ist = get_ltacctx (snd args) in tclCLAUSES ist (ssrrewritetac ist (fst args)) clauses ] END (** The "unlock" tactic *) let pr_unlockarg (occ, t) = pr_occ occ ++ pr_term t let pr_ssrunlockarg _ _ _ = pr_unlockarg ARGUMENT EXTEND ssrunlockarg TYPED AS ssrocc * ssrterm PRINTED BY pr_ssrunlockarg | [ "{" ssrocc(occ) "}" ssrterm(t) ] -> [ occ, t ] | [ ssrterm(t) ] -> [ None, t ] END let pr_ssrunlockargs _ _ _ (args, _) = pr_list spc pr_unlockarg args ARGUMENT EXTEND ssrunlockargs TYPED AS ssrunlockarg list * ltacctx PRINTED BY pr_ssrunlockargs | [ ssrunlockarg_list(args) ] -> [ args, rawltacctx ] END let unfoldtac occ ko t kt gl = let cl, c = pf_fill_occ_term gl occ (fst (strip_unfold_term t kt)) in let cl' = subst1 (pf_unfoldn [(true, [1]), get_evalref c] gl c) cl in let f = if ko = None then Closure.betaiotazeta else Closure.betaiota in convert_concl (pf_reduce (Reductionops.clos_norm_flags f) gl cl') gl let unlocktac ist args = let utac (occ, gt) gl = unfoldtac occ occ (interp_term ist gl gt) (fst gt) gl in let locked = mkSsrConst "locked" in let key = mkSsrConst "master_key" in let ktacs = [ (fun gl -> unfoldtac None None (project gl,locked) '(' gl); simplest_newcase key ] in tclTHENLIST (List.map utac args @ ktacs) TACTIC EXTEND ssrunlock | [ "unlock" ssrunlockargs(args) ssrclauses(clauses) ] -> [ let ist = get_ltacctx (snd args) in tclCLAUSES ist (unlocktac ist (fst args)) clauses ] END (** 8. Forward chaining tactics (pose, set, have, suffice, wlog) *) (** Defined identifier *) type ssrfwdid = identifier let pr_ssrfwdid _ _ _ id = pr_spc () ++ pr_id id (* We use a primitive parser for the head identifier of forward *) (* tactis to avoid syntactic conflicts with basic Coq tactics. *) ARGUMENT EXTEND ssrfwdid TYPED AS ident PRINTED BY pr_ssrfwdid | [ "Qed" ] -> [ Util.anomaly "Grammar placeholder match" ] END let accept_ssrfwdid strm = match Compat.get_tok (stream_nth 0 strm) with | Tok.IDENT id -> accept_before_syms_or_any_id [":"; ":="; "("] strm | _ -> raise Stream.Failure let test_ssrfwdid = Gram.Entry.of_parser "test_ssrfwdid" accept_ssrfwdid GEXTEND Gram GLOBAL: ssrfwdid; ssrfwdid: [[ test_ssrfwdid; id = Prim.ident -> id ]]; END (** Definition value formatting *) (* We use an intermediate structure to correctly render the binder list *) (* abbreviations. We use a list of hints to extract the binders and *) (* base term from a term, for the two first levels of representation of *) (* of constr terms. *) type 'term ssrbind = | Bvar of name | Bdecl of name list * 'term | Bdef of name * 'term option * 'term | Bstruct of name | Bcast of 'term let pr_binder prl = function | Bvar x -> pr_name x | Bdecl (xs, t) -> str "(" ++ pr_list pr_spc pr_name xs ++ str " : " ++ prl t ++ str ")" | Bdef (x, None, v) -> str "(" ++ pr_name x ++ str " := " ++ prl v ++ str ")" | Bdef (x, Some t, v) -> str "(" ++ pr_name x ++ str " : " ++ prl t ++ str " := " ++ prl v ++ str ")" | Bstruct x -> str "{struct " ++ pr_name x ++ str "}" | Bcast t -> str ": " ++ prl t type 'term ssrbindval = 'term ssrbind list * 'term type ssrbindfmt = | BFvar | BFdecl of int (* #xs *) | BFcast (* final cast *) | BFdef of bool (* has cast? *) | BFrec of bool * bool (* has struct? * has cast? *) let rec mkBstruct i = function | Bvar x :: b -> if i = 0 then [Bstruct x] else mkBstruct (i - 1) b | Bdecl (xs, _) :: b -> let i' = i - List.length xs in if i' < 0 then [Bstruct (List.nth xs i)] else mkBstruct i' b | _ :: b -> mkBstruct i b | [] -> [] let rec format_local_binders h0 bl0 = match h0, bl0 with | BFvar :: h, LocalRawAssum ([_, x], _, _) :: bl -> Bvar x :: format_local_binders h bl | BFdecl _ :: h, LocalRawAssum (lxs, _, t) :: bl -> Bdecl (List.map snd lxs, t) :: format_local_binders h bl | BFdef false :: h, LocalRawDef ((_, x), v) :: bl -> Bdef (x, None, v) :: format_local_binders h bl | BFdef true :: h, LocalRawDef ((_, x), CCast (_, v, CastConv (_, t))) :: bl -> Bdef (x, Some t, v) :: format_local_binders h bl | _ -> [] let rec format_constr_expr h0 c0 = match h0, c0 with | BFvar :: h, CLambdaN (_, [[_, x], _, _], c) -> let bs, c' = format_constr_expr h c in Bvar x :: bs, c' | BFdecl _:: h, CLambdaN (_, [lxs, _, t], c) -> let bs, c' = format_constr_expr h c in Bdecl (List.map snd lxs, t) :: bs, c' | BFdef false :: h, CLetIn(_, (_, x), v, c) -> let bs, c' = format_constr_expr h c in Bdef (x, None, v) :: bs, c' | BFdef true :: h, CLetIn(_, (_, x), CCast (_, v, CastConv (_, t)), c) -> let bs, c' = format_constr_expr h c in Bdef (x, Some t, v) :: bs, c' | [BFcast], CCast (_, c, CastConv (_, t)) -> [Bcast t], c | BFrec (has_str, has_cast) :: h, CFix (_, _, [_, (Some locn, CStructRec), bl, t, c]) -> let bs = format_local_binders h bl in let bstr = if has_str then [Bstruct (Name (snd locn))] else [] in bs @ bstr @ (if has_cast then [Bcast t] else []), c | BFrec (_, has_cast) :: h, CCoFix (_, _, [_, bl, t, c]) -> format_local_binders h bl @ (if has_cast then [Bcast t] else []), c | _, c -> [], c let rec format_glob_decl h0 d0 = match h0, d0 with | BFvar :: h, (x, _, None, _) :: d -> Bvar x :: format_glob_decl h d | BFdecl 1 :: h, (x, _, None, t) :: d -> Bdecl ([x], t) :: format_glob_decl h d | BFdecl n :: h, (x, _, None, t) :: d when n > 1 -> begin match format_glob_decl (BFdecl (n - 1) :: h) d with | Bdecl (xs, _) :: bs -> Bdecl (x :: xs, t) :: bs | bs -> Bdecl ([x], t) :: bs end | BFdef false :: h, (x, _, Some v, _) :: d -> Bdef (x, None, v) :: format_glob_decl h d | BFdef true:: h, (x, _, Some (GCast (_, v, CastConv (_, t))), _) :: d -> Bdef (x, Some t, v) :: format_glob_decl h d | _, (x, _, None, t) :: d -> Bdecl ([x], t) :: format_glob_decl [] d | _, (x, _, Some v, _) :: d -> Bdef (x, None, v) :: format_glob_decl [] d | _, [] -> [] let rec format_glob_constr h0 c0 = match h0, c0 with | BFvar :: h, GLambda (_, x, _, _, c) -> let bs, c' = format_glob_constr h c in Bvar x :: bs, c' | BFdecl 1 :: h, GLambda (_, x, _, t, c) -> let bs, c' = format_glob_constr h c in Bdecl ([x], t) :: bs, c' | BFdecl n :: h, GLambda (_, x, _, t, c) when n > 1 -> begin match format_glob_constr (BFdecl (n - 1) :: h) c with | Bdecl (xs, _) :: bs, c' -> Bdecl (x :: xs, t) :: bs, c' | _ -> [Bdecl ([x], t)], c end | BFdef false :: h, GLetIn(_, x, v, c) -> let bs, c' = format_glob_constr h c in Bdef (x, None, v) :: bs, c' | BFdef true :: h, GLetIn(_, x, GCast (_, v, CastConv (_, t)), c) -> let bs, c' = format_glob_constr h c in Bdef (x, Some t, v) :: bs, c' | [BFcast], GCast (_, c, CastConv(_, t)) -> [Bcast t], c | BFrec (has_str, has_cast) :: h, GRec (_, f, _, bl, t, c) when Array.length c = 1 -> let bs = format_glob_decl h bl.(0) in let bstr = match has_str, f with | true, GFix ([|Some i, GStructRec|], _) -> mkBstruct i bs | _ -> [] in bs @ bstr @ (if has_cast then [Bcast t.(0)] else []), c.(0) | _, c -> [], c (** Forward chaining argument *) (* There are three kinds of forward definitions: *) (* - Hint: type only, cast to Type, may have proof hint. *) (* - Have: type option + value, no space before type *) (* - Pose: binders + value, space before binders. *) type ssrfwdkind = FwdHint of string | FwdHave | FwdPose type ssrfwdfmt = ssrfwdkind * ssrbindfmt list let pr_fwdkind = function FwdHint s -> str (s ^ " ") | _ -> str " :=" ++ spc () let pr_fwdfmt (fk, _ : ssrfwdfmt) = pr_fwdkind fk let wit_ssrfwdfmt, globwit_ssrfwdfmt, rawwit_ssrfwdfmt = add_genarg "ssrfwdfmt" pr_fwdfmt (* type ssrfwd = ssrfwdfmt * ssrterm *) let mkFwdVal fk c = ((fk, []), mk_term ' ' c), rawltacctx let mkssrFwdVal fk c = ((fk, []), (c,None)), rawltacctx let mkFwdCast fk loc t c = ((fk, [BFcast]), mk_term ' ' (CCast (loc, c, dC t))), rawltacctx let mkssrFwdCast fk loc t c = ((fk, [BFcast]), (c, Some t)), rawltacctx let mkFwdHint s t = let loc = constr_loc t in mkFwdCast (FwdHint s) loc t (CHole (loc, None)) let pr_gen_fwd prval prc prlc fk (bs, c) = let prc s = str s ++ spc () ++ prval prc prlc c in match fk, bs with | FwdHint s, [Bcast t] -> str s ++ spc () ++ prlc t | FwdHint s, _ -> prc (s ^ "(* typeof *)") | FwdHave, [Bcast t] -> str ":" ++ spc () ++ prlc t ++ prc " :=" | _, [] -> prc " :=" | _, _ -> spc () ++ pr_list spc (pr_binder prlc) bs ++ prc " :=" let pr_fwd_guarded prval prval' = function | ((fk, h), (_, (_, Some c))), _ -> pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c) | ((fk, h), (_, (c, None))), _ -> pr_gen_fwd prval' pr_glob_constr prl_glob_constr fk (format_glob_constr h c) let pr_unguarded prc prlc = prlc let pr_fwd = pr_fwd_guarded pr_unguarded pr_unguarded let pr_ssrfwd _ _ _ = pr_fwd ARGUMENT EXTEND ssrfwd TYPED AS (ssrfwdfmt * ssrterm) * ltacctx PRINTED BY pr_ssrfwd | [ ":=" lconstr(c) ] -> [ mkFwdVal FwdPose c ] | [ ":" lconstr (t) ":=" lconstr(c) ] -> [ mkFwdCast FwdPose loc t c ] END (** Independent parsing for binders *) (* The pose, pose fix, and pose cofix tactics use these internally to *) (* parse argument fragments. *) let pr_ssrbvar prc _ _ v = prc v ARGUMENT EXTEND ssrbvar TYPED AS constr PRINTED BY pr_ssrbvar | [ ident(id) ] -> [ mkCVar loc id ] | [ "_" ] -> [ CHole (loc, None) ] END let bvar_lname = function | CRef (Ident (loc, id)) -> loc, Name id | c -> constr_loc c, Anonymous let pr_ssrbinder prc _ _ (_, c) = prc c ARGUMENT EXTEND ssrbinder TYPED AS ssrfwdfmt * constr PRINTED BY pr_ssrbinder | [ ssrbvar(bv) ] -> [ let xloc, _ as x = bvar_lname bv in (FwdPose, [BFvar]), CLambdaN (loc,[[x],Default Explicit,CHole (xloc,None)],CHole (loc,None)) ] | [ "(" ssrbvar(bv) ")" ] -> [ let xloc, _ as x = bvar_lname bv in (FwdPose, [BFvar]), CLambdaN (loc,[[x],Default Explicit,CHole (xloc,None)],CHole (loc,None)) ] | [ "(" ssrbvar(bv) ":" lconstr(t) ")" ] -> [ let x = bvar_lname bv in (FwdPose, [BFdecl 1]), CLambdaN (loc, [[x], Default Explicit, t], CHole (loc, None)) ] | [ "(" ssrbvar(bv) ne_ssrbvar_list(bvs) ":" lconstr(t) ")" ] -> [ let xs = List.map bvar_lname (bv :: bvs) in let n = List.length xs in (FwdPose, [BFdecl n]), CLambdaN (loc, [xs, Default Explicit, t], CHole (loc, None)) ] | [ "(" ssrbvar(id) ":" lconstr(t) ":=" lconstr(v) ")" ] -> [ let loc' = Util.join_loc (constr_loc t) (constr_loc v) in let v' = CCast (loc', v, dC t) in (FwdPose,[BFdef true]), CLetIn (loc,bvar_lname id, v',CHole (loc,None)) ] | [ "(" ssrbvar(id) ":=" lconstr(v) ")" ] -> [ (FwdPose,[BFdef false]), CLetIn (loc,bvar_lname id, v,CHole (loc,None)) ] END GEXTEND Gram GLOBAL: ssrbinder; ssrbinder: [ [ ["of" | "&"]; c = operconstr LEVEL "99" -> (FwdPose, [BFvar]), CLambdaN (loc,[[loc,Anonymous],Default Explicit,c],CHole (loc,None)) ] ]; END let rec binders_fmts = function | ((_, h), _) :: bs -> h @ binders_fmts bs | _ -> [] let push_binders c2 bs = let loc2 = constr_loc c2 in let mkloc loc1 = Util.join_loc loc1 loc2 in let rec loop ty c = function | (_, CLambdaN (loc1, b, _)) :: bs when ty -> CProdN (mkloc loc1, b, loop ty c bs) | (_, CLambdaN (loc1, b, _)) :: bs -> CLambdaN (mkloc loc1, b, loop ty c bs) | (_, CLetIn (loc1, x, v, _)) :: bs -> CLetIn (mkloc loc1, x, v, loop ty c bs) | [] -> c | _ -> anomaly "binder not a lambda nor a let in" in match c2 with | CCast (x, ct, CastConv (y, cty)) -> (CCast (x, loop false ct bs, CastConv (y, loop true cty bs))) | ct -> loop false ct bs let rec fix_binders = function | (_, CLambdaN (_, [xs, _, t], _)) :: bs -> LocalRawAssum (xs, Default Explicit, t) :: fix_binders bs | (_, CLetIn (_, x, v, _)) :: bs -> LocalRawDef (x, v) :: fix_binders bs | _ -> [] let pr_ssrstruct _ _ _ = function | Some id -> str "{struct " ++ pr_id id ++ str "}" | None -> mt () ARGUMENT EXTEND ssrstruct TYPED AS ident option PRINTED BY pr_ssrstruct | [ "{" "struct" ident(id) "}" ] -> [ Some id ] | [ ] -> [ None ] END (** The "pose" tactic *) (* The plain pose form. *) let bind_fwd bs = function | ((fk, h), (ck, (rc, Some c))), ctx -> ((fk,binders_fmts bs @ h), (ck,(rc,Some (push_binders c bs)))), ctx | fwd -> fwd ARGUMENT EXTEND ssrposefwd TYPED AS ssrfwd PRINTED BY pr_ssrfwd | [ ssrbinder_list(bs) ssrfwd(fwd) ] -> [ bind_fwd bs fwd ] END (* The pose fix form. *) let pr_ssrfixfwd _ _ _ (id, fwd) = str " fix " ++ pr_id id ++ pr_fwd fwd let bvar_locid = function | CRef (Ident (loc, id)) -> loc, id | _ -> Util.error "Missing identifier after \"(co)fix\"" ARGUMENT EXTEND ssrfixfwd TYPED AS ident * ssrfwd PRINTED BY pr_ssrfixfwd | [ "fix" ssrbvar(bv) ssrbinder_list(bs) ssrstruct(sid) ssrfwd(fwd) ] -> [ let (_, id) as lid = bvar_locid bv in let ((fk, h), (ck, (rc, oc))), ctx = fwd in let c = Option.get oc in let has_cast, t', c' = match format_constr_expr h c with | [Bcast t'], c' -> true, t', c' | _ -> false, CHole (constr_loc c, None), c in let lb = fix_binders bs in let has_struct, i = let rec loop = function (l', Name id') :: _ when sid = Some id' -> true, (l', id') | [l', Name id'] when sid = None -> false, (l', id') | _ :: bn -> loop bn | [] -> Util.error "Bad structural argument" in loop (names_of_local_assums lb) in let h' = BFrec (has_struct, has_cast) :: binders_fmts bs in let fix = CFix (loc, lid, [lid, (Some i, CStructRec), lb, t', c']) in id, (((fk, h'), (ck, (rc, Some fix))), ctx) ] END (* The pose cofix form. *) let pr_ssrcofixfwd _ _ _ (id, fwd) = str " cofix " ++ pr_id id ++ pr_fwd fwd ARGUMENT EXTEND ssrcofixfwd TYPED AS ssrfixfwd PRINTED BY pr_ssrcofixfwd | [ "cofix" ssrbvar(bv) ssrbinder_list(bs) ssrfwd(fwd) ] -> [ let _, id as lid = bvar_locid bv in let ((fk, h), (ck, (rc, oc))), ctx = fwd in let c = Option.get oc in let has_cast, t', c' = match format_constr_expr h c with | [Bcast t'], c' -> true, t', c' | _ -> false, CHole (constr_loc c, None), c in let h' = BFrec (false, has_cast) :: binders_fmts bs in let cofix = CCoFix (loc, lid, [lid, fix_binders bs, t', c']) in id, (((fk, h'), (ck, (rc, Some cofix))), ctx) ] END let ssrposetac (id, ((_, t), ctx)) gl = posetac id (pf_abs_ssrterm (get_ltacctx ctx) gl t) gl let prof_ssrposetac = mk_profiler "ssrposetac";; let ssrposetac arg gl = prof_ssrposetac.profile (ssrposetac arg) gl;; TACTIC EXTEND ssrpose | [ "pose" ssrfixfwd(ffwd) ] -> [ ssrposetac ffwd ] | [ "pose" ssrcofixfwd(ffwd) ] -> [ ssrposetac ffwd ] | [ "pose" ssrfwdid(id) ssrposefwd(fwd) ] -> [ ssrposetac (id, fwd) ] END (** The "set" tactic *) (* type ssrsetfwd = ssrfwd * ssrdocc *) let guard_setrhs s i = s.[i] = '{' let pr_setrhs occ prc prlc c = if occ = nodocc then pr_guarded guard_setrhs prlc c else pr_docc occ ++ prc c let pr_fwd_guarded prval prval' = function | ((fk, h), (_, (_, Some c))), _ -> pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c) | ((fk, h), (_, (c, None))), _ -> pr_gen_fwd prval' pr_glob_constr prl_glob_constr fk (format_glob_constr h c) (* This does not print the type, it should be fixed... *) let pr_ssrsetfwd _ _ _ ((((fk,_),(t,_)),ctx), docc) = pr_gen_fwd (fun _ _ -> pr_cpattern) (fun _ -> mt()) (fun _ -> mt()) fk ([Bcast ()],t) ARGUMENT EXTEND ssrsetfwd TYPED AS ((ssrfwdfmt * (lcpattern * ssrterm option)) * ltacctx) * ssrdocc PRINTED BY pr_ssrsetfwd | [ ":" lconstr(t) ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> [ mkssrFwdCast FwdPose loc (mk_lterm t) c, mkocc occ ] | [ ":" lconstr(t) ":=" lcpattern(c) ] -> [ mkssrFwdCast FwdPose loc (mk_lterm t) c, nodocc ] | [ ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> [ mkssrFwdVal FwdPose c, mkocc occ ] | [ ":=" lcpattern(c) ] -> [ mkssrFwdVal FwdPose c, nodocc ] END let ssrsettac ist id (((_, (pat, pty)), _), (_, occ)) gl = let pat = interp_cpattern ist gl pat (Option.map snd pty) in let cl, sigma, env = pf_concl gl, project gl, pf_env gl in let c, cl = try fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1 with NoMatch -> redex_of_pattern pat, cl in if occur_existential c then errorstrm(str"The pattern"++spc()++ pr_constr_pat c++spc()++str"did not match and has holes."++spc()++ str"Did you mean pose?") else let c, cty = match kind_of_term c with | Cast(t, DEFAULTcast, ty) -> t, ty | _ -> c, pf_type_of gl c in let cl' = mkLetIn (Name id, c, cty, cl) in tclTHEN (convert_concl cl') (introid id) gl TACTIC EXTEND ssrset | [ "set" ssrfwdid(id) ssrsetfwd(fwd) ssrclauses(clauses) ] -> [ let (_, ctx), _ = fwd in let ist = get_ltacctx ctx in tclCLAUSES ist (ssrsettac ist id fwd) clauses ] END (** The "have" tactic *) (* type ssrhavefwd = ssrfwd * ssrhint *) let pr_ssrhavefwd _ _ prt (fwd, hint) = pr_fwd fwd ++ pr_hint prt hint ARGUMENT EXTEND ssrhavefwd TYPED AS ssrfwd * ssrhint PRINTED BY pr_ssrhavefwd | [ ":" lconstr(t) ssrhint(hint) ] -> [ mkFwdHint ":" t, hint ] | [ ":" lconstr(t) ":=" lconstr(c) ] -> [ mkFwdCast FwdHave loc t c, nohint ] | [ ":=" lconstr(c) ] -> [ mkFwdVal FwdHave c, nohint ] END let intro_id_to_binder = List.map (function | IpatId id -> let xloc, _ as x = bvar_lname (mkCVar dummy_loc id) in (FwdPose, [BFvar]), CLambdaN (dummy_loc, [[x], Default Explicit, CHole (xloc, None)], CHole (dummy_loc, None)) | _ -> anomaly "non-id accepted as binder") let binder_to_intro_id = List.map (function | (FwdPose, [BFvar]), CLambdaN (_,[ids,_,_],_) | (FwdPose, [BFdecl _]), CLambdaN (_,[ids,_,_],_) -> List.map (function (_, Name id) -> IpatId id | _ -> IpatAnon) ids | (FwdPose, [BFdef _]), CLetIn (_,(_,Name id),_,_) -> [IpatId id] | (FwdPose, [BFdef _]), CLetIn (_,(_,Anonymous),_,_) -> [IpatAnon] | _ -> anomaly "ssrbinder is not a binder") let pr_ssrhavefwdwbinders _ _ prt (hpats, (fwd, hint)) = pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint ARGUMENT EXTEND ssrhavefwdwbinders TYPED AS ssrhpats * (ssrfwd * ssrhint) PRINTED BY pr_ssrhavefwdwbinders | [ ssrhpats(pats) ssrbinder_list(bs) ssrhavefwd(fwd) ] -> [ let ((clr, pats), binders), simpl = pats in let allbs = intro_id_to_binder binders @ bs in let allbinders = binders @ List.flatten (binder_to_intro_id bs) in (((clr, pats), allbinders), simpl), (bind_fwd allbs (fst fwd), snd fwd) ] END (* Tactic. *) let havegentac ist t gl = let c = pf_abs_ssrterm ist gl t in apply_type (mkArrow (pf_type_of gl c) (pf_concl gl)) [c] gl let havetac ((((clr, pats), binders), simpl), ((((fk, _), t), ctx), hint)) suff namefst gl = let ist, concl = get_ltacctx ctx, pf_concl gl in let itac_c = introstac ~ist (IpatSimpl(clr,Nop) :: pats) in let itac, id, clr = introstac ~ist pats, tclIDTAC, cleartac clr in let binderstac n = let rec aux = function 0 -> [] | n -> IpatAnon :: aux (n-1) in tclTHEN (if binders <> [] then introstac ~ist (aux n) else tclIDTAC) (introstac ~ist binders) in let simpltac = introstac ~ist simpl in let hint = hinttac ist true hint in let cuttac t gl = basecuttac "ssr_have" t gl in let mkt t = mk_term ' ' t in let mkl t = (' ', (t, None)) in let interp t = pf_abs_ssrterm ist gl t in let interp_ty t = let a,b,_ = pf_interp_ty ist gl t in a, b in let ct, cty, hole, loc = match t with | _, (_, Some (CCast (loc, ct, CastConv (_, cty)))) -> mkt ct, mkt cty, mkt (mkCHole dummy_loc), loc | _, (_, Some ct) -> mkt ct, mkt (mkCHole dummy_loc), mkt (mkCHole dummy_loc), dummy_loc | _, (GCast (loc, ct, CastConv (_, cty)), None) -> mkl ct, mkl cty, mkl mkRHole, loc | _, (t, None) -> mkl t, mkl mkRHole, mkl mkRHole, dummy_loc in let cut, sol, itac1, itac2 = match fk, namefst, suff with | FwdHave, true, true -> errorstrm (str"Suff have does not accept a proof term") | FwdHave, false, true -> let cty = combineCG cty hole (mkCArrow loc) mkRArrow in let t = interp (combineCG ct cty (mkCCast loc) mkRCast) in let ty = pf_type_of gl t in let ctx, _ = decompose_prod_n 1 ty in let assert_is_conv gl = try convert_concl (compose_prod ctx concl) gl with _ -> errorstrm (str "Given proof term is not of type " ++ pr_constr (mkArrow (mkVar (id_of_string "_")) concl)) in ty, tclTHEN assert_is_conv (apply t), id, itac_c | FwdHave, false, false -> let t = interp (combineCG ct cty (mkCCast loc) mkRCast) in pf_type_of gl t, apply t, id, tclTHEN itac_c simpltac | _, true, true -> mkArrow (snd (interp_ty cty)) concl, hint, itac, clr | _, false, true -> mkArrow (snd (interp_ty cty)) concl, hint, id, itac_c | _, false, false -> let n, cty = interp_ty cty in cty, tclTHEN (binderstac n) hint, id, tclTHEN itac_c simpltac | _, true, false -> assert false in tclTHENS (cuttac cut) [ tclTHEN sol itac1; itac2 ] gl let prof_havetac = mk_profiler "havetac";; let havetac arg a b gl = prof_havetac.profile (havetac arg a b) gl;; TACTIC EXTEND ssrhave | [ "have" ssrhavefwdwbinders(fwd) ] -> [ havetac fwd false false ] END TACTIC EXTEND ssrhavesuff | [ "have" "suff" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> [ havetac (pats, fwd) true false ] END TACTIC EXTEND ssrhavesuffices | [ "have" "suffices" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> [ havetac (pats, fwd) true false ] END TACTIC EXTEND ssrsuffhave | [ "suff" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> [ havetac (pats, fwd) true true ] END TACTIC EXTEND ssrsufficeshave | [ "suffices" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> [ havetac (pats, fwd) true true ] END (** The "suffice" tactic *) ARGUMENT EXTEND ssrsufffwd TYPED AS ssrhpats * (ssrfwd * ssrhint) PRINTED BY pr_ssrhavefwdwbinders | [ ssrhpats(pats) ssrbinder_list(bs) ":" lconstr(t) ssrhint(hint) ] -> [ let ((clr, pats), binders), simpl = pats in let allbs = intro_id_to_binder binders @ bs in let allbinders = binders @ List.flatten (binder_to_intro_id bs) in let fwd = mkFwdHint ":" t in (((clr, pats), allbinders), simpl), (bind_fwd allbs fwd, hint) ] END let sufftac ((((clr, pats),binders),simpl), (((_, c), ctx), hint)) = let ist = get_ltacctx ctx in let htac = tclTHEN (introstac ~ist pats) (hinttac ist true hint) in let c = match c with | (a, (b, Some (CCast (_, _, CastConv (_, cty))))) -> a, (b, Some cty) | (a, (GCast (_, _, CastConv (_, cty)), None)) -> a, (cty, None) | _ -> anomaly "suff: ssr cast hole deleted by typecheck" in let ctac gl = basecuttac "ssr_suff" (pi2 (pf_interp_ty ist gl c)) gl in tclTHENS ctac [htac; tclTHEN (cleartac clr) (introstac ~ist (binders@simpl))] TACTIC EXTEND ssrsuff | [ "suff" ssrsufffwd(fwd) ] -> [ sufftac fwd ] END TACTIC EXTEND ssrsuffices | [ "suffices" ssrsufffwd(fwd) ] -> [ sufftac fwd ] END (** The "wlog" (Without Loss Of Generality) tactic *) (* type ssrwlogfwd = ssrwgen list * ssrfwd *) let pr_ssrwlogfwd _ _ _ (gens, t) = str ":" ++ pr_list mt pr_wgen gens ++ spc() ++ pr_fwd t ARGUMENT EXTEND ssrwlogfwd TYPED AS ssrwgen list * ssrfwd PRINTED BY pr_ssrwlogfwd | [ ":" ssrwgen_list(gens) "/" lconstr(t) ] -> [ gens, mkFwdHint "/" t] END let destProd_or_LetIn c = match kind_of_term c with | Prod (n,ty,c) -> (n,None,ty), c | LetIn (n,bo,ty,c) -> (n,Some bo,ty), c | _ -> raise (Invalid_argument "destProd_or_LetIn") let wlogtac (((clr0, pats),_),_) (gens, ((_, ct), ctx)) hint suff ghave gl = let ist = get_ltacctx ctx in let mkabs gen = abs_wgen false ist gl (fun x -> x) gen in let mkclr gen clrs = clr_of_wgen gen clrs in let mkpats = function | _, Some ((SsrHyp (_,x), _), _) -> fun pats -> IpatId x :: pats | _ -> fun x -> x in let ct = match ct with | (a, (b, Some (CCast (_, _, CastConv (_, cty))))) -> a, (b, Some cty) | (a, (GCast (_, _, CastConv (_, cty)), None)) -> a, (cty, None) | _ -> anomaly "wlog: ssr cast hole deleted by typecheck" in let cut_implies_goal = not (suff || ghave <> `NoGen) in let c, args, ct = let gens = List.filter (function _, Some _ -> true | _ -> false) gens in let concl = pf_concl gl in let c = mkProp in let c = if cut_implies_goal then mkArrow c concl else c in let args, c = List.fold_right mkabs gens ([],c) in let env, _ = List.fold_left (fun (env, c) _ -> let rd, c = destProd_or_LetIn c in Environ.push_rel rd env, c) (pf_env gl, c) gens in let sigma, ev = Evarutil.new_evar (project gl) env Term.mkProp in let k, _ = Term.destEvar ev in let fake_gl = { Evd.it = Goal.build k; Evd.sigma = sigma } in let _, ct, _ = pf_interp_ty ist fake_gl ct in let rec subst c g s = match kind_of_term c, g with | Prod(Anonymous,_,c), [] -> mkProd(Anonymous, subst_vars s ct, c) | Sort _, [] -> subst_vars s ct | LetIn(Name id as n,b,ty,c), _::g -> mkLetIn (n,b,ty,subst c g (id::s)) | Prod(Name id as n,ty,c), _::g -> mkProd (n,ty,subst c g (id::s)) | _ -> anomaly "SSR: wlog: subst" in subst c gens [], args, ct in let tacipat pats = introstac ~ist pats in let tacigens = tclTHEN (tclTHENLIST(List.rev(List.fold_right mkclr gens [cleartac clr0]))) (introstac ~ist (List.fold_right mkpats gens [])) in let hinttac = hinttac ist true hint in let cut_kind, fst_goal_tac, snd_goal_tac = match suff, ghave with | true, `NoGen -> "ssr_wlog", tclTHEN hinttac (tacipat pats), tacigens | false, `NoGen -> "ssr_wlog", hinttac, tclTHEN tacigens (tacipat pats) | true, `Gen _ -> assert false | false, `Gen id -> if gens = [] then errorstrm(str"gen have requires some generalizations"); let id, name_general_hyp, cleanup, pats = match id, pats with | None, (IpatId id as ip)::pats -> Some id, tacipat [ip], tclIDTAC, pats | None, _ -> None, tclIDTAC, tclIDTAC, pats | Some (Some id),_ -> Some id, introid id, tclIDTAC, pats | Some _,_ -> let id = mk_anon_id "tmp" gl in Some id, introid id, clear [id], pats in let tac_specialize = match id with | None -> tclIDTAC | Some id -> if pats = [] then tclIDTAC else let args = Array.of_list args in tclTHENS (basecuttac "ssr_have" ct) [apply (mkApp (mkVar id,args)); tclIDTAC] in "ssr_have", (if hint = nohint then tacigens else hinttac), tclTHENLIST [name_general_hyp; tac_specialize; tacipat pats; cleanup] in tclTHENS (basecuttac cut_kind c) [fst_goal_tac; snd_goal_tac] gl TACTIC EXTEND ssrwlog | [ "wlog" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> [ wlogtac pats fwd hint false `NoGen ] END TACTIC EXTEND ssrwlogs | [ "wlog" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> [ wlogtac pats fwd hint true `NoGen ] END TACTIC EXTEND ssrwlogss | [ "wlog" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> [ wlogtac pats fwd hint true `NoGen ] END TACTIC EXTEND ssrwithoutloss | [ "without" "loss" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> [ wlogtac pats fwd hint false `NoGen ] END TACTIC EXTEND ssrwithoutlosss | [ "without" "loss" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> [ wlogtac pats fwd hint true `NoGen ] END TACTIC EXTEND ssrwithoutlossss | [ "without" "loss" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> [ wlogtac pats fwd hint true `NoGen ] END (* Generally have *) let pr_idcomma _ _ _ = function | None -> mt() | Some None -> str"_, " | Some (Some id) -> pr_id id ++ str", " ARGUMENT EXTEND ssr_idcomma TYPED AS ident option option PRINTED BY pr_idcomma | [ ] -> [ None ] END let accept_idcomma strm = match Compat.get_tok (stream_nth 0 strm) with | Tok.IDENT _ | Tok.KEYWORD "_" -> accept_before_syms [","] strm | _ -> raise Stream.Failure let test_idcomma = Gram.Entry.of_parser "test_idcomma" accept_idcomma GEXTEND Gram GLOBAL: ssr_idcomma; ssr_idcomma: [ [ test_idcomma; ip = [ id = IDENT -> Some (id_of_string id) | "_" -> None ]; "," -> Some ip ] ]; END TACTIC EXTEND ssrgenhave | [ "gen" "have" ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> [ wlogtac pats fwd hint false (`Gen id) ] END TACTIC EXTEND ssrgenhave2 | [ "generally" "have" ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> [ wlogtac pats fwd hint false (`Gen id) ] END (** Canonical Structure alias *) let def_body : Vernacexpr.definition_expr Gram.Entry.e = Obj.magic (Grammar.Entry.find (Obj.magic gallina_ext) "vernac:def_body") in GEXTEND Gram GLOBAL: gallina_ext; gallina_ext: (* Canonical structure *) [[ IDENT "Canonical"; qid = Constr.global -> Vernacexpr.VernacCanonical (AN qid) | IDENT "Canonical"; ntn = Prim.by_notation -> Vernacexpr.VernacCanonical (ByNotation ntn) | IDENT "Canonical"; qid = Constr.global; d = def_body -> let s = coerce_reference_to_id qid in Vernacexpr.VernacDefinition ((Decl_kinds.Global,Decl_kinds.CanonicalStructure), (dummy_loc,s),(d ), (fun _ -> Recordops.declare_canonical_structure)) (*It seems there is not need for these: (* Canonical structure *) | IDENT "Canonical"; IDENT "Structure"; qid = Constr.global -> Vernacexpr.VernacCanonical (AN qid) | IDENT "Canonical"; IDENT "Structure"; ntn = Prim.by_notation -> Vernacexpr.VernacCanonical (ByNotation ntn) | IDENT "Canonical"; IDENT "Structure"; qid = Constr.global; d = def_body -> let s = coerce_reference_to_id qid in Vernacexpr.VernacDefinition ((Decl_kinds.Global,false,Decl_kinds.CanonicalStructure), (dummy_loc,s),(d ), (fun _ -> Recordops.declare_canonical_structure)) *) ]]; END (** 9. Keyword compatibility fixes. *) (* Coq v8.1 notation uses "by" and "of" quasi-keywords, i.e., reserved *) (* identifiers used as keywords. This is incompatible with ssreflect.v *) (* which makes "by" and "of" true keywords, because of technicalities *) (* in the internal lexer-parser API of Coq. We patch this here by *) (* adding new parsing rules that recognize the new keywords. *) (* To make matters worse, the Coq grammar for tactics fails to *) (* export the non-terminals we need to patch. Fortunately, the CamlP5 *) (* API provides a backdoor access (with loads of Obj.magic trickery). *) (* Coq v8.3 defines "by" as a keyword, some hacks are not needed any *) (* longer and thus comment out. Such comments are marked with v8.3 *) let tac_ent = List.fold_left Grammar.Entry.find (Obj.magic simple_tactic) in let hypident_ent = tac_ent ["clause_dft_all"; "in_clause"; "hypident_occ"; "hypident"] in let id_or_meta : Obj.t Gram.Entry.e = Obj.magic (Grammar.Entry.find hypident_ent "id_or_meta") in (*v8.3 let by_tactic : raw_tactic_expr Gram.Entry.e = Obj.magic (tac_ent ["by_tactic"]) in let opt_by_tactic : raw_tactic_expr option Gram.Entry.e = Obj.magic (tac_ent ["opt_by_tactic"]) in *) let hypident : (Obj.t * hyp_location_flag) Gram.Entry.e = Obj.magic hypident_ent in GEXTEND Gram GLOBAL: (*v8.3 opt_by_tactic by_tactic*) hypident; (*v8.3 opt_by_tactic: [ [ "by"; tac = tactic_expr LEVEL "3" -> Some tac ] ]; by_tactic: [ [ "by"; tac = tactic_expr LEVEL "3" -> TacComplete tac ] ]; *) hypident: [ [ "("; IDENT "type"; "of"; id = id_or_meta; ")" -> id, InHypTypeOnly | "("; IDENT "value"; "of"; id = id_or_meta; ")" -> id, InHypValueOnly ] ]; END GEXTEND Gram GLOBAL: hloc (*v8.3 by_arg_tac*); hloc: [ [ "in"; "("; "Type"; "of"; id = ident; ")" -> HypLocation ((Util.dummy_loc,id), InHypTypeOnly) | "in"; "("; IDENT "Value"; "of"; id = ident; ")" -> HypLocation ((Util.dummy_loc,id), InHypValueOnly) ] ]; (*v8.3 by_arg_tac: [ [ "by"; tac = tactic_expr LEVEL "3" -> Some tac ] ]; *) END (*v8.3 open Rewrite let pr_ssrrelattr prc _ _ (a, c) = pr_id a ++ str " proved by " ++ prc c ARGUMENT EXTEND ssrrelattr TYPED AS ident * constr PRINTED BY pr_ssrrelattr [ ident(a) "proved" "by" constr(c) ] -> [ a, c ] END GEXTEND Gram GLOBAL: ssrrelattr; ssrrelattr: [[ a = ident; IDENT "proved"; "by"; c = constr -> a, c ]]; END let rec ssr_add_relation n d b deq pf_refl pf_sym pf_trans = function | [] -> Rewrite.declare_relation ~binders:b d deq n pf_refl pf_sym pf_trans | (aid, c) :: al -> match string_of_id aid with | "reflexivity" when pf_refl = None -> ssr_add_relation n d b deq (Some c) pf_sym pf_trans al | "symmetry" when pf_sym = None -> ssr_add_relation n d b deq pf_refl (Some c) pf_trans al | "transitivity" when pf_trans = None -> ssr_add_relation n d b deq pf_refl pf_sym (Some c) al | a -> Util.error ("bad attribute \"" ^ a ^ "\" in Add Relation") VERNAC COMMAND EXTEND SsrAddRelation [ "Add" "Relation" constr(d) constr(deq) ssrrelattr_list(al) "as" ident(n) ] -> [ ssr_add_relation n d [] deq None None None al ] END VERNAC COMMAND EXTEND SsrAddParametricRelation [ "Add" "Parametric" "Relation" binders_let(b) ":" constr(d) constr(deq) ssrrelattr_list(al) "as" ident(n) ] -> [ ssr_add_relation n d b deq None None None al ] END *) (* We wipe out all the keywords generated by the grammar rules we defined. *) (* The user is supposed to Require Import ssreflect or Require ssreflect *) (* and Import ssreflectSsrSyntax to obtain these keywords and as a *) (* consequence the extended ssreflect grammar. *) let () = Lexer.unfreeze frozen_lexer ;; (* vim: set filetype=ocaml foldmethod=marker: *) ssreflect-1.5/README0000644000175000017500000000341312175455021013220 0ustar garesgares THE SSREFLECT EXTENSION FOR THE COQ SYSTEM ------------------------------------------ INSTALLATION ============ See the file INSTALL for installation procedure. DOCUMENTATION ============= The documentation of the ssreflect tactics, a brief description of the libraries contained in the theories/ directory of the archive, and a detailed list of the changes made in the last release is available as an Inria Research Report at http://hal.inria.fr/inria-00258384 This online version may be updated between two successive releases of ssreflect. TUTORIAL ======== A brief tutorial of the ssreflect extension can be found in the doc/ directory. The doc/tutorial.v file contains the code of the examples presented in this document. This tutorial is also available as an Inria Technical Report at http://hal.inria.fr/inria-00407778 This online version may be updated between two successive releases of ssreflect. AVAILABILITY ============ Ssreflect files are available at: http://www.msr-inria.inria.fr/Projects/math-components THE SSREFLECT DISCUSSION LIST ============================= The ssreflect list is meant to be a standard way to discuss questions about the ssreflect extension and its use. To subscribe to ssreflect@msr-inria.inria.fr please send an email to sympa@msr-inria.inria.fr, whose title is "subscribe ssreflect". LICENSING ========= This program is free software; you can redistribute it and/or modify it under the terms of the CeCILL B FREE SOFTWARE LICENSE. You should have received a copy of the CeCILL B License with this Kit, in the file named "CeCILL-B". If not, visit http://www.cecill.info ssreflect-1.5/Make0000644000175000017500000000152612175453667013161 0ustar garesgares### Uncomment for static linking ## #-custom "$(COQBIN)coqmktop -coqlib `$(COQBIN)coqtop -where` -opt -o bin/ssrcoq src/ssrmatching.cmx src/ssreflect.cmx" "src/ssrmatching.cmx src/ssreflect.cmx" bin/ssrcoq #-custom "$(COQBIN)coqmktop -coqlib `$(COQBIN)coqtop -where` -o bin/ssrcoq.byte src/ssrmatching.cmo src/ssreflect.cmo" "src/ssrmatching.cmo src/ssreflect.cmo" bin/ssrcoq.byte #-custom "$(SSRCOQ) $(COQFLAGS) -compile $*" "%.v $(SSRCOQ)" "%.vo" #SSRCOQ = bin/ssrcoq ## ## What follows should be left untouched by the final user of ssreflect -R theories Ssreflect -I src CAMLP4OPTIONS = "-loc loc" src/ssrmatching.mli src/ssrmatching.ml4 src/ssreflect.ml4 theories/ssreflect.mllib theories/ssrmatching.v theories/ssreflect.v theories/ssrfun.v theories/ssrbool.v theories/eqtype.v theories/ssrnat.v theories/seq.v theories/choice.v theories/fintype.v ssreflect-1.5/CeCILL-B0000644000175000017500000005262312175455021013444 0ustar garesgaresCeCILL-B FREE SOFTWARE LICENSE AGREEMENT Notice This Agreement is a Free Software license agreement that is the result of discussions between its authors in order to ensure compliance with the two main principles guiding its drafting: * firstly, compliance with the principles governing the distribution of Free Software: access to source code, broad rights granted to users, * secondly, the election of a governing law, French law, with which it is conformant, both as regards the law of torts and intellectual property law, and the protection that it offers to both authors and holders of the economic rights over software. The authors of the CeCILL-B (for Ce[a] C[nrs] I[nria] L[ogiciel] L[ibre]) license are: Commissariat à l'Energie Atomique - CEA, a public scientific, technical and industrial research establishment, having its principal place of business at 25 rue Leblanc, immeuble Le Ponant D, 75015 Paris, France. Centre National de la Recherche Scientifique - CNRS, a public scientific and technological establishment, having its principal place of business at 3 rue Michel-Ange, 75794 Paris cedex 16, France. Institut National de Recherche en Informatique et en Automatique - INRIA, a public scientific and technological establishment, having its principal place of business at Domaine de Voluceau, Rocquencourt, BP 105, 78153 Le Chesnay cedex, France. Preamble This Agreement is an open source software license intended to give users significant freedom to modify and redistribute the software licensed hereunder. The exercising of this freedom is conditional upon a strong obligation of giving credits for everybody that distributes a software incorporating a software ruled by the current license so as all contributions to be properly identified and acknowledged. In consideration of access to the source code and the rights to copy, modify and redistribute granted by the license, users are provided only with a limited warranty and the software's author, the holder of the economic rights, and the successive licensors only have limited liability. In this respect, the risks associated with loading, using, modifying and/or developing or reproducing the software by the user are brought to the user's attention, given its Free Software status, which may make it complicated to use, with the result that its use is reserved for developers and experienced professionals having in-depth computer knowledge. Users are therefore encouraged to load and test the suitability of the software as regards their requirements in conditions enabling the security of their systems and/or data to be ensured and, more generally, to use and operate it in the same conditions of security. This Agreement may be freely reproduced and published, provided it is not altered, and that no provisions are either added or removed herefrom. This Agreement may apply to any or all software for which the holder of the economic rights decides to submit the use thereof to its provisions. Article 1 - DEFINITIONS For the purpose of this Agreement, when the following expressions commence with a capital letter, they shall have the following meaning: Agreement: means this license agreement, and its possible subsequent versions and annexes. Software: means the software in its Object Code and/or Source Code form and, where applicable, its documentation, "as is" when the Licensee accepts the Agreement. Initial Software: means the Software in its Source Code and possibly its Object Code form and, where applicable, its documentation, "as is" when it is first distributed under the terms and conditions of the Agreement. Modified Software: means the Software modified by at least one Contribution. Source Code: means all the Software's instructions and program lines to which access is required so as to modify the Software. Object Code: means the binary files originating from the compilation of the Source Code. Holder: means the holder(s) of the economic rights over the Initial Software. Licensee: means the Software user(s) having accepted the Agreement. Contributor: means a Licensee having made at least one Contribution. Licensor: means the Holder, or any other individual or legal entity, who distributes the Software under the Agreement. Contribution: means any or all modifications, corrections, translations, adaptations and/or new functions integrated into the Software by any or all Contributors, as well as any or all Internal Modules. Module: means a set of sources files including their documentation that enables supplementary functions or services in addition to those offered by the Software. External Module: means any or all Modules, not derived from the Software, so that this Module and the Software run in separate address spaces, with one calling the other when they are run. Internal Module: means any or all Module, connected to the Software so that they both execute in the same address space. Parties: mean both the Licensee and the Licensor. These expressions may be used both in singular and plural form. Article 2 - PURPOSE The purpose of the Agreement is the grant by the Licensor to the Licensee of a non-exclusive, transferable and worldwide license for the Software as set forth in Article 5 hereinafter for the whole term of the protection granted by the rights over said Software. Article 3 - ACCEPTANCE 3.1 The Licensee shall be deemed as having accepted the terms and conditions of this Agreement upon the occurrence of the first of the following events: * (i) loading the Software by any or all means, notably, by downloading from a remote server, or by loading from a physical medium; * (ii) the first time the Licensee exercises any of the rights granted hereunder. 3.2 One copy of the Agreement, containing a notice relating to the characteristics of the Software, to the limited warranty, and to the fact that its use is restricted to experienced users has been provided to the Licensee prior to its acceptance as set forth in Article 3.1 hereinabove, and the Licensee hereby acknowledges that it has read and understood it. Article 4 - EFFECTIVE DATE AND TERM 4.1 EFFECTIVE DATE The Agreement shall become effective on the date when it is accepted by the Licensee as set forth in Article 3.1. 4.2 TERM The Agreement shall remain in force for the entire legal term of protection of the economic rights over the Software. Article 5 - SCOPE OF RIGHTS GRANTED The Licensor hereby grants to the Licensee, who accepts, the following rights over the Software for any or all use, and for the term of the Agreement, on the basis of the terms and conditions set forth hereinafter. Besides, if the Licensor owns or comes to own one or more patents protecting all or part of the functions of the Software or of its components, the Licensor undertakes not to enforce the rights granted by these patents against successive Licensees using, exploiting or modifying the Software. If these patents are transferred, the Licensor undertakes to have the transferees subscribe to the obligations set forth in this paragraph. 5.1 RIGHT OF USE The Licensee is authorized to use the Software, without any limitation as to its fields of application, with it being hereinafter specified that this comprises: 1. permanent or temporary reproduction of all or part of the Software by any or all means and in any or all form. 2. loading, displaying, running, or storing the Software on any or all medium. 3. entitlement to observe, study or test its operation so as to determine the ideas and principles behind any or all constituent elements of said Software. This shall apply when the Licensee carries out any or all loading, displaying, running, transmission or storage operation as regards the Software, that it is entitled to carry out hereunder. 5.2 ENTITLEMENT TO MAKE CONTRIBUTIONS The right to make Contributions includes the right to translate, adapt, arrange, or make any or all modifications to the Software, and the right to reproduce the resulting software. The Licensee is authorized to make any or all Contributions to the Software provided that it includes an explicit notice that it is the author of said Contribution and indicates the date of the creation thereof. 5.3 RIGHT OF DISTRIBUTION In particular, the right of distribution includes the right to publish, transmit and communicate the Software to the general public on any or all medium, and by any or all means, and the right to market, either in consideration of a fee, or free of charge, one or more copies of the Software by any means. The Licensee is further authorized to distribute copies of the modified or unmodified Software to third parties according to the terms and conditions set forth hereinafter. 5.3.1 DISTRIBUTION OF SOFTWARE WITHOUT MODIFICATION The Licensee is authorized to distribute true copies of the Software in Source Code or Object Code form, provided that said distribution complies with all the provisions of the Agreement and is accompanied by: 1. a copy of the Agreement, 2. a notice relating to the limitation of both the Licensor's warranty and liability as set forth in Articles 8 and 9, and that, in the event that only the Object Code of the Software is redistributed, the Licensee allows effective access to the full Source Code of the Software at a minimum during the entire period of its distribution of the Software, it being understood that the additional cost of acquiring the Source Code shall not exceed the cost of transferring the data. 5.3.2 DISTRIBUTION OF MODIFIED SOFTWARE If the Licensee makes any Contribution to the Software, the resulting Modified Software may be distributed under a license agreement other than this Agreement subject to compliance with the provisions of Article 5.3.4. 5.3.3 DISTRIBUTION OF EXTERNAL MODULES When the Licensee has developed an External Module, the terms and conditions of this Agreement do not apply to said External Module, that may be distributed under a separate license agreement. 5.3.4 CREDITS Any Licensee who may distribute a Modified Software hereby expressly agrees to: 1. indicate in the related documentation that it is based on the Software licensed hereunder, and reproduce the intellectual property notice for the Software, 2. ensure that written indications of the Software intended use, intellectual property notice and license hereunder are included in easily accessible format from the Modified Software interface, 3. mention, on a freely accessible website describing the Modified Software, at least throughout the distribution term thereof, that it is based on the Software licensed hereunder, and reproduce the Software intellectual property notice, 4. where it is distributed to a third party that may distribute a Modified Software without having to make its source code available, make its best efforts to ensure that said third party agrees to comply with the obligations set forth in this Article . If the Software, whether or not modified, is distributed with an External Module designed for use in connection with the Software, the Licensee shall submit said External Module to the foregoing obligations. 5.3.5 COMPATIBILITY WITH THE CeCILL AND CeCILL-C LICENSES Where a Modified Software contains a Contribution subject to the CeCILL license, the provisions set forth in Article 5.3.4 shall be optional. A Modified Software may be distributed under the CeCILL-C license. In such a case the provisions set forth in Article 5.3.4 shall be optional. Article 6 - INTELLECTUAL PROPERTY 6.1 OVER THE INITIAL SOFTWARE The Holder owns the economic rights over the Initial Software. Any or all use of the Initial Software is subject to compliance with the terms and conditions under which the Holder has elected to distribute its work and no one shall be entitled to modify the terms and conditions for the distribution of said Initial Software. The Holder undertakes that the Initial Software will remain ruled at least by this Agreement, for the duration set forth in Article 4.2. 6.2 OVER THE CONTRIBUTIONS The Licensee who develops a Contribution is the owner of the intellectual property rights over this Contribution as defined by applicable law. 6.3 OVER THE EXTERNAL MODULES The Licensee who develops an External Module is the owner of the intellectual property rights over this External Module as defined by applicable law and is free to choose the type of agreement that shall govern its distribution. 6.4 JOINT PROVISIONS The Licensee expressly undertakes: 1. not to remove, or modify, in any manner, the intellectual property notices attached to the Software; 2. to reproduce said notices, in an identical manner, in the copies of the Software modified or not. The Licensee undertakes not to directly or indirectly infringe the intellectual property rights of the Holder and/or Contributors on the Software and to take, where applicable, vis-à-vis its staff, any and all measures required to ensure respect of said intellectual property rights of the Holder and/or Contributors. Article 7 - RELATED SERVICES 7.1 Under no circumstances shall the Agreement oblige the Licensor to provide technical assistance or maintenance services for the Software. However, the Licensor is entitled to offer this type of services. The terms and conditions of such technical assistance, and/or such maintenance, shall be set forth in a separate instrument. Only the Licensor offering said maintenance and/or technical assistance services shall incur liability therefor. 7.2 Similarly, any Licensor is entitled to offer to its licensees, under its sole responsibility, a warranty, that shall only be binding upon itself, for the redistribution of the Software and/or the Modified Software, under terms and conditions that it is free to decide. Said warranty, and the financial terms and conditions of its application, shall be subject of a separate instrument executed between the Licensor and the Licensee. Article 8 - LIABILITY 8.1 Subject to the provisions of Article 8.2, the Licensee shall be entitled to claim compensation for any direct loss it may have suffered from the Software as a result of a fault on the part of the relevant Licensor, subject to providing evidence thereof. 8.2 The Licensor's liability is limited to the commitments made under this Agreement and shall not be incurred as a result of in particular: (i) loss due the Licensee's total or partial failure to fulfill its obligations, (ii) direct or consequential loss that is suffered by the Licensee due to the use or performance of the Software, and (iii) more generally, any consequential loss. In particular the Parties expressly agree that any or all pecuniary or business loss (i.e. loss of data, loss of profits, operating loss, loss of customers or orders, opportunity cost, any disturbance to business activities) or any or all legal proceedings instituted against the Licensee by a third party, shall constitute consequential loss and shall not provide entitlement to any or all compensation from the Licensor. Article 9 - WARRANTY 9.1 The Licensee acknowledges that the scientific and technical state-of-the-art when the Software was distributed did not enable all possible uses to be tested and verified, nor for the presence of possible defects to be detected. In this respect, the Licensee's attention has been drawn to the risks associated with loading, using, modifying and/or developing and reproducing the Software which are reserved for experienced users. The Licensee shall be responsible for verifying, by any or all means, the suitability of the product for its requirements, its good working order, and for ensuring that it shall not cause damage to either persons or properties. 9.2 The Licensor hereby represents, in good faith, that it is entitled to grant all the rights over the Software (including in particular the rights set forth in Article 5). 9.3 The Licensee acknowledges that the Software is supplied "as is" by the Licensor without any other express or tacit warranty, other than that provided for in Article 9.2 and, in particular, without any warranty as to its commercial value, its secured, safe, innovative or relevant nature. Specifically, the Licensor does not warrant that the Software is free from any error, that it will operate without interruption, that it will be compatible with the Licensee's own equipment and software configuration, nor that it will meet the Licensee's requirements. 9.4 The Licensor does not either expressly or tacitly warrant that the Software does not infringe any third party intellectual property right relating to a patent, software or any other property right. Therefore, the Licensor disclaims any and all liability towards the Licensee arising out of any or all proceedings for infringement that may be instituted in respect of the use, modification and redistribution of the Software. Nevertheless, should such proceedings be instituted against the Licensee, the Licensor shall provide it with technical and legal assistance for its defense. Such technical and legal assistance shall be decided on a case-by-case basis between the relevant Licensor and the Licensee pursuant to a memorandum of understanding. The Licensor disclaims any and all liability as regards the Licensee's use of the name of the Software. No warranty is given as regards the existence of prior rights over the name of the Software or as regards the existence of a trademark. Article 10 - TERMINATION 10.1 In the event of a breach by the Licensee of its obligations hereunder, the Licensor may automatically terminate this Agreement thirty (30) days after notice has been sent to the Licensee and has remained ineffective. 10.2 A Licensee whose Agreement is terminated shall no longer be authorized to use, modify or distribute the Software. However, any licenses that it may have granted prior to termination of the Agreement shall remain valid subject to their having been granted in compliance with the terms and conditions hereof. Article 11 - MISCELLANEOUS 11.1 EXCUSABLE EVENTS Neither Party shall be liable for any or all delay, or failure to perform the Agreement, that may be attributable to an event of force majeure, an act of God or an outside cause, such as defective functioning or interruptions of the electricity or telecommunications networks, network paralysis following a virus attack, intervention by government authorities, natural disasters, water damage, earthquakes, fire, explosions, strikes and labor unrest, war, etc. 11.2 Any failure by either Party, on one or more occasions, to invoke one or more of the provisions hereof, shall under no circumstances be interpreted as being a waiver by the interested Party of its right to invoke said provision(s) subsequently. 11.3 The Agreement cancels and replaces any or all previous agreements, whether written or oral, between the Parties and having the same purpose, and constitutes the entirety of the agreement between said Parties concerning said purpose. No supplement or modification to the terms and conditions hereof shall be effective as between the Parties unless it is made in writing and signed by their duly authorized representatives. 11.4 In the event that one or more of the provisions hereof were to conflict with a current or future applicable act or legislative text, said act or legislative text shall prevail, and the Parties shall make the necessary amendments so as to comply with said act or legislative text. All other provisions shall remain effective. Similarly, invalidity of a provision of the Agreement, for any reason whatsoever, shall not cause the Agreement as a whole to be invalid. 11.5 LANGUAGE The Agreement is drafted in both French and English and both versions are deemed authentic. Article 12 - NEW VERSIONS OF THE AGREEMENT 12.1 Any person is authorized to duplicate and distribute copies of this Agreement. 12.2 So as to ensure coherence, the wording of this Agreement is protected and may only be modified by the authors of the License, who reserve the right to periodically publish updates or new versions of the Agreement, each with a separate number. These subsequent versions may address new issues encountered by Free Software. 12.3 Any Software distributed under a given version of the Agreement may only be subsequently distributed under the same version of the Agreement or a subsequent version. Article 13 - GOVERNING LAW AND JURISDICTION 13.1 The Agreement is governed by French law. The Parties agree to endeavor to seek an amicable solution to any disagreements or disputes that may arise during the performance of the Agreement. 13.2 Failing an amicable solution within two (2) months as from their occurrence, and unless emergency proceedings are necessary, the disagreements or disputes shall be referred to the Paris Courts having jurisdiction, by the more diligent Party. Version 1.0 dated 2006-09-05. ssreflect-1.5/theories/0000755000175000017500000000000012175455021014161 5ustar garesgaresssreflect-1.5/theories/seq.v0000644000175000017500000025507212175455021015153 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import ssreflect ssrfun ssrbool eqtype ssrnat. (******************************************************************************) (* The seq type is the ssreflect type for sequences; it is an alias for the *) (* standard Coq list type. The ssreflect library equips it with many *) (* operations, as well as eqType and predType (and, later, choiceType) *) (* structures. The operations are geared towards reflection: they generally *) (* expect and provide boolean predicates, e.g., the membership predicate *) (* expects an eqType. To avoid any confusion we do not Import the Coq List *) (* module. *) (* As there is no true subtyping in Coq, we don't use a type for non-empty *) (* sequences; rather, we pass explicitly the head and tail of the sequence. *) (* The empty sequence is especially bothersome for subscripting, since it *) (* forces us to pass a default value. This default value can often be hidden *) (* by a notation. *) (* Here is the list of seq operations: *) (* ** Constructors: *) (* seq T == the type of sequences with items of type T *) (* bitseq == seq bool *) (* [::], nil, Nil T == the empty sequence (of type T) *) (* x :: s, cons x s, Cons T x s == the sequence x followed by s (of type T) *) (* [:: x] == the singleton sequence *) (* [:: x_0; ...; x_n] == the explicit sequence of the x_i *) (* [:: x_0, ..., x_n & s] == the sequence of the x_i, followed by s *) (* rcons s x == the sequence s, followed by x *) (* All of the above, except rcons, can be used in patterns. We define a view *) (* lastP and and induction principle last_ind that can be used to decompose *) (* or traverse a sequence in a right to left order. The view lemma lastP has *) (* a dependent family type, so the ssreflect tactic case/lastP: p => [|p' x] *) (* will generate two subgoals in which p has been replaced by [::] and by *) (* rcons p' x, respectively. *) (* ** Factories: *) (* nseq n x == a sequence of n x's *) (* ncons n x s == a sequence of n x's, followed by s *) (* seqn n x_0 ... x_n-1 == the sequence of the x_i (can be partially applied) *) (* iota m n == the sequence m, m + 1, ..., m + n - 1 *) (* mkseq f n == the sequence f 0, f 1, ..., f (n - 1) *) (* ** Sequential access: *) (* head x0 s == the head (zero'th item) of s if s is non-empty, else x0 *) (* ohead s == None if s is empty, else Some x where x is the head of s *) (* behead s == s minus its head, i.e., s' if s = x :: s', else [::] *) (* last x s == the last element of x :: s (which is non-empty) *) (* belast x s == x :: s minus its last item *) (* ** Dimensions: *) (* size s == the number of items (length) in s *) (* shape ss == the sequence of sizes of the items of the sequence of *) (* sequences ss *) (* ** Random access: *) (* nth x0 s i == the item i of s (numbered from 0), or x0 if s does *) (* not have at least i+1 items (i.e., size x <= i) *) (* s`_i == standard notation for nth x0 s i for a default x0, *) (* e.g., 0 for rings. *) (* set_nth x0 s i y == s where item i has been changed to y; if s does not *) (* have an item i it is first padded with copieds of x0 *) (* to size i+1. *) (* incr_nth s i == the nat sequence s with item i incremented (s is *) (* first padded with 0's to size i+1, if needed). *) (* ** Predicates: *) (* nilp s == s is [::] *) (* := (size s == 0) *) (* x \in s == x appears in s (this requires an eqType for T) *) (* index x s == the first index at which x appears in s, or size s if *) (* x \notin s *) (* has p s == the (applicative, boolean) predicate p holds for some *) (* item in s *) (* all p s == p holds for all items in s *) (* find p s == the index of the first item in s for which p holds, *) (* or size s if no such item is found *) (* count p s == the number of items of s for which p holds *) (* constant s == all items in s are identical (trivial if s = [::]) *) (* uniq s == all the items in s are pairwise different *) (* subseq s1 s2 == s1 is a subsequence of s2, i.e., s1 = mask m s2 for *) (* some m : bitseq (see below). *) (* perm_eq s1 s2 == s2 is a permutation of s1, i.e., s1 and s2 have the *) (* items (with the same repetitions), but possibly in a *) (* different order. *) (* perm_eql s1 s2 <-> s1 and s2 behave identically on the left of perm_eq *) (* perm_eqr s1 s2 <-> s1 and s2 behave identically on the rightt of perm_eq *) (* These left/right transitive versions of perm_eq make it easier to chain *) (* a sequence of equivalences. *) (* ** Filtering: *) (* filter p s == the subsequence of s consisting of all the items *) (* for which the (boolean) predicate p holds *) (* subfilter s : seq sT == when sT has a subType p structure, the sequence *) (* of items of type sT corresponding to items of s *) (* for which p holds. *) (* rem x s == the subsequence of s, where the first occurrence *) (* of x has been removed (compare filter (predC1 x) s *) (* where ALL occurrences of x are removed). *) (* undup s == the subsequence of s containing only the first *) (* occurrence of each item in s, i.e., s with all *) (* duplicates removed. *) (* mask m s == the subsequence of s selected by m : bitseq, with *) (* item i of s selected by bit i in m (extra items or *) (* bits are ignored. *) (* ** Surgery: *) (* s1 ++ s2, cat s1 s2 == the concatenation of s1 and s2. *) (* take n s == the sequence containing only the first n items of s *) (* (or all of s if size s <= n). *) (* drop n s == s minus its first n items ([::] if size s <= n) *) (* rot n s == s rotated left n times (or s if size s <= n). *) (* := drop n s ++ take n s *) (* rotr n s == s rotated right n times (or s if size s <= n). *) (* rev s == the (linear time) reversal of s. *) (* catrev s1 s2 == the reversal of s1 followed by s2 (this is the *) (* recursive form of rev). *) (* ** Iterators: for s == [:: x_1, ..., x_n], t == [:: y_1, ..., y_m], *) (* map f s == the sequence [:: f x_1, ..., f x_n]. *) (* allpairs f s t == the sequence of all the f x y, with x and y drawn from *) (* s and t, respectievly, in row-major order. *) (* := [:: f x_1 y_1; ...; f x_1 y_m; f x_2 y_1; ...; f x_n y_m] *) (* pmap pf s == the sequence [:: y_i1, ..., y_ik] where i1 < ... < ik, *) (* pf x_i = Some y_i, and pf x_j = None iff j is not in *) (* {i1, ..., ik}. *) (* foldr f a s == the right fold of s by f (i.e., the natural iterator). *) (* := f x_1 (f x_2 ... (f x_n a)) *) (* sumn s == x_1 + (x_2 + ... + (x_n + 0)) (when s : seq nat). *) (* foldl f a s == the left fold of s by f. *) (* := f (f ... (f a x_1) ... x_n-1) x_n *) (* scanl f a s == the sequence of partial accumulators of foldl f a s. *) (* := [:: f a x_1; ...; foldl f a s] *) (* pairmap f a s == the sequence of f applied to consecutie items in a :: s. *) (* := [:: f a x_1; f x_1 x_2; ...; f x_n-1 x_n] *) (* zip s t == itemwise pairing of s and t (dropping any extra items). *) (* := [:: (x_1, y_1); ...; (x_mn, y_mn)] with mn = minn n m. *) (* unzip1 s == [:: (x_1).1; ...; (x_n).1] when s : seq (S * T). *) (* unzip2 s == [:: (x_1).2; ...; (x_n).2] when s : seq (S * T). *) (* flatten s == x_1 ++ ... ++ x_n ++ [::] when s : seq (seq T). *) (* reshape r s == s reshaped into a sequence of sequences whose sizes are *) (* given by r (trucating if s is too long or too short). *) (* := [:: [:: x_1; ...; x_r1]; *) (* [:: x_(r1 + 1); ...; x_(r0 + r1)]; *) (* ...; *) (* [:: x_(r1 + ... + r(k-1) + 1); ...; x_(r0 + ... rk)]] *) (* ** Notation for manifest comprehensions: *) (* [seq x <- s | C] := filter (fun x => C) s. *) (* [seq E | x <- s] := map (fun x => E) s. *) (* [seq E | x <- s, y <- t] := allpairs (fun x y => E) s t. *) (* [seq x <- s | C1 & C2] := [seq x <- s | C1 && C2]. *) (* [seq E | x <- s & C] := [seq E | x <- [seq x | C]]. *) (* --> The above allow optional type casts on the eigenvariables, as in *) (* [seq x : T <- s | C] or [seq E | x : T <- s, y : U <- t]. The cast may be *) (* needed as type inference considers E or C before s. *) (* We are quite systematic in providing lemmas to rewrite any composition *) (* of two operations. "rev", whose simplifications are not natural, is *) (* protected with nosimpl. *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Delimit Scope seq_scope with SEQ. Open Scope seq_scope. (* Inductive seq (T : Type) : Type := Nil | Cons of T & seq T. *) Notation seq := list. Prenex Implicits cons. Notation Cons T := (@cons T) (only parsing). Notation Nil T := (@nil T) (only parsing). Bind Scope seq_scope with list. Arguments Scope cons [type_scope _ seq_scope]. (* As :: and ++ are (improperly) declared in Init.datatypes, we only rebind *) (* them here. *) Infix "::" := cons : seq_scope. (* GG - this triggers a camlp4 warning, as if this Notation had been Reserved *) Notation "[ :: ]" := nil (at level 0, format "[ :: ]") : seq_scope. Notation "[ :: x1 ]" := (x1 :: [::]) (at level 0, format "[ :: x1 ]") : seq_scope. Notation "[ :: x & s ]" := (x :: s) (at level 0, only parsing) : seq_scope. Notation "[ :: x1 , x2 , .. , xn & s ]" := (x1 :: x2 :: .. (xn :: s) ..) (at level 0, format "'[hv' [ :: '[' x1 , '/' x2 , '/' .. , '/' xn ']' '/ ' & s ] ']'" ) : seq_scope. Notation "[ :: x1 ; x2 ; .. ; xn ]" := (x1 :: x2 :: .. [:: xn] ..) (at level 0, format "[ :: '[' x1 ; '/' x2 ; '/' .. ; '/' xn ']' ]" ) : seq_scope. Section Sequences. Variable n0 : nat. (* numerical parameter for take, drop et al *) Variable T : Type. (* must come before the implicit Type *) Variable x0 : T. (* default for head/nth *) Implicit Types x y z : T. Implicit Types m n : nat. Implicit Type s : seq T. Fixpoint size s := if s is _ :: s' then (size s').+1 else 0. Lemma size0nil s : size s = 0 -> s = [::]. Proof. by case: s. Qed. Definition nilp s := size s == 0. Lemma nilP s : reflect (s = [::]) (nilp s). Proof. by case: s => [|x s]; constructor. Qed. Definition ohead s := if s is x :: _ then Some x else None. Definition head s := if s is x :: _ then x else x0. Definition behead s := if s is _ :: s' then s' else [::]. Lemma size_behead s : size (behead s) = (size s).-1. Proof. by case: s. Qed. (* Factories *) Definition ncons n x := iter n (cons x). Definition nseq n x := ncons n x [::]. Lemma size_ncons n x s : size (ncons n x s) = n + size s. Proof. by elim: n => //= n ->. Qed. Lemma size_nseq n x : size (nseq n x) = n. Proof. by rewrite size_ncons addn0. Qed. (* n-ary, dependently typed constructor. *) Fixpoint seqn_type n := if n is n'.+1 then T -> seqn_type n' else seq T. Fixpoint seqn_rec f n : seqn_type n := if n is n'.+1 return seqn_type n then fun x => seqn_rec (fun s => f (x :: s)) n' else f [::]. Definition seqn := seqn_rec id. (* Sequence catenation "cat". *) Fixpoint cat s1 s2 := if s1 is x :: s1' then x :: s1' ++ s2 else s2 where "s1 ++ s2" := (cat s1 s2) : seq_scope. Lemma cat0s s : [::] ++ s = s. Proof. by []. Qed. Lemma cat1s x s : [:: x] ++ s = x :: s. Proof. by []. Qed. Lemma cat_cons x s1 s2 : (x :: s1) ++ s2 = x :: s1 ++ s2. Proof. by []. Qed. Lemma cat_nseq n x s : nseq n x ++ s = ncons n x s. Proof. by elim: n => //= n ->. Qed. Lemma cats0 s : s ++ [::] = s. Proof. by elim: s => //= x s ->. Qed. Lemma catA s1 s2 s3 : s1 ++ s2 ++ s3 = (s1 ++ s2) ++ s3. Proof. by elim: s1 => //= x s1 ->. Qed. Lemma size_cat s1 s2 : size (s1 ++ s2) = size s1 + size s2. Proof. by elim: s1 => //= x s1 ->. Qed. (* last, belast, rcons, and last induction. *) Fixpoint rcons s z := if s is x :: s' then x :: rcons s' z else [:: z]. Lemma rcons_cons x s z : rcons (x :: s) z = x :: rcons s z. Proof. by []. Qed. Lemma cats1 s z : s ++ [:: z] = rcons s z. Proof. by elim: s => //= x s ->. Qed. Fixpoint last x s := if s is x' :: s' then last x' s' else x. Fixpoint belast x s := if s is x' :: s' then x :: (belast x' s') else [::]. Lemma lastI x s : x :: s = rcons (belast x s) (last x s). Proof. by elim: s x => [|y s IHs] x //=; rewrite IHs. Qed. Lemma last_cons x y s : last x (y :: s) = last y s. Proof. by []. Qed. Lemma size_rcons s x : size (rcons s x) = (size s).+1. Proof. by rewrite -cats1 size_cat addnC. Qed. Lemma size_belast x s : size (belast x s) = size s. Proof. by elim: s x => [|y s IHs] x //=; rewrite IHs. Qed. Lemma last_cat x s1 s2 : last x (s1 ++ s2) = last (last x s1) s2. Proof. by elim: s1 x => [|y s1 IHs] x //=; rewrite IHs. Qed. Lemma last_rcons x s z : last x (rcons s z) = z. Proof. by rewrite -cats1 last_cat. Qed. Lemma belast_cat x s1 s2 : belast x (s1 ++ s2) = belast x s1 ++ belast (last x s1) s2. Proof. by elim: s1 x => [|y s1 IHs] x //=; rewrite IHs. Qed. Lemma belast_rcons x s z : belast x (rcons s z) = x :: s. Proof. by rewrite lastI -!cats1 belast_cat. Qed. Lemma cat_rcons x s1 s2 : rcons s1 x ++ s2 = s1 ++ x :: s2. Proof. by rewrite -cats1 -catA. Qed. Lemma rcons_cat x s1 s2 : rcons (s1 ++ s2) x = s1 ++ rcons s2 x. Proof. by rewrite -!cats1 catA. Qed. CoInductive last_spec : seq T -> Type := | LastNil : last_spec [::] | LastRcons s x : last_spec (rcons s x). Lemma lastP s : last_spec s. Proof. case: s => [|x s]; [left | rewrite lastI; right]. Qed. Lemma last_ind P : P [::] -> (forall s x, P s -> P (rcons s x)) -> forall s, P s. Proof. move=> Hnil Hlast s; rewrite -(cat0s s). elim: s [::] Hnil => [|x s2 IHs] s1 Hs1; first by rewrite cats0. by rewrite -cat_rcons; auto. Qed. (* Sequence indexing. *) Fixpoint nth s n {struct n} := if s is x :: s' then if n is n'.+1 then @nth s' n' else x else x0. Fixpoint set_nth s n y {struct n} := if s is x :: s' then if n is n'.+1 then x :: @set_nth s' n' y else y :: s' else ncons n x0 [:: y]. Lemma nth0 s : nth s 0 = head s. Proof. by []. Qed. Lemma nth_default s n : size s <= n -> nth s n = x0. Proof. by elim: s n => [|x s IHs] [|n]; try exact (IHs n). Qed. Lemma nth_nil n : nth [::] n = x0. Proof. by case: n. Qed. Lemma last_nth x s : last x s = nth (x :: s) (size s). Proof. by elim: s x => [|y s IHs] x /=. Qed. Lemma nth_last s : nth s (size s).-1 = last x0 s. Proof. by case: s => //= x s; rewrite last_nth. Qed. Lemma nth_behead s n : nth (behead s) n = nth s n.+1. Proof. by case: s n => [|x s] [|n]. Qed. Lemma nth_cat s1 s2 n : nth (s1 ++ s2) n = if n < size s1 then nth s1 n else nth s2 (n - size s1). Proof. by elim: s1 n => [|x s1 IHs] [|n]; try exact (IHs n). Qed. Lemma nth_rcons s x n : nth (rcons s x) n = if n < size s then nth s n else if n == size s then x else x0. Proof. by elim: s n => [|y s IHs] [|n] //=; rewrite ?nth_nil ?IHs. Qed. Lemma nth_ncons m x s n : nth (ncons m x s) n = if n < m then x else nth s (n - m). Proof. by elim: m n => [|m IHm] [|n] //=; exact: IHm. Qed. Lemma nth_nseq m x n : nth (nseq m x) n = (if n < m then x else x0). Proof. by elim: m n => [|m IHm] [|n] //=; exact: IHm. Qed. Lemma eq_from_nth s1 s2 : size s1 = size s2 -> (forall i, i < size s1 -> nth s1 i = nth s2 i) -> s1 = s2. Proof. elim: s1 s2 => [|x1 s1 IHs1] [|x2 s2] //= [eq_sz] eq_s12. rewrite [x1](eq_s12 0) // (IHs1 s2) // => i; exact: (eq_s12 i.+1). Qed. Lemma size_set_nth s n y : size (set_nth s n y) = maxn n.+1 (size s). Proof. elim: s n => [|x s IHs] [|n] //=. - by rewrite size_ncons addn1 maxn0. - by rewrite maxnE subn1. by rewrite IHs -add1n addn_maxr. Qed. Lemma set_nth_nil n y : set_nth [::] n y = ncons n x0 [:: y]. Proof. by case: n. Qed. Lemma nth_set_nth s n y : nth (set_nth s n y) =1 [eta nth s with n |-> y]. Proof. elim: s n => [|x s IHs] [|n] [|m] //=; rewrite ?nth_nil ?IHs // nth_ncons eqSS. case: ltngtP => // [lt_nm | ->]; last by rewrite subnn. by rewrite nth_default // subn_gt0. Qed. Lemma set_set_nth s n1 y1 n2 y2 (s2 := set_nth s n2 y2) : set_nth (set_nth s n1 y1) n2 y2 = if n1 == n2 then s2 else set_nth s2 n1 y1. Proof. have [-> | ne_n12] := altP eqP. apply: eq_from_nth => [|i _]; first by rewrite !size_set_nth maxnA maxnn. by do 2!rewrite !nth_set_nth /=; case: eqP. apply: eq_from_nth => [|i _]; first by rewrite !size_set_nth maxnCA. do 2!rewrite !nth_set_nth /=; case: eqP => // ->. by rewrite eq_sym -if_neg ne_n12. Qed. (* find, count, has, all. *) Section SeqFind. Variable a : pred T. Fixpoint find s := if s is x :: s' then if a x then 0 else (find s').+1 else 0. Fixpoint filter s := if s is x :: s' then if a x then x :: filter s' else filter s' else [::]. Fixpoint count s := if s is x :: s' then a x + count s' else 0. Fixpoint has s := if s is x :: s' then a x || has s' else false. Fixpoint all s := if s is x :: s' then a x && all s' else true. Lemma count_filter s : count s = size (filter s). Proof. by elim: s => //= x s ->; case (a x). Qed. Lemma has_count s : has s = (0 < count s). Proof. by elim: s => //= x s ->; case (a x). Qed. Lemma count_size s : count s <= size s. Proof. by elim: s => //= x s; case: (a x); last exact: leqW. Qed. Lemma all_count s : all s = (count s == size s). Proof. elim: s => //= x s; case: (a x) => _ //=. by rewrite add0n eqn_leq andbC ltnNge count_size. Qed. Lemma filter_all s : all (filter s). Proof. by elim: s => //= x s IHs; case: ifP => //= ->. Qed. Lemma all_filterP s : reflect (filter s = s) (all s). Proof. apply: (iffP idP) => [| <-]; last exact: filter_all. by elim: s => //= x s IHs /andP[-> Hs]; rewrite IHs. Qed. Lemma filter_id s : filter (filter s) = filter s. Proof. by apply/all_filterP; exact: filter_all. Qed. Lemma has_find s : has s = (find s < size s). Proof. by elim: s => //= x s IHs; case (a x); rewrite ?leqnn. Qed. Lemma find_size s : find s <= size s. Proof. by elim: s => //= x s IHs; case (a x). Qed. Lemma find_cat s1 s2 : find (s1 ++ s2) = if has s1 then find s1 else size s1 + find s2. Proof. by elim: s1 => //= x s1 IHs; case: (a x) => //; rewrite IHs (fun_if succn). Qed. Lemma has_nil : has [::] = false. Proof. by []. Qed. Lemma has_seq1 x : has [:: x] = a x. Proof. exact: orbF. Qed. Lemma has_seqb (b : bool) x : has (nseq b x) = b && a x. Proof. by case: b => //=; exact: orbF. Qed. Lemma all_nil : all [::] = true. Proof. by []. Qed. Lemma all_seq1 x : all [:: x] = a x. Proof. exact: andbT. Qed. Lemma nth_find s : has s -> a (nth s (find s)). Proof. by elim: s => //= x s IHs; case Hx: (a x). Qed. Lemma before_find s i : i < find s -> a (nth s i) = false. Proof. by elim: s i => //= x s IHs; case Hx: (a x) => [|] // [|i] //; exact: (IHs i). Qed. Lemma filter_cat s1 s2 : filter (s1 ++ s2) = filter s1 ++ filter s2. Proof. by elim: s1 => //= x s1 ->; case (a x). Qed. Lemma filter_rcons s x : filter (rcons s x) = if a x then rcons (filter s) x else filter s. Proof. by rewrite -!cats1 filter_cat /=; case (a x); rewrite /= ?cats0. Qed. Lemma count_cat s1 s2 : count (s1 ++ s2) = count s1 + count s2. Proof. by rewrite !count_filter filter_cat size_cat. Qed. Lemma has_cat s1 s2 : has (s1 ++ s2) = has s1 || has s2. Proof. by elim: s1 => [|x s1 IHs] //=; rewrite IHs orbA. Qed. Lemma has_rcons s x : has (rcons s x) = a x || has s. Proof. by rewrite -cats1 has_cat has_seq1 orbC. Qed. Lemma all_cat s1 s2 : all (s1 ++ s2) = all s1 && all s2. Proof. by elim: s1 => [|x s1 IHs] //=; rewrite IHs andbA. Qed. Lemma all_rcons s x : all (rcons s x) = a x && all s. Proof. by rewrite -cats1 all_cat all_seq1 andbC. Qed. End SeqFind. Lemma eq_find a1 a2 : a1 =1 a2 -> find a1 =1 find a2. Proof. by move=> Ea; elim=> //= x s IHs; rewrite Ea IHs. Qed. Lemma eq_filter a1 a2 : a1 =1 a2 -> filter a1 =1 filter a2. Proof. by move=> Ea; elim=> //= x s IHs; rewrite Ea IHs. Qed. Lemma eq_count a1 a2 : a1 =1 a2 -> count a1 =1 count a2. Proof. by move=> Ea s; rewrite !count_filter (eq_filter Ea). Qed. Lemma eq_has a1 a2 : a1 =1 a2 -> has a1 =1 has a2. Proof. by move=> Ea s; rewrite !has_count (eq_count Ea). Qed. Lemma eq_all a1 a2 : a1 =1 a2 -> all a1 =1 all a2. Proof. by move=> Ea s; rewrite !all_count (eq_count Ea). Qed. Section SubPred. Variable (a1 a2 : pred T). Hypothesis s12 : subpred a1 a2. Lemma sub_find s : find a2 s <= find a1 s. Proof. by elim: s => //= x s IHs; case: ifP => // /(contraFF (@s12 x))->. Qed. Lemma sub_has s : has a1 s -> has a2 s. Proof. by rewrite !has_find; exact: leq_ltn_trans (sub_find s). Qed. Lemma sub_count s : count a1 s <= count a2 s. Proof. by elim: s => //= x s; apply: leq_add; case a1x: (a1 x); rewrite // s12. Qed. Lemma sub_all s : all a1 s -> all a2 s. Proof. by rewrite !all_count !eqn_leq !count_size => /leq_trans-> //; exact: sub_count. Qed. End SubPred. Lemma filter_pred0 s : filter pred0 s = [::]. Proof. by elim: s. Qed. Lemma filter_predT s : filter predT s = s. Proof. by elim: s => //= x s ->. Qed. Lemma filter_predI a1 a2 s : filter (predI a1 a2) s = filter a1 (filter a2 s). Proof. elim: s => //= x s IHs; rewrite andbC IHs. by case: (a2 x) => //=; case (a1 x). Qed. Lemma count_pred0 s : count pred0 s = 0. Proof. by rewrite count_filter filter_pred0. Qed. Lemma count_predT s : count predT s = size s. Proof. by rewrite count_filter filter_predT. Qed. Lemma count_predUI a1 a2 s : count (predU a1 a2) s + count (predI a1 a2) s = count a1 s + count a2 s. Proof. elim: s => //= x s IHs; rewrite /= addnCA -addnA IHs addnA addnC. by rewrite -!addnA; do 2 nat_congr; case (a1 x); case (a2 x). Qed. Lemma count_predC a s : count a s + count (predC a) s = size s. Proof. by elim: s => //= x s IHs; rewrite addnCA -addnA IHs addnA addn_negb. Qed. Lemma has_pred0 s : has pred0 s = false. Proof. by rewrite has_count count_pred0. Qed. Lemma has_predT s : has predT s = (0 < size s). Proof. by rewrite has_count count_predT. Qed. Lemma has_predC a s : has (predC a) s = ~~ all a s. Proof. by elim: s => //= x s ->; case (a x). Qed. Lemma has_predU a1 a2 s : has (predU a1 a2) s = has a1 s || has a2 s. Proof. by elim: s => //= x s ->; rewrite -!orbA; do !bool_congr. Qed. Lemma all_pred0 s : all pred0 s = (size s == 0). Proof. by rewrite all_count count_pred0 eq_sym. Qed. Lemma all_predT s : all predT s. Proof. by rewrite all_count count_predT. Qed. Lemma all_predC a s : all (predC a) s = ~~ has a s. Proof. by elim: s => //= x s ->; case (a x). Qed. Lemma all_predI a1 a2 s : all (predI a1 a2) s = all a1 s && all a2 s. Proof. apply: (can_inj negbK); rewrite negb_and -!has_predC -has_predU. by apply: eq_has => x; rewrite /= negb_and. Qed. (* Surgery: drop, take, rot, rotr. *) Fixpoint drop n s {struct s} := match s, n with | _ :: s', n'.+1 => drop n' s' | _, _ => s end. Lemma drop_behead : drop n0 =1 iter n0 behead. Proof. by elim: n0 => [|n IHn] [|x s] //; rewrite iterSr -IHn. Qed. Lemma drop0 s : drop 0 s = s. Proof. by case: s. Qed. Lemma drop1 : drop 1 =1 behead. Proof. by case=> [|x [|y s]]. Qed. Lemma drop_oversize n s : size s <= n -> drop n s = [::]. Proof. by elim: n s => [|n IHn] [|x s]; try exact (IHn s). Qed. Lemma drop_size s : drop (size s) s = [::]. Proof. by rewrite drop_oversize // leqnn. Qed. Lemma drop_cons x s : drop n0 (x :: s) = if n0 is n.+1 then drop n s else x :: s. Proof. by []. Qed. Lemma size_drop s : size (drop n0 s) = size s - n0. Proof. by elim: s n0 => [|x s IHs] [|n]; try exact (IHs n). Qed. Lemma drop_cat s1 s2 : drop n0 (s1 ++ s2) = if n0 < size s1 then drop n0 s1 ++ s2 else drop (n0 - size s1) s2. Proof. by elim: s1 n0 => [|x s1 IHs] [|n]; try exact (IHs n). Qed. Lemma drop_size_cat n s1 s2 : size s1 = n -> drop n (s1 ++ s2) = s2. Proof. by move <-; elim: s1 => //=; rewrite drop0. Qed. Lemma nconsK n x : cancel (ncons n x) (drop n). Proof. by elim: n => //; case. Qed. Fixpoint take n s {struct s} := match s, n with | x :: s', n'.+1 => x :: take n' s' | _, _ => [::] end. Lemma take0 s : take 0 s = [::]. Proof. by case: s. Qed. Lemma take_oversize n s : size s <= n -> take n s = s. Proof. by elim: n s => [|n IHn] [|x s] Hsn; try (congr cons; apply: IHn). Qed. Lemma take_size s : take (size s) s = s. Proof. by rewrite take_oversize // leqnn. Qed. Lemma take_cons x s : take n0 (x :: s) = if n0 is n.+1 then x :: (take n s) else [::]. Proof. by []. Qed. Lemma drop_rcons s : n0 <= size s -> forall x, drop n0 (rcons s x) = rcons (drop n0 s) x. Proof. by elim: s n0 => [|y s IHs] [|n]; try exact (IHs n). Qed. Lemma cat_take_drop s : take n0 s ++ drop n0 s = s. Proof. by elim: s n0 => [|x s IHs] [|n]; try exact (congr1 _ (IHs n)). Qed. Lemma size_takel s : n0 <= size s -> size (take n0 s) = n0. Proof. by move/subnKC; rewrite -{2}(cat_take_drop s) size_cat size_drop => /addIn. Qed. Lemma size_take s : size (take n0 s) = if n0 < size s then n0 else size s. Proof. have [le_sn | lt_ns] := leqP (size s) n0; first by rewrite take_oversize. by rewrite size_takel // ltnW. Qed. Lemma take_cat s1 s2 : take n0 (s1 ++ s2) = if n0 < size s1 then take n0 s1 else s1 ++ take (n0 - size s1) s2. Proof. elim: s1 n0 => [|x s1 IHs] [|n] //=. by rewrite ltnS subSS -(fun_if (cons x)) -IHs. Qed. Lemma take_size_cat n s1 s2 : size s1 = n -> take n (s1 ++ s2) = s1. Proof. by move <-; elim: s1 => [|x s1 IHs]; rewrite ?take0 //= IHs. Qed. Lemma takel_cat s1 : n0 <= size s1 -> forall s2, take n0 (s1 ++ s2) = take n0 s1. Proof. move=> Hn0 s2; rewrite take_cat ltn_neqAle Hn0 andbT. by case: (n0 =P size s1) => //= ->; rewrite subnn take0 cats0 take_size. Qed. Lemma nth_drop s i : nth (drop n0 s) i = nth s (n0 + i). Proof. have [lt_n0_s | le_s_n0] := ltnP n0 (size s). rewrite -{2}[s]cat_take_drop nth_cat size_take lt_n0_s /= addKn. by rewrite ltnNge leq_addr. rewrite !nth_default //; first exact: leq_trans (leq_addr _ _). by rewrite size_drop (eqnP le_s_n0). Qed. Lemma nth_take i : i < n0 -> forall s, nth (take n0 s) i = nth s i. Proof. move=> lt_i_n0 s; case lt_n0_s: (n0 < size s). by rewrite -{2}[s]cat_take_drop nth_cat size_take lt_n0_s /= lt_i_n0. by rewrite -{1}[s]cats0 take_cat lt_n0_s /= cats0. Qed. (* drop_nth and take_nth below do NOT use the default n0, because the "n" *) (* can be inferred from the condition, whereas the nth default value x0 *) (* will have to be given explicitly (and this will provide "d" as well). *) Lemma drop_nth n s : n < size s -> drop n s = nth s n :: drop n.+1 s. Proof. by elim: s n => [|x s IHs] [|n] Hn //=; rewrite ?drop0 1?IHs. Qed. Lemma take_nth n s : n < size s -> take n.+1 s = rcons (take n s) (nth s n). Proof. by elim: s n => [|x s IHs] //= [|n] Hn /=; rewrite ?take0 -?IHs. Qed. (* Rotation *) Definition rot n s := drop n s ++ take n s. Lemma rot0 s : rot 0 s = s. Proof. by rewrite /rot drop0 take0 cats0. Qed. Lemma size_rot s : size (rot n0 s) = size s. Proof. by rewrite -{2}[s]cat_take_drop /rot !size_cat addnC. Qed. Lemma rot_oversize n s : size s <= n -> rot n s = s. Proof. by move=> le_s_n; rewrite /rot take_oversize ?drop_oversize. Qed. Lemma rot_size s : rot (size s) s = s. Proof. exact: rot_oversize. Qed. Lemma has_rot s a : has a (rot n0 s) = has a s. Proof. by rewrite has_cat orbC -has_cat cat_take_drop. Qed. Lemma rot_size_cat s1 s2 : rot (size s1) (s1 ++ s2) = s2 ++ s1. Proof. by rewrite /rot take_size_cat ?drop_size_cat. Qed. Definition rotr n s := rot (size s - n) s. Lemma rotK : cancel (rot n0) (rotr n0). Proof. move=> s; rewrite /rotr size_rot -size_drop {2}/rot. by rewrite rot_size_cat cat_take_drop. Qed. Lemma rot_inj : injective (rot n0). Proof. exact (can_inj rotK). Qed. Lemma rot1_cons x s : rot 1 (x :: s) = rcons s x. Proof. by rewrite /rot /= take0 drop0 -cats1. Qed. (* (efficient) reversal *) Fixpoint catrev s1 s2 := if s1 is x :: s1' then catrev s1' (x :: s2) else s2. End Sequences. (* rev must be defined outside a Section because Coq's end of section *) (* "cooking" removes the nosimpl guard. *) Definition rev T (s : seq T) := nosimpl (catrev s [::]). Implicit Arguments nilP [T s]. Implicit Arguments all_filterP [T a s]. Prenex Implicits size nilP head ohead behead last rcons belast. Prenex Implicits cat take drop rev rot rotr. Prenex Implicits find count nth all has filter all_filterP. Infix "++" := cat : seq_scope. Notation "[ 'seq' x <- s | C ]" := (filter (fun x => C%B) s) (at level 0, x at level 99, format "[ '[hv' 'seq' x <- s '/ ' | C ] ']'") : seq_scope. Notation "[ 'seq' x <- s | C1 & C2 ]" := [seq x <- s | C1 && C2] (at level 0, x at level 99, format "[ '[hv' 'seq' x <- s '/ ' | C1 '/ ' & C2 ] ']'") : seq_scope. Notation "[ 'seq' x : T <- s | C ]" := (filter (fun x : T => C%B) s) (at level 0, x at level 99, only parsing). Notation "[ 'seq' x : T <- s | C1 & C2 ]" := [seq x : T <- s | C1 && C2] (at level 0, x at level 99, only parsing). Section Rev. Variable T : Type. Implicit Types s t : seq T. Lemma catrev_catl s t u : catrev (s ++ t) u = catrev t (catrev s u). Proof. by elim: s u => /=. Qed. Lemma catrev_catr s t u : catrev s (t ++ u) = catrev s t ++ u. Proof. by elim: s t => //= x s IHs t; rewrite -IHs. Qed. Lemma catrevE s t : catrev s t = rev s ++ t. Proof. by rewrite -catrev_catr. Qed. Lemma rev_cons x s : rev (x :: s) = rcons (rev s) x. Proof. by rewrite -cats1 -catrevE. Qed. Lemma size_rev s : size (rev s) = size s. Proof. by elim: s => // x s IHs; rewrite rev_cons size_rcons IHs. Qed. Lemma rev_cat s t : rev (s ++ t) = rev t ++ rev s. Proof. by rewrite -catrev_catr -catrev_catl. Qed. Lemma rev_rcons s x : rev (rcons s x) = x :: rev s. Proof. by rewrite -cats1 rev_cat. Qed. Lemma revK : involutive (@rev T). Proof. by elim=> //= x s IHs; rewrite rev_cons rev_rcons IHs. Qed. Lemma nth_rev x0 n s : n < size s -> nth x0 (rev s) n = nth x0 s (size s - n.+1). Proof. elim/last_ind: s => // s x IHs in n *. rewrite rev_rcons size_rcons ltnS subSS -cats1 nth_cat /=. case: n => [|n] lt_n_s; first by rewrite subn0 ltnn subnn. by rewrite -{2}(subnK lt_n_s) -addSnnS leq_addr /= -IHs. Qed. End Rev. (* Equality and eqType for seq. *) Section EqSeq. Variables (n0 : nat) (T : eqType) (x0 : T). Notation Local nth := (nth x0). Implicit Type s : seq T. Implicit Types x y z : T. Fixpoint eqseq s1 s2 {struct s2} := match s1, s2 with | [::], [::] => true | x1 :: s1', x2 :: s2' => (x1 == x2) && eqseq s1' s2' | _, _ => false end. Lemma eqseqP : Equality.axiom eqseq. Proof. move; elim=> [|x1 s1 IHs] [|x2 s2]; do [by constructor | simpl]. case: (x1 =P x2) => [<-|neqx]; last by right; case. by apply: (iffP (IHs s2)) => [<-|[]]. Qed. Canonical seq_eqMixin := EqMixin eqseqP. Canonical seq_eqType := Eval hnf in EqType (seq T) seq_eqMixin. Lemma eqseqE : eqseq = eq_op. Proof. by []. Qed. Lemma eqseq_cons x1 x2 s1 s2 : (x1 :: s1 == x2 :: s2) = (x1 == x2) && (s1 == s2). Proof. by []. Qed. Lemma eqseq_cat s1 s2 s3 s4 : size s1 = size s2 -> (s1 ++ s3 == s2 ++ s4) = (s1 == s2) && (s3 == s4). Proof. elim: s1 s2 => [|x1 s1 IHs] [|x2 s2] //= [sz12]. by rewrite !eqseq_cons -andbA IHs. Qed. Lemma eqseq_rcons s1 s2 x1 x2 : (rcons s1 x1 == rcons s2 x2) = (s1 == s2) && (x1 == x2). Proof. elim: s1 s2 => [|y1 s1 IHs] [|y2 s2] /=; rewrite ?eqseq_cons ?andbT ?andbF // ?IHs 1?andbA // andbC; by [case s2 | case s1]. Qed. Lemma has_filter a s : has a s = (filter a s != [::]). Proof. by rewrite has_count count_filter; case (filter a s). Qed. Lemma size_eq0 s : (size s == 0) = (s == [::]). Proof. apply/eqP/eqP=> [|-> //]; exact: size0nil. Qed. (* mem_seq and index. *) (* mem_seq defines a predType for seq. *) Fixpoint mem_seq (s : seq T) := if s is y :: s' then xpredU1 y (mem_seq s') else xpred0. Definition eqseq_class := seq T. Identity Coercion seq_of_eqseq : eqseq_class >-> seq. Coercion pred_of_eq_seq (s : eqseq_class) : pred_class := [eta mem_seq s]. Canonical seq_predType := @mkPredType T (seq T) pred_of_eq_seq. (* The line below makes mem_seq a canonical instance of topred. *) Canonical mem_seq_predType := mkPredType mem_seq. Lemma in_cons y s x : (x \in y :: s) = (x == y) || (x \in s). Proof. by []. Qed. Lemma in_nil x : (x \in [::]) = false. Proof. by []. Qed. Lemma mem_seq1 x y : (x \in [:: y]) = (x == y). Proof. by rewrite in_cons orbF. Qed. (* to be repeated after the Section discharge. *) Let inE := (mem_seq1, in_cons, inE). Lemma mem_seq2 x y1 y2 : (x \in [:: y1; y2]) = xpred2 y1 y2 x. Proof. by rewrite !inE. Qed. Lemma mem_seq3 x y1 y2 y3 : (x \in [:: y1; y2; y3]) = xpred3 y1 y2 y3 x. Proof. by rewrite !inE. Qed. Lemma mem_seq4 x y1 y2 y3 y4 : (x \in [:: y1; y2; y3; y4]) = xpred4 y1 y2 y3 y4 x. Proof. by rewrite !inE. Qed. Lemma mem_cat x s1 s2 : (x \in s1 ++ s2) = (x \in s1) || (x \in s2). Proof. by elim: s1 => //= y s1 IHs; rewrite !inE /= -orbA -IHs. Qed. Lemma mem_rcons s y : rcons s y =i y :: s. Proof. by move=> x; rewrite -cats1 /= mem_cat mem_seq1 orbC in_cons. Qed. Lemma mem_head x s : x \in x :: s. Proof. exact: predU1l. Qed. Lemma mem_last x s : last x s \in x :: s. Proof. by rewrite lastI mem_rcons mem_head. Qed. Lemma mem_behead s : {subset behead s <= s}. Proof. by case: s => // y s x; exact: predU1r. Qed. Lemma mem_belast s y : {subset belast y s <= y :: s}. Proof. by move=> x ys'x; rewrite lastI mem_rcons mem_behead. Qed. Lemma mem_nth s n : n < size s -> nth s n \in s. Proof. by elim: s n => [|x s IHs] // [_|n sz_s]; rewrite ?mem_head // mem_behead ?IHs. Qed. Lemma mem_take s x : x \in take n0 s -> x \in s. Proof. by move=> s0x; rewrite -(cat_take_drop n0 s) mem_cat /= s0x. Qed. Lemma mem_drop s x : x \in drop n0 s -> x \in s. Proof. by move=> s0'x; rewrite -(cat_take_drop n0 s) mem_cat /= s0'x orbT. Qed. Lemma mem_rev s : rev s =i s. Proof. by move=> y; elim: s => // x s IHs; rewrite rev_cons /= mem_rcons in_cons IHs. Qed. Section Filters. Variable a : pred T. Lemma hasP s : reflect (exists2 x, x \in s & a x) (has a s). Proof. elim: s => [|y s IHs] /=; first by right; case. case ay: (a y); first by left; exists y; rewrite ?mem_head. apply: (iffP IHs) => [] [x ysx ax]; exists x => //; first exact: mem_behead. by case: (predU1P ysx) ax => [->|//]; rewrite ay. Qed. Lemma hasPn s : reflect (forall x, x \in s -> ~~ a x) (~~ has a s). Proof. apply: (iffP idP) => not_a_s => [x s_x|]. by apply: contra not_a_s => a_x; apply/hasP; exists x. by apply/hasP=> [[x s_x]]; apply/negP; exact: not_a_s. Qed. Lemma allP s : reflect (forall x, x \in s -> a x) (all a s). Proof. elim: s => [|x s IHs]; first by left. rewrite /= andbC; case: IHs => IHs /=. apply: (iffP idP) => [Hx y|]; last by apply; exact: mem_head. by case/predU1P=> [->|Hy]; auto. by right=> H; case IHs => y Hy; apply H; exact: mem_behead. Qed. Lemma allPn s : reflect (exists2 x, x \in s & ~~ a x) (~~ all a s). Proof. elim: s => [|x s IHs]; first by right=> [[x Hx _]]. rewrite /= andbC negb_and; case: IHs => IHs /=. by left; case: IHs => y Hy Hay; exists y; first exact: mem_behead. apply: (iffP idP) => [|[y]]; first by exists x; rewrite ?mem_head. by case/predU1P=> [-> // | s_y not_a_y]; case: IHs; exists y. Qed. Lemma mem_filter x s : (x \in filter a s) = a x && (x \in s). Proof. rewrite andbC; elim: s => //= y s IHs. rewrite (fun_if (fun s' : seq T => x \in s')) !in_cons {}IHs. by case: eqP => [->|_]; case (a y); rewrite /= ?andbF. Qed. End Filters. Section EqIn. Variables a1 a2 : pred T. Lemma eq_in_filter s : {in s, a1 =1 a2} -> filter a1 s = filter a2 s. Proof. elim: s => //= x s IHs eq_a. rewrite eq_a ?mem_head ?IHs // => y s_y; apply: eq_a; exact: mem_behead. Qed. Lemma eq_in_find s : {in s, a1 =1 a2} -> find a1 s = find a2 s. Proof. elim: s => //= x s IHs eq_a12; rewrite eq_a12 ?mem_head // IHs // => y s'y. by rewrite eq_a12 // mem_behead. Qed. Lemma eq_in_count s : {in s, a1 =1 a2} -> count a1 s = count a2 s. Proof. by move/eq_in_filter=> eq_a12; rewrite !count_filter eq_a12. Qed. Lemma eq_in_all s : {in s, a1 =1 a2} -> all a1 s = all a2 s. Proof. by move=> eq_a12; rewrite !all_count eq_in_count. Qed. Lemma eq_in_has s : {in s, a1 =1 a2} -> has a1 s = has a2 s. Proof. by move/eq_in_filter=> eq_a12; rewrite !has_filter eq_a12. Qed. End EqIn. Lemma eq_has_r s1 s2 : s1 =i s2 -> has^~ s1 =1 has^~ s2. Proof. move=> Es12 a; apply/(hasP a s1)/(hasP a s2) => [] [x Hx Hax]; by exists x; rewrite // ?Es12 // -Es12. Qed. Lemma eq_all_r s1 s2 : s1 =i s2 -> all^~ s1 =1 all^~ s2. Proof. by move=> Es12 a; apply/(allP a s1)/(allP a s2) => Hs x Hx; apply: Hs; rewrite Es12 in Hx *. Qed. Lemma has_sym s1 s2 : has (mem s1) s2 = has (mem s2) s1. Proof. by apply/(hasP _ s2)/(hasP _ s1) => [] [x]; exists x. Qed. Lemma has_pred1 x s : has (pred1 x) s = (x \in s). Proof. by rewrite -(eq_has (mem_seq1^~ x)) (has_sym [:: x]) /= orbF. Qed. (* Constant sequences, i.e., the image of nseq. *) Definition constant s := if s is x :: s' then all (pred1 x) s' else true. Lemma all_pred1P x s : reflect (s = nseq (size s) x) (all (pred1 x) s). Proof. elim: s => [|y s IHs] /=; first by left. case: eqP => [->{y} | ne_xy]; last by right=> [] [? _]; case ne_xy. by apply: (iffP IHs) => [<- //| []]. Qed. Lemma all_pred1_constant x s : all (pred1 x) s -> constant s. Proof. by case: s => //= y s /andP[/eqP->]. Qed. Lemma all_pred1_nseq x y n : all (pred1 x) (nseq n y) = (n == 0) || (x == y). Proof. case: n => //= n; rewrite eq_sym; apply: andb_idr => /eqP->{x}. by elim: n => //= n ->; rewrite eqxx. Qed. Lemma constant_nseq n x : constant (nseq n x). Proof. by case: n => //= n; rewrite all_pred1_nseq eqxx orbT. Qed. (* Uses x0 *) Lemma constantP s : reflect (exists x, s = nseq (size s) x) (constant s). Proof. apply: (iffP idP) => [| [x ->]]; last exact: constant_nseq. case: s => [|x s] /=; first by exists x0. by move/all_pred1P=> def_s; exists x; rewrite -def_s. Qed. (* Duplicate-freenes. *) Fixpoint uniq s := if s is x :: s' then (x \notin s') && uniq s' else true. Lemma cons_uniq x s : uniq (x :: s) = (x \notin s) && uniq s. Proof. by []. Qed. Lemma cat_uniq s1 s2 : uniq (s1 ++ s2) = [&& uniq s1, ~~ has (mem s1) s2 & uniq s2]. Proof. elim: s1 => [|x s1 IHs]; first by rewrite /= has_pred0. by rewrite has_sym /= mem_cat !negb_or has_sym IHs -!andbA; do !bool_congr. Qed. Lemma uniq_catC s1 s2 : uniq (s1 ++ s2) = uniq (s2 ++ s1). Proof. by rewrite !cat_uniq has_sym andbCA andbA andbC. Qed. Lemma uniq_catCA s1 s2 s3 : uniq (s1 ++ s2 ++ s3) = uniq (s2 ++ s1 ++ s3). Proof. by rewrite !catA -!(uniq_catC s3) !(cat_uniq s3) uniq_catC !has_cat orbC. Qed. Lemma rcons_uniq s x : uniq (rcons s x) = (x \notin s) && uniq s. Proof. by rewrite -cats1 uniq_catC. Qed. Lemma filter_uniq s a : uniq s -> uniq (filter a s). Proof. elim: s => [|x s IHs] //= /andP[Hx Hs]; case (a x); auto. by rewrite /= mem_filter /= (negbTE Hx) andbF; auto. Qed. Lemma rot_uniq s : uniq (rot n0 s) = uniq s. Proof. by rewrite /rot uniq_catC cat_take_drop. Qed. Lemma rev_uniq s : uniq (rev s) = uniq s. Proof. elim: s => // x s IHs. by rewrite rev_cons -cats1 cat_uniq /= andbT andbC mem_rev orbF IHs. Qed. Lemma count_uniq_mem s x : uniq s -> count (pred1 x) s = (x \in s). Proof. elim: s => //= [y s IHs] /andP[/negbTE Hy /IHs-> {IHs}]. by rewrite in_cons eq_sym; case: eqP => // ->; rewrite Hy. Qed. Lemma filter_pred1_uniq s x : uniq s -> x \in s -> filter (pred1 x) s = [:: x]. Proof. move=> uniq_s s_x; rewrite (all_pred1P _ _ (filter_all _ _)). by rewrite -count_filter count_uniq_mem ?s_x. Qed. (* Removing duplicates *) Fixpoint undup s := if s is x :: s' then if x \in s' then undup s' else x :: undup s' else [::]. Lemma size_undup s : size (undup s) <= size s. Proof. by elim: s => //= x s IHs; case: (x \in s) => //=; exact: ltnW. Qed. Lemma mem_undup s : undup s =i s. Proof. move=> x; elim: s => //= y s IHs. by case Hy: (y \in s); rewrite in_cons IHs //; case: eqP => // ->. Qed. Lemma undup_uniq s : uniq (undup s). Proof. by elim: s => //= x s IHs; case s_x: (x \in s); rewrite //= mem_undup s_x. Qed. Lemma undup_id s : uniq s -> undup s = s. Proof. by elim: s => //= x s IHs /andP[/negbTE-> /IHs->]. Qed. Lemma ltn_size_undup s : (size (undup s) < size s) = ~~ uniq s. Proof. by elim: s => //= x s IHs; case Hx: (x \in s); rewrite //= ltnS size_undup. Qed. Lemma filter_undup p s : filter p (undup s) = undup (filter p s). Proof. elim: s => //= x s IHs; rewrite (fun_if undup) fun_if /= mem_filter /=. by rewrite (fun_if (filter p)) /= IHs; case: ifP => -> //=; exact: if_same. Qed. (* Lookup *) Definition index x := find (pred1 x). Lemma index_size x s : index x s <= size s. Proof. by rewrite /index find_size. Qed. Lemma index_mem x s : (index x s < size s) = (x \in s). Proof. by rewrite -has_pred1 has_find. Qed. Lemma nth_index x s : x \in s -> nth s (index x s) = x. Proof. by rewrite -has_pred1 => /(nth_find x0)/eqP. Qed. Lemma index_cat x s1 s2 : index x (s1 ++ s2) = if x \in s1 then index x s1 else size s1 + index x s2. Proof. by rewrite /index find_cat has_pred1. Qed. Lemma index_uniq i s : i < size s -> uniq s -> index (nth s i) s = i. Proof. elim: s i => [|x s IHs] //= [|i]; rewrite /= ?eqxx // ltnS => lt_i_s. case/andP=> not_s_x /(IHs i)-> {IHs}//. by case: eqP not_s_x => // ->; rewrite mem_nth. Qed. Lemma index_head x s : index x (x :: s) = 0. Proof. by rewrite /= eqxx. Qed. Lemma index_last x s : uniq (x :: s) -> index (last x s) (x :: s) = size s. Proof. rewrite lastI rcons_uniq -cats1 index_cat size_belast. by case: ifP => //=; rewrite eqxx addn0. Qed. Lemma nth_uniq s i j : i < size s -> j < size s -> uniq s -> (nth s i == nth s j) = (i == j). Proof. move=> lt_i_s lt_j_s Us; apply/eqP/eqP=> [eq_sij|-> //]. by rewrite -(index_uniq lt_i_s Us) eq_sij index_uniq. Qed. Lemma mem_rot s : rot n0 s =i s. Proof. by move=> x; rewrite -{2}(cat_take_drop n0 s) !mem_cat /= orbC. Qed. Lemma eqseq_rot s1 s2 : (rot n0 s1 == rot n0 s2) = (s1 == s2). Proof. by apply: inj_eq; exact: rot_inj. Qed. CoInductive rot_to_spec s x := RotToSpec i s' of rot i s = x :: s'. Lemma rot_to s x : x \in s -> rot_to_spec s x. Proof. move=> s_x; pose i := index x s; exists i (drop i.+1 s ++ take i s). rewrite -cat_cons {}/i; congr cat; elim: s s_x => //= y s IHs. by rewrite eq_sym in_cons; case: eqP => // -> _; rewrite drop0. Qed. End EqSeq. Definition inE := (mem_seq1, in_cons, inE). Prenex Implicits mem_seq1 uniq undup index. Implicit Arguments eqseqP [T x y]. Implicit Arguments hasP [T a s]. Implicit Arguments hasPn [T a s]. Implicit Arguments allP [T a s]. Implicit Arguments allPn [T a s]. Prenex Implicits eqseqP hasP hasPn allP allPn. Section NseqthTheory. Lemma nthP (T : eqType) (s : seq T) x x0 : reflect (exists2 i, i < size s & nth x0 s i = x) (x \in s). Proof. apply: (iffP idP) => [|[n Hn <-]]; last by apply mem_nth. by exists (index x s); [rewrite index_mem | apply nth_index]. Qed. Variable T : Type. Lemma has_nthP (a : pred T) s x0 : reflect (exists2 i, i < size s & a (nth x0 s i)) (has a s). Proof. elim: s => [|x s IHs] /=; first by right; case. case nax: (a x); first by left; exists 0. by apply: (iffP IHs) => [[i]|[[|i]]]; [exists i.+1 | rewrite nax | exists i]. Qed. Lemma all_nthP (a : pred T) s x0 : reflect (forall i, i < size s -> a (nth x0 s i)) (all a s). Proof. rewrite -(eq_all (fun x => negbK (a x))) all_predC. case: (has_nthP _ _ x0) => [na_s | a_s]; [right=> a_s | left=> i lti]. by case: na_s => i lti; rewrite a_s. by apply/idPn=> na_si; case: a_s; exists i. Qed. End NseqthTheory. Lemma set_nth_default T s (y0 x0 : T) n : n < size s -> nth x0 s n = nth y0 s n. Proof. by elim: s n => [|y s' IHs] [|n] /=; auto. Qed. Lemma headI T s (x : T) : rcons s x = head x s :: behead (rcons s x). Proof. by case: s. Qed. Implicit Arguments nthP [T s x]. Implicit Arguments has_nthP [T a s]. Implicit Arguments all_nthP [T a s]. Prenex Implicits nthP has_nthP all_nthP. Definition bitseq := seq bool. Canonical bitseq_eqType := Eval hnf in [eqType of bitseq]. Canonical bitseq_predType := Eval hnf in [predType of bitseq]. (* Incrementing the ith nat in a seq nat, padding with 0's if needed. This *) (* allows us to use nat seqs as bags of nats. *) Fixpoint incr_nth v i {struct i} := if v is n :: v' then if i is i'.+1 then n :: incr_nth v' i' else n.+1 :: v' else ncons i 0 [:: 1]. Lemma nth_incr_nth v i j : nth 0 (incr_nth v i) j = (i == j) + nth 0 v j. Proof. elim: v i j => [|n v IHv] [|i] [|j] //=; rewrite ?eqSS ?addn0 //; try by case j. elim: i j => [|i IHv] [|j] //=; rewrite ?eqSS //; by case j. Qed. Lemma size_incr_nth v i : size (incr_nth v i) = if i < size v then size v else i.+1. Proof. elim: v i => [|n v IHv] [|i] //=; first by rewrite size_ncons /= addn1. rewrite IHv; exact: fun_if. Qed. (* equality up to permutation *) Section PermSeq. Variable T : eqType. Implicit Type s : seq T. Definition same_count1 s1 s2 x := count (pred1 x) s1 == count (pred1 x) s2. Definition perm_eq s1 s2 := all (same_count1 s1 s2) (s1 ++ s2). Lemma perm_eqP s1 s2 : reflect (count^~ s1 =1 count^~ s2) (perm_eq s1 s2). Proof. apply: (iffP allP) => [eq_cnt1 a | eq_cnt x _]; last exact/eqP. elim: {a}_.+1 {-2}a (ltnSn (count a (s1 ++ s2))) => // n IHn a le_an. case: (posnP (count a (s1 ++ s2))). by move/eqP; rewrite count_cat addn_eq0; do 2!case: eqP => // ->. rewrite -has_count => /hasP[x s12x a_x]; pose a' := predD1 a x. have cnt_a' s : count a s = count (pred1 x) s + count a' s. rewrite count_filter -(count_predC (pred1 x)) 2!count_filter. rewrite -!filter_predI -!count_filter; congr (_ + _). by apply: eq_count => y /=; case: eqP => // ->. rewrite !cnt_a' (eqnP (eq_cnt1 _ s12x)) (IHn a') // -ltnS. apply: leq_trans le_an. by rewrite ltnS cnt_a' -add1n leq_add2r -has_count has_pred1. Qed. Lemma perm_eq_refl s : perm_eq s s. Proof. exact/perm_eqP. Qed. Hint Resolve perm_eq_refl. Lemma perm_eq_sym : symmetric perm_eq. Proof. by move=> s1 s2; apply/perm_eqP/perm_eqP=> ? ?. Qed. Lemma perm_eq_trans : transitive perm_eq. Proof. move=> s2 s1 s3; move/perm_eqP=> eq12; move/perm_eqP=> eq23. by apply/perm_eqP=> a; rewrite eq12. Qed. Notation perm_eql s1 s2 := (perm_eq s1 =1 perm_eq s2). Notation perm_eqr s1 s2 := (perm_eq^~ s1 =1 perm_eq^~ s2). Lemma perm_eqlE s1 s2 : perm_eql s1 s2 -> perm_eq s1 s2. Proof. by move->. Qed. Lemma perm_eqlP s1 s2 : reflect (perm_eql s1 s2) (perm_eq s1 s2). Proof. apply: (iffP idP) => [eq12 s3 | -> //]. apply/idP/idP; last exact: perm_eq_trans. by rewrite -!(perm_eq_sym s3); move/perm_eq_trans; apply. Qed. Lemma perm_eqrP s1 s2 : reflect (perm_eqr s1 s2) (perm_eq s1 s2). Proof. apply: (iffP idP) => [/perm_eqlP eq12 s3| <- //]. by rewrite !(perm_eq_sym s3) eq12. Qed. Lemma perm_catC s1 s2 : perm_eql (s1 ++ s2) (s2 ++ s1). Proof. by apply/perm_eqlP; apply/perm_eqP=> a; rewrite !count_cat addnC. Qed. Lemma perm_cat2l s1 s2 s3 : perm_eq (s1 ++ s2) (s1 ++ s3) = perm_eq s2 s3. Proof. apply/perm_eqP/perm_eqP=> eq23 a; apply/eqP; by move/(_ a)/eqP: eq23; rewrite !count_cat eqn_add2l. Qed. Lemma perm_cons x s1 s2 : perm_eq (x :: s1) (x :: s2) = perm_eq s1 s2. Proof. exact: (perm_cat2l [::x]). Qed. Lemma perm_cat2r s1 s2 s3 : perm_eq (s2 ++ s1) (s3 ++ s1) = perm_eq s2 s3. Proof. by do 2!rewrite perm_eq_sym perm_catC; exact: perm_cat2l. Qed. Lemma perm_catAC s1 s2 s3 : perm_eql ((s1 ++ s2) ++ s3) ((s1 ++ s3) ++ s2). Proof. by apply/perm_eqlP; rewrite -!catA perm_cat2l perm_catC. Qed. Lemma perm_catCA s1 s2 s3 : perm_eql (s1 ++ s2 ++ s3) (s2 ++ s1 ++ s3). Proof. by apply/perm_eqlP; rewrite !catA perm_cat2r perm_catC. Qed. Lemma perm_rcons x s : perm_eql (rcons s x) (x :: s). Proof. by move=> /= s2; rewrite -cats1 perm_catC. Qed. Lemma perm_rot n s : perm_eql (rot n s) s. Proof. by move=> /= s2; rewrite perm_catC cat_take_drop. Qed. Lemma perm_rotr n s : perm_eql (rotr n s) s. Proof. exact: perm_rot. Qed. Lemma perm_filterC a s : perm_eql (filter a s ++ filter (predC a) s) s. Proof. apply/perm_eqlP; elim: s => //= x s IHs. by case: (a x); last rewrite /= -cat1s perm_catCA; rewrite perm_cons. Qed. Lemma perm_eq_mem s1 s2 : perm_eq s1 s2 -> s1 =i s2. Proof. by move/perm_eqP=> eq12 x; rewrite -!has_pred1 !has_count eq12. Qed. Lemma perm_eq_size s1 s2 : perm_eq s1 s2 -> size s1 = size s2. Proof. by move/perm_eqP=> eq12; rewrite -!count_predT eq12. Qed. Lemma perm_eq_small s1 s2 : size s2 <= 1 -> perm_eq s1 s2 -> s1 = s2. Proof. move=> s2_le1 eqs12; move/perm_eq_size: eqs12 s2_le1 (perm_eq_mem eqs12). by case: s2 s1 => [|x []] // [|y []] // _ _ /(_ x); rewrite !inE eqxx => /eqP->. Qed. Lemma uniq_leq_size s1 s2 : uniq s1 -> {subset s1 <= s2} -> size s1 <= size s2. Proof. elim: s1 s2 => //= x s1 IHs s2 /andP[not_s1x Us1] /allP/=/andP[s2x /allP ss12]. have [i s3 def_s2] := rot_to s2x; rewrite -(size_rot i s2) def_s2. apply: IHs => // y s1y; have:= ss12 y s1y. by rewrite -(mem_rot i) def_s2 inE (negPf (memPn _ y s1y)). Qed. Lemma leq_size_uniq s1 s2 : uniq s1 -> {subset s1 <= s2} -> size s2 <= size s1 -> uniq s2. Proof. elim: s1 s2 => [[] | x s1 IHs s2] // Us1x; have /andP[not_s1x Us1] := Us1x. case/allP/andP=> /rot_to[i s3 def_s2] /allP ss12 le_s21. rewrite -(rot_uniq i) -(size_rot i) def_s2 /= in le_s21 *. have ss13 y (s1y : y \in s1): y \in s3. by have:= ss12 y s1y; rewrite -(mem_rot i) def_s2 inE (negPf (memPn _ y s1y)). rewrite IHs // andbT; apply: contraL _ le_s21 => s3x; rewrite -leqNgt. by apply/(uniq_leq_size Us1x)/allP; rewrite /= s3x; exact/allP. Qed. Lemma uniq_size_uniq s1 s2 : uniq s1 -> s1 =i s2 -> uniq s2 = (size s2 == size s1). Proof. move=> Us1 eqs12; apply/idP/idP=> [Us2 | /eqP eq_sz12]. by rewrite eqn_leq !uniq_leq_size // => y; rewrite eqs12. by apply: (leq_size_uniq Us1) => [y|]; rewrite (eqs12, eq_sz12). Qed. Lemma leq_size_perm s1 s2 : uniq s1 -> {subset s1 <= s2} -> size s2 <= size s1 -> s1 =i s2 /\ size s1 = size s2. Proof. move=> Us1 ss12 le_s21; have Us2: uniq s2 := leq_size_uniq Us1 ss12 le_s21. suffices: s1 =i s2 by split; last by apply/eqP; rewrite -uniq_size_uniq. move=> x; apply/idP/idP=> [/ss12// | s2x]; apply: contraLR le_s21 => not_s1x. rewrite -ltnNge (@uniq_leq_size (x :: s1)) /= ?not_s1x //. by apply/allP; rewrite /= s2x; exact/allP. Qed. Lemma perm_uniq s1 s2 : s1 =i s2 -> size s1 = size s2 -> uniq s1 = uniq s2. Proof. by move=> Es12 Hs12; apply/idP/idP => Us; rewrite (uniq_size_uniq Us) ?Hs12 ?eqxx. Qed. Lemma perm_eq_uniq s1 s2 : perm_eq s1 s2 -> uniq s1 = uniq s2. Proof. move=> eq_s12; apply: perm_uniq; [exact: perm_eq_mem | exact: perm_eq_size]. Qed. Lemma uniq_perm_eq s1 s2 : uniq s1 -> uniq s2 -> s1 =i s2 -> perm_eq s1 s2. Proof. move=> Us1 Us2 eq12; apply/allP=> x _; apply/eqP. by rewrite !count_uniq_mem ?eq12. Qed. Lemma count_mem_uniq s : (forall x, count (pred1 x) s = (x \in s)) -> uniq s. Proof. move=> count1_s; have Uus := undup_uniq s. suffices: perm_eq s (undup s) by move/perm_eq_uniq->. by apply/allP=> x _; apply/eqP; rewrite (count_uniq_mem x Uus) mem_undup. Qed. Lemma catCA_perm_ind P : (forall s1 s2 s3, P (s1 ++ s2 ++ s3) -> P (s2 ++ s1 ++ s3)) -> (forall s1 s2, perm_eq s1 s2 -> P s1 -> P s2). Proof. move=> PcatCA s1 s2 eq_s12; rewrite -[s1]cats0 -[s2]cats0. elim: s2 nil => [| x s2 IHs] s3 in s1 eq_s12 *. by case: s1 {eq_s12}(perm_eq_size eq_s12). have /rot_to[i s' def_s1]: x \in s1 by rewrite (perm_eq_mem eq_s12) mem_head. rewrite -(cat_take_drop i s1) -catA => /PcatCA. rewrite catA -/(rot i s1) def_s1 /= -cat1s => /PcatCA/IHs/PcatCA; apply. by rewrite -(perm_cons x) -def_s1 perm_rot. Qed. Lemma catCA_perm_subst R F : (forall s1 s2 s3, F (s1 ++ s2 ++ s3) = F (s2 ++ s1 ++ s3) :> R) -> (forall s1 s2, perm_eq s1 s2 -> F s1 = F s2). Proof. move=> FcatCA s1 s2 /catCA_perm_ind => ind_s12. by apply: (ind_s12 (eq _ \o F)) => //= *; rewrite FcatCA. Qed. End PermSeq. Notation perm_eql s1 s2 := (perm_eq s1 =1 perm_eq s2). Notation perm_eqr s1 s2 := (perm_eq^~ s1 =1 perm_eq^~ s2). Implicit Arguments perm_eqP [T s1 s2]. Implicit Arguments perm_eqlP [T s1 s2]. Implicit Arguments perm_eqrP [T s1 s2]. Prenex Implicits perm_eq perm_eqP perm_eqlP perm_eqrP. Hint Resolve perm_eq_refl. Section RotrLemmas. Variables (n0 : nat) (T : Type) (T' : eqType). Implicit Type s : seq T. Lemma size_rotr s : size (rotr n0 s) = size s. Proof. by rewrite size_rot. Qed. Lemma mem_rotr (s : seq T') : rotr n0 s =i s. Proof. by move=> x; rewrite mem_rot. Qed. Lemma rotr_size_cat s1 s2 : rotr (size s2) (s1 ++ s2) = s2 ++ s1. Proof. by rewrite /rotr size_cat addnK rot_size_cat. Qed. Lemma rotr1_rcons x s : rotr 1 (rcons s x) = x :: s. Proof. by rewrite -rot1_cons rotK. Qed. Lemma has_rotr a s : has a (rotr n0 s) = has a s. Proof. by rewrite has_rot. Qed. Lemma rotr_uniq (s : seq T') : uniq (rotr n0 s) = uniq s. Proof. by rewrite rot_uniq. Qed. Lemma rotrK : cancel (@rotr T n0) (rot n0). Proof. move=> s; have [lt_n0s | ge_n0s] := ltnP n0 (size s). by rewrite -{1}(subKn (ltnW lt_n0s)) -{1}[size s]size_rotr; exact: rotK. by rewrite -{2}(rot_oversize ge_n0s) /rotr (eqnP ge_n0s) rot0. Qed. Lemma rotr_inj : injective (@rotr T n0). Proof. exact (can_inj rotrK). Qed. Lemma rev_rot s : rev (rot n0 s) = rotr n0 (rev s). Proof. rewrite /rotr size_rev -{3}(cat_take_drop n0 s) {1}/rot !rev_cat. by rewrite -size_drop -size_rev rot_size_cat. Qed. Lemma rev_rotr s : rev (rotr n0 s) = rot n0 (rev s). Proof. apply: canRL rotrK _; rewrite {1}/rotr size_rev size_rotr /rotr {2}/rot rev_cat. set m := size s - n0; rewrite -{1}(@size_takel m _ _ (leq_subr _ _)). by rewrite -size_rev rot_size_cat -rev_cat cat_take_drop. Qed. End RotrLemmas. Section RotCompLemmas. Variable T : Type. Implicit Type s : seq T. Lemma rot_addn m n s : m + n <= size s -> rot (m + n) s = rot m (rot n s). Proof. move=> sz_s; rewrite {1}/rot -[take _ s](cat_take_drop n). rewrite 5!(catA, =^~ rot_size_cat) !cat_take_drop. by rewrite size_drop !size_takel ?leq_addl ?addnK. Qed. Lemma rotS n s : n < size s -> rot n.+1 s = rot 1 (rot n s). Proof. exact: (@rot_addn 1). Qed. Lemma rot_add_mod m n s : n <= size s -> m <= size s -> rot m (rot n s) = rot (if m + n <= size s then m + n else m + n - size s) s. Proof. move=> Hn Hm; case: leqP => [/rot_addn // | /ltnW Hmn]; symmetry. by rewrite -{2}(rotK n s) /rotr -rot_addn size_rot addnBA ?subnK ?addnK. Qed. Lemma rot_rot m n s : rot m (rot n s) = rot n (rot m s). Proof. case: (ltnP (size s) m) => Hm. by rewrite !(@rot_oversize T m) ?size_rot 1?ltnW. case: (ltnP (size s) n) => Hn. by rewrite !(@rot_oversize T n) ?size_rot 1?ltnW. by rewrite !rot_add_mod 1?addnC. Qed. Lemma rot_rotr m n s : rot m (rotr n s) = rotr n (rot m s). Proof. by rewrite {2}/rotr size_rot rot_rot. Qed. Lemma rotr_rotr m n s : rotr m (rotr n s) = rotr n (rotr m s). Proof. by rewrite /rotr !size_rot rot_rot. Qed. End RotCompLemmas. Section Mask. Variables (n0 : nat) (T : Type). Implicit Types (m : bitseq) (s : seq T). Fixpoint mask m s {struct m} := match m, s with | b :: m', x :: s' => if b then x :: mask m' s' else mask m' s' | _, _ => [::] end. Lemma mask_false s n : mask (nseq n false) s = [::]. Proof. by elim: s n => [|x s IHs] [|n] /=. Qed. Lemma mask_true s n : size s <= n -> mask (nseq n true) s = s. Proof. by elim: s n => [|x s IHs] [|n] //= Hn; congr (_ :: _); apply: IHs. Qed. Lemma mask0 m : mask m [::] = [::]. Proof. by case: m. Qed. Lemma mask1 b x : mask [:: b] [:: x] = nseq b x. Proof. by case: b. Qed. Lemma mask_cons b m x s : mask (b :: m) (x :: s) = nseq b x ++ mask m s. Proof. by case: b. Qed. Lemma size_mask m s : size m = size s -> size (mask m s) = count id m. Proof. by elim: m s => [|b m IHm] [|x s] //= [Hs]; case: b; rewrite /= IHm. Qed. Lemma mask_cat m1 s1 : size m1 = size s1 -> forall m2 s2, mask (m1 ++ m2) (s1 ++ s2) = mask m1 s1 ++ mask m2 s2. Proof. move=> Hm1 m2 s2; elim: m1 s1 Hm1 => [|b1 m1 IHm] [|x1 s1] //= [Hm1]. by rewrite (IHm _ Hm1); case b1. Qed. Lemma has_mask_cons a b m x s : has a (mask (b :: m) (x :: s)) = b && a x || has a (mask m s). Proof. by case: b. Qed. Lemma has_mask a m s : has a (mask m s) -> has a s. Proof. elim: m s => [|b m IHm] [|x s] //; rewrite has_mask_cons /= andbC. by case: (a x) => //= /IHm. Qed. Lemma mask_rot m s : size m = size s -> mask (rot n0 m) (rot n0 s) = rot (count id (take n0 m)) (mask m s). Proof. move=> Hs. have Hsn0: size (take n0 m) = size (take n0 s) by rewrite !size_take Hs. rewrite -(size_mask Hsn0) {1 2}/rot mask_cat ?size_drop ?Hs //. rewrite -{4}(cat_take_drop n0 m) -{4}(cat_take_drop n0 s) mask_cat //. by rewrite rot_size_cat. Qed. Lemma resize_mask m s : {m1 | size m1 = size s & mask m s = mask m1 s}. Proof. by exists (take (size s) m ++ nseq (size s - size m) false); elim: s m => [|x s IHs] [|b m] //=; rewrite (size_nseq, mask_false, IHs). Qed. End Mask. Section EqMask. Variables (n0 : nat) (T : eqType). Implicit Types (s : seq T) (m : bitseq). Lemma mem_mask_cons x b m y s : (x \in mask (b :: m) (y :: s)) = b && (x == y) || (x \in mask m s). Proof. by case: b. Qed. Lemma mem_mask x m s : x \in mask m s -> x \in s. Proof. by rewrite -!has_pred1 => /has_mask. Qed. Lemma mask_uniq s : uniq s -> forall m, uniq (mask m s). Proof. elim: s => [|x s IHs] Uxs [|b m] //=. case: b Uxs => //= /andP[s'x Us]; rewrite {}IHs // andbT. by apply: contra s'x; exact: mem_mask. Qed. Lemma mem_mask_rot m s : size m = size s -> mask (rot n0 m) (rot n0 s) =i mask m s. Proof. by move=> Hm x; rewrite mask_rot // mem_rot. Qed. End EqMask. Section Subseq. Variable T : eqType. Implicit Type s : seq T. Fixpoint subseq s1 s2 := if s2 is y :: s2' then if s1 is x :: s1' then subseq (if x == y then s1' else s1) s2' else true else s1 == [::]. Lemma sub0seq s : subseq [::] s. Proof. by case: s. Qed. Lemma subseq0 s : subseq s [::] = (s == [::]). Proof. by []. Qed. Lemma subseqP s1 s2 : reflect (exists2 m, size m = size s2 & s1 = mask m s2) (subseq s1 s2). Proof. elim: s2 s1 => [|y s2 IHs2] [|x s1]. - by left; exists [::]. - by right; do 2!case. - by left; exists (nseq (size s2).+1 false); rewrite ?size_nseq //= mask_false. apply: {IHs2}(iffP (IHs2 _)) => [] [m sz_m def_s1]. by exists ((x == y) :: m); rewrite /= ?sz_m // -def_s1; case: eqP => // ->. case: eqP => [_ | ne_xy]; last first. by case: m def_s1 sz_m => [//|[m []//|m]] -> [<-]; exists m. pose i := index true m; have def_m_i: take i m = nseq (size (take i m)) false. apply/all_pred1P; apply/(all_nthP true) => j. rewrite size_take ltnNge geq_min negb_or -ltnNge; case/andP=> lt_j_i _. rewrite nth_take //= -negb_add addbF -addbT -negb_eqb. by rewrite [_ == _](before_find _ lt_j_i). have lt_i_m: i < size m. rewrite ltnNge; apply/negP=> le_m_i; rewrite take_oversize // in def_m_i. by rewrite def_m_i mask_false in def_s1. rewrite size_take lt_i_m in def_m_i. exists (take i m ++ drop i.+1 m). rewrite size_cat size_take size_drop lt_i_m. by rewrite sz_m in lt_i_m *; rewrite subnKC. rewrite {s1 def_s1}[s1](congr1 behead def_s1). rewrite -[s2](cat_take_drop i) -{1}[m](cat_take_drop i) {}def_m_i -cat_cons. have sz_i_s2: size (take i s2) = i by apply: size_takel; rewrite sz_m in lt_i_m. rewrite lastI cat_rcons !mask_cat ?size_nseq ?size_belast ?mask_false //=. by rewrite (drop_nth true) // nth_index -?index_mem. Qed. Lemma mask_subseq m s : subseq (mask m s) s. Proof. by apply/subseqP; have [m1] := resize_mask m s; exists m1. Qed. Lemma subseq_trans : transitive subseq. Proof. move=> _ _ s /subseqP[m2 _ ->] /subseqP[m1 _ ->]. elim: s => [|x s IHs] in m2 m1 *; first by rewrite !mask0. case: m1 => [|[] m1]; first by rewrite mask0. case: m2 => [|[] m2] //; first by rewrite /= eqxx IHs. case/subseqP: (IHs m2 m1) => m sz_m def_s; apply/subseqP. by exists (false :: m); rewrite //= sz_m. case/subseqP: (IHs m2 m1) => m sz_m def_s; apply/subseqP. by exists (false :: m); rewrite //= sz_m. Qed. Lemma subseq_refl s : subseq s s. Proof. by elim: s => //= x s IHs; rewrite eqxx. Qed. Hint Resolve subseq_refl. Lemma cat_subseq s1 s2 s3 s4 : subseq s1 s3 -> subseq s2 s4 -> subseq (s1 ++ s2) (s3 ++ s4). Proof. case/subseqP=> m1 sz_m1 ->; case/subseqP=> m2 sz_m2 ->; apply/subseqP. by exists (m1 ++ m2); rewrite ?size_cat ?mask_cat ?sz_m1 ?sz_m2. Qed. Lemma prefix_subseq s1 s2 : subseq s1 (s1 ++ s2). Proof. by rewrite -{1}[s1]cats0 cat_subseq ?sub0seq. Qed. Lemma suffix_subseq s1 s2 : subseq s2 (s1 ++ s2). Proof. by rewrite -{1}[s2]cat0s cat_subseq ?sub0seq. Qed. Lemma mem_subseq s1 s2 : subseq s1 s2 -> {subset s1 <= s2}. Proof. by case/subseqP=> m _ -> x; exact: mem_mask. Qed. Lemma sub1seq x s : subseq [:: x] s = (x \in s). Proof. by elim: s => //= y s; rewrite inE; case: (x == y); rewrite ?sub0seq. Qed. Lemma size_subseq s1 s2 : subseq s1 s2 -> size s1 <= size s2. Proof. by case/subseqP=> m sz_m ->; rewrite size_mask -sz_m ?count_size. Qed. Lemma size_subseq_leqif s1 s2 : subseq s1 s2 -> size s1 <= size s2 ?= iff (s1 == s2). Proof. move=> sub12; split; first exact: size_subseq. apply/idP/eqP=> [|-> //]; case/subseqP: sub12 => m sz_m ->{s1}. rewrite size_mask -sz_m // -all_count -(eq_all eqb_id). by move/(@all_pred1P _ true)->; rewrite sz_m mask_true. Qed. Lemma subseq_cons s x : subseq s (x :: s). Proof. by exact: (@cat_subseq [::] s [:: x]). Qed. Lemma subseq_rcons s x : subseq s (rcons s x). Proof. by rewrite -{1}[s]cats0 -cats1 cat_subseq. Qed. Lemma subseq_uniq s1 s2 : subseq s1 s2 -> uniq s2 -> uniq s1. Proof. by case/subseqP=> m _ -> Us2; exact: mask_uniq. Qed. End Subseq. Prenex Implicits subseq. Implicit Arguments subseqP [T s1 s2]. Hint Resolve subseq_refl. Section Rem. Variables (T : eqType) (x : T). Fixpoint rem s := if s is y :: t then (if y == x then t else y :: rem t) else s. Lemma rem_id s : x \notin s -> rem s = s. Proof. by elim: s => //= y s IHs /norP[neq_yx /IHs->]; rewrite eq_sym (negbTE neq_yx). Qed. Lemma perm_to_rem s : x \in s -> perm_eq s (x :: rem s). Proof. elim: s => // y s IHs; rewrite inE /= eq_sym perm_eq_sym. case: eqP => [-> // | _ /IHs]. by rewrite (perm_catCA [:: x] [:: y]) perm_cons perm_eq_sym. Qed. Lemma size_rem s : x \in s -> size (rem s) = (size s).-1. Proof. by move/perm_to_rem/perm_eq_size->. Qed. Lemma rem_subseq s : subseq (rem s) s. Proof. elim: s => //= y s IHs; rewrite eq_sym. by case: ifP => _; [exact: subseq_cons | rewrite eqxx]. Qed. Lemma rem_uniq s : uniq s -> uniq (rem s). Proof. by apply: subseq_uniq; exact: rem_subseq. Qed. Lemma mem_rem s : {subset rem s <= s}. Proof. exact: mem_subseq (rem_subseq s). Qed. Lemma rem_filter s : uniq s -> rem s = filter (predC1 x) s. Proof. elim: s => //= y s IHs /andP[not_s_y /IHs->]. by case: eqP => //= <-; apply/esym/all_filterP; rewrite all_predC has_pred1. Qed. Lemma mem_rem_uniq s : uniq s -> rem s =i [predD1 s & x]. Proof. by move/rem_filter=> -> y; rewrite mem_filter. Qed. End Rem. Section Map. Variables (n0 : nat) (T1 : Type) (x1 : T1). Variables (T2 : Type) (x2 : T2) (f : T1 -> T2). Fixpoint map s := if s is x :: s' then f x :: map s' else [::]. Lemma map_cons x s : map (x :: s) = f x :: map s. Proof. by []. Qed. Lemma map_nseq x : map (nseq n0 x) = nseq n0 (f x). Proof. by elim: n0 => // *; congr (_ :: _). Qed. Lemma map_cat s1 s2 : map (s1 ++ s2) = map s1 ++ map s2. Proof. by elim: s1 => [|x s1 IHs] //=; rewrite IHs. Qed. Lemma size_map s : size (map s) = size s. Proof. by elim: s => //= x s ->. Qed. Lemma behead_map s : behead (map s) = map (behead s). Proof. by case: s. Qed. Lemma nth_map n s : n < size s -> nth x2 (map s) n = f (nth x1 s n). Proof. by elim: s n => [|x s IHs] [|n]; try exact (IHs n). Qed. Lemma map_rcons s x : map (rcons s x) = rcons (map s) (f x). Proof. by rewrite -!cats1 map_cat. Qed. Lemma last_map s x : last (f x) (map s) = f (last x s). Proof. by elim: s x => /=. Qed. Lemma belast_map s x : belast (f x) (map s) = map (belast x s). Proof. by elim: s x => //= y s IHs x; rewrite IHs. Qed. Lemma filter_map a s : filter a (map s) = map (filter (preim f a) s). Proof. by elim: s => //= x s IHs; rewrite (fun_if map) /= IHs. Qed. Lemma find_map a s : find a (map s) = find (preim f a) s. Proof. by elim: s => //= x s ->. Qed. Lemma has_map a s : has a (map s) = has (preim f a) s. Proof. by elim: s => //= x s ->. Qed. Lemma all_map a s : all a (map s) = all (preim f a) s. Proof. by elim: s => //= x s ->. Qed. Lemma count_map a s : count a (map s) = count (preim f a) s. Proof. by elim: s => //= x s ->. Qed. Lemma map_take s : map (take n0 s) = take n0 (map s). Proof. by elim: n0 s => [|n IHn] [|x s] //=; rewrite IHn. Qed. Lemma map_drop s : map (drop n0 s) = drop n0 (map s). Proof. by elim: n0 s => [|n IHn] [|x s] //=; rewrite IHn. Qed. Lemma map_rot s : map (rot n0 s) = rot n0 (map s). Proof. by rewrite /rot map_cat map_take map_drop. Qed. Lemma map_rotr s : map (rotr n0 s) = rotr n0 (map s). Proof. by apply: canRL (@rotK n0 T2) _; rewrite -map_rot rotrK. Qed. Lemma map_rev s : map (rev s) = rev (map s). Proof. by elim: s => //= x s IHs; rewrite !rev_cons -!cats1 map_cat IHs. Qed. Lemma map_mask m s : map (mask m s) = mask m (map s). Proof. by elim: m s => [|[|] m IHm] [|x p] //=; rewrite IHm. Qed. Lemma inj_map : injective f -> injective map. Proof. by move=> injf; elim=> [|y1 s1 IHs] [|y2 s2] //= [/injf-> /IHs->]. Qed. End Map. Notation "[ 'seq' E | i <- s ]" := (map (fun i => E) s) (at level 0, E at level 99, i ident, format "[ '[hv' 'seq' E '/ ' | i <- s ] ']'") : seq_scope. Notation "[ 'seq' E | i <- s & C ]" := [seq E | i <- [seq i <- s | C]] (at level 0, E at level 99, i ident, format "[ '[hv' 'seq' E '/ ' | i <- s '/ ' & C ] ']'") : seq_scope. Notation "[ 'seq' E | i : T <- s ]" := (map (fun i : T => E) s) (at level 0, E at level 99, i ident, only parsing) : seq_scope. Notation "[ 'seq' E | i : T <- s & C ]" := [seq E | i : T <- [seq i : T <- s | C]] (at level 0, E at level 99, i ident, only parsing) : seq_scope. Lemma filter_mask T a (s : seq T) : filter a s = mask (map a s) s. Proof. by elim: s => //= x s <-; case: (a x). Qed. Section FilterSubseq. Variable T : eqType. Implicit Types (s : seq T) (a : pred T). Lemma filter_subseq a s : subseq (filter a s) s. Proof. by apply/subseqP; exists (map a s); rewrite ?size_map ?filter_mask. Qed. Lemma subseq_filter s1 s2 a : subseq s1 (filter a s2) = all a s1 && subseq s1 s2. Proof. elim: s2 s1 => [|x s2 IHs] [|y s1] //=; rewrite ?andbF ?sub0seq //. by case a_x: (a x); rewrite /= !IHs /=; case: eqP => // ->; rewrite a_x. Qed. Lemma subseq_uniqP s1 s2 : uniq s2 -> reflect (s1 = filter (mem s1) s2) (subseq s1 s2). Proof. move=> uniq_s2; apply: (iffP idP) => [ss12 | ->]; last exact: filter_subseq. apply/eqP; rewrite -size_subseq_leqif ?subseq_filter ?(introT allP) //. apply/eqP/esym/perm_eq_size. rewrite uniq_perm_eq ?filter_uniq ?(subseq_uniq ss12) // => x. by rewrite mem_filter; apply: andb_idr; exact: (mem_subseq ss12). Qed. Lemma perm_to_subseq s1 s2 : subseq s1 s2 -> {s3 | perm_eq s2 (s1 ++ s3)}. Proof. elim Ds2: s2 s1 => [|y s2' IHs] [|x s1] //=; try by exists s2; rewrite Ds2. case: eqP => [-> | _] /IHs[s3 perm_s2] {IHs}. by exists s3; rewrite perm_cons. by exists (rcons s3 y); rewrite -cat_cons -perm_rcons -!cats1 catA perm_cat2r. Qed. End FilterSubseq. Implicit Arguments subseq_uniqP [T s1 s2]. Section EqMap. Variables (n0 : nat) (T1 : eqType) (x1 : T1). Variables (T2 : eqType) (x2 : T2) (f : T1 -> T2). Implicit Type s : seq T1. Lemma map_f s x : x \in s -> f x \in map f s. Proof. elim: s => [|y s IHs] //=. by case/predU1P=> [->|Hx]; [exact: predU1l | apply: predU1r; auto]. Qed. Lemma mapP s y : reflect (exists2 x, x \in s & y = f x) (y \in map f s). Proof. elim: s => [|x s IHs]; first by right; case. rewrite /= in_cons eq_sym; case Hxy: (f x == y). by left; exists x; [rewrite mem_head | rewrite (eqP Hxy)]. apply: (iffP IHs) => [[x' Hx' ->]|[x' Hx' Dy]]. by exists x'; first exact: predU1r. by move: Dy Hxy => ->; case/predU1P: Hx' => [->|]; [rewrite eqxx | exists x']. Qed. Lemma map_uniq s : uniq (map f s) -> uniq s. Proof. elim: s => //= x s IHs /andP[not_sfx /IHs->]; rewrite andbT. by apply: contra not_sfx => sx; apply/mapP; exists x. Qed. Lemma map_inj_in_uniq s : {in s &, injective f} -> uniq (map f s) = uniq s. Proof. elim: s => //= x s IHs //= injf; congr (~~ _ && _). apply/mapP/idP=> [[y sy /injf] | ]; last by exists x. by rewrite mem_head mem_behead // => ->. apply: IHs => y z sy sz; apply: injf => //; exact: predU1r. Qed. Lemma map_subseq s1 s2 : subseq s1 s2 -> subseq (map f s1) (map f s2). Proof. case/subseqP=> m sz_m ->; apply/subseqP. by exists m; rewrite ?size_map ?map_mask. Qed. Lemma nth_index_map s x0 x : {in s &, injective f} -> x \in s -> nth x0 s (index (f x) (map f s)) = x. Proof. elim: s => //= y s IHs inj_f s_x; rewrite (inj_in_eq inj_f) ?mem_head //. move: s_x; rewrite inE eq_sym; case: eqP => [-> | _] //=; apply: IHs. by apply: sub_in2 inj_f => z; exact: predU1r. Qed. Lemma perm_map s t : perm_eq s t -> perm_eq (map f s) (map f t). Proof. by move/perm_eqP=> Est; apply/perm_eqP=> a; rewrite !count_map Est. Qed. Hypothesis Hf : injective f. Lemma mem_map s x : (f x \in map f s) = (x \in s). Proof. by apply/mapP/idP=> [[y Hy /Hf->] //|]; exists x. Qed. Lemma index_map s x : index (f x) (map f s) = index x s. Proof. by rewrite /index; elim: s => //= y s IHs; rewrite (inj_eq Hf) IHs. Qed. Lemma map_inj_uniq s : uniq (map f s) = uniq s. Proof. apply: map_inj_in_uniq; exact: in2W. Qed. End EqMap. Implicit Arguments mapP [T1 T2 f s y]. Prenex Implicits mapP. Lemma map_of_seq (T1 : eqType) T2 (s : seq T1) (fs : seq T2) (y0 : T2) : {f | uniq s -> size fs = size s -> map f s = fs}. Proof. exists (fun x => nth y0 fs (index x s)) => uAs eq_sz. apply/esym/(@eq_from_nth _ y0); rewrite ?size_map eq_sz // => i ltis. have x0 : T1 by [case: (s) ltis]; by rewrite (nth_map x0) // index_uniq. Qed. Section MapComp. Variable T1 T2 T3 : Type. Lemma map_id (s : seq T1) : map id s = s. Proof. by elim: s => //= x s ->. Qed. Lemma eq_map (f1 f2 : T1 -> T2) : f1 =1 f2 -> map f1 =1 map f2. Proof. by move=> Ef; elim=> //= x s ->; rewrite Ef. Qed. Lemma map_comp (f1 : T2 -> T3) (f2 : T1 -> T2) s : map (f1 \o f2) s = map f1 (map f2 s). Proof. by elim: s => //= x s ->. Qed. Lemma mapK (f1 : T1 -> T2) (f2 : T2 -> T1) : cancel f1 f2 -> cancel (map f1) (map f2). Proof. by move=> eq_f12; elim=> //= x s ->; rewrite eq_f12. Qed. End MapComp. Lemma eq_in_map (T1 : eqType) T2 (f1 f2 : T1 -> T2) (s : seq T1) : {in s, f1 =1 f2} <-> map f1 s = map f2 s. Proof. elim: s => //= x s IHs; split=> [eqf12 | [f12x /IHs eqf12]]; last first. by move=> y /predU1P[-> | /eqf12]. rewrite eqf12 ?mem_head //; congr (_ :: _). by apply/IHs=> y s_y; rewrite eqf12 // mem_behead. Qed. Lemma map_id_in (T : eqType) f (s : seq T) : {in s, f =1 id} -> map f s = s. Proof. by move/eq_in_map->; apply: map_id. Qed. (* Map a partial function *) Section Pmap. Variables (aT rT : Type) (f : aT -> option rT) (g : rT -> aT). Fixpoint pmap s := if s is x :: s' then let r := pmap s' in oapp (cons^~ r) r (f x) else [::]. Lemma map_pK : pcancel g f -> cancel (map g) pmap. Proof. by move=> gK; elim=> //= x s ->; rewrite gK. Qed. Lemma size_pmap s : size (pmap s) = count [eta f] s. Proof. by elim: s => //= x s <-; case: (f _). Qed. Lemma pmapS_filter s : map some (pmap s) = map f (filter [eta f] s). Proof. by elim: s => //= x s; case fx: (f x) => //= [u] <-; congr (_ :: _). Qed. Hypothesis fK : ocancel f g. Lemma pmap_filter s : map g (pmap s) = filter [eta f] s. Proof. by elim: s => //= x s <-; rewrite -{3}(fK x); case: (f _). Qed. End Pmap. Section EqPmap. Variables (aT rT : eqType) (f : aT -> option rT) (g : rT -> aT). Lemma eq_pmap (f1 f2 : aT -> option rT) : f1 =1 f2 -> pmap f1 =1 pmap f2. Proof. by move=> Ef; elim=> //= x s ->; rewrite Ef. Qed. Lemma mem_pmap s u : (u \in pmap f s) = (Some u \in map f s). Proof. by elim: s => //= x s IHs; rewrite in_cons -IHs; case: (f x). Qed. Hypothesis fK : ocancel f g. Lemma can2_mem_pmap : pcancel g f -> forall s u, (u \in pmap f s) = (g u \in s). Proof. by move=> gK s u; rewrite -(mem_map (pcan_inj gK)) pmap_filter // mem_filter gK. Qed. Lemma pmap_uniq s : uniq s -> uniq (pmap f s). Proof. by move/(filter_uniq [eta f]); rewrite -(pmap_filter fK); exact: map_uniq. Qed. End EqPmap. Section PmapSub. Variables (T : Type) (p : pred T) (sT : subType p). Lemma size_pmap_sub s : size (pmap (insub : T -> option sT) s) = count p s. Proof. by rewrite size_pmap (eq_count (isSome_insub _)). Qed. End PmapSub. Section EqPmapSub. Variables (T : eqType) (p : pred T) (sT : subType p). Let insT : T -> option sT := insub. Lemma mem_pmap_sub s u : (u \in pmap insT s) = (val u \in s). Proof. apply: (can2_mem_pmap (insubK _)); exact: valK. Qed. Lemma pmap_sub_uniq s : uniq s -> uniq (pmap insT s). Proof. exact: (pmap_uniq (insubK _)). Qed. End EqPmapSub. (* Index sequence *) Fixpoint iota m n := if n is n'.+1 then m :: iota m.+1 n' else [::]. Lemma size_iota m n : size (iota m n) = n. Proof. by elim: n m => //= n IHn m; rewrite IHn. Qed. Lemma iota_add m n1 n2 : iota m (n1 + n2) = iota m n1 ++ iota (m + n1) n2. Proof. by elim: n1 m => //= [|n1 IHn1] m; rewrite ?addn0 // -addSnnS -IHn1. Qed. Lemma iota_addl m1 m2 n : iota (m1 + m2) n = map (addn m1) (iota m2 n). Proof. by elim: n m2 => //= n IHn m2; rewrite -addnS IHn. Qed. Lemma nth_iota m n i : i < n -> nth 0 (iota m n) i = m + i. Proof. by move/subnKC <-; rewrite addSnnS iota_add nth_cat size_iota ltnn subnn. Qed. Lemma mem_iota m n i : (i \in iota m n) = (m <= i) && (i < m + n). Proof. elim: n m => [|n IHn] /= m; first by rewrite addn0 ltnNge andbN. rewrite -addSnnS leq_eqVlt in_cons eq_sym. by case: eqP => [->|_]; [rewrite leq_addr | exact: IHn]. Qed. Lemma iota_uniq m n : uniq (iota m n). Proof. by elim: n m => //= n IHn m; rewrite mem_iota ltnn /=. Qed. (* Making a sequence of a specific length, using indexes to compute items. *) Section MakeSeq. Variables (T : Type) (x0 : T). Definition mkseq f n : seq T := map f (iota 0 n). Lemma size_mkseq f n : size (mkseq f n) = n. Proof. by rewrite size_map size_iota. Qed. Lemma eq_mkseq f g : f =1 g -> mkseq f =1 mkseq g. Proof. by move=> Efg n; exact: eq_map Efg _. Qed. Lemma nth_mkseq f n i : i < n -> nth x0 (mkseq f n) i = f i. Proof. by move=> Hi; rewrite (nth_map 0) ?nth_iota ?size_iota. Qed. Lemma mkseq_nth s : mkseq (nth x0 s) (size s) = s. Proof. by apply: (@eq_from_nth _ x0); rewrite size_mkseq // => i Hi; rewrite nth_mkseq. Qed. End MakeSeq. Section MakeEqSeq. Variable T : eqType. Lemma mkseq_uniq (f : nat -> T) n : injective f -> uniq (mkseq f n). Proof. by move/map_inj_uniq->; apply: iota_uniq. Qed. Lemma perm_eq_iotaP {s t : seq T} x0 (It := iota 0 (size t)) : reflect (exists2 Is, perm_eq Is It & s = map (nth x0 t) Is) (perm_eq s t). Proof. apply: (iffP idP) => [Est | [Is eqIst ->]]; last first. by rewrite -{2}[t](mkseq_nth x0) perm_map. elim: t => [|x t IHt] in s It Est *. by rewrite (perm_eq_small _ Est) //; exists [::]. have /rot_to[k s1 Ds]: x \in s by rewrite (perm_eq_mem Est) mem_head. have [|Is1 eqIst1 Ds1] := IHt s1; first by rewrite -(perm_cons x) -Ds perm_rot. exists (rotr k (0 :: map succn Is1)). by rewrite perm_rot /It /= perm_cons (iota_addl 1) perm_map. by rewrite map_rotr /= -map_comp -(@eq_map _ _ (nth x0 t)) // -Ds1 -Ds rotK. Qed. End MakeEqSeq. Implicit Arguments perm_eq_iotaP [[T] [s] [t]]. Section FoldRight. Variables (T R : Type) (f : T -> R -> R) (z0 : R). Fixpoint foldr s := if s is x :: s' then f x (foldr s') else z0. End FoldRight. Section FoldRightComp. Variables (T1 T2 : Type) (h : T1 -> T2). Variables (R : Type) (f : T2 -> R -> R) (z0 : R). Lemma foldr_cat s1 s2 : foldr f z0 (s1 ++ s2) = foldr f (foldr f z0 s2) s1. Proof. by elim: s1 => //= x s1 ->. Qed. Lemma foldr_map s : foldr f z0 (map h s) = foldr (fun x z => f (h x) z) z0 s. Proof. by elim: s => //= x s ->. Qed. End FoldRightComp. (* Quick characterization of the null sequence. *) Definition sumn := foldr addn 0. Lemma sumn_nseq x n : sumn (nseq n x) = x * n. Proof. by rewrite mulnC; elim: n => //= n ->. Qed. Lemma sumn_cat s1 s2 : sumn (s1 ++ s2) = sumn s1 + sumn s2. Proof. by elim: s1 => //= x s1 ->; rewrite addnA. Qed. Lemma natnseq0P s : reflect (s = nseq (size s) 0) (sumn s == 0). Proof. apply: (iffP idP) => [|->]; last by rewrite sumn_nseq. by elim: s => //= x s IHs; rewrite addn_eq0 => /andP[/eqP-> /IHs <-]. Qed. Section FoldLeft. Variables (T R : Type) (f : R -> T -> R). Fixpoint foldl z s := if s is x :: s' then foldl (f z x) s' else z. Lemma foldl_rev z s : foldl z (rev s) = foldr (fun x z => f z x) z s. Proof. elim/last_ind: s z => [|s x IHs] z //=. by rewrite rev_rcons -cats1 foldr_cat -IHs. Qed. Lemma foldl_cat z s1 s2 : foldl z (s1 ++ s2) = foldl (foldl z s1) s2. Proof. by rewrite -(revK (s1 ++ s2)) foldl_rev rev_cat foldr_cat -!foldl_rev !revK. Qed. End FoldLeft. Section Scan. Variables (T1 : Type) (x1 : T1) (T2 : Type) (x2 : T2). Variables (f : T1 -> T1 -> T2) (g : T1 -> T2 -> T1). Fixpoint pairmap x s := if s is y :: s' then f x y :: pairmap y s' else [::]. Lemma size_pairmap x s : size (pairmap x s) = size s. Proof. by elim: s x => //= y s IHs x; rewrite IHs. Qed. Lemma pairmap_cat x s1 s2 : pairmap x (s1 ++ s2) = pairmap x s1 ++ pairmap (last x s1) s2. Proof. by elim: s1 x => //= y s1 IHs1 x; rewrite IHs1. Qed. Lemma nth_pairmap s n : n < size s -> forall x, nth x2 (pairmap x s) n = f (nth x1 (x :: s) n) (nth x1 s n). Proof. by elim: s n => [|y s IHs] [|n] //= Hn x; apply: IHs. Qed. Fixpoint scanl x s := if s is y :: s' then let x' := g x y in x' :: scanl x' s' else [::]. Lemma size_scanl x s : size (scanl x s) = size s. Proof. by elim: s x => //= y s IHs x; rewrite IHs. Qed. Lemma scanl_cat x s1 s2 : scanl x (s1 ++ s2) = scanl x s1 ++ scanl (foldl g x s1) s2. Proof. by elim: s1 x => //= y s1 IHs1 x; rewrite IHs1. Qed. Lemma nth_scanl s n : n < size s -> forall x, nth x1 (scanl x s) n = foldl g x (take n.+1 s). Proof. by elim: s n => [|y s IHs] [|n] Hn x //=; rewrite ?take0 ?IHs. Qed. Lemma scanlK : (forall x, cancel (g x) (f x)) -> forall x, cancel (scanl x) (pairmap x). Proof. by move=> Hfg x s; elim: s x => //= y s IHs x; rewrite Hfg IHs. Qed. Lemma pairmapK : (forall x, cancel (f x) (g x)) -> forall x, cancel (pairmap x) (scanl x). Proof. by move=> Hgf x s; elim: s x => //= y s IHs x; rewrite Hgf IHs. Qed. End Scan. Prenex Implicits mask map pmap foldr foldl scanl pairmap. Section Zip. Variables S T : Type. Fixpoint zip (s : seq S) (t : seq T) {struct t} := match s, t with | x :: s', y :: t' => (x, y) :: zip s' t' | _, _ => [::] end. Definition unzip1 := map (@fst S T). Definition unzip2 := map (@snd S T). Lemma zip_unzip s : zip (unzip1 s) (unzip2 s) = s. Proof. by elim: s => [|[x y] s /= ->]. Qed. Lemma unzip1_zip s t : size s <= size t -> unzip1 (zip s t) = s. Proof. by elim: s t => [|x s IHs] [|y t] //= le_s_t; rewrite IHs. Qed. Lemma unzip2_zip s t : size t <= size s -> unzip2 (zip s t) = t. Proof. by elim: s t => [|x s IHs] [|y t] //= le_t_s; rewrite IHs. Qed. Lemma size1_zip s t : size s <= size t -> size (zip s t) = size s. Proof. by elim: s t => [|x s IHs] [|y t] //= Hs; rewrite IHs. Qed. Lemma size2_zip s t : size t <= size s -> size (zip s t) = size t. Proof. by elim: s t => [|x s IHs] [|y t] //= Hs; rewrite IHs. Qed. Lemma size_zip s t : size (zip s t) = minn (size s) (size t). Proof. by elim: s t => [|x s IHs] [|t2 t] //=; rewrite IHs -add1n addn_minr. Qed. Lemma zip_cat s1 s2 t1 t2 : size s1 = size t1 -> zip (s1 ++ s2) (t1 ++ t2) = zip s1 t1 ++ zip s2 t2. Proof. by move/eqP; elim: s1 t1 => [|x s IHs] [|y t] //= /IHs->. Qed. Lemma nth_zip x y s t i : size s = size t -> nth (x, y) (zip s t) i = (nth x s i, nth y t i). Proof. by move/eqP; elim: i s t => [|i IHi] [|y1 s1] [|y2 t] //= /IHi->. Qed. Lemma nth_zip_cond p s t i : nth p (zip s t) i = (if i < size (zip s t) then (nth p.1 s i, nth p.2 t i) else p). Proof. rewrite size_zip ltnNge geq_min. by elim: s t i => [|x s IHs] [|y t] [|i] //=; rewrite ?orbT -?IHs. Qed. Lemma zip_rcons s1 s2 z1 z2 : size s1 = size s2 -> zip (rcons s1 z1) (rcons s2 z2) = rcons (zip s1 s2) (z1, z2). Proof. by move=> eq_sz; rewrite -!cats1 zip_cat //= eq_sz. Qed. Lemma rev_zip s1 s2 : size s1 = size s2 -> rev (zip s1 s2) = zip (rev s1) (rev s2). Proof. elim: s1 s2 => [|x s1 IHs] [|y s2] //= [eq_sz]. by rewrite !rev_cons zip_rcons ?IHs ?size_rev. Qed. End Zip. Prenex Implicits zip unzip1 unzip2. Section Flatten. Variable T : Type. Implicit Types (s : seq T) (ss : seq (seq T)). Definition flatten := foldr cat (Nil T). Definition shape := map (@size T). Fixpoint reshape sh s := if sh is n :: sh' then take n s :: reshape sh' (drop n s) else [::]. Lemma size_flatten ss : size (flatten ss) = sumn (shape ss). Proof. by elim: ss => //= s ss <-; rewrite size_cat. Qed. Lemma flatten_cat ss1 ss2 : flatten (ss1 ++ ss2) = flatten ss1 ++ flatten ss2. Proof. by elim: ss1 => //= s ss1 ->; rewrite catA. Qed. Lemma flattenK ss : reshape (shape ss) (flatten ss) = ss. Proof. by elim: ss => //= s ss IHss; rewrite take_size_cat ?drop_size_cat ?IHss. Qed. Lemma reshapeKr sh s : size s <= sumn sh -> flatten (reshape sh s) = s. Proof. elim: sh s => [[]|n sh IHsh] //= s sz_s; rewrite IHsh ?cat_take_drop //. by rewrite size_drop leq_subLR. Qed. Lemma reshapeKl sh s : size s >= sumn sh -> shape (reshape sh s) = sh. Proof. elim: sh s => [[]|n sh IHsh] //= s sz_s. rewrite size_takel; last exact: leq_trans (leq_addr _ _) sz_s. by rewrite IHsh // -(leq_add2l n) size_drop -maxnE leq_max sz_s orbT. Qed. End Flatten. Prenex Implicits flatten shape reshape. Section EqFlatten. Variables S T : eqType. Lemma flattenP (A : seq (seq T)) x : reflect (exists2 s, s \in A & x \in s) (x \in flatten A). Proof. elim: A => /= [|s A /iffP IH_A]; [by right; case | rewrite mem_cat]. have [s_x|s'x] := @idP (x \in s); first by left; exists s; rewrite ?mem_head. by apply: IH_A => [[t] | [t /predU1P[->|]]]; exists t; rewrite // mem_behead. Qed. Implicit Arguments flattenP [A x]. Lemma flatten_mapP (A : S -> seq T) s y : reflect (exists2 x, x \in s & y \in A x) (y \in flatten (map A s)). Proof. apply: (iffP flattenP) => [[_ /mapP[x sx ->]] | [x sx]] Axy; first by exists x. by exists (A x); rewrite ?map_f. Qed. End EqFlatten. Implicit Arguments flattenP [T A x]. Implicit Arguments flatten_mapP [S T A s y]. Section AllPairs. Variables (S T R : Type) (f : S -> T -> R). Implicit Types (s : seq S) (t : seq T). Definition allpairs s t := foldr (fun x => cat (map (f x) t)) [::] s. Lemma size_allpairs s t : size (allpairs s t) = size s * size t. Proof. by elim: s => //= x s IHs; rewrite size_cat size_map IHs. Qed. Lemma allpairs_cat s1 s2 t : allpairs (s1 ++ s2) t = allpairs s1 t ++ allpairs s2 t. Proof. by elim: s1 => //= x s1 ->; rewrite catA. Qed. End AllPairs. Prenex Implicits allpairs. Notation "[ 'seq' E | i <- s , j <- t ]" := (allpairs (fun i j => E) s t) (at level 0, E at level 99, i ident, j ident, format "[ '[hv' 'seq' E '/ ' | i <- s , '/ ' j <- t ] ']'") : seq_scope. Notation "[ 'seq' E | i : T <- s , j : U <- t ]" := (allpairs (fun (i : T) (j : U) => E) s t) (at level 0, E at level 99, i ident, j ident, only parsing) : seq_scope. Section EqAllPairs. Variables S T : eqType. Implicit Types (R : eqType) (s : seq S) (t : seq T). Lemma allpairsP R (f : S -> T -> R) s t z : reflect (exists p, [/\ p.1 \in s, p.2 \in t & z = f p.1 p.2]) (z \in allpairs f s t). Proof. elim: s => [|x s IHs /=]; first by right=> [[p []]]. rewrite mem_cat; have [fxt_z | not_fxt_z] := altP mapP. by left; have [y t_y ->] := fxt_z; exists (x, y); rewrite mem_head. apply: (iffP IHs) => [] [[x' y] /= [s_x' t_y def_z]]; exists (x', y). by rewrite !inE predU1r. by have [def_x' | //] := predU1P s_x'; rewrite def_z def_x' map_f in not_fxt_z. Qed. Lemma mem_allpairs R (f : S -> T -> R) s1 t1 s2 t2 : s1 =i s2 -> t1 =i t2 -> allpairs f s1 t1 =i allpairs f s2 t2. Proof. move=> eq_s eq_t z. by apply/allpairsP/allpairsP=> [] [p fpz]; exists p; rewrite eq_s eq_t in fpz *. Qed. Lemma allpairs_catr R (f : S -> T -> R) s t1 t2 : allpairs f s (t1 ++ t2) =i allpairs f s t1 ++ allpairs f s t2. Proof. move=> z; rewrite mem_cat. apply/allpairsP/orP=> [[p [sP1]]|]. by rewrite mem_cat; case/orP; [left | right]; apply/allpairsP; exists p. by case=> /allpairsP[p [sp1 sp2 ->]]; exists p; rewrite mem_cat sp2 ?orbT. Qed. Lemma allpairs_uniq R (f : S -> T -> R) s t : uniq s -> uniq t -> {in [seq (x, y) | x <- s, y <- t] &, injective (prod_curry f)} -> uniq (allpairs f s t). Proof. move=> Us Ut inj_f; have: all (mem s) s by exact/allP. elim: {-2}s Us => //= x s1 IHs /andP[s1'x Us1] /andP[sx1 ss1]. rewrite cat_uniq {}IHs // andbT map_inj_in_uniq ?Ut // => [|y1 y2 *]. apply/hasPn=> _ /allpairsP[z [s1z tz ->]]; apply/mapP=> [[y ty Dy]]. suffices [Dz1 _]: (z.1, z.2) = (x, y) by rewrite -Dz1 s1z in s1'x. apply: inj_f => //; apply/allpairsP; last by exists (x, y). by have:= allP ss1 _ s1z; exists z. suffices: (x, y1) = (x, y2) by case. by apply: inj_f => //; apply/allpairsP; [exists (x, y1) | exists (x, y2)]. Qed. End EqAllPairs. ssreflect-1.5/theories/choice.v0000644000175000017500000006607312175455021015616 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. (******************************************************************************) (* This file contains the definitions of: *) (* choiceType == interface for types with a choice operator *) (* countType == interface for countable types *) (* subCountType == interface for types that are both subType and countType. *) (* xchoose exP == a standard x such that P x, given exP : exists x : T, P x *) (* when T is a choiceType. The choice depends only on the *) (* extent of P (in particular, it is independent of exP). *) (* choose P x0 == if P x0, a standard x such that P x. *) (* pickle x == a nat encoding the value x : T, where T is a countType. *) (* unpickle n == a partial inverse to pickle: unpickle (pickle x) = Some x *) (* pickle_inv n == a sharp partial inverse to pickle pickle_inv n = Some x *) (* if and only if pickle x = n. *) (* [choiceType of T for cT] == clone for T of the choiceType cT. *) (* [choiceType of T] == clone for T of the choiceType inferred for T. *) (* [countType of T for cT] == clone for T of the countType cT. *) (* [count Type of T] == clone for T of the countType inferred for T. *) (* [choiceMixin of T by <:] == Choice mixin for T when T has a subType p *) (* structure with p : pred cT and cT has a Choice *) (* structure; the corresponding structure is Canonical.*) (* [countMixin of T by <:] == Count mixin for a subType T of a countType. *) (* PcanChoiceMixin fK == Choice mixin for T, given f : T -> cT where cT has *) (* a Choice structure, a left inverse partial function *) (* g and fK : pcancel f g. *) (* CanChoiceMixin fK == Choice mixin for T, given f : T -> cT, g and *) (* fK : cancel f g. *) (* PcanCountMixin fK == Count mixin for T, given f : T -> cT where cT has *) (* a Countable structure, a left inverse partial *) (* function g and fK : pcancel f g. *) (* CanCountMixin fK == Count mixin for T, given f : T -> cT, g and *) (* fK : cancel f g. *) (* GenTree.tree T == generic n-ary tree type with nat-labeled nodes and *) (* T-labeled nodes. It is equipped with canonical *) (* eqType, choiceType, and countType instances, and so *) (* can be used to similarly equip simple datatypes *) (* by using the mixins above. *) (* In addition to the lemmas relevant to these definitions, this file also *) (* contains definitions of a Canonical choiceType and countType instances for *) (* all basic datatypes (e.g., nat, bool, subTypes, pairs, sums, etc.). *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. (* Technical definitions about coding and decoding of nat sequances, which *) (* are used below to define various Canonical instances of the choice and *) (* countable interfaces. *) Module CodeSeq. (* Goedel-style one-to-one encoding of seq nat into nat. *) (* The code for [:: n1; ...; nk] has binary representation *) (* 1 0 ... 0 1 ... 1 0 ... 0 1 0 ... 0 *) (* <-----> <-----> <-----> *) (* nk 0s n2 0s n1 0s *) Definition code := foldr (fun n m => 2 ^ n * m.*2.+1) 0. Fixpoint decode_rec (v q r : nat) {struct q} := match q, r with | 0, _ => [:: v] | q'.+1, 0 => v :: [rec 0, q', q'] | q'.+1, 1 => [rec v.+1, q', q'] | q'.+1, r'.+2 => [rec v, q', r'] end where "[ 'rec' v , q , r ]" := (decode_rec v q r). Definition decode n := if n is 0 then [::] else [rec 0, n.-1, n.-1]. Lemma decodeK : cancel decode code. Proof. have m2s: forall n, n.*2 - n = n by move=> n; rewrite -addnn addnK. case=> //= n; rewrite -[n.+1]mul1n -(expn0 2) -{3}[n]m2s. elim: n {2 4}n {1 3}0 => [|q IHq] [|[|r]] v //=; rewrite {}IHq ?mul1n ?m2s //. by rewrite expnSr -mulnA mul2n. Qed. Lemma codeK : cancel code decode. Proof. elim=> //= v s IHs; rewrite -[_ * _]prednK ?muln_gt0 ?expn_gt0 //=. rewrite -{3}[v]addn0; elim: v {1 4}0 => [|v IHv {IHs}] q. rewrite mul1n /= -{1}addnn -{4}IHs; move: (_ s) {IHs} => n. by elim: {1 3}n => //=; case: n. rewrite expnS -mulnA mul2n -{1}addnn -[_ * _]prednK ?muln_gt0 ?expn_gt0 //. by rewrite doubleS addSn /= addSnnS; elim: {-2}_.-1 => //=. Qed. Lemma ltn_code s : all (fun j => j < code s) s. Proof. elim: s => //= i s IHs; rewrite -[_.+1]muln1 leq_mul 1?ltn_expl //=. apply: sub_all IHs => j /leqW lejs; rewrite -[j.+1]mul1n leq_mul ?expn_gt0 //. by rewrite ltnS -[j]mul1n -mul2n leq_mul. Qed. Lemma gtn_decode n : all (ltn^~ n) (decode n). Proof. by rewrite -{1}[n]decodeK ltn_code. Qed. End CodeSeq. Section OtherEncodings. (* Miscellaneous encodings: option T -c-> seq T, T1 * T2 -c-> {i : T1 & T2} *) (* T1 + T2 -c-> option T1 * option T2, unit -c-> bool; bool -c-> nat is *) (* already covered in ssrnat by the nat_of_bool coercion, the odd predicate, *) (* and their "cancellation" lemma oddb. We use these encodings to propagate *) (* canonical structures through these type constructors so that ultimately *) (* all Choice and Countable instanced derive from nat and the seq and sigT *) (* constructors. *) Variables T T1 T2 : Type. Definition seq_of_opt := @oapp T _ (nseq 1) [::]. Lemma seq_of_optK : cancel seq_of_opt ohead. Proof. by case. Qed. Definition tag_of_pair (p : T1 * T2) := @Tagged T1 p.1 (fun _ => T2) p.2. Definition pair_of_tag (u : {i : T1 & T2}) := (tag u, tagged u). Lemma tag_of_pairK : cancel tag_of_pair pair_of_tag. Proof. by case. Qed. Lemma pair_of_tagK : cancel pair_of_tag tag_of_pair. Proof. by case. Qed. Definition opair_of_sum (s : T1 + T2) := match s with inl x => (Some x, None) | inr y => (None, Some y) end. Definition sum_of_opair p := oapp (some \o @inr T1 T2) (omap (@inl _ T2) p.1) p.2. Lemma opair_of_sumK : pcancel opair_of_sum sum_of_opair. Proof. by case. Qed. Lemma bool_of_unitK : cancel (fun _ => true) (fun _ => tt). Proof. by case. Qed. End OtherEncodings. (* Generic variable-arity tree type, providing an encoding target for *) (* miscellaneous user datatypes. The GenTree.tree type can be combined with *) (* a sigT type to model multi-sorted concrete datatypes. *) Module GenTree. Section Def. Variable T : Type. Unset Elimination Schemes. Inductive tree := Leaf of T | Node of nat & seq tree. Definition tree_rect K IH_leaf IH_node := fix loop t : K t := match t with | Leaf x => IH_leaf x | Node n f0 => let fix iter_pair f : foldr (fun t => prod (K t)) unit f := if f is t :: f' then (loop t, iter_pair f') else tt in IH_node n f0 (iter_pair f0) end. Definition tree_rec (K : tree -> Set) := @tree_rect K. Definition tree_ind K IH_leaf IH_node := fix loop t : K t : Prop := match t with | Leaf x => IH_leaf x | Node n f0 => let fix iter_conj f : foldr (fun t => and (K t)) True f := if f is t :: f' then conj (loop t) (iter_conj f') else Logic.I in IH_node n f0 (iter_conj f0) end. Fixpoint encode t : seq (nat + T) := match t with | Leaf x => [:: inr _ x] | Node n f => inl _ n.+1 :: rcons (flatten (map encode f)) (inl _ 0) end. Definition decode_step c fs := match c with | inr x => (Leaf x :: fs.1, fs.2) | inl 0 => ([::], fs.1 :: fs.2) | inl n.+1 => (Node n fs.1 :: head [::] fs.2, behead fs.2) end. Definition decode c := ohead (foldr decode_step ([::], [::]) c).1. Lemma codeK : pcancel encode decode. Proof. move=> t; rewrite /decode; set fs := (_, _). suffices ->: foldr decode_step fs (encode t) = (t :: fs.1, fs.2) by []. elim: t => //= n f IHt in (fs) *; elim: f IHt => //= t f IHf []. by rewrite rcons_cat foldr_cat => -> /= /IHf[-> -> ->]. Qed. End Def. End GenTree. Implicit Arguments GenTree.codeK []. Definition tree_eqMixin (T : eqType) := PcanEqMixin (GenTree.codeK T). Canonical tree_eqType (T : eqType) := EqType (GenTree.tree T) (tree_eqMixin T). (* Structures for Types with a choice function, and for Types with countably *) (* many elements. The two concepts are closely linked: we indeed make *) (* Countable a subclass of Choice, as countable choice is valid in CiC. This *) (* apparent redundancy is needed to ensure the consistency of the Canonical *) (* inference, as the canonical Choice for a given type may differ from the *) (* countable choice for its canonical Countable structure, e.g., for options. *) (* The Choice interface exposes two choice functions; for T : choiceType *) (* and P : pred T, we provide: *) (* xchoose : (exists x, P x) -> T *) (* choose : pred T -> T -> T *) (* While P (xchoose exP) will always hold, P (choose P x0) will be true if *) (* and only if P x0 holds. Both xchoose and choose are extensional in P and *) (* do not depend on the witness exP or x0 (provided P x0 holds). Note that *) (* xchoose is slightly more powerful, but less convenient to use. *) (* Howver, neither choose nor xchoose are composable: it would not be *) (* be possible to extend the Choice structure to arbitrary pairs using only *) (* these functions, for instance. Internally, the interfaces provides a *) (* subtly stronger operation, Choice.InternalTheory.find, which performs a *) (* limited search using an integer parameter only rather than a full value as *) (* [x]choose does. This is neither a restriction in the constructive setting *) (* (where all types are concrete and hence countable), nor in an axiomatic *) (* such as the Coq reals library, where the axiom of choice and excluded *) (* middle suppress the need for guidance. Nevertheless this operation is just *) (* what is needed to make the Choice interface compose. *) (* The Countable interface provides three functions; for T : countType we *) (* geth pickle : T -> nat, and unpickle, pickle_inv : nat -> option T. *) (* The functions provide an effective embedding of T in nat: unpickle is a *) (* left inverse to pickle, which satisfies pcancel pickle unpickle, i.e., *) (* unpickle \o pickle =1 some; pickle_inv is a more precise inverse for which *) (* we also have ocancel pickle_inv pickle. Both unpickle and pickle need to *) (* partial functions, to allow for possibly empty types such as {x | P x}. *) (* The names of these functions underline the correspondence with the *) (* notion of "Serializable" types in programming languages. *) (* Finally, we need to provide a join class to let type inference unify *) (* subType and countType class constraints, e.g., for a countable subType of *) (* an uncountable choiceType (the issue does not arise earlier with eqType or *) (* choiceType because in practice the base type of an Equality/Choice subType *) (* is always an Equality/Choice Type). *) Module Choice. Section ClassDef. Record mixin_of T := Mixin { find : pred T -> nat -> option T; _ : forall P n x, find P n = Some x -> P x; _ : forall P : pred T, (exists x, P x) -> exists n, find P n; _ : forall P Q : pred T, P =1 Q -> find P =1 find Q }. Record class_of T := Class {base : Equality.class_of T; mixin : mixin_of T}. Local Coercion base : class_of >-> Equality.class_of. Structure type := Pack {sort; _ : class_of sort; _ : Type}. Local Coercion sort : type >-> Sortclass. Variables (T : Type) (cT : type). Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c. Definition clone c of phant_id class c := @Pack T c T. Let xT := let: Pack T _ _ := cT in T. Notation xclass := (class : class_of xT). Definition pack m := fun b bT & phant_id (Equality.class bT) b => Pack (@Class T b m) T. (* Inheritance *) Definition eqType := @Equality.Pack cT xclass xT. End ClassDef. Module Import Exports. Coercion base : class_of >-> Equality.class_of. Coercion sort : type >-> Sortclass. Coercion eqType : type >-> Equality.type. Canonical eqType. Notation choiceType := type. Notation choiceMixin := mixin_of. Notation ChoiceType T m := (@pack T m _ _ id). Notation "[ 'choiceType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) (at level 0, format "[ 'choiceType' 'of' T 'for' cT ]") : form_scope. Notation "[ 'choiceType' 'of' T ]" := (@clone T _ _ id) (at level 0, format "[ 'choiceType' 'of' T ]") : form_scope. End Exports. Module InternalTheory. Section InternalTheory. (* Inner choice function. *) Definition find T := find (mixin (class T)). Variable T : choiceType. Implicit Types P Q : pred T. Lemma correct P n x : find P n = Some x -> P x. Proof. by case: T => _ [_ []] //= in P n x *. Qed. Lemma complete P : (exists x, P x) -> (exists n, find P n). Proof. by case: T => _ [_ []] //= in P *. Qed. Lemma extensional P Q : P =1 Q -> find P =1 find Q. Proof. by case: T => _ [_ []] //= in P Q *. Qed. Fact xchoose_subproof P exP : {x | find P (ex_minn (@complete P exP)) = Some x}. Proof. by case: (ex_minnP (complete exP)) => n; case: (find P n) => // x; exists x. Qed. End InternalTheory. End InternalTheory. End Choice. Export Choice.Exports. Section ChoiceTheory. Implicit Type T : choiceType. Import Choice.InternalTheory CodeSeq. Local Notation dc := decode. Section OneType. Variable T : choiceType. Implicit Types P Q : pred T. Definition xchoose P exP := sval (@xchoose_subproof T P exP). Lemma xchooseP P exP : P (@xchoose P exP). Proof. by rewrite /xchoose; case: (xchoose_subproof exP) => x /= /correct. Qed. Lemma eq_xchoose P Q exP exQ : P =1 Q -> @xchoose P exP = @xchoose Q exQ. Proof. rewrite /xchoose => eqPQ. case: (xchoose_subproof exP) => x; case: (xchoose_subproof exQ) => y /=. case: ex_minnP => n; case: ex_minnP => m. rewrite -(extensional eqPQ) {1}(extensional eqPQ). move=> Qm minPm Pn minQn; suffices /eqP->: m == n by move=> -> []. by rewrite eqn_leq minQn ?minPm. Qed. Lemma sigW P : (exists x, P x) -> {x | P x}. Proof. by move=> exP; exists (xchoose exP); exact: xchooseP. Qed. Lemma sig2W P Q : (exists2 x, P x & Q x) -> {x | P x & Q x}. Proof. move=> exPQ; have [|x /andP[]] := @sigW (predI P Q); last by exists x. by have [x Px Qx] := exPQ; exists x; exact/andP. Qed. Lemma sig_eqW (vT : eqType) (lhs rhs : T -> vT) : (exists x, lhs x = rhs x) -> {x | lhs x = rhs x}. Proof. move=> exP; suffices [x /eqP Ex]: {x | lhs x == rhs x} by exists x. by apply: sigW; have [x /eqP Ex] := exP; exists x. Qed. Lemma sig2_eqW (vT : eqType) (P : pred T) (lhs rhs : T -> vT) : (exists2 x, P x & lhs x = rhs x) -> {x | P x & lhs x = rhs x}. Proof. move=> exP; suffices [x Px /eqP Ex]: {x | P x & lhs x == rhs x} by exists x. by apply: sig2W; have [x Px /eqP Ex] := exP; exists x. Qed. Definition choose P x0 := if insub x0 : {? x | P x} is Some (exist x Px) then xchoose (ex_intro [eta P] x Px) else x0. Lemma chooseP P x0 : P x0 -> P (choose P x0). Proof. by move=> Px0; rewrite /choose insubT xchooseP. Qed. Lemma choose_id P x0 y0 : P x0 -> P y0 -> choose P x0 = choose P y0. Proof. by move=> Px0 Py0; rewrite /choose !insubT /=; exact: eq_xchoose. Qed. Lemma eq_choose P Q : P =1 Q -> choose P =1 choose Q. Proof. rewrite /choose => eqPQ x0. do [case: insubP; rewrite eqPQ] => [[x Px] Qx0 _| ?]; last by rewrite insubN. by rewrite insubT; exact: eq_xchoose. Qed. Section CanChoice. Variables (sT : Type) (f : sT -> T). Lemma PcanChoiceMixin f' : pcancel f f' -> choiceMixin sT. Proof. move=> fK; pose liftP sP := [pred x | oapp sP false (f' x)]. pose sf sP := [fun n => obind f' (find (liftP sP) n)]. exists sf => [sP n x | sP [y sPy] | sP sQ eqPQ n] /=. - by case Df: (find _ n) => //= [?] Dx; have:= correct Df; rewrite /= Dx. - have [|n Pn] := @complete T (liftP sP); first by exists (f y); rewrite /= fK. exists n; case Df: (find _ n) Pn => //= [x] _. by have:= correct Df => /=; case: (f' x). by congr (obind _ _); apply: extensional => x /=; case: (f' x) => /=. Qed. Definition CanChoiceMixin f' (fK : cancel f f') := PcanChoiceMixin (can_pcan fK). End CanChoice. Section SubChoice. Variables (P : pred T) (sT : subType P). Definition sub_choiceMixin := PcanChoiceMixin (@valK T P sT). Definition sub_choiceClass := @Choice.Class sT (sub_eqMixin sT) sub_choiceMixin. Canonical sub_choiceType := Choice.Pack sub_choiceClass sT. End SubChoice. Fact seq_choiceMixin : choiceMixin (seq T). Proof. pose r f := [fun xs => fun x : T => f (x :: xs) : option (seq T)]. pose fix f sP ns xs {struct ns} := if ns is n :: ns1 then let fr := r (f sP ns1) xs in obind fr (find fr n) else if sP xs then Some xs else None. exists (fun sP nn => f sP (dc nn) nil) => [sP n ys | sP [ys] | sP sQ eqPQ n]. - elim: {n}(dc n) nil => [|n ns IHs] xs /=; first by case: ifP => // sPxs [<-]. by case: (find _ n) => //= [x]; apply: IHs. - rewrite -(cats0 ys); elim/last_ind: ys nil => [|ys y IHs] xs /=. by move=> sPxs; exists 0; rewrite /= sPxs. rewrite cat_rcons => /IHs[n1 sPn1] {IHs}. have /complete[n]: exists z, f sP (dc n1) (z :: xs) by exists y. case Df: (find _ n)=> // [x] _; exists (code (n :: dc n1)). by rewrite codeK /= Df /= (correct Df). elim: {n}(dc n) nil => [|n ns IHs] xs /=; first by rewrite eqPQ. rewrite (@extensional _ _ (r (f sQ ns) xs)) => [|x]; last by rewrite IHs. by case: find => /=. Qed. Canonical seq_choiceType := Eval hnf in ChoiceType (seq T) seq_choiceMixin. End OneType. Section TagChoice. Variables (I : choiceType) (T_ : I -> choiceType). Fact tagged_choiceMixin : choiceMixin {i : I & T_ i}. Proof. pose mkT i (x : T_ i) := Tagged T_ x. pose ft tP n i := omap (mkT i) (find (tP \o mkT i) n). pose fi tP ni nt := obind (ft tP nt) (find (ft tP nt) ni). pose f tP n := if dc n is [:: ni; nt] then fi tP ni nt else None. exists f => [tP n u | tP [[i x] tPxi] | sP sQ eqPQ n]. - rewrite /f /fi; case: (dc n) => [|ni [|nt []]] //=. case: (find _ _) => //= [i]; rewrite /ft. by case Df: (find _ _) => //= [x] [<-]; have:= correct Df. - have /complete[nt tPnt]: exists y, (tP \o mkT i) y by exists x. have{tPnt}: exists j, ft tP nt j by exists i; rewrite /ft; case: find tPnt. case/complete=> ni tPn; exists (code [:: ni; nt]); rewrite /f codeK /fi. by case Df: find tPn => //= [j] _; have:= correct Df. rewrite /f /fi; case: (dc n) => [|ni [|nt []]] //=. rewrite (@extensional _ _ (ft sQ nt)) => [|i]. by case: find => //= i; congr (omap _ _); apply: extensional => x /=. by congr (omap _ _); apply: extensional => x /=. Qed. Canonical tagged_choiceType := Eval hnf in ChoiceType {i : I & T_ i} tagged_choiceMixin. End TagChoice. Fact nat_choiceMixin : choiceMixin nat. Proof. pose f := [fun (P : pred nat) n => if P n then Some n else None]. exists f => [P n m | P [n Pn] | P Q eqPQ n] /=; last by rewrite eqPQ. by case: ifP => // Pn [<-]. by exists n; rewrite Pn. Qed. Canonical nat_choiceType := Eval hnf in ChoiceType nat nat_choiceMixin. Definition bool_choiceMixin := CanChoiceMixin oddb. Canonical bool_choiceType := Eval hnf in ChoiceType bool bool_choiceMixin. Canonical bitseq_choiceType := Eval hnf in [choiceType of bitseq]. Definition unit_choiceMixin := CanChoiceMixin bool_of_unitK. Canonical unit_choiceType := Eval hnf in ChoiceType unit unit_choiceMixin. Definition option_choiceMixin T := CanChoiceMixin (@seq_of_optK T). Canonical option_choiceType T := Eval hnf in ChoiceType (option T) (option_choiceMixin T). Definition sig_choiceMixin T (P : pred T) : choiceMixin {x | P x} := sub_choiceMixin _. Canonical sig_choiceType T (P : pred T) := Eval hnf in ChoiceType {x | P x} (sig_choiceMixin P). Definition prod_choiceMixin T1 T2 := CanChoiceMixin (@tag_of_pairK T1 T2). Canonical prod_choiceType T1 T2 := Eval hnf in ChoiceType (T1 * T2) (prod_choiceMixin T1 T2). Definition sum_choiceMixin T1 T2 := PcanChoiceMixin (@opair_of_sumK T1 T2). Canonical sum_choiceType T1 T2 := Eval hnf in ChoiceType (T1 + T2) (sum_choiceMixin T1 T2). Definition tree_choiceMixin T := PcanChoiceMixin (GenTree.codeK T). Canonical tree_choiceType T := ChoiceType (GenTree.tree T) (tree_choiceMixin T). End ChoiceTheory. Prenex Implicits xchoose choose. Notation "[ 'choiceMixin' 'of' T 'by' <: ]" := (sub_choiceMixin _ : choiceMixin T) (at level 0, format "[ 'choiceMixin' 'of' T 'by' <: ]") : form_scope. Module Countable. Record mixin_of (T : Type) : Type := Mixin { pickle : T -> nat; unpickle : nat -> option T; pickleK : pcancel pickle unpickle }. Definition EqMixin T m := PcanEqMixin (@pickleK T m). Definition ChoiceMixin T m := PcanChoiceMixin (@pickleK T m). Section ClassDef. Record class_of T := Class { base : Choice.class_of T; mixin : mixin_of T }. Local Coercion base : class_of >-> Choice.class_of. Structure type : Type := Pack {sort : Type; _ : class_of sort; _ : Type}. Local Coercion sort : type >-> Sortclass. Variables (T : Type) (cT : type). Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c. Definition clone c of phant_id class c := @Pack T c T. Let xT := let: Pack T _ _ := cT in T. Notation xclass := (class : class_of xT). Definition pack m := fun bT b & phant_id (Choice.class bT) b => Pack (@Class T b m) T. Definition eqType := @Equality.Pack cT xclass xT. Definition choiceType := @Choice.Pack cT xclass xT. End ClassDef. Module Exports. Coercion base : class_of >-> Choice.class_of. Coercion mixin : class_of >-> mixin_of. Coercion sort : type >-> Sortclass. Coercion eqType : type >-> Equality.type. Canonical eqType. Coercion choiceType : type >-> Choice.type. Canonical choiceType. Notation countType := type. Notation CountType T m := (@pack T m _ _ id). Notation CountMixin := Mixin. Notation CountChoiceMixin := ChoiceMixin. Notation "[ 'countType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) (at level 0, format "[ 'countType' 'of' T 'for' cT ]") : form_scope. Notation "[ 'countType' 'of' T ]" := (@clone T _ _ id) (at level 0, format "[ 'countType' 'of' T ]") : form_scope. End Exports. End Countable. Export Countable.Exports. Definition unpickle T := Countable.unpickle (Countable.class T). Definition pickle T := Countable.pickle (Countable.class T). Implicit Arguments unpickle [T]. Prenex Implicits pickle unpickle. Section CountableTheory. Variable T : countType. Lemma pickleK : @pcancel nat T pickle unpickle. Proof. exact: Countable.pickleK. Qed. Definition pickle_inv n := obind (fun x : T => if pickle x == n then Some x else None) (unpickle n). Lemma pickle_invK : ocancel pickle_inv pickle. Proof. by rewrite /pickle_inv => n; case def_x: (unpickle n) => //= [x]; case: eqP. Qed. Lemma pickleK_inv : pcancel pickle pickle_inv. Proof. by rewrite /pickle_inv => x; rewrite pickleK /= eqxx. Qed. Lemma pcan_pickleK sT f f' : @pcancel T sT f f' -> pcancel (pickle \o f) (pcomp f' unpickle). Proof. by move=> fK x; rewrite /pcomp pickleK /= fK. Qed. Definition PcanCountMixin sT f f' (fK : pcancel f f') := @CountMixin sT _ _ (pcan_pickleK fK). Definition CanCountMixin sT f f' (fK : cancel f f') := @PcanCountMixin sT _ _ (can_pcan fK). Definition sub_countMixin P sT := PcanCountMixin (@valK T P sT). Definition pickle_seq s := CodeSeq.code (map (@pickle T) s). Definition unpickle_seq n := Some (pmap (@unpickle T) (CodeSeq.decode n)). Lemma pickle_seqK : pcancel pickle_seq unpickle_seq. Proof. by move=> s; rewrite /unpickle_seq CodeSeq.codeK (map_pK pickleK). Qed. Definition seq_countMixin := CountMixin pickle_seqK. Canonical seq_countType := Eval hnf in CountType (seq T) seq_countMixin. End CountableTheory. Notation "[ 'countMixin' 'of' T 'by' <: ]" := (sub_countMixin _ : Countable.mixin_of T) (at level 0, format "[ 'countMixin' 'of' T 'by' <: ]") : form_scope. Section SubCountType. Variables (T : choiceType) (P : pred T). Import Countable. Structure subCountType : Type := SubCountType {subCount_sort :> subType P; _ : mixin_of subCount_sort}. Coercion sub_countType (sT : subCountType) := Eval hnf in pack (let: SubCountType _ m := sT return mixin_of sT in m) id. Canonical sub_countType. Definition pack_subCountType U := fun sT cT & sub_sort sT * sort cT -> U * U => fun b m & phant_id (Class b m) (class cT) => @SubCountType sT m. End SubCountType. (* This assumes that T has both countType and subType structures. *) Notation "[ 'subCountType' 'of' T ]" := (@pack_subCountType _ _ T _ _ id _ _ id) (at level 0, format "[ 'subCountType' 'of' T ]") : form_scope. Section TagCountType. Variables (I : countType) (T_ : I -> countType). Definition pickle_tagged (u : {i : I & T_ i}) := CodeSeq.code [:: pickle (tag u); pickle (tagged u)]. Definition unpickle_tagged s := if CodeSeq.decode s is [:: ni; nx] then obind (fun i => omap (@Tagged I i T_) (unpickle nx)) (unpickle ni) else None. Lemma pickle_taggedK : pcancel pickle_tagged unpickle_tagged. Proof. by case=> i x; rewrite /unpickle_tagged CodeSeq.codeK /= pickleK /= pickleK. Qed. Definition tag_countMixin := CountMixin pickle_taggedK. Canonical tag_countType := Eval hnf in CountType {i : I & T_ i} tag_countMixin. End TagCountType. (* The remaining Canonicals for standard datatypes. *) Section CountableDataTypes. Implicit Type T : countType. Lemma nat_pickleK : pcancel id (@Some nat). Proof. by []. Qed. Definition nat_countMixin := CountMixin nat_pickleK. Canonical nat_countType := Eval hnf in CountType nat nat_countMixin. Definition bool_countMixin := CanCountMixin oddb. Canonical bool_countType := Eval hnf in CountType bool bool_countMixin. Canonical bitseq_countType := Eval hnf in [countType of bitseq]. Definition unit_countMixin := CanCountMixin bool_of_unitK. Canonical unit_countType := Eval hnf in CountType unit unit_countMixin. Definition option_countMixin T := CanCountMixin (@seq_of_optK T). Canonical option_countType T := Eval hnf in CountType (option T) (option_countMixin T). Definition sig_countMixin T (P : pred T) := [countMixin of {x | P x} by <:]. Canonical sig_countType T (P : pred T) := Eval hnf in CountType {x | P x} (sig_countMixin P). Canonical sig_subCountType T (P : pred T) := Eval hnf in [subCountType of {x | P x}]. Definition prod_countMixin T1 T2 := CanCountMixin (@tag_of_pairK T1 T2). Canonical prod_countType T1 T2 := Eval hnf in CountType (T1 * T2) (prod_countMixin T1 T2). Definition sum_countMixin T1 T2 := PcanCountMixin (@opair_of_sumK T1 T2). Canonical sum_countType T1 T2 := Eval hnf in CountType (T1 + T2) (sum_countMixin T1 T2). Definition tree_countMixin T := PcanCountMixin (GenTree.codeK T). Canonical tree_countType T := CountType (GenTree.tree T) (tree_countMixin T). End CountableDataTypes. ssreflect-1.5/theories/ssreflect.v0000644000175000017500000005227312175455021016353 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import Bool. (* For bool_scope delimiter 'bool'. *) Require Import ssrmatching. Set SsrAstVersion. (******************************************************************************) (* This file is the Gallina part of the ssreflect plugin implementation. *) (* Files that use the ssreflect plugin should always Require ssreflect and *) (* either Import ssreflect or Import ssreflect.SsrSyntax. *) (* Part of the contents of this file is technical and will only interest *) (* advanced developers; in addition the following are defined: *) (* [the str of v by f] == the Canonical s : str such that f s = v. *) (* [the str of v] == the Canonical s : str that coerces to v. *) (* argumentType c == the T such that c : forall x : T, P x. *) (* returnType c == the R such that c : T -> R. *) (* {type of c for s} == P s where c : forall x : T, P x. *) (* phantom T v == singleton type with inhabitant Phantom T v. *) (* phant T == singleton type with inhabitant Phant v. *) (* =^~ r == the converse of rewriting rule r (e.g., in a *) (* rewrite multirule). *) (* unkeyed t == t, but treated as an unkeyed matching pattern by *) (* the ssreflect matching algorithm. *) (* nosimpl t == t, but on the right-hand side of Definition C := *) (* nosimpl disables expansion of C by /=. *) (* locked t == t, but locked t is not convertible to t. *) (* locked_with k t == t, but not convertible to t or locked_with k' t *) (* unless k = k' (with k : unit). Coq type-checking *) (* will be much more efficient if locked_with with a *) (* bespoke k is used for sealed definitions. *) (* unlockable v == interface for sealed constant definitions of v. *) (* Unlockable def == the unlockable that registers def : C = v. *) (* [unlockable of C] == a clone for C of the canonical unlockable for the *) (* definition of C (e.g., if it uses locked_with). *) (* [unlockable fun C] == [unlockable of C] with the expansion forced to be *) (* an explicit lambda expression. *) (* -> The usage pattern for ADT operations is: *) (* Definition foo_def x1 .. xn := big_foo_expression. *) (* Fact foo_key : unit. Proof. by []. Qed. *) (* Definition foo := locked_with foo_key foo_def. *) (* Canonical foo_unlockable := [unlockable fun foo]. *) (* This minimizes the comparison overhead for foo, while still allowing *) (* rewrite unlock to expose big_foo_expression. *) (* More information about these definitions and their use can be found in the *) (* ssreflect manual, and in specific comments below. *) (******************************************************************************) Global Set Asymmetric Patterns. Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Module SsrSyntax. (* Declare Ssr keywords: 'is' 'of' '//' '/=' and '//='. We also declare the *) (* parsing level 8, as a workaround for a notation grammar factoring problem. *) (* Arguments of application-style notations (at level 10) should be declared *) (* at level 8 rather than 9 or the camlp5 grammar will not factor properly. *) Reserved Notation "(* x 'is' y 'of' z 'isn't' // /= //= *)" (at level 8). Reserved Notation "(* 69 *)" (at level 69). End SsrSyntax. Export SsrMatchingSyntax. Export SsrSyntax. (* Make the general "if" into a notation, so that we can override it below. *) (* The notations are "only parsing" because the Coq decompiler will not *) (* recognize the expansion of the boolean if; using the default printer *) (* avoids a spurrious trailing %GEN_IF. *) Delimit Scope general_if_scope with GEN_IF. Notation "'if' c 'then' v1 'else' v2" := (if c then v1 else v2) (at level 200, c, v1, v2 at level 200, only parsing) : general_if_scope. Notation "'if' c 'return' t 'then' v1 'else' v2" := (if c return t then v1 else v2) (at level 200, c, t, v1, v2 at level 200, only parsing) : general_if_scope. Notation "'if' c 'as' x 'return' t 'then' v1 'else' v2" := (if c as x return t then v1 else v2) (at level 200, c, t, v1, v2 at level 200, x ident, only parsing) : general_if_scope. (* Force boolean interpretation of simple if expressions. *) Delimit Scope boolean_if_scope with BOOL_IF. Notation "'if' c 'return' t 'then' v1 'else' v2" := (if c%bool is true in bool return t then v1 else v2) : boolean_if_scope. Notation "'if' c 'then' v1 'else' v2" := (if c%bool is true in bool return _ then v1 else v2) : boolean_if_scope. Notation "'if' c 'as' x 'return' t 'then' v1 'else' v2" := (if c%bool is true as x in bool return t then v1 else v2) : boolean_if_scope. Open Scope boolean_if_scope. (* To allow a wider variety of notations without reserving a large number of *) (* of identifiers, the ssreflect library systematically uses "forms" to *) (* enclose complex mixfix syntax. A "form" is simply a mixfix expression *) (* enclosed in square brackets and introduced by a keyword: *) (* [keyword ... ] *) (* Because the keyword follows a bracket it does not need to be reserved. *) (* Non-ssreflect libraries that do not respect the form syntax (e.g., the Coq *) (* Lists library) should be loaded before ssreflect so that their notations *) (* do not mask all ssreflect forms. *) Delimit Scope form_scope with FORM. Open Scope form_scope. (* Allow overloading of the cast (x : T) syntax, put whitespace around the *) (* ":" symbol to avoid lexical clashes (and for consistency with the parsing *) (* precedence of the notation, which binds less tightly than application), *) (* and put printing boxes that print the type of a long definition on a *) (* separate line rather than force-fit it at the right margin. *) Notation "x : T" := (x : T) (at level 100, right associativity, format "'[hv' x '/ ' : T ']'") : core_scope. (* Allow the casual use of notations like nat * nat for explicit Type *) (* declarations. Note that (nat * nat : Type) is NOT equivalent to *) (* (nat * nat)%type, whose inferred type is legacy type "Set". *) Notation "T : 'Type'" := (T%type : Type) (at level 100, only parsing) : core_scope. (* Allow similarly Prop annotation for, e.g., rewrite multirules. *) Notation "P : 'Prop'" := (P%type : Prop) (at level 100, only parsing) : core_scope. (* Syntax for referring to canonical structures: *) (* [the struct_type of proj_val by proj_fun] *) (* This form denotes the Canonical instance s of the Structure type *) (* struct_type whose proj_fun projection is proj_val, i.e., such that *) (* proj_fun s = proj_val. *) (* Typically proj_fun will be A record field accessors of struct_type, but *) (* this need not be the case; it can be, for instance, a field of a record *) (* type to which struct_type coerces; proj_val will likewise be coerced to *) (* the return type of proj_fun. In all but the simplest cases, proj_fun *) (* should be eta-expanded to allow for the insertion of implicit arguments. *) (* In the common case where proj_fun itself is a coercion, the "by" part *) (* can be omitted entirely; in this case it is inferred by casting s to the *) (* inferred type of proj_val. Obviously the latter can be fixed by using an *) (* explicit cast on proj_val, and it is highly recommended to do so when the *) (* return type intended for proj_fun is "Type", as the type inferred for *) (* proj_val may vary because of sort polymorphism (it could be Set or Prop). *) (* Note when using the [the _ of _] form to generate a substructure from a *) (* telescopes-style canonical hierarchy (implementing inheritance with *) (* coercions), one should always project or coerce the value to the BASE *) (* structure, because Coq will only find a Canonical derived structure for *) (* the Canonical base structure -- not for a base structure that is specific *) (* to proj_value. *) Module TheCanonical. CoInductive put vT sT (v1 v2 : vT) (s : sT) := Put. Definition get vT sT v s (p : @put vT sT v v s) := let: Put := p in s. Definition get_by vT sT of sT -> vT := @get vT sT. End TheCanonical. Import TheCanonical. (* Note: no export. *) Notation "[ 'the' sT 'of' v 'by' f ]" := (@get_by _ sT f _ _ ((fun v' (s : sT) => Put v' (f s) s) v _)) (at level 0, only parsing) : form_scope. Notation "[ 'the' sT 'of' v ]" := (get ((fun s : sT => Put v (*coerce*)s s) _)) (at level 0, only parsing) : form_scope. (* The following are "format only" versions of the above notations. Since Coq *) (* doesn't provide this facility, we fake it by splitting the "the" keyword. *) (* We need to do this to prevent the formatter from being be thrown off by *) (* application collapsing, coercion insertion and beta reduction in the right *) (* hand side of the notations above. *) Notation "[ 'th' 'e' sT 'of' v 'by' f ]" := (@get_by _ sT f v _ _) (at level 0, format "[ 'th' 'e' sT 'of' v 'by' f ]") : form_scope. Notation "[ 'th' 'e' sT 'of' v ]" := (@get _ sT v _ _) (at level 0, format "[ 'th' 'e' sT 'of' v ]") : form_scope. (* We would like to recognize Notation "[ 'th' 'e' sT 'of' v : 'Type' ]" := (@get Type sT v _ _) (at level 0, format "[ 'th' 'e' sT 'of' v : 'Type' ]") : form_scope. *) (* Helper notation for canonical structure inheritance support. *) (* This is a workaround for the poor interaction between delta reduction and *) (* canonical projections in Coq's unification algorithm, by which transparent *) (* definitions hide canonical instances, i.e., in *) (* Canonical a_type_struct := @Struct a_type ... *) (* Definition my_type := a_type. *) (* my_type doesn't effectively inherit the struct structure from a_type. Our *) (* solution is to redeclare the instance as follows *) (* Canonical my_type_struct := Eval hnf in [struct of my_type]. *) (* The special notation [str of _] must be defined for each Strucure "str" *) (* with constructor "Str", typically as follows *) (* Definition clone_str s := *) (* let: Str _ x y ... z := s return {type of Str for s} -> str in *) (* fun k => k _ x y ... z. *) (* Notation "[ 'str' 'of' T 'for' s ]" := (@clone_str s (@Str T)) *) (* (at level 0, format "[ 'str' 'of' T 'for' s ]") : form_scope. *) (* Notation "[ 'str' 'of' T ]" := (repack_str (fun x => @Str T x)) *) (* (at level 0, format "[ 'str' 'of' T ]") : form_scope. *) (* The notation for the match return predicate is defined below; the eta *) (* expansion in the second form serves both to distinguish it from the first *) (* and to avoid the delta reduction problem. *) (* There are several variations on the notation and the definition of the *) (* the "clone" function, for telescopes, mixin classes, and join (multiple *) (* inheritance) classes. We describe a different idiom for clones in ssrfun; *) (* it uses phantom types (see below) and static unification; see fintype and *) (* ssralg for examples. *) Definition argumentType T P & forall x : T, P x := T. Definition dependentReturnType T P & forall x : T, P x := P. Definition returnType aT rT & aT -> rT := rT. Notation "{ 'type' 'of' c 'for' s }" := (dependentReturnType c s) (at level 0, format "{ 'type' 'of' c 'for' s }") : type_scope. (* A generic "phantom" type (actually, a unit type with a phantom parameter). *) (* This type can be used for type definitions that require some Structure *) (* on one of their parameters, to allow Coq to infer said structure so it *) (* does not have to be supplied explicitly or via the "[the _ of _]" notation *) (* (the latter interacts poorly with other Notation). *) (* The definition of a (co)inductive type with a parameter p : p_type, that *) (* needs to use the operations of a structure *) (* Structure p_str : Type := p_Str {p_repr :> p_type; p_op : p_repr -> ...} *) (* should be given as *) (* Inductive indt_type (p : p_str) := Indt ... . *) (* Definition indt_of (p : p_str) & phantom p_type p := indt_type p. *) (* Notation "{ 'indt' p }" := (indt_of (Phantom p)). *) (* Definition indt p x y ... z : {indt p} := @Indt p x y ... z. *) (* Notation "[ 'indt' x y ... z ]" := (indt x y ... z). *) (* That is, the concrete type and its constructor should be shadowed by *) (* definitions that use a phantom argument to infer and display the true *) (* value of p (in practice, the "indt" constructor often performs additional *) (* functions, like "locking" the representation -- see below). *) (* We also define a simpler version ("phant" / "Phant") of phantom for the *) (* common case where p_type is Type. *) CoInductive phantom T (p : T) := Phantom. Implicit Arguments phantom []. Implicit Arguments Phantom []. CoInductive phant (p : Type) := Phant. (* Internal tagging used by the implementation of the ssreflect elim. *) Definition protect_term (A : Type) (x : A) : A := x. (* The ssreflect idiom for a non-keyed pattern: *) (* - unkeyed t wiil match any subterm that unifies with t, regardless of *) (* whether it displays the same head symbol as t. *) (* - unkeyed t a b will match any application of a term f unifying with t, *) (* to two arguments unifying with with a and b, repectively, regardless of *) (* apparent head symbols. *) (* - unkeyed x where x is a variable will match any subterm with the same *) (* type as x (when x would raise the 'indeterminate pattern' error). *) Notation unkeyed x := (let flex := x in flex). (* Ssreflect converse rewrite rule rule idiom. *) Definition ssr_converse R (r : R) := (Logic.I, r). Notation "=^~ r" := (ssr_converse r) (at level 100) : form_scope. (* Term tagging (user-level). *) (* The ssreflect library uses four strengths of term tagging to restrict *) (* convertibility during type checking: *) (* nosimpl t simplifies to t EXCEPT in a definition; more precisely, given *) (* Definition foo := nosimpl bar, foo (or foo t') will NOT be expanded by *) (* the /= and //= switches unless it is in a forcing context (e.g., in *) (* match foo t' with ... end, foo t' will be reduced if this allows the *) (* match to be reduced). Note that nosimpl bar is simply notation for a *) (* a term that beta-iota reduces to bar; hence rewrite /foo will replace *) (* foo by bar, and rewrite -/foo will replace bar by foo. *) (* CAVEAT: nosimpl should not be used inside a Section, because the end of *) (* section "cooking" removes the iota redex. *) (* locked t is provably equal to t, but is not convertible to t; 'locked' *) (* provides support for selective rewriting, via the lock t : t = locked t *) (* Lemma, and the ssreflect unlock tactic. *) (* locked_with k t is equal but not convertible to t, much like locked t, *) (* but supports explicit tagging with a value k : unit. This is used to *) (* mitigate a flaw in the term comparison heuristic of the Coq kernel, *) (* which treats all terms of the form locked t as equal and conpares their *) (* arguments recursively, leading to an exponential blowup of comparison. *) (* For this reason locked_with should be used rather than locked when *) (* defining ADT operations. The unlock tactic does not support locked_with *) (* but the unlock rewrite rule does, via the unlockable interface. *) (* we also use Module Type ascription to create truly opaque constants, *) (* because simple expansion of constants to reveal an unreducible term *) (* doubles the time complexity of a negative comparison. Such opaque *) (* constants can be expanded generically with the unlock rewrite rule. *) (* See the definition of card and subset in fintype for examples of this. *) Notation nosimpl t := (let: tt := tt in t). Lemma master_key : unit. Proof. exact tt. Qed. Definition locked A := let: tt := master_key in fun x : A => x. Lemma lock A x : x = locked x :> A. Proof. unlock; reflexivity. Qed. (* Needed for locked predicates, in particular for eqType's. *) Lemma not_locked_false_eq_true : locked false <> true. Proof. unlock; discriminate. Qed. (* The basic closing tactic "done". *) Ltac done := trivial; hnf; intros; solve [ do ![solve [trivial | apply: sym_equal; trivial] | discriminate | contradiction | split] | case not_locked_false_eq_true; assumption | match goal with H : ~ _ |- _ => solve [case H; trivial] end ]. (* To unlock opaque constants. *) Structure unlockable T v := Unlockable {unlocked : T; _ : unlocked = v}. Lemma unlock T x C : @unlocked T x C = x. Proof. by case: C. Qed. Notation "[ 'unlockable' 'of' C ]" := (@Unlockable _ _ C (unlock _)) (at level 0, format "[ 'unlockable' 'of' C ]") : form_scope. Notation "[ 'unlockable' 'fun' C ]" := (@Unlockable _ (fun _ => _) C (unlock _)) (at level 0, format "[ 'unlockable' 'fun' C ]") : form_scope. (* Generic keyed constant locking. *) (* The argument order ensures that k is always compared before T. *) Definition locked_with k := let: tt := k in fun T x => x : T. (* This can be used as a cheap alternative to cloning the unlockable instance *) (* below, but with caution as unkeyed matching can be expensive. *) Lemma locked_withE T k x : unkeyed (locked_with k x) = x :> T. Proof. by case: k. Qed. (* Intensionaly, this instance will not apply to locked u. *) Canonical locked_with_unlockable T k x := @Unlockable T x (locked_with k x) (locked_withE k x). (* More accurate variant of unlock, and safer alternative to locked_withE. *) Lemma unlock_with T k x : unlocked (locked_with_unlockable k x) = x :> T. Proof. exact: unlock. Qed. (* The internal lemmas for the have tactics. *) Definition ssr_have Plemma Pgoal (step : Plemma) rest : Pgoal := rest step. Implicit Arguments ssr_have [Pgoal]. Definition ssr_suff Plemma Pgoal step (rest : Plemma) : Pgoal := step rest. Implicit Arguments ssr_suff [Pgoal]. Definition ssr_wlog := ssr_suff. Implicit Arguments ssr_wlog [Pgoal]. (* Internal N-ary congruence lemmas for the congr tactic. *) Fixpoint nary_congruence_statement (n : nat) : (forall B, (B -> B -> Prop) -> Prop) -> Prop := match n with | O => fun k => forall B, k B (fun x1 x2 : B => x1 = x2) | S n' => let k' A B e (f1 f2 : A -> B) := forall x1 x2, x1 = x2 -> (e (f1 x1) (f2 x2) : Prop) in fun k => forall A, nary_congruence_statement n' (fun B e => k _ (k' A B e)) end. Lemma nary_congruence n (k := fun B e => forall y : B, (e y y : Prop)) : nary_congruence_statement n k. Proof. have: k _ _ := _; rewrite {1}/k. elim: n k => [|n IHn] k k_P /= A; first exact: k_P. by apply: IHn => B e He; apply: k_P => f x1 x2 <-. Qed. Lemma ssr_congr_arrow Plemma Pgoal : Plemma = Pgoal -> Plemma -> Pgoal. Proof. by move->. Qed. Implicit Arguments ssr_congr_arrow []. (* View lemmas that don't use reflection. *) Section ApplyIff. Variables P Q : Prop. Hypothesis eqPQ : P <-> Q. Lemma iffLR : P -> Q. Proof. by case: eqPQ. Qed. Lemma iffRL : Q -> P. Proof. by case: eqPQ. Qed. Lemma iffLRn : ~P -> ~Q. Proof. by move=> nP tQ; case: nP; case: eqPQ tQ. Qed. Lemma iffRLn : ~Q -> ~P. Proof. by move=> nQ tP; case: nQ; case: eqPQ tP. Qed. End ApplyIff. Hint View for move/ iffLRn|2 iffRLn|2 iffLR|2 iffRL|2. Hint View for apply/ iffRLn|2 iffLRn|2 iffRL|2 iffLR|2. ssreflect-1.5/theories/eqtype.v0000644000175000017500000007612312175455021015670 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import ssreflect ssrfun ssrbool. (******************************************************************************) (* This file defines two "base" combinatorial interfaces: *) (* eqType == the structure for types with a decidable equality. *) (* Equality mixins can be made Canonical to allow generic *) (* folding of equality predicates. *) (* subType p == the structure for types isomorphic to {x : T | p x} with *) (* p : pred T for some type T. *) (* The eqType interface supports the following operations: *) (* x == y <=> x compares equal to y (this is a boolean test). *) (* x == y :> T <=> x == y at type T. *) (* x != y <=> x and y compare unequal. *) (* x != y :> T <=> " " " " at type T. *) (* x =P y :: a proof of reflect (x = y) (x == y); this coerces *) (* to x == y -> x = y. *) (* comparable T <-> equality on T is decidable *) (* := forall x y : T, decidable (x = y) *) (* comparableClass compT == eqType mixin/class for compT : comparable T. *) (* pred1 a == the singleton predicate [pred x | x == a]. *) (* pred2, pred3, pred4 == pair, triple, quad predicates. *) (* predC1 a == [pred x | x != a]. *) (* [predU1 a & A] == [pred x | (x == a) || (x \in A)]. *) (* [predD1 A & a] == [pred x | x != a & x \in A]. *) (* predU1 a P, predD1 P a == applicative versions of the above. *) (* frel f == the relation associated with f : T -> T. *) (* := [rel x y | f x == y]. *) (* invariant k f == elements of T whose k-class is f-invariant. *) (* := [pred x | k (f x) == k x] with f : T -> T. *) (* [fun x : T => e0 with a1 |-> e1, .., a_n |-> e_n] *) (* [eta f with a1 |-> e1, .., a_n |-> e_n] == *) (* the auto-expanding function that maps x = a_i to e_i, and other values *) (* of x to e0 (resp. f x). In the first form the `: T' is optional and x *) (* can occur in a_i or e_i. *) (* Equality on an eqType is proof-irrelevant (lemma eq_irrelevance). *) (* The eqType interface is implemented for most standard datatypes: *) (* bool, unit, void, option, prod (denoted A * B), sum (denoted A + B), *) (* sig (denoted {x | P}), sigT (denoted {i : I & T}). We also define *) (* tagged_as u v == v cast as T_(tag u) if tag v == tag u, else u. *) (* -> We have u == v <=> (tag u == tag v) && (tagged u == tagged_as u v). *) (* The subType interface supports the following operations: *) (* val == the generic injection from a subType S of T into T. *) (* For example, if u : {x : T | P}, then val u : T. *) (* val is injective because P is proof-irrelevant (P is in bool, *) (* and the is_true coercion expands to P = true). *) (* valP == the generic proof of P (val u) for u : subType P. *) (* Sub x Px == the generic constructor for a subType P; Px is a proof of P x *) (* and P should be inferred from the expected return type. *) (* insub x == the generic partial projection of T into a subType S of T. *) (* This returns an option S; if S : subType P then *) (* insub x = Some u with val u = x if P x, *) (* None if ~~ P x *) (* The insubP lemma encapsulates this dichotomy. *) (* P should be infered from the expected return type. *) (* innew x == total (non-option) variant of insub when P = predT. *) (* {? x | P} == option {x | P} (syntax for casting insub x). *) (* insubd u0 x == the generic projection with default value u0. *) (* := odflt u0 (insub x). *) (* insigd A0 x == special case of insubd for S == {x | x \in A}, where A0 is *) (* a proof of x0 \in A. *) (* insub_eq x == transparent version of insub x that expands to Some/None *) (* when P x can evaluate. *) (* The subType P interface is most often implemented using one of: *) (* [subType for S_val] *) (* where S_val : S -> T is the first projection of a type S isomorphic to *) (* {x : T | P}. *) (* [newType for S_val] *) (* where S_val : S -> T is the projection of a type S isomorphic to *) (* wrapped T; in this case P must be predT. *) (* [subType for S_val by Srect], [newType for S_val by Srect] *) (* variants of the above where the eliminator is explicitly provided. *) (* Here S no longer needs to be syntactically identical to {x | P x} or *) (* wrapped T, but it must have a derived constructor S_Sub statisfying an *) (* eliminator Srect identical to the one the Coq Inductive command would *) (* have generated, and S_val (S_Sub x Px) (resp. S_val (S_sub x) for the *) (* newType form) must be convertible to x. *) (* variant of the above when S is a wrapper type for T (so P = predT). *) (* [subType of S], [subType of S for S_val] *) (* clones the canonical subType structure for S; if S_val is specified, *) (* then it replaces the inferred projector. *) (* Subtypes inherit the eqType structure of their base types; the generic *) (* structure should be explicitly instantiated using the *) (* [eqMixin of S by <:] *) (* construct to declare the Equality mixin; this pattern is repeated for all *) (* the combinatorial interfaces (Choice, Countable, Finite). *) (* More generally, the eqType structure can be transfered by (partial) *) (* injections, using: *) (* InjEqMixin injf == an Equality mixin for T, using an f : T -> eT where *) (* eT has an eqType structure and injf : injective f. *) (* PcanEqMixin fK == an Equality mixin similarly derived from f and a left *) (* inverse partial function g and fK : pcancel f g. *) (* CanEqMixin fK == an Equality mixin similarly derived from f and a left *) (* inverse function g and fK : cancel f g. *) (* We add the following to the standard suffixes documented in ssrbool.v: *) (* 1, 2, 3, 4 -- explicit enumeration predicate for 1 (singleton), 2, 3, or *) (* 4 values. *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Module Equality. Definition axiom T (e : rel T) := forall x y, reflect (x = y) (e x y). Structure mixin_of T := Mixin {op : rel T; _ : axiom op}. Notation class_of := mixin_of (only parsing). Section ClassDef. Structure type := Pack {sort; _ : class_of sort; _ : Type}. Local Coercion sort : type >-> Sortclass. Variables (T : Type) (cT : type). Definition class := let: Pack _ c _ := cT return class_of cT in c. Definition pack c := @Pack T c T. Definition clone := fun c & cT -> T & phant_id (pack c) cT => pack c. End ClassDef. Module Exports. Coercion sort : type >-> Sortclass. Notation eqType := type. Notation EqMixin := Mixin. Notation EqType T m := (@pack T m). Notation "[ 'eqMixin' 'of' T ]" := (class _ : mixin_of T) (at level 0, format "[ 'eqMixin' 'of' T ]") : form_scope. Notation "[ 'eqType' 'of' T 'for' C ]" := (@clone T C _ idfun id) (at level 0, format "[ 'eqType' 'of' T 'for' C ]") : form_scope. Notation "[ 'eqType' 'of' T ]" := (@clone T _ _ id id) (at level 0, format "[ 'eqType' 'of' T ]") : form_scope. End Exports. End Equality. Export Equality.Exports. Definition eq_op T := Equality.op (Equality.class T). Lemma eqE T x : eq_op x = Equality.op (Equality.class T) x. Proof. by []. Qed. Lemma eqP T : Equality.axiom (@eq_op T). Proof. by case: T => ? []. Qed. Implicit Arguments eqP [T x y]. Delimit Scope eq_scope with EQ. Open Scope eq_scope. Notation "x == y" := (eq_op x y) (at level 70, no associativity) : bool_scope. Notation "x == y :> T" := ((x : T) == (y : T)) (at level 70, y at next level) : bool_scope. Notation "x != y" := (~~ (x == y)) (at level 70, no associativity) : bool_scope. Notation "x != y :> T" := (~~ (x == y :> T)) (at level 70, y at next level) : bool_scope. Notation "x =P y" := (eqP : reflect (x = y) (x == y)) (at level 70, no associativity) : eq_scope. Notation "x =P y :> T" := (eqP : reflect (x = y :> T) (x == y :> T)) (at level 70, y at next level, no associativity) : eq_scope. Prenex Implicits eq_op eqP. Lemma eq_refl (T : eqType) (x : T) : x == x. Proof. exact/eqP. Qed. Notation eqxx := eq_refl. Lemma eq_sym (T : eqType) (x y : T) : (x == y) = (y == x). Proof. exact/eqP/eqP. Qed. Hint Resolve eq_refl eq_sym. Section Contrapositives. Variable T : eqType. Implicit Types (A : pred T) (b : bool) (x : T). Lemma contraTeq b x y : (x != y -> ~~ b) -> b -> x = y. Proof. by move=> imp hyp; apply/eqP; apply: contraTT hyp. Qed. Lemma contraNeq b x y : (x != y -> b) -> ~~ b -> x = y. Proof. by move=> imp hyp; apply/eqP; apply: contraNT hyp. Qed. Lemma contraFeq b x y : (x != y -> b) -> b = false -> x = y. Proof. by move=> imp /negbT; apply: contraNeq. Qed. Lemma contraTneq b x y : (x = y -> ~~ b) -> b -> x != y. Proof. by move=> imp; apply: contraTN => /eqP. Qed. Lemma contraNneq b x y : (x = y -> b) -> ~~ b -> x != y. Proof. by move=> imp; apply: contraNN => /eqP. Qed. Lemma contraFneq b x y : (x = y -> b) -> b = false -> x != y. Proof. by move=> imp /negbT; apply: contraNneq. Qed. Lemma contra_eqN b x y : (b -> x != y) -> x = y -> ~~ b. Proof. by move=> imp /eqP; apply: contraL. Qed. Lemma contra_eqF b x y : (b -> x != y) -> x = y -> b = false. Proof. by move=> imp /eqP; apply: contraTF. Qed. Lemma contra_eqT b x y : (~~ b -> x != y) -> x = y -> b. Proof. by move=> imp /eqP; apply: contraLR. Qed. Lemma contra_eq x1 y1 x2 y2 : (x2 != y2 -> x1 != y1) -> x1 = y1 -> x2 = y2. Proof. by move=> imp /eqP; apply: contraTeq. Qed. Lemma contra_neq x1 y1 x2 y2 : (x2 = y2 -> x1 = y1) -> x1 != y1 -> x2 != y2. Proof. by move=> imp; apply: contraNneq => /imp->. Qed. Lemma memPn A x : reflect {in A, forall y, y != x} (x \notin A). Proof. apply: (iffP idP) => [notDx y | notDx]; first by apply: contraTneq => ->. exact: contraL (notDx x) _. Qed. Lemma memPnC A x : reflect {in A, forall y, x != y} (x \notin A). Proof. by apply: (iffP (memPn A x)) => A'x y /A'x; rewrite eq_sym. Qed. Lemma ifN_eq R x y vT vF : x != y -> (if x == y then vT else vF) = vF :> R. Proof. exact: ifN. Qed. Lemma ifN_eqC R x y vT vF : x != y -> (if y == x then vT else vF) = vF :> R. Proof. by rewrite eq_sym; apply: ifN. Qed. End Contrapositives. Implicit Arguments memPn [T A x]. Implicit Arguments memPnC [T A x]. Theorem eq_irrelevance (T : eqType) x y : forall e1 e2 : x = y :> T, e1 = e2. Proof. pose proj z e := if x =P z is ReflectT e0 then e0 else e. suff: injective (proj y) by rewrite /proj => injp e e'; apply: injp; case: eqP. pose join (e : x = _) := etrans (esym e). apply: can_inj (join x y (proj x (erefl x))) _. by case: y /; case: _ / (proj x _). Qed. Corollary eq_axiomK (T : eqType) (x : T) : all_equal_to (erefl x). Proof. move=> eq_x_x; exact: eq_irrelevance. Qed. (* We use the module system to circumvent a silly limitation that *) (* forbids using the same constant to coerce to different targets. *) Module Type EqTypePredSig. Parameter sort : eqType -> predArgType. End EqTypePredSig. Module MakeEqTypePred (eqmod : EqTypePredSig). Coercion eqmod.sort : eqType >-> predArgType. End MakeEqTypePred. Module Export EqTypePred := MakeEqTypePred Equality. Lemma unit_eqP : Equality.axiom (fun _ _ : unit => true). Proof. by do 2!case; left. Qed. Definition unit_eqMixin := EqMixin unit_eqP. Canonical unit_eqType := Eval hnf in EqType unit unit_eqMixin. (* Comparison for booleans. *) (* This is extensionally equal, but not convertible to Bool.eqb. *) Definition eqb b := addb (~~ b). Lemma eqbP : Equality.axiom eqb. Proof. by do 2!case; constructor. Qed. Canonical bool_eqMixin := EqMixin eqbP. Canonical bool_eqType := Eval hnf in EqType bool bool_eqMixin. Lemma eqbE : eqb = eq_op. Proof. by []. Qed. Lemma bool_irrelevance (x y : bool) (E E' : x = y) : E = E'. Proof. exact: eq_irrelevance. Qed. Lemma negb_add b1 b2 : ~~ (b1 (+) b2) = (b1 == b2). Proof. by rewrite -addNb. Qed. Lemma negb_eqb b1 b2 : (b1 != b2) = b1 (+) b2. Proof. by rewrite -addNb negbK. Qed. Lemma eqb_id b : (b == true) = b. Proof. by case: b. Qed. Lemma eqbF_neg b : (b == false) = ~~ b. Proof. by case: b. Qed. Lemma eqb_negLR b1 b2 : (~~ b1 == b2) = (b1 == ~~ b2). Proof. by case: b1; case: b2. Qed. (* Equality-based predicates. *) Notation xpred1 := (fun a1 x => x == a1). Notation xpred2 := (fun a1 a2 x => (x == a1) || (x == a2)). Notation xpred3 := (fun a1 a2 a3 x => [|| x == a1, x == a2 | x == a3]). Notation xpred4 := (fun a1 a2 a3 a4 x => [|| x == a1, x == a2, x == a3 | x == a4]). Notation xpredU1 := (fun a1 (p : pred _) x => (x == a1) || p x). Notation xpredC1 := (fun a1 x => x != a1). Notation xpredD1 := (fun (p : pred _) a1 x => (x != a1) && p x). Section EqPred. Variable T : eqType. Definition pred1 (a1 : T) := SimplPred (xpred1 a1). Definition pred2 (a1 a2 : T) := SimplPred (xpred2 a1 a2). Definition pred3 (a1 a2 a3 : T) := SimplPred (xpred3 a1 a2 a3). Definition pred4 (a1 a2 a3 a4 : T) := SimplPred (xpred4 a1 a2 a3 a4). Definition predU1 (a1 : T) p := SimplPred (xpredU1 a1 p). Definition predC1 (a1 : T) := SimplPred (xpredC1 a1). Definition predD1 p (a1 : T) := SimplPred (xpredD1 p a1). Lemma pred1E : pred1 =2 eq_op. Proof. move=> x y; exact: eq_sym. Qed. Variables (T2 : eqType) (x y : T) (z u : T2) (b : bool). Lemma predU1P : reflect (x = y \/ b) ((x == y) || b). Proof. apply: (iffP orP) => [] []; by [right | move/eqP; left]. Qed. Lemma pred2P : reflect (x = y \/ z = u) ((x == y) || (z == u)). Proof. by apply: (iffP orP) => [] [] /eqP; by [left | right]. Qed. Lemma predD1P : reflect (x <> y /\ b) ((x != y) && b). Proof. by apply: (iffP andP)=> [] [] // /eqP. Qed. Lemma predU1l : x = y -> (x == y) || b. Proof. by move->; rewrite eqxx. Qed. Lemma predU1r : b -> (x == y) || b. Proof. by move->; rewrite orbT. Qed. Lemma eqVneq : {x = y} + {x != y}. Proof. by case: eqP; [left | right]. Qed. End EqPred. Implicit Arguments predU1P [T x y b]. Implicit Arguments pred2P [T T2 x y z u]. Implicit Arguments predD1P [T x y b]. Prenex Implicits pred1 pred2 pred3 pred4 predU1 predC1 predD1 predU1P. Notation "[ 'predU1' x & A ]" := (predU1 x [mem A]) (at level 0, format "[ 'predU1' x & A ]") : fun_scope. Notation "[ 'predD1' A & x ]" := (predD1 [mem A] x) (at level 0, format "[ 'predD1' A & x ]") : fun_scope. (* Lemmas for reflected equality and functions. *) Section EqFun. Section Exo. Variables (aT rT : eqType) (D : pred aT) (f : aT -> rT) (g : rT -> aT). Lemma inj_eq : injective f -> forall x y, (f x == f y) = (x == y). Proof. by move=> inj_f x y; apply/eqP/eqP=> [|-> //]; exact: inj_f. Qed. Lemma can_eq : cancel f g -> forall x y, (f x == f y) = (x == y). Proof. move/can_inj; exact: inj_eq. Qed. Lemma bij_eq : bijective f -> forall x y, (f x == f y) = (x == y). Proof. move/bij_inj; apply: inj_eq. Qed. Lemma can2_eq : cancel f g -> cancel g f -> forall x y, (f x == y) = (x == g y). Proof. by move=> fK gK x y; rewrite -{1}[y]gK; exact: can_eq. Qed. Lemma inj_in_eq : {in D &, injective f} -> {in D &, forall x y, (f x == f y) = (x == y)}. Proof. by move=> inj_f x y Dx Dy; apply/eqP/eqP=> [|-> //]; exact: inj_f. Qed. Lemma can_in_eq : {in D, cancel f g} -> {in D &, forall x y, (f x == f y) = (x == y)}. Proof. by move/can_in_inj; exact: inj_in_eq. Qed. End Exo. Section Endo. Variable T : eqType. Definition frel f := [rel x y : T | f x == y]. Lemma inv_eq f : involutive f -> forall x y : T, (f x == y) = (x == f y). Proof. by move=> fK; exact: can2_eq. Qed. Lemma eq_frel f f' : f =1 f' -> frel f =2 frel f'. Proof. by move=> eq_f x y; rewrite /= eq_f. Qed. End Endo. Variable aT : Type. (* The invariant of an function f wrt a projection k is the pred of points *) (* that have the same projection as their image. *) Definition invariant (rT : eqType) f (k : aT -> rT) := [pred x | k (f x) == k x]. Variables (rT1 rT2 : eqType) (f : aT -> aT) (h : rT1 -> rT2) (k : aT -> rT1). Lemma invariant_comp : subpred (invariant f k) (invariant f (h \o k)). Proof. by move=> x eq_kfx; rewrite /= (eqP eq_kfx). Qed. Lemma invariant_inj : injective h -> invariant f (h \o k) =1 invariant f k. Proof. move=> inj_h x; exact: (inj_eq inj_h). Qed. End EqFun. Prenex Implicits frel. (* The coercion to rel must be explicit for derived Notations to unparse. *) Notation coerced_frel f := (rel_of_simpl_rel (frel f)) (only parsing). Section FunWith. Variables (aT : eqType) (rT : Type). CoInductive fun_delta : Type := FunDelta of aT & rT. Definition fwith x y (f : aT -> rT) := [fun z => if z == x then y else f z]. Definition app_fdelta df f z := let: FunDelta x y := df in if z == x then y else f z. End FunWith. Prenex Implicits fwith. Notation "x |-> y" := (FunDelta x y) (at level 190, no associativity, format "'[hv' x '/ ' |-> y ']'") : fun_delta_scope. Delimit Scope fun_delta_scope with FUN_DELTA. Arguments Scope app_fdelta [_ type_scope fun_delta_scope _ _]. Notation "[ 'fun' z : T => F 'with' d1 , .. , dn ]" := (SimplFunDelta (fun z : T => app_fdelta d1%FUN_DELTA .. (app_fdelta dn%FUN_DELTA (fun _ => F)) ..)) (at level 0, z ident, only parsing) : fun_scope. Notation "[ 'fun' z => F 'with' d1 , .. , dn ]" := (SimplFunDelta (fun z => app_fdelta d1%FUN_DELTA .. (app_fdelta dn%FUN_DELTA (fun _ => F)) ..)) (at level 0, z ident, format "'[hv' [ '[' 'fun' z => '/ ' F ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" ) : fun_scope. Notation "[ 'eta' f 'with' d1 , .. , dn ]" := (SimplFunDelta (fun _ => app_fdelta d1%FUN_DELTA .. (app_fdelta dn%FUN_DELTA f) ..)) (at level 0, z ident, format "'[hv' [ '[' 'eta' '/ ' f ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" ) : fun_scope. (* Various EqType constructions. *) Section ComparableType. Variable T : Type. Definition comparable := forall x y : T, decidable (x = y). Hypothesis Hcompare : comparable. Definition compareb x y : bool := Hcompare x y. Lemma compareP : Equality.axiom compareb. Proof. by move=> x y; exact: sumboolP. Qed. Definition comparableClass := EqMixin compareP. End ComparableType. Definition eq_comparable (T : eqType) : comparable T := fun x y => decP (x =P y). Section SubType. Variables (T : Type) (P : pred T). Structure subType : Type := SubType { sub_sort :> Type; val : sub_sort -> T; Sub : forall x, P x -> sub_sort; _ : forall K (_ : forall x Px, K (@Sub x Px)) u, K u; _ : forall x Px, val (@Sub x Px) = x }. Implicit Arguments Sub [s]. Lemma vrefl : forall x, P x -> x = x. Proof. by []. Qed. Definition vrefl_rect := vrefl. Definition clone_subType U v := fun sT & sub_sort sT -> U => fun c Urec cK (sT' := @SubType U v c Urec cK) & phant_id sT' sT => sT'. Variable sT : subType. CoInductive Sub_spec : sT -> Type := SubSpec x Px : Sub_spec (Sub x Px). Lemma SubP u : Sub_spec u. Proof. by case: sT Sub_spec SubSpec u => T' _ C rec /= _. Qed. Lemma SubK x Px : @val sT (Sub x Px) = x. Proof. by case: sT. Qed. Definition insub x := if @idP (P x) is ReflectT Px then @Some sT (Sub x Px) else None. Definition insubd u0 x := odflt u0 (insub x). CoInductive insub_spec x : option sT -> Type := | InsubSome u of P x & val u = x : insub_spec x (Some u) | InsubNone of ~~ P x : insub_spec x None. Lemma insubP x : insub_spec x (insub x). Proof. by rewrite /insub; case: {-}_ / idP; [left; rewrite ?SubK | right; exact/negP]. Qed. Lemma insubT x Px : insub x = Some (Sub x Px). Proof. case: insubP; last by case/negP. case/SubP=> y Py _ def_x; rewrite -def_x SubK in Px *. congr (Some (Sub _ _)); exact: bool_irrelevance. Qed. Lemma insubF x : P x = false -> insub x = None. Proof. by move/idP; case: insubP. Qed. Lemma insubN x : ~~ P x -> insub x = None. Proof. by move/negPf/insubF. Qed. Lemma isSome_insub : ([eta insub] : pred T) =1 P. Proof. by apply: fsym => x; case: insubP => // /negPf. Qed. Lemma insubK : ocancel insub (@val _). Proof. by move=> x; case: insubP. Qed. Lemma valP (u : sT) : P (val u). Proof. by case/SubP: u => x Px; rewrite SubK. Qed. Lemma valK : pcancel (@val _) insub. Proof. case/SubP=> x Px; rewrite SubK; exact: insubT. Qed. Lemma val_inj : injective (@val sT). Proof. exact: pcan_inj valK. Qed. Lemma valKd u0 : cancel (@val _) (insubd u0). Proof. by move=> u; rewrite /insubd valK. Qed. Lemma val_insubd u0 x : val (insubd u0 x) = if P x then x else val u0. Proof. by rewrite /insubd; case: insubP => [u -> | /negPf->]. Qed. Lemma insubdK u0 : {in P, cancel (insubd u0) (@val _)}. Proof. by move=> x Px; rewrite /= val_insubd [P x]Px. Qed. Definition insub_eq x := let Some_sub Px := Some (Sub x Px : sT) in let None_sub _ := None in (if P x as Px return P x = Px -> _ then Some_sub else None_sub) (erefl _). Lemma insub_eqE : insub_eq =1 insub. Proof. rewrite /insub_eq /insub => x; case: {2 3}_ / idP (erefl _) => // Px Px'. by congr (Some _); apply: val_inj; rewrite !SubK. Qed. End SubType. Implicit Arguments SubType [T P]. Implicit Arguments Sub [T P s]. Implicit Arguments vrefl [T P]. Implicit Arguments vrefl_rect [T P]. Implicit Arguments clone_subType [T P sT c Urec cK]. Implicit Arguments insub [T P sT]. Implicit Arguments insubT [T sT x]. Implicit Arguments val_inj [T P sT]. Prenex Implicits val Sub vrefl vrefl_rect insub insubd val_inj. Local Notation inlined_sub_rect := (fun K K_S u => let (x, Px) as u return K u := u in K_S x Px). Local Notation inlined_new_rect := (fun K K_S u => let (x) as u return K u := u in K_S x). Notation "[ 'subType' 'for' v ]" := (SubType _ v _ inlined_sub_rect vrefl_rect) (at level 0, only parsing) : form_scope. Notation "[ 'sub' 'Type' 'for' v ]" := (SubType _ v _ _ vrefl_rect) (at level 0, format "[ 'sub' 'Type' 'for' v ]") : form_scope. Notation "[ 'subType' 'for' v 'by' rec ]" := (SubType _ v _ rec vrefl) (at level 0, format "[ 'subType' 'for' v 'by' rec ]") : form_scope. Notation "[ 'subType' 'of' U 'for' v ]" := (clone_subType U v id idfun) (at level 0, format "[ 'subType' 'of' U 'for' v ]") : form_scope. (* Notation "[ 'subType' 'for' v ]" := (clone_subType _ v id idfun) (at level 0, format "[ 'subType' 'for' v ]") : form_scope. *) Notation "[ 'subType' 'of' U ]" := (clone_subType U _ id id) (at level 0, format "[ 'subType' 'of' U ]") : form_scope. Definition NewType T U v c Urec := let Urec' P IH := Urec P (fun x : T => IH x isT : P _) in SubType U v (fun x _ => c x) Urec'. Implicit Arguments NewType [T U]. Notation "[ 'newType' 'for' v ]" := (NewType v _ inlined_new_rect vrefl_rect) (at level 0, only parsing) : form_scope. Notation "[ 'new' 'Type' 'for' v ]" := (NewType v _ _ vrefl_rect) (at level 0, format "[ 'new' 'Type' 'for' v ]") : form_scope. Notation "[ 'newType' 'for' v 'by' rec ]" := (NewType v _ rec vrefl) (at level 0, format "[ 'newType' 'for' v 'by' rec ]") : form_scope. Definition innew T nT x := @Sub T predT nT x (erefl true). Implicit Arguments innew [T nT]. Prenex Implicits innew. Lemma innew_val T nT : cancel val (@innew T nT). Proof. by move=> u; apply: val_inj; exact: SubK. Qed. (* Prenex Implicits and renaming. *) Notation sval := (@proj1_sig _ _). Notation "@ 'sval'" := (@proj1_sig) (at level 10, format "@ 'sval'"). Section SigProj. Variables (T : Type) (P Q : T -> Prop). Lemma svalP : forall u : sig P, P (sval u). Proof. by case. Qed. Definition s2val (u : sig2 P Q) := let: exist2 x _ _ := u in x. Lemma s2valP u : P (s2val u). Proof. by case: u. Qed. Lemma s2valP' u : Q (s2val u). Proof. by case: u. Qed. End SigProj. Prenex Implicits svalP s2val s2valP s2valP'. Canonical sig_subType T (P : pred T) : subType [eta P] := Eval hnf in [subType for @sval T [eta [eta P]]]. (* Shorthand for sigma types over collective predicates. *) Notation "{ x 'in' A }" := {x | x \in A} (at level 0, x at level 99, format "{ x 'in' A }") : type_scope. Notation "{ x 'in' A | P }" := {x | (x \in A) && P} (at level 0, x at level 99, format "{ x 'in' A | P }") : type_scope. (* Shorthand for the return type of insub. *) Notation "{ ? x : T | P }" := (option {x : T | is_true P}) (at level 0, x at level 99, only parsing) : type_scope. Notation "{ ? x | P }" := {? x : _ | P} (at level 0, x at level 99, format "{ ? x | P }") : type_scope. Notation "{ ? x 'in' A }" := {? x | x \in A} (at level 0, x at level 99, format "{ ? x 'in' A }") : type_scope. Notation "{ ? x 'in' A | P }" := {? x | (x \in A) && P} (at level 0, x at level 99, format "{ ? x 'in' A | P }") : type_scope. (* A variant of injection with default that infers a collective predicate *) (* from the membership proof for the default value. *) Definition insigd T (A : mem_pred T) x (Ax : in_mem x A) := insubd (exist [eta A] x Ax). (* There should be a rel definition for the subType equality op, but this *) (* seems to cause the simpl tactic to diverge on expressions involving == *) (* on 4+ nested subTypes in a "strict" position (e.g., after ~~). *) (* Definition feq f := [rel x y | f x == f y]. *) Section TransferEqType. Variables (T : Type) (eT : eqType) (f : T -> eT). Lemma inj_eqAxiom : injective f -> Equality.axiom (fun x y => f x == f y). Proof. by move=> f_inj x y; apply: (iffP eqP) => [|-> //]; exact: f_inj. Qed. Definition InjEqMixin f_inj := EqMixin (inj_eqAxiom f_inj). Definition PcanEqMixin g (fK : pcancel f g) := InjEqMixin (pcan_inj fK). Definition CanEqMixin g (fK : cancel f g) := InjEqMixin (can_inj fK). End TransferEqType. Section SubEqType. Variables (T : eqType) (P : pred T) (sT : subType P). Notation Local ev_ax := (fun T v => @Equality.axiom T (fun x y => v x == v y)). Lemma val_eqP : ev_ax sT val. Proof. exact: inj_eqAxiom val_inj. Qed. Definition sub_eqMixin := EqMixin val_eqP. Canonical sub_eqType := Eval hnf in EqType sT sub_eqMixin. Definition SubEqMixin := (let: SubType _ v _ _ _ as sT' := sT return ev_ax sT' val -> Equality.class_of sT' in fun vP : ev_ax _ v => EqMixin vP ) val_eqP. Lemma val_eqE (u v : sT) : (val u == val v) = (u == v). Proof. by []. Qed. End SubEqType. Implicit Arguments val_eqP [T P sT x y]. Prenex Implicits val_eqP. Notation "[ 'eqMixin' 'of' T 'by' <: ]" := (SubEqMixin _ : Equality.class_of T) (at level 0, format "[ 'eqMixin' 'of' T 'by' <: ]") : form_scope. Section SigEqType. Variables (T : eqType) (P : pred T). Definition sig_eqMixin := Eval hnf in [eqMixin of {x | P x} by <:]. Canonical sig_eqType := Eval hnf in EqType {x | P x} sig_eqMixin. End SigEqType. Section ProdEqType. Variable T1 T2 : eqType. Definition pair_eq := [rel u v : T1 * T2 | (u.1 == v.1) && (u.2 == v.2)]. Lemma pair_eqP : Equality.axiom pair_eq. Proof. move=> [x1 x2] [y1 y2] /=; apply: (iffP andP) => [[]|[<- <-]] //=. by do 2!move/eqP->. Qed. Definition prod_eqMixin := EqMixin pair_eqP. Canonical prod_eqType := Eval hnf in EqType (T1 * T2) prod_eqMixin. Lemma pair_eqE : pair_eq = eq_op :> rel _. Proof. by []. Qed. Lemma xpair_eqE (x1 y1 : T1) (x2 y2 : T2) : ((x1, x2) == (y1, y2)) = ((x1 == y1) && (x2 == y2)). Proof. by []. Qed. Lemma pair_eq1 (u v : T1 * T2) : u == v -> u.1 == v.1. Proof. by case/andP. Qed. Lemma pair_eq2 (u v : T1 * T2) : u == v -> u.2 == v.2. Proof. by case/andP. Qed. End ProdEqType. Implicit Arguments pair_eqP [T1 T2]. Prenex Implicits pair_eqP. Definition predX T1 T2 (p1 : pred T1) (p2 : pred T2) := [pred z | p1 z.1 & p2 z.2]. Notation "[ 'predX' A1 & A2 ]" := (predX [mem A1] [mem A2]) (at level 0, format "[ 'predX' A1 & A2 ]") : fun_scope. Section OptionEqType. Variable T : eqType. Definition opt_eq (u v : option T) : bool := oapp (fun x => oapp (eq_op x) false v) (~~ v) u. Lemma opt_eqP : Equality.axiom opt_eq. Proof. case=> [x|] [y|] /=; by [constructor | apply: (iffP eqP) => [|[]] ->]. Qed. Canonical option_eqMixin := EqMixin opt_eqP. Canonical option_eqType := Eval hnf in EqType (option T) option_eqMixin. End OptionEqType. Definition tag := projS1. Definition tagged I T_ : forall u, T_(tag u) := @projS2 I [eta T_]. Definition Tagged I i T_ x := @existS I [eta T_] i x. Implicit Arguments Tagged [I i]. Prenex Implicits tag tagged Tagged. Section TaggedAs. Variables (I : eqType) (T_ : I -> Type). Implicit Types u v : {i : I & T_ i}. Definition tagged_as u v := if tag u =P tag v is ReflectT eq_uv then eq_rect_r T_ (tagged v) eq_uv else tagged u. Lemma tagged_asE u x : tagged_as u (Tagged T_ x) = x. Proof. by rewrite /tagged_as /=; case: eqP => // eq_uu; rewrite [eq_uu]eq_axiomK. Qed. End TaggedAs. Section TagEqType. Variables (I : eqType) (T_ : I -> eqType). Implicit Types u v : {i : I & T_ i}. Definition tag_eq u v := (tag u == tag v) && (tagged u == tagged_as u v). Lemma tag_eqP : Equality.axiom tag_eq. Proof. rewrite /tag_eq => [] [i x] [j] /=. case: eqP => [<-|Hij] y; last by right; case. by apply: (iffP eqP) => [->|<-]; rewrite tagged_asE. Qed. Canonical tag_eqMixin := EqMixin tag_eqP. Canonical tag_eqType := Eval hnf in EqType {i : I & T_ i} tag_eqMixin. Lemma tag_eqE : tag_eq = eq_op. Proof. by []. Qed. Lemma eq_tag u v : u == v -> tag u = tag v. Proof. by move/eqP->. Qed. Lemma eq_Tagged u x :(u == Tagged _ x) = (tagged u == x). Proof. by rewrite -tag_eqE /tag_eq eqxx tagged_asE. Qed. End TagEqType. Implicit Arguments tag_eqP [I T_ x y]. Prenex Implicits tag_eqP. Section SumEqType. Variables T1 T2 : eqType. Implicit Types u v : T1 + T2. Definition sum_eq u v := match u, v with | inl x, inl y | inr x, inr y => x == y | _, _ => false end. Lemma sum_eqP : Equality.axiom sum_eq. Proof. case=> x [] y /=; by [right | apply: (iffP eqP) => [->|[->]]]. Qed. Canonical sum_eqMixin := EqMixin sum_eqP. Canonical sum_eqType := Eval hnf in EqType (T1 + T2) sum_eqMixin. Lemma sum_eqE : sum_eq = eq_op. Proof. by []. Qed. End SumEqType. Implicit Arguments sum_eqP [T1 T2 x y]. Prenex Implicits sum_eqP. ssreflect-1.5/theories/ssrmatching.v0000644000175000017500000000205412175455021016673 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Declare ML Module "ssreflect". Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Module SsrMatchingSyntax. (* Reserve the notation for rewrite patterns so that the user is not allowed *) (* to declare it at a different level. *) Reserved Notation "( a 'in' b )" (at level 0). Reserved Notation "( a 'as' b )" (at level 0). Reserved Notation "( a 'in' b 'in' c )" (at level 0). Reserved Notation "( a 'as' b 'in' c )" (at level 0). (* Notation to define shortcuts for the "X in t" part of a pattern. *) Notation "( X 'in' t )" := (_ : fun X => t) : ssrpatternscope. Delimit Scope ssrpatternscope with pattern. (* Some shortcuts for recurrent "X in t" parts. *) Notation RHS := (X in _ = X)%pattern. Notation LHS := (X in X = _)%pattern. End SsrMatchingSyntax. Export SsrMatchingSyntax. ssreflect-1.5/theories/ssrbool.v0000644000175000017500000024730612175455021016047 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import ssreflect ssrfun. (******************************************************************************) (* A theory of boolean predicates and operators. A large part of this file is *) (* concerned with boolean reflection. *) (* Definitions and notations: *) (* is_true b == the coercion of b : bool to Prop (:= b = true). *) (* This is just input and displayed as `b''. *) (* reflect P b == the reflection inductive predicate, asserting *) (* that the logical proposition P : prop with the *) (* formula b : bool. Lemmas asserting reflect P b *) (* are often referred to as "views". *) (* iffP, appP, sameP, rwP :: lemmas for direct manipulation of reflection *) (* views: iffP is used to prove reflection from *) (* logical equivalence, appP to compose views, and *) (* sameP and rwP to perform boolean and setoid *) (* rewriting. *) (* elimT :: coercion reflect >-> Funclass, which allows the *) (* direct application of `reflect' views to *) (* boolean assertions. *) (* decidable P <-> P is effectively decidable (:= {P} + {~ P}. *) (* contra, contraL, ... :: contraposition lemmas. *) (* altP my_viewP :: natural alternative for reflection; given *) (* lemma myvieP: reflect my_Prop my_formula, *) (* have [myP | not_myP] := altP my_viewP. *) (* generates two subgoals, in which my_formula has *) (* been replaced by true and false, resp., with *) (* new assumptions myP : my_Prop and *) (* not_myP: ~~ my_formula. *) (* Caveat: my_formula must be an APPLICATION, not *) (* a variable, constant, let-in, etc. (due to the *) (* poor behaviour of dependent index matching). *) (* boolP my_formula :: boolean disjunction, equivalent to *) (* altP (idP my_formula) but circumventing the *) (* dependent index capture issue; destructing *) (* boolP my_formula generates two subgoals with *) (* assumtions my_formula and ~~ myformula. As *) (* with altP, my_formula must be an application. *) (* unless C P <-> hP : P may be assumed when proving P. *) (* := (P -> C) -> C (Pierce's law). *) (* This is slightly weaker but easier to use than *) (* P \/ C when P C : Prop. *) (* classically P <-> hP : P can be assumed when proving is_true b *) (* := forall b : bool, (P -> b) -> b. *) (* This is equivalent to ~ (~ P) when P : Prop. *) (* a && b == the boolean conjunction of a and b. *) (* a || b == then boolean disjunction of a and b. *) (* a ==> b == the boolean implication of b by a. *) (* ~~ a == the boolean negation of a. *) (* a (+) b == the boolean exclusive or (or sum) of a and b. *) (* [ /\ P1 , P2 & P3 ] == multiway logical conjunction, up to 5 terms. *) (* [ \/ P1 , P2 | P3 ] == multiway logical disjunction, up to 4 terms. *) (* [&& a, b, c & d] == iterated, right associative boolean conjunction *) (* with arbitrary arity. *) (* [|| a, b, c | d] == iterated, right associative boolean disjunction *) (* with arbitrary arity. *) (* [==> a, b, c => d] == iterated, right associative boolean implication *) (* with arbitrary arity. *) (* and3P, ... == specific reflection lemmas for iterated *) (* connectives. *) (* andTb, orbAC, ... == systematic names for boolean connective *) (* properties (see suffix conventions below). *) (* prop_congr == a tactic to move a boolean equality from *) (* its coerced form in Prop to the equality *) (* in bool. *) (* bool_congr == resolution tactic for blindly weeding out *) (* like terms from boolean equalities (can fail). *) (* This file provides a theory of boolean predicates and relations: *) (* pred T == the type of bool predicates (:= T -> bool). *) (* simpl_pred T == the type of simplifying bool predicates, using *) (* the simpl_fun from ssrfun.v. *) (* rel T == the type of bool relations. *) (* := T -> pred T or T -> T -> bool. *) (* simpl_rel T == type of simplifying relations. *) (* predType == the generic predicate interface, supported for *) (* for lists and sets. *) (* pred_class == a coercion class for the predType projection to *) (* pred; declaring a coercion to pred_class is an *) (* alternative way of equipping a type with a *) (* predType structure, which interoperates better *) (* with coercion subtyping. This is used, e.g., *) (* for finite sets, so that finite groups inherit *) (* the membership operation by coercing to sets. *) (* If P is a predicate the proposition "x satisfies P" can be written *) (* applicatively as (P x), or using an explicit connective as (x \in P); in *) (* the latter case we say that P is a "collective" predicate. We use A, B *) (* rather than P, Q for collective predicates: *) (* x \in A == x satisfies the (collective) predicate A. *) (* x \notin A == x doesn't satisfy the (collective) predicate A. *) (* The pred T type can be used as a generic predicate type for either kind, *) (* but the two kinds of predicates should not be confused. When a "generic" *) (* pred T value of one type needs to be passed as the other the following *) (* conversions should be used explicitly: *) (* SimplPred P == a (simplifying) applicative equivalent of P. *) (* mem A == an applicative equivalent of A: *) (* mem A x simplifies to x \in A. *) (* Alternatively one can use the syntax for explicit simplifying predicates *) (* and relations (in the following x is bound in E): *) (* [pred x | E] == simplifying (see ssrfun) predicate x => E. *) (* [pred x : T | E] == predicate x => T, with a cast on the argument. *) (* [pred : T | P] == constant predicate P on type T. *) (* [pred x | E1 & E2] == [pred x | E1 && E2]; an x : T cast is allowed. *) (* [pred x in A] == [pred x | x in A]. *) (* [pred x in A | E] == [pred x | x in A & E]. *) (* [pred x in A | E1 & E2] == [pred x in A | E1 && E2]. *) (* [predU A & B] == union of two collective predicates A and B. *) (* [predI A & B] == intersection of collective predicates A and B. *) (* [predD A & B] == difference of collective predicates A and B. *) (* [predC A] == complement of the collective predicate A. *) (* [preim f of A] == preimage under f of the collective predicate A. *) (* predU P Q, ... == union, etc of applicative predicates. *) (* pred0 == the empty predicate. *) (* predT == the total (always true) predicate. *) (* if T : predArgType, then T coerces to predT. *) (* {: T} == T cast to predArgType (e.g., {: bool * nat}) *) (* In the following, x and y are bound in E: *) (* [rel x y | E] == simplifying relation x, y => E. *) (* [rel x y : T | E] == simplifying relation with arguments cast. *) (* [rel x y in A & B | E] == [rel x y | [&& x \in A, y \in B & E]]. *) (* [rel x y in A & B] == [rel x y | (x \in A) && (y \in B)]. *) (* [rel x y in A | E] == [rel x y in A & A | E]. *) (* [rel x y in A] == [rel x y in A & A]. *) (* relU R S == union of relations R and S. *) (* Explicit values of type pred T (i.e., lamdba terms) should always be used *) (* applicatively, while values of collection types implementing the predType *) (* interface, such as sequences or sets should always be used as collective *) (* predicates. Defined constants and functions of type pred T or simpl_pred T *) (* as well as the explicit simpl_pred T values described below, can generally *) (* be used either way. Note however that x \in A will not auto-simplify when *) (* A is an explicit simpl_pred T value; the generic simplification rule inE *) (* must be used (when A : pred T, the unfold_in rule can be used). Constants *) (* of type pred T with an explicit simpl_pred value do not auto-simplify when *) (* used applicatively, but can still be expanded with inE. This behavior can *) (* be controlled as follows: *) (* Let A : collective_pred T := [pred x | ... ]. *) (* The collective_pred T type is just an alias for pred T, but this cast *) (* stops rewrite inE from expanding the definition of A, thus treating A *) (* into an abstract collection (unfold_in or in_collective can be used to *) (* expand manually). *) (* Let A : applicative_pred T := [pred x | ...]. *) (* This cast causes inE to turn x \in A into the applicative A x form; *) (* A will then have to unfolded explicitly with the /A rule. This will *) (* also apply to any definition that reduces to A (e.g., Let B := A). *) (* Canonical A_app_pred := ApplicativePred A. *) (* This declaration, given after definition of A, similarly causes inE to *) (* turn x \in A into A x, but in addition allows the app_predE rule to *) (* turn A x back into x \in A; it can be used for any definition of type *) (* pred T, which makes it especially useful for ambivalent predicates *) (* as the relational transitive closure connect, that are used in both *) (* applicative and collective styles. *) (* Purely for aesthetics, we provide a subtype of collective predicates: *) (* qualifier q T == a pred T pretty-printing wrapper. An A : qualifier q T *) (* coerces to pred_class and thus behaves as a collective *) (* predicate, but x \in A and x \notin A are displayed as: *) (* x \is A and x \isn't A when q = 0, *) (* x \is a A and x \isn't a A when q = 1, *) (* x \is an A and x \isn't an A when q = 2, respectively. *) (* [qualify x | P] := Qualifier 0 (fun x => P), constructor for the above. *) (* [qualify x : T | P], [qualify a x | P], [qualify an X | P], etc. *) (* variants of the above with type constraints and different *) (* values of q. *) (* We provide an internal interface to support attaching properties (such as *) (* being multiplicative) to predicates: *) (* pred_key p == phantom type that will serve as a support for properties *) (* to be attached to p : pred_class; instances should be *) (* created with Fact/Qed so as to be opaque. *) (* KeyedPred k_p == an instance of the interface structure that attaches *) (* (k_p : pred_key P) to P; the structure projection is a *) (* coercion to pred_class. *) (* KeyedQualifier k_q == an instance of the interface structure that attaches *) (* (k_q : pred_key q) to (q : qualifier n T). *) (* DefaultPredKey p == a default value for pred_key p; the vernacular command *) (* Import DefaultKeying attaches this key to all predicates *) (* that are not explicitly keyed. *) (* Keys can be used to attach properties to predicates, qualifiers and *) (* generic nouns in a way that allows them to be used tranparently. The key *) (* projection of a predicate property structure such as unsignedPred should *) (* be a pred_key, not a pred, and corresponding lemmas will have the form *) (* Lemma rpredN R S (oppS : @opprPred R S) (kS : keyed_pred oppS) : *) (* {mono -%R: x / x \in kS}. *) (* Because x \in kS will be displayed as x \in S (or x \is S, etc), the *) (* canonical instance of opprPred will not normally be exposed (it will also *) (* be erased by /= simplification). In addition each predicate structure *) (* should have a DefaultPredKey Canonical instance that simply issues the *) (* property as a proof obligation (which can be caught by the Prop-irrelevant *) (* feature of the ssreflect plugin). *) (* Some properties of predicates and relations: *) (* A =i B <-> A and B are extensionally equivalent. *) (* {subset A <= B} <-> A is a (collective) subpredicate of B. *) (* subpred P Q <-> P is an (applicative) subpredicate or Q. *) (* subrel R S <-> R is a subrelation of S. *) (* In the following R is in rel T: *) (* reflexive R <-> R is reflexive. *) (* irreflexive R <-> R is irreflexive. *) (* symmetric R <-> R (in rel T) is symmetric (equation). *) (* pre_symmetric R <-> R is symmetric (implication). *) (* antisymmetric R <-> R is antisymmetric. *) (* total R <-> R is total. *) (* transitive R <-> R is transitive. *) (* left_transitive R <-> R is a congruence on its left hand side. *) (* right_transitive R <-> R is a congruence on its right hand side. *) (* equivalence_rel R <-> R is an equivalence relation. *) (* Localization of (Prop) predicates; if P1 is convertible to forall x, Qx, *) (* P2 to forall x y, Qxy and P3 to forall x y z, Qxyz : *) (* {for y, P1} <-> Qx{y / x}. *) (* {in A, P1} <-> forall x, x \in A -> Qx. *) (* {in A1 & A2, P2} <-> forall x y, x \in A1 -> y \in A2 -> Qxy. *) (* {in A &, P2} <-> forall x y, x \in A -> y \in A -> Qxy. *) (* {in A1 & A2 & A3, Q3} <-> forall x y z, *) (* x \in A1 -> y \in A2 -> z \in A3 -> Qxyz. *) (* {in A1 & A2 &, Q3} == {in A1 & A2 & A2, Q3}. *) (* {in A1 && A3, Q3} == {in A1 & A1 & A3, Q3}. *) (* {in A &&, Q3} == {in A & A & A, Q3}. *) (* {in A, bijective f} == f has a right inverse in A. *) (* {on C, P1} == forall x, (f x) \in C -> Qx *) (* when P1 is also convertible to Pf f. *) (* {on C &, P2} == forall x y, f x \in C -> f y \in C -> Qxy *) (* when P2 is also convertible to Pf f. *) (* {on C, P1' & g} == forall x, (f x) \in cd -> Qx *) (* when P1' is convertible to Pf f *) (* and P1' g is convertible to forall x, Qx. *) (* {on C, bijective f} == f has a right inverse on C. *) (* This file extends the lemma name suffix conventions of ssrfun as follows: *) (* A -- associativity, as in andbA : associative andb. *) (* AC -- right commutativity. *) (* ACA -- self-interchange (inner commutativity), e.g., *) (* orbACA : (a || b) || (c || d) = (a || c) || (b || d). *) (* b -- a boolean argument, as in andbb : idempotent andb. *) (* C -- commutativity, as in andbC : commutative andb, *) (* or predicate complement, as in predC. *) (* CA -- left commutativity. *) (* D -- predicate difference, as in predD. *) (* E -- elimination, as in negbEf : ~~ b = false -> b. *) (* F or f -- boolean false, as in andbF : b && false = false. *) (* I -- left/right injectivity, as in addbI : right_injective addb, *) (* or predicate intersection, as in predI. *) (* l -- a left-hand operation, as andb_orl : left_distributive andb orb. *) (* N or n -- boolean negation, as in andbN : a && (~~ a) = false. *) (* P -- a characteristic property, often a reflection lemma, as in *) (* andP : reflect (a /\ b) (a && b). *) (* r -- a right-hand operation, as orb_andr : rightt_distributive orb andb. *) (* T or t -- boolean truth, as in andbT: right_id true andb. *) (* U -- predicate union, as in predU. *) (* W -- weakening, as in in1W : {in D, forall x, P} -> forall x, P. *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Reserved Notation "~~ b" (at level 35, right associativity). Reserved Notation "b ==> c" (at level 55, right associativity). Reserved Notation "b1 (+) b2" (at level 50, left associativity). Reserved Notation "x \in A" (at level 70, format "'[hv' x '/ ' \in A ']'", no associativity). Reserved Notation "x \notin A" (at level 70, format "'[hv' x '/ ' \notin A ']'", no associativity). Reserved Notation "p1 =i p2" (at level 70, format "'[hv' p1 '/ ' =i p2 ']'", no associativity). (* We introduce a number of n-ary "list-style" notations that share a common *) (* format, namely *) (* [op arg1, arg2, ... last_separator last_arg] *) (* This usually denotes a right-associative applications of op, e.g., *) (* [&& a, b, c & d] denotes a && (b && (c && d)) *) (* The last_separator must be a non-operator token. Here we use &, | or =>; *) (* our default is &, but we try to match the intended meaning of op. The *) (* separator is a workaround for limitations of the parsing engine; the same *) (* limitations mean the separator cannot be omitted even when last_arg can. *) (* The Notation declarations are complicated by the separate treatment for *) (* some fixed arities (binary for bool operators, and all arities for Prop *) (* operators). *) (* We also use the square brackets in comprehension-style notations *) (* [type var separator expr] *) (* where "type" is the type of the comprehension (e.g., pred) and "separator" *) (* is | or => . It is important that in other notations a leading square *) (* bracket [ is always by an operator symbol or a fixed identifier. *) Reserved Notation "[ /\ P1 & P2 ]" (at level 0, only parsing). Reserved Notation "[ /\ P1 , P2 & P3 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 ']' '/ ' & P3 ] ']'"). Reserved Notation "[ /\ P1 , P2 , P3 & P4 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 ']' '/ ' & P4 ] ']'"). Reserved Notation "[ /\ P1 , P2 , P3 , P4 & P5 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 ']' '/ ' & P5 ] ']'"). Reserved Notation "[ \/ P1 | P2 ]" (at level 0, only parsing). Reserved Notation "[ \/ P1 , P2 | P3 ]" (at level 0, format "'[hv' [ \/ '[' P1 , '/' P2 ']' '/ ' | P3 ] ']'"). Reserved Notation "[ \/ P1 , P2 , P3 | P4 ]" (at level 0, format "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 ']' '/ ' | P4 ] ']'"). Reserved Notation "[ && b1 & c ]" (at level 0, only parsing). Reserved Notation "[ && b1 , b2 , .. , bn & c ]" (at level 0, format "'[hv' [ && '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/ ' & c ] ']'"). Reserved Notation "[ || b1 | c ]" (at level 0, only parsing). Reserved Notation "[ || b1 , b2 , .. , bn | c ]" (at level 0, format "'[hv' [ || '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/ ' | c ] ']'"). Reserved Notation "[ ==> b1 => c ]" (at level 0, only parsing). Reserved Notation "[ ==> b1 , b2 , .. , bn => c ]" (at level 0, format "'[hv' [ ==> '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/' => c ] ']'"). Reserved Notation "[ 'pred' : T => E ]" (at level 0, format "'[hv' [ 'pred' : T => '/ ' E ] ']'"). Reserved Notation "[ 'pred' x => E ]" (at level 0, x at level 8, format "'[hv' [ 'pred' x => '/ ' E ] ']'"). Reserved Notation "[ 'pred' x : T => E ]" (at level 0, x at level 8, format "'[hv' [ 'pred' x : T => '/ ' E ] ']'"). Reserved Notation "[ 'rel' x y => E ]" (at level 0, x, y at level 8, format "'[hv' [ 'rel' x y => '/ ' E ] ']'"). Reserved Notation "[ 'rel' x y : T => E ]" (at level 0, x, y at level 8, format "'[hv' [ 'rel' x y : T => '/ ' E ] ']'"). (* Shorter delimiter *) Delimit Scope bool_scope with B. Open Scope bool_scope. (* An alternative to xorb that behaves somewhat better wrt simplification. *) Definition addb b := if b then negb else id. (* Notation for && and || is declared in Init.Datatypes. *) Notation "~~ b" := (negb b) : bool_scope. Notation "b ==> c" := (implb b c) : bool_scope. Notation "b1 (+) b2" := (addb b1 b2) : bool_scope. (* Constant is_true b := b = true is defined in Init.Datatypes. *) Coercion is_true : bool >-> Sortclass. (* Prop *) Lemma prop_congr : forall b b' : bool, b = b' -> b = b' :> Prop. Proof. by move=> b b' ->. Qed. Ltac prop_congr := apply: prop_congr. (* Lemmas for trivial. *) Lemma is_true_true : true. Proof. by []. Qed. Lemma not_false_is_true : ~ false. Proof. by []. Qed. Lemma is_true_locked_true : locked true. Proof. by unlock. Qed. Hint Resolve is_true_true not_false_is_true is_true_locked_true. (* Shorter names. *) Definition isT := is_true_true. Definition notF := not_false_is_true. (* Negation lemmas. *) (* We generally take NEGATION as the standard form of a false condition: *) (* negative boolean hypotheses should be of the form ~~ b, rather than ~ b or *) (* b = false, as much as possible. *) Lemma negbT b : b = false -> ~~ b. Proof. by case: b. Qed. Lemma negbTE b : ~~ b -> b = false. Proof. by case: b. Qed. Lemma negbF b : (b : bool) -> ~~ b = false. Proof. by case: b. Qed. Lemma negbFE b : ~~ b = false -> b. Proof. by case: b. Qed. Lemma negbK : involutive negb. Proof. by case. Qed. Lemma negbNE b : ~~ ~~ b -> b. Proof. by case: b. Qed. Lemma negb_inj : injective negb. Proof. exact: can_inj negbK. Qed. Lemma negbLR b c : b = ~~ c -> ~~ b = c. Proof. exact: canLR negbK. Qed. Lemma negbRL b c : ~~ b = c -> b = ~~ c. Proof. exact: canRL negbK. Qed. Lemma contra (c b : bool) : (c -> b) -> ~~ b -> ~~ c. Proof. by case: b => //; case: c. Qed. Definition contraNN := contra. Lemma contraL (c b : bool) : (c -> ~~ b) -> b -> ~~ c. Proof. by case: b => //; case: c. Qed. Definition contraTN := contraL. Lemma contraR (c b : bool) : (~~ c -> b) -> ~~ b -> c. Proof. by case: b => //; case: c. Qed. Definition contraNT := contraR. Lemma contraLR (c b : bool) : (~~ c -> ~~ b) -> b -> c. Proof. by case: b => //; case: c. Qed. Definition contraTT := contraLR. Lemma contraT b : (~~ b -> false) -> b. Proof. by case: b => // ->. Qed. Lemma wlog_neg b : (~~ b -> b) -> b. Proof. by case: b => // ->. Qed. Lemma contraFT (c b : bool) : (~~ c -> b) -> b = false -> c. Proof. by move/contraR=> notb_c /negbT. Qed. Lemma contraFN (c b : bool) : (c -> b) -> b = false -> ~~ c. Proof. by move/contra=> notb_notc /negbT. Qed. Lemma contraTF (c b : bool) : (c -> ~~ b) -> b -> c = false. Proof. by move/contraL=> b_notc /b_notc/negbTE. Qed. Lemma contraNF (c b : bool) : (c -> b) -> ~~ b -> c = false. Proof. by move/contra=> notb_notc /notb_notc/negbTE. Qed. Lemma contraFF (c b : bool) : (c -> b) -> b = false -> c = false. Proof. by move/contraFN=> bF_notc /bF_notc/negbTE. Qed. (* Coercion of sum-style datatypes into bool, which makes it possible *) (* to use ssr's boolean if rather than Coq's "generic" if. *) Coercion isSome T (u : option T) := if u is Some _ then true else false. Coercion is_inl A B (u : A + B) := if u is inl _ then true else false. Coercion is_left A B (u : {A} + {B}) := if u is left _ then true else false. Coercion is_inleft A B (u : A + {B}) := if u is inleft _ then true else false. Prenex Implicits isSome is_inl is_left is_inleft. Definition decidable P := {P} + {~ P}. (* Lemmas for ifs with large conditions, which allow reasoning about the *) (* condition without repeating it inside the proof (the latter IS *) (* preferable when the condition is short). *) (* Usage : *) (* if the goal contains (if cond then ...) = ... *) (* case: ifP => Hcond. *) (* generates two subgoal, with the assumption Hcond : cond = true/false *) (* Rewrite if_same eliminates redundant ifs *) (* Rewrite (fun_if f) moves a function f inside an if *) (* Rewrite if_arg moves an argument inside a function-valued if *) Section BoolIf. Variables (A B : Type) (x : A) (f : A -> B) (b : bool) (vT vF : A). CoInductive if_spec (not_b : Prop) : bool -> A -> Set := | IfSpecTrue of b : if_spec not_b true vT | IfSpecFalse of not_b : if_spec not_b false vF. Lemma ifP : if_spec (b = false) b (if b then vT else vF). Proof. by case def_b: b; constructor. Qed. Lemma ifPn : if_spec (~~ b) b (if b then vT else vF). Proof. by case def_b: b; constructor; rewrite ?def_b. Qed. Lemma ifT : b -> (if b then vT else vF) = vT. Proof. by move->. Qed. Lemma ifF : b = false -> (if b then vT else vF) = vF. Proof. by move->. Qed. Lemma ifN : ~~ b -> (if b then vT else vF) = vF. Proof. by move/negbTE->. Qed. Lemma if_same : (if b then vT else vT) = vT. Proof. by case b. Qed. Lemma if_neg : (if ~~ b then vT else vF) = if b then vF else vT. Proof. by case b. Qed. Lemma fun_if : f (if b then vT else vF) = if b then f vT else f vF. Proof. by case b. Qed. Lemma if_arg (fT fF : A -> B) : (if b then fT else fF) x = if b then fT x else fF x. Proof. by case b. Qed. (* Turning a boolean "if" form into an application. *) Definition if_expr := if b then vT else vF. Lemma ifE : (if b then vT else vF) = if_expr. Proof. by []. Qed. End BoolIf. (* The reflection predicate. *) Inductive reflect (P : Prop) : bool -> Set := | ReflectT of P : reflect P true | ReflectF of ~ P : reflect P false. (* Core (internal) reflection lemmas, used for the three kinds of views. *) Section ReflectCore. Variables (P Q : Prop) (b c : bool). Hypothesis Hb : reflect P b. Lemma introNTF : (if c then ~ P else P) -> ~~ b = c. Proof. by case c; case Hb. Qed. Lemma introTF : (if c then P else ~ P) -> b = c. Proof. by case c; case Hb. Qed. Lemma elimNTF : ~~ b = c -> if c then ~ P else P. Proof. by move <-; case Hb. Qed. Lemma elimTF : b = c -> if c then P else ~ P. Proof. by move <-; case Hb. Qed. Lemma equivPif : (Q -> P) -> (P -> Q) -> if b then Q else ~ Q. Proof. by case Hb; auto. Qed. Lemma xorPif : Q \/ P -> ~ (Q /\ P) -> if b then ~ Q else Q. Proof. by case Hb => [? _ H ? | ? H _]; case: H. Qed. End ReflectCore. (* Internal negated reflection lemmas *) Section ReflectNegCore. Variables (P Q : Prop) (b c : bool). Hypothesis Hb : reflect P (~~ b). Lemma introTFn : (if c then ~ P else P) -> b = c. Proof. by move/(introNTF Hb) <-; case b. Qed. Lemma elimTFn : b = c -> if c then ~ P else P. Proof. by move <-; apply: (elimNTF Hb); case b. Qed. Lemma equivPifn : (Q -> P) -> (P -> Q) -> if b then ~ Q else Q. Proof. rewrite -if_neg; exact: equivPif. Qed. Lemma xorPifn : Q \/ P -> ~ (Q /\ P) -> if b then Q else ~ Q. Proof. rewrite -if_neg; exact: xorPif. Qed. End ReflectNegCore. (* User-oriented reflection lemmas *) Section Reflect. Variables (P Q : Prop) (b b' c : bool). Hypotheses (Pb : reflect P b) (Pb' : reflect P (~~ b')). Lemma introT : P -> b. Proof. exact: introTF true _. Qed. Lemma introF : ~ P -> b = false. Proof. exact: introTF false _. Qed. Lemma introN : ~ P -> ~~ b. Proof. exact: introNTF true _. Qed. Lemma introNf : P -> ~~ b = false. Proof. exact: introNTF false _. Qed. Lemma introTn : ~ P -> b'. Proof. exact: introTFn true _. Qed. Lemma introFn : P -> b' = false. Proof. exact: introTFn false _. Qed. Lemma elimT : b -> P. Proof. exact: elimTF true _. Qed. Lemma elimF : b = false -> ~ P. Proof. exact: elimTF false _. Qed. Lemma elimN : ~~ b -> ~P. Proof. exact: elimNTF true _. Qed. Lemma elimNf : ~~ b = false -> P. Proof. exact: elimNTF false _. Qed. Lemma elimTn : b' -> ~ P. Proof. exact: elimTFn true _. Qed. Lemma elimFn : b' = false -> P. Proof. exact: elimTFn false _. Qed. Lemma introP : (b -> Q) -> (~~ b -> ~ Q) -> reflect Q b. Proof. by case b; constructor; auto. Qed. Lemma iffP : (P -> Q) -> (Q -> P) -> reflect Q b. Proof. by case: Pb; constructor; auto. Qed. Lemma equivP : (P <-> Q) -> reflect Q b. Proof. by case; exact: iffP. Qed. Lemma sumboolP (decQ : decidable Q) : reflect Q decQ. Proof. by case: decQ; constructor. Qed. Lemma appP : reflect Q b -> P -> Q. Proof. by move=> Qb; move/introT; case: Qb. Qed. Lemma sameP : reflect P c -> b = c. Proof. case; [exact: introT | exact: introF]. Qed. Lemma decPcases : if b then P else ~ P. Proof. by case Pb. Qed. Definition decP : decidable P. by case: b decPcases; [left | right]. Defined. Lemma rwP : P <-> b. Proof. by split; [exact: introT | exact: elimT]. Qed. Lemma rwP2 : reflect Q b -> (P <-> Q). Proof. by move=> Qb; split=> ?; [exact: appP | apply: elimT; case: Qb]. Qed. (* Predicate family to reflect excluded middle in bool. *) CoInductive alt_spec : bool -> Type := | AltTrue of P : alt_spec true | AltFalse of ~~ b : alt_spec false. Lemma altP : alt_spec b. Proof. by case def_b: b / Pb; constructor; rewrite ?def_b. Qed. End Reflect. Hint View for move/ elimTF|3 elimNTF|3 elimTFn|3 introT|2 introTn|2 introN|2. Hint View for apply/ introTF|3 introNTF|3 introTFn|3 elimT|2 elimTn|2 elimN|2. Hint View for apply// equivPif|3 xorPif|3 equivPifn|3 xorPifn|3. (* Allow the direct application of a reflection lemma to a boolean assertion. *) Coercion elimT : reflect >-> Funclass. (* Pierce's law, a weak form of classical reasoning. *) Definition unless condition property := (property -> condition) -> condition. Lemma bind_unless C P {Q} : unless C P -> unless (unless C Q) P. Proof. by move=> haveP suffPQ suffQ; apply: haveP => /suffPQ; exact. Qed. Lemma unless_contra b C : (~~ b -> C) -> unless C b. Proof. by case: b => [_ haveC | haveC _]; exact: haveC. Qed. (* Classical reasoning becomes directly accessible for any bool subgoal. *) (* Note that we cannot use "unless" here for lack of universe polymorphism. *) Definition classically P : Prop := forall b : bool, (P -> b) -> b. Lemma classicP : forall P : Prop, classically P <-> ~ ~ P. Proof. move=> P; split=> [cP nP | nnP [] // nP]; last by case nnP; move/nP. by have: P -> false; [move/nP | move/cP]. Qed. Lemma classic_bind : forall P Q, (P -> classically Q) -> (classically P -> classically Q). Proof. by move=> P Q IH IH_P b IH_Q; apply: IH_P; move/IH; exact. Qed. Lemma classic_EM : forall P, classically (decidable P). Proof. by move=> P [] // IH; apply IH; right => ?; apply: notF (IH _); left. Qed. Lemma classic_imply : forall P Q, (P -> classically Q) -> classically (P -> Q). Proof. move=> P Q IH [] // notPQ; apply notPQ; move/IH=> hQ; case: notF. by apply: hQ => hQ; case: notF; exact: notPQ. Qed. Lemma classic_pick : forall T P, classically ({x : T | P x} + (forall x, ~ P x)). Proof. move=> T P [] // IH; apply IH; right=> x Px; case: notF. by apply: IH; left; exists x. Qed. (* List notations for wider connectives; the Prop connectives have a fixed *) (* width so as to avoid iterated destruction (we go up to width 5 for /\, and *) (* width 4 for or. The bool connectives have arbitrary widths, but denote *) (* expressions that associate to the RIGHT. This is consistent with the right *) (* associativity of list expressions and thus more convenient in most proofs. *) Inductive and3 (P1 P2 P3 : Prop) : Prop := And3 of P1 & P2 & P3. Inductive and4 (P1 P2 P3 P4 : Prop) : Prop := And4 of P1 & P2 & P3 & P4. Inductive and5 (P1 P2 P3 P4 P5 : Prop) : Prop := And5 of P1 & P2 & P3 & P4 & P5. Inductive or3 (P1 P2 P3 : Prop) : Prop := Or31 of P1 | Or32 of P2 | Or33 of P3. Inductive or4 (P1 P2 P3 P4 : Prop) : Prop := Or41 of P1 | Or42 of P2 | Or43 of P3 | Or44 of P4. Notation "[ /\ P1 & P2 ]" := (and P1 P2) (only parsing) : type_scope. Notation "[ /\ P1 , P2 & P3 ]" := (and3 P1 P2 P3) : type_scope. Notation "[ /\ P1 , P2 , P3 & P4 ]" := (and4 P1 P2 P3 P4) : type_scope. Notation "[ /\ P1 , P2 , P3 , P4 & P5 ]" := (and5 P1 P2 P3 P4 P5) : type_scope. Notation "[ \/ P1 | P2 ]" := (or P1 P2) (only parsing) : type_scope. Notation "[ \/ P1 , P2 | P3 ]" := (or3 P1 P2 P3) : type_scope. Notation "[ \/ P1 , P2 , P3 | P4 ]" := (or4 P1 P2 P3 P4) : type_scope. Notation "[ && b1 & c ]" := (b1 && c) (only parsing) : bool_scope. Notation "[ && b1 , b2 , .. , bn & c ]" := (b1 && (b2 && .. (bn && c) .. )) : bool_scope. Notation "[ || b1 | c ]" := (b1 || c) (only parsing) : bool_scope. Notation "[ || b1 , b2 , .. , bn | c ]" := (b1 || (b2 || .. (bn || c) .. )) : bool_scope. Notation "[ ==> b1 , b2 , .. , bn => c ]" := (b1 ==> (b2 ==> .. (bn ==> c) .. )) : bool_scope. Notation "[ ==> b1 => c ]" := (b1 ==> c) (only parsing) : bool_scope. Section AllAnd. Variables (T : Type) (P1 P2 P3 P4 P5 : T -> Prop). Local Notation a P := (forall x, P x). Lemma all_and2 (hP : forall x, [/\ P1 x & P2 x]) : [/\ a P1 & a P2]. Proof. by split=> x; case: (hP x). Qed. Lemma all_and3 (hP : forall x, [/\ P1 x, P2 x & P3 x]) : [/\ a P1, a P2 & a P3]. Proof. by split=> x; case: (hP x). Qed. Lemma all_and4 (hP : forall x, [/\ P1 x, P2 x, P3 x & P4 x]) : [/\ a P1, a P2, a P3 & a P4]. Proof. by split=> x; case: (hP x). Qed. Lemma all_and5 (hP : forall x, [/\ P1 x, P2 x, P3 x, P4 x & P5 x]) : [/\ a P1, a P2, a P3, a P4 & a P5]. Proof. by split=> x; case: (hP x). Qed. End AllAnd. Lemma pair_andP P Q : P /\ Q <-> P * Q. Proof. by split; case. Qed. Section ReflectConnectives. Variable b1 b2 b3 b4 b5 : bool. Lemma idP : reflect b1 b1. Proof. by case b1; constructor. Qed. Lemma boolP : alt_spec b1 b1 b1. Proof. exact: (altP idP). Qed. Lemma idPn : reflect (~~ b1) (~~ b1). Proof. by case b1; constructor. Qed. Lemma negP : reflect (~ b1) (~~ b1). Proof. by case b1; constructor; auto. Qed. Lemma negPn : reflect b1 (~~ ~~ b1). Proof. by case b1; constructor. Qed. Lemma negPf : reflect (b1 = false) (~~ b1). Proof. by case b1; constructor. Qed. Lemma andP : reflect (b1 /\ b2) (b1 && b2). Proof. by case b1; case b2; constructor=> //; case. Qed. Lemma and3P : reflect [/\ b1, b2 & b3] [&& b1, b2 & b3]. Proof. by case b1; case b2; case b3; constructor; try by case. Qed. Lemma and4P : reflect [/\ b1, b2, b3 & b4] [&& b1, b2, b3 & b4]. Proof. by case b1; case b2; case b3; case b4; constructor; try by case. Qed. Lemma and5P : reflect [/\ b1, b2, b3, b4 & b5] [&& b1, b2, b3, b4 & b5]. Proof. by case b1; case b2; case b3; case b4; case b5; constructor; try by case. Qed. Lemma orP : reflect (b1 \/ b2) (b1 || b2). Proof. by case b1; case b2; constructor; auto; case. Qed. Lemma or3P : reflect [\/ b1, b2 | b3] [|| b1, b2 | b3]. Proof. case b1; first by constructor; constructor 1. case b2; first by constructor; constructor 2. case b3; first by constructor; constructor 3. by constructor; case. Qed. Lemma or4P : reflect [\/ b1, b2, b3 | b4] [|| b1, b2, b3 | b4]. Proof. case b1; first by constructor; constructor 1. case b2; first by constructor; constructor 2. case b3; first by constructor; constructor 3. case b4; first by constructor; constructor 4. by constructor; case. Qed. Lemma nandP : reflect (~~ b1 \/ ~~ b2) (~~ (b1 && b2)). Proof. by case b1; case b2; constructor; auto; case; auto. Qed. Lemma norP : reflect (~~ b1 /\ ~~ b2) (~~ (b1 || b2)). Proof. by case b1; case b2; constructor; auto; case; auto. Qed. Lemma implyP : reflect (b1 -> b2) (b1 ==> b2). Proof. by case b1; case b2; constructor; auto. Qed. End ReflectConnectives. Implicit Arguments idP [b1]. Implicit Arguments idPn [b1]. Implicit Arguments negP [b1]. Implicit Arguments negPn [b1]. Implicit Arguments negPf [b1]. Implicit Arguments andP [b1 b2]. Implicit Arguments and3P [b1 b2 b3]. Implicit Arguments and4P [b1 b2 b3 b4]. Implicit Arguments and5P [b1 b2 b3 b4 b5]. Implicit Arguments orP [b1 b2]. Implicit Arguments or3P [b1 b2 b3]. Implicit Arguments or4P [b1 b2 b3 b4]. Implicit Arguments nandP [b1 b2]. Implicit Arguments norP [b1 b2]. Implicit Arguments implyP [b1 b2]. Prenex Implicits idP idPn negP negPn negPf. Prenex Implicits andP and3P and4P and5P orP or3P or4P nandP norP implyP. (* Shorter, more systematic names for the boolean connectives laws. *) Lemma andTb : left_id true andb. Proof. by []. Qed. Lemma andFb : left_zero false andb. Proof. by []. Qed. Lemma andbT : right_id true andb. Proof. by case. Qed. Lemma andbF : right_zero false andb. Proof. by case. Qed. Lemma andbb : idempotent andb. Proof. by case. Qed. Lemma andbC : commutative andb. Proof. by do 2!case. Qed. Lemma andbA : associative andb. Proof. by do 3!case. Qed. Lemma andbCA : left_commutative andb. Proof. by do 3!case. Qed. Lemma andbAC : right_commutative andb. Proof. by do 3!case. Qed. Lemma andbACA : interchange andb andb. Proof. by do 4!case. Qed. Lemma orTb : forall b, true || b. Proof. by []. Qed. Lemma orFb : left_id false orb. Proof. by []. Qed. Lemma orbT : forall b, b || true. Proof. by case. Qed. Lemma orbF : right_id false orb. Proof. by case. Qed. Lemma orbb : idempotent orb. Proof. by case. Qed. Lemma orbC : commutative orb. Proof. by do 2!case. Qed. Lemma orbA : associative orb. Proof. by do 3!case. Qed. Lemma orbCA : left_commutative orb. Proof. by do 3!case. Qed. Lemma orbAC : right_commutative orb. Proof. by do 3!case. Qed. Lemma orbACA : interchange orb orb. Proof. by do 4!case. Qed. Lemma andbN b : b && ~~ b = false. Proof. by case: b. Qed. Lemma andNb b : ~~ b && b = false. Proof. by case: b. Qed. Lemma orbN b : b || ~~ b = true. Proof. by case: b. Qed. Lemma orNb b : ~~ b || b = true. Proof. by case: b. Qed. Lemma andb_orl : left_distributive andb orb. Proof. by do 3!case. Qed. Lemma andb_orr : right_distributive andb orb. Proof. by do 3!case. Qed. Lemma orb_andl : left_distributive orb andb. Proof. by do 3!case. Qed. Lemma orb_andr : right_distributive orb andb. Proof. by do 3!case. Qed. Lemma andb_idl (a b : bool) : (b -> a) -> a && b = b. Proof. by case: a; case: b => // ->. Qed. Lemma andb_idr (a b : bool) : (a -> b) -> a && b = a. Proof. by case: a; case: b => // ->. Qed. Lemma andb_id2l (a b c : bool) : (a -> b = c) -> a && b = a && c. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma andb_id2r (a b c : bool) : (b -> a = c) -> a && b = c && b. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma orb_idl (a b : bool) : (a -> b) -> a || b = b. Proof. by case: a; case: b => // ->. Qed. Lemma orb_idr (a b : bool) : (b -> a) -> a || b = a. Proof. by case: a; case: b => // ->. Qed. Lemma orb_id2l (a b c : bool) : (~~ a -> b = c) -> a || b = a || c. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma orb_id2r (a b c : bool) : (~~ b -> a = c) -> a || b = c || b. Proof. by case: a; case: b; case: c => // ->. Qed. Lemma negb_and (a b : bool) : ~~ (a && b) = ~~ a || ~~ b. Proof. by case: a; case: b. Qed. Lemma negb_or (a b : bool) : ~~ (a || b) = ~~ a && ~~ b. Proof. by case: a; case: b. Qed. (* Pseudo-cancellation -- i.e, absorbtion *) Lemma andbK a b : a && b || a = a. Proof. by case: a; case: b. Qed. Lemma andKb a b : a || b && a = a. Proof. by case: a; case: b. Qed. Lemma orbK a b : (a || b) && a = a. Proof. by case: a; case: b. Qed. Lemma orKb a b : a && (b || a) = a. Proof. by case: a; case: b. Qed. (* Imply *) Lemma implybT b : b ==> true. Proof. by case: b. Qed. Lemma implybF b : (b ==> false) = ~~ b. Proof. by case: b. Qed. Lemma implyFb b : false ==> b. Proof. by []. Qed. Lemma implyTb b : (true ==> b) = b. Proof. by []. Qed. Lemma implybb b : b ==> b. Proof. by case: b. Qed. Lemma negb_imply a b : ~~ (a ==> b) = a && ~~ b. Proof. by case: a; case: b. Qed. Lemma implybE a b : (a ==> b) = ~~ a || b. Proof. by case: a; case: b. Qed. Lemma implyNb a b : (~~ a ==> b) = a || b. Proof. by case: a; case: b. Qed. Lemma implybN a b : (a ==> ~~ b) = (b ==> ~~ a). Proof. by case: a; case: b. Qed. Lemma implybNN a b : (~~ a ==> ~~ b) = b ==> a. Proof. by case: a; case: b. Qed. Lemma implyb_idl (a b : bool) : (~~ a -> b) -> (a ==> b) = b. Proof. by case: a; case: b => // ->. Qed. Lemma implyb_idr (a b : bool) : (b -> ~~ a) -> (a ==> b) = ~~ a. Proof. by case: a; case: b => // ->. Qed. Lemma implyb_id2l (a b c : bool) : (a -> b = c) -> (a ==> b) = (a ==> c). Proof. by case: a; case: b; case: c => // ->. Qed. (* Addition (xor) *) Lemma addFb : left_id false addb. Proof. by []. Qed. Lemma addbF : right_id false addb. Proof. by case. Qed. Lemma addbb : self_inverse false addb. Proof. by case. Qed. Lemma addbC : commutative addb. Proof. by do 2!case. Qed. Lemma addbA : associative addb. Proof. by do 3!case. Qed. Lemma addbCA : left_commutative addb. Proof. by do 3!case. Qed. Lemma addbAC : right_commutative addb. Proof. by do 3!case. Qed. Lemma addbACA : interchange addb addb. Proof. by do 4!case. Qed. Lemma andb_addl : left_distributive andb addb. Proof. by do 3!case. Qed. Lemma andb_addr : right_distributive andb addb. Proof. by do 3!case. Qed. Lemma addKb : left_loop id addb. Proof. by do 2!case. Qed. Lemma addbK : right_loop id addb. Proof. by do 2!case. Qed. Lemma addIb : left_injective addb. Proof. by do 3!case. Qed. Lemma addbI : right_injective addb. Proof. by do 3!case. Qed. Lemma addTb b : true (+) b = ~~ b. Proof. by []. Qed. Lemma addbT b : b (+) true = ~~ b. Proof. by case: b. Qed. Lemma addbN a b : a (+) ~~ b = ~~ (a (+) b). Proof. by case: a; case: b. Qed. Lemma addNb a b : ~~ a (+) b = ~~ (a (+) b). Proof. by case: a; case: b. Qed. Lemma addbP a b : reflect (~~ a = b) (a (+) b). Proof. by case: a; case: b; constructor. Qed. Implicit Arguments addbP [a b]. (* Resolution tactic for blindly weeding out common terms from boolean *) (* equalities. When faced with a goal of the form (andb/orb/addb b1 b2) = b3 *) (* they will try to locate b1 in b3 and remove it. This can fail! *) Ltac bool_congr := match goal with | |- (?X1 && ?X2 = ?X3) => first [ symmetry; rewrite -1?(andbC X1) -?(andbCA X1); congr 1 (andb X1); symmetry | case: (X1); [ rewrite ?andTb ?andbT // | by rewrite ?andbF /= ] ] | |- (?X1 || ?X2 = ?X3) => first [ symmetry; rewrite -1?(orbC X1) -?(orbCA X1); congr 1 (orb X1); symmetry | case: (X1); [ by rewrite ?orbT //= | rewrite ?orFb ?orbF ] ] | |- (?X1 (+) ?X2 = ?X3) => symmetry; rewrite -1?(addbC X1) -?(addbCA X1); congr 1 (addb X1); symmetry | |- (~~ ?X1 = ?X2) => congr 1 negb end. (******************************************************************************) (* Predicates, i.e., packaged functions to bool. *) (* - pred T, the basic type for predicates over a type T, is simply an alias *) (* for T -> bool. *) (* We actually distinguish two kinds of predicates, which we call applicative *) (* and collective, based on the syntax used to test them at some x in T: *) (* - For an applicative predicate P, one uses prefix syntax: *) (* P x *) (* Also, most operations on applicative predicates use prefix syntax as *) (* well (e.g., predI P Q). *) (* - For a collective predicate A, one uses infix syntax: *) (* x \in A *) (* and all operations on collective predicates use infix syntax as well *) (* (e.g., [predI A & B]). *) (* There are only two kinds of applicative predicates: *) (* - pred T, the alias for T -> bool mentioned above *) (* - simpl_pred T, an alias for simpl_fun T bool with a coercion to pred T *) (* that auto-simplifies on application (see ssrfun). *) (* On the other hand, the set of collective predicate types is open-ended via *) (* - predType T, a Structure that can be used to put Canonical collective *) (* predicate interpretation on other types, such as lists, tuples, *) (* finite sets, etc. *) (* Indeed, we define such interpretations for applicative predicate types, *) (* which can therefore also be used with the infix syntax, e.g., *) (* x \in predI P Q *) (* Moreover these infix forms are convertible to their prefix counterpart *) (* (e.g., predI P Q x which in turn simplifies to P x && Q x). The converse *) (* is not true, however; collective predicate types cannot, in general, be *) (* general, be used applicatively, because of the "uniform inheritance" *) (* restriction on implicit coercions. *) (* However, we do define an explicit generic coercion *) (* - mem : forall (pT : predType), pT -> mem_pred T *) (* where mem_pred T is a variant of simpl_pred T that preserves the infix *) (* syntax, i.e., mem A x auto-simplifies to x \in A. *) (* Indeed, the infix "collective" operators are notation for a prefix *) (* operator with arguments of type mem_pred T or pred T, applied to coerced *) (* collective predicates, e.g., *) (* Notation "x \in A" := (in_mem x (mem A)). *) (* This prevents the variability in the predicate type from interfering with *) (* the application of generic lemmas. Moreover this also makes it much easier *) (* to define generic lemmas, because the simplest type -- pred T -- can be *) (* used as the type of generic collective predicates, provided one takes care *) (* not to use it applicatively; this avoids the burden of having to declare a *) (* different predicate type for each predicate parameter of each section or *) (* lemma. *) (* This trick is made possible by the fact that the constructor of the *) (* mem_pred T type aligns the unification process, forcing a generic *) (* "collective" predicate A : pred T to unify with the actual collective B, *) (* which mem has coerced to pred T via an internal, hidden implicit coercion, *) (* supplied by the predType structure for B. Users should take care not to *) (* inadvertently "strip" (mem B) down to the coerced B, since this will *) (* expose the internal coercion: Coq will display a term B x that cannot be *) (* typed as such. The topredE lemma can be used to restore the x \in B *) (* syntax in this case. While -topredE can conversely be used to change *) (* x \in P into P x, it is safer to use the inE and memE lemmas instead, as *) (* they do not run the risk of exposing internal coercions. As a consequence *) (* it is better to explicitly cast a generic applicative pred T to simpl_pred *) (* using the SimplPred constructor, when it is used as a collective predicate *) (* (see, e.g., Lemma eq_big in bigop). *) (* We also sometimes "instantiate" the predType structure by defining a *) (* coercion to the sort of the predPredType structure. This works better for *) (* types such as {set T} that have subtypes that coerce to them, since the *) (* same coercion will be inserted by the application of mem. It also lets us *) (* turn any Type aT : predArgType into the total predicate over that type, *) (* i.e., fun _: aT => true. This allows us to write, e.g., #|'I_n| for the *) (* cardinal of the (finite) type of integers less than n. *) (* Collective predicates have a specific extensional equality, *) (* - A =i B, *) (* while applicative predicates use the extensional equality of functions, *) (* - P =1 Q *) (* The two forms are convertible, however. *) (* We lift boolean operations to predicates, defining: *) (* - predU (union), predI (intersection), predC (complement), *) (* predD (difference), and preim (preimage, i.e., composition) *) (* For each operation we define three forms, typically: *) (* - predU : pred T -> pred T -> simpl_pred T *) (* - [predU A & B], a Notation for predU (mem A) (mem B) *) (* - xpredU, a Notation for the lambda-expression inside predU, *) (* which is mostly useful as an argument of =1, since it exposes the head *) (* head constant of the expression to the ssreflect matching algorithm. *) (* The syntax for the preimage of a collective predicate A is *) (* - [preim f of A] *) (* Finally, the generic syntax for defining a simpl_pred T is *) (* - [pred x : T | P(x)], [pred x | P(x)], [pred x in A | P(x)], etc. *) (* We also support boolean relations, but only the applicative form, with *) (* types *) (* - rel T, an alias for T -> pred T *) (* - simpl_rel T, an auto-simplifying version, and syntax *) (* [rel x y | P(x,y)], [rel x y in A & B | P(x,y)], etc. *) (* The notation [rel of fA] can be used to coerce a function returning a *) (* collective predicate to one returning pred T. *) (* Finally, note that there is specific support for ambivalent predicates *) (* that can work in either style, as per this file's head descriptor. *) (******************************************************************************) Definition pred T := T -> bool. Identity Coercion fun_of_pred : pred >-> Funclass. Definition rel T := T -> pred T. Identity Coercion fun_of_rel : rel >-> Funclass. Notation xpred0 := (fun _ => false). Notation xpredT := (fun _ => true). Notation xpredI := (fun (p1 p2 : pred _) x => p1 x && p2 x). Notation xpredU := (fun (p1 p2 : pred _) x => p1 x || p2 x). Notation xpredC := (fun (p : pred _) x => ~~ p x). Notation xpredD := (fun (p1 p2 : pred _) x => ~~ p2 x && p1 x). Notation xpreim := (fun f (p : pred _) x => p (f x)). Notation xrelU := (fun (r1 r2 : rel _) x y => r1 x y || r2 x y). Section Predicates. Variables T : Type. Definition subpred (p1 p2 : pred T) := forall x, p1 x -> p2 x. Definition subrel (r1 r2 : rel T) := forall x y, r1 x y -> r2 x y. Definition simpl_pred := simpl_fun T bool. Definition applicative_pred := pred T. Definition collective_pred := pred T. Definition SimplPred (p : pred T) : simpl_pred := SimplFun p. Coercion pred_of_simpl (p : simpl_pred) : pred T := fun_of_simpl p. Coercion applicative_pred_of_simpl (p : simpl_pred) : applicative_pred := fun_of_simpl p. Coercion collective_pred_of_simpl (p : simpl_pred) : collective_pred := fun x => (let: SimplFun f := p in fun _ => f x) x. (* Note: applicative_of_simpl is convertible to pred_of_simpl, while *) (* collective_of_simpl is not. *) Definition pred0 := SimplPred xpred0. Definition predT := SimplPred xpredT. Definition predI p1 p2 := SimplPred (xpredI p1 p2). Definition predU p1 p2 := SimplPred (xpredU p1 p2). Definition predC p := SimplPred (xpredC p). Definition predD p1 p2 := SimplPred (xpredD p1 p2). Definition preim rT f (d : pred rT) := SimplPred (xpreim f d). Definition simpl_rel := simpl_fun T (pred T). Definition SimplRel (r : rel T) : simpl_rel := [fun x => r x]. Coercion rel_of_simpl_rel (r : simpl_rel) : rel T := fun x y => r x y. Definition relU r1 r2 := SimplRel (xrelU r1 r2). Lemma subrelUl r1 r2 : subrel r1 (relU r1 r2). Proof. by move=> *; apply/orP; left. Qed. Lemma subrelUr r1 r2 : subrel r2 (relU r1 r2). Proof. by move=> *; apply/orP; right. Qed. CoInductive mem_pred := Mem of pred T. Definition isMem pT topred mem := mem = (fun p : pT => Mem [eta topred p]). Structure predType := PredType { pred_sort :> Type; topred : pred_sort -> pred T; _ : {mem | isMem topred mem} }. Definition mkPredType pT toP := PredType (exist (@isMem pT toP) _ (erefl _)). Canonical predPredType := Eval hnf in @mkPredType (pred T) id. Canonical simplPredType := Eval hnf in mkPredType pred_of_simpl. Canonical boolfunPredType := Eval hnf in @mkPredType (T -> bool) id. Coercion pred_of_mem mp : pred_sort predPredType := let: Mem p := mp in [eta p]. Canonical memPredType := Eval hnf in mkPredType pred_of_mem. Definition clone_pred U := fun pT & pred_sort pT -> U => fun a mP (pT' := @PredType U a mP) & phant_id pT' pT => pT'. End Predicates. Implicit Arguments pred0 [T]. Implicit Arguments predT [T]. Prenex Implicits pred0 predT predI predU predC predD preim relU. Notation "[ 'pred' : T | E ]" := (SimplPred (fun _ : T => E%B)) (at level 0, format "[ 'pred' : T | E ]") : fun_scope. Notation "[ 'pred' x | E ]" := (SimplPred (fun x => E%B)) (at level 0, x ident, format "[ 'pred' x | E ]") : fun_scope. Notation "[ 'pred' x | E1 & E2 ]" := [pred x | E1 && E2 ] (at level 0, x ident, format "[ 'pred' x | E1 & E2 ]") : fun_scope. Notation "[ 'pred' x : T | E ]" := (SimplPred (fun x : T => E%B)) (at level 0, x ident, only parsing) : fun_scope. Notation "[ 'pred' x : T | E1 & E2 ]" := [pred x : T | E1 && E2 ] (at level 0, x ident, only parsing) : fun_scope. Notation "[ 'rel' x y | E ]" := (SimplRel (fun x y => E%B)) (at level 0, x ident, y ident, format "[ 'rel' x y | E ]") : fun_scope. Notation "[ 'rel' x y : T | E ]" := (SimplRel (fun x y : T => E%B)) (at level 0, x ident, y ident, only parsing) : fun_scope. Notation "[ 'predType' 'of' T ]" := (@clone_pred _ T _ id _ _ id) (at level 0, format "[ 'predType' 'of' T ]") : form_scope. (* This redundant coercion lets us "inherit" the simpl_predType canonical *) (* instance by declaring a coercion to simpl_pred. This hack is the only way *) (* to put a predType structure on a predArgType. We use simpl_pred rather *) (* than pred to ensure that /= removes the identity coercion. Note that the *) (* coercion will never be used directly for simpl_pred, since the canonical *) (* instance should always be resolved. *) Notation pred_class := (pred_sort (predPredType _)). Coercion sort_of_simpl_pred T (p : simpl_pred T) : pred_class := p : pred T. (* This lets us use some types as a synonym for their universal predicate. *) (* Unfortunately, this won't work for existing types like bool, unless we *) (* redefine bool, true, false and all bool ops. *) Definition predArgType := Type. Bind Scope type_scope with predArgType. Identity Coercion sort_of_predArgType : predArgType >-> Sortclass. Coercion pred_of_argType (T : predArgType) : simpl_pred T := predT. Notation "{ : T }" := (T%type : predArgType) (at level 0, format "{ : T }") : type_scope. (* These must be defined outside a Section because "cooking" kills the *) (* nosimpl tag. *) Definition mem T (pT : predType T) : pT -> mem_pred T := nosimpl (let: PredType _ _ (exist mem _) := pT return pT -> _ in mem). Definition in_mem T x mp := nosimpl pred_of_mem T mp x. Prenex Implicits mem. Coercion pred_of_mem_pred T mp := [pred x : T | in_mem x mp]. Definition eq_mem T p1 p2 := forall x : T, in_mem x p1 = in_mem x p2. Definition sub_mem T p1 p2 := forall x : T, in_mem x p1 -> in_mem x p2. Typeclasses Opaque eq_mem. Lemma sub_refl T (p : mem_pred T) : sub_mem p p. Proof. by []. Qed. Implicit Arguments sub_refl [[T] [p]]. Notation "x \in A" := (in_mem x (mem A)) : bool_scope. Notation "x \in A" := (in_mem x (mem A)) : bool_scope. Notation "x \notin A" := (~~ (x \in A)) : bool_scope. Notation "A =i B" := (eq_mem (mem A) (mem B)) : type_scope. Notation "{ 'subset' A <= B }" := (sub_mem (mem A) (mem B)) (at level 0, A, B at level 69, format "{ '[hv' 'subset' A '/ ' <= B ']' }") : type_scope. Notation "[ 'mem' A ]" := (pred_of_simpl (pred_of_mem_pred (mem A))) (at level 0, only parsing) : fun_scope. Notation "[ 'rel' 'of' fA ]" := (fun x => [mem (fA x)]) (at level 0, format "[ 'rel' 'of' fA ]") : fun_scope. Notation "[ 'predI' A & B ]" := (predI [mem A] [mem B]) (at level 0, format "[ 'predI' A & B ]") : fun_scope. Notation "[ 'predU' A & B ]" := (predU [mem A] [mem B]) (at level 0, format "[ 'predU' A & B ]") : fun_scope. Notation "[ 'predD' A & B ]" := (predD [mem A] [mem B]) (at level 0, format "[ 'predD' A & B ]") : fun_scope. Notation "[ 'predC' A ]" := (predC [mem A]) (at level 0, format "[ 'predC' A ]") : fun_scope. Notation "[ 'preim' f 'of' A ]" := (preim f [mem A]) (at level 0, format "[ 'preim' f 'of' A ]") : fun_scope. Notation "[ 'pred' x 'in' A ]" := [pred x | x \in A] (at level 0, x ident, format "[ 'pred' x 'in' A ]") : fun_scope. Notation "[ 'pred' x 'in' A | E ]" := [pred x | x \in A & E] (at level 0, x ident, format "[ 'pred' x 'in' A | E ]") : fun_scope. Notation "[ 'pred' x 'in' A | E1 & E2 ]" := [pred x | x \in A & E1 && E2 ] (at level 0, x ident, format "[ 'pred' x 'in' A | E1 & E2 ]") : fun_scope. Notation "[ 'rel' x y 'in' A & B | E ]" := [rel x y | (x \in A) && (y \in B) && E] (at level 0, x ident, y ident, format "[ 'rel' x y 'in' A & B | E ]") : fun_scope. Notation "[ 'rel' x y 'in' A & B ]" := [rel x y | (x \in A) && (y \in B)] (at level 0, x ident, y ident, format "[ 'rel' x y 'in' A & B ]") : fun_scope. Notation "[ 'rel' x y 'in' A | E ]" := [rel x y in A & A | E] (at level 0, x ident, y ident, format "[ 'rel' x y 'in' A | E ]") : fun_scope. Notation "[ 'rel' x y 'in' A ]" := [rel x y in A & A] (at level 0, x ident, y ident, format "[ 'rel' x y 'in' A ]") : fun_scope. Section simpl_mem. Variables (T : Type) (pT : predType T). Implicit Types (x : T) (p : pred T) (sp : simpl_pred T) (pp : pT). (* Bespoke structures that provide fine-grained control over matching the *) (* various forms of the \in predicate; note in particular the different forms *) (* of hoisting that are used. We had to work around several bugs in the *) (* implementation of unification, notably improper expansion of telescope *) (* projections and overwriting of a variable assignment by a later *) (* unification (probably due to conversion cache cross-talk). *) Structure manifest_applicative_pred p := ManifestApplicativePred { manifest_applicative_pred_value :> pred T; _ : manifest_applicative_pred_value = p }. Definition ApplicativePred p := ManifestApplicativePred (erefl p). Canonical applicative_pred_applicative sp := ApplicativePred (applicative_pred_of_simpl sp). Structure manifest_simpl_pred p := ManifestSimplPred { manifest_simpl_pred_value :> simpl_pred T; _ : manifest_simpl_pred_value = SimplPred p }. Canonical expose_simpl_pred p := ManifestSimplPred (erefl (SimplPred p)). Structure manifest_mem_pred p := ManifestMemPred { manifest_mem_pred_value :> mem_pred T; _ : manifest_mem_pred_value= Mem [eta p] }. Canonical expose_mem_pred p := @ManifestMemPred p _ (erefl _). Structure applicative_mem_pred p := ApplicativeMemPred {applicative_mem_pred_value :> manifest_mem_pred p}. Canonical check_applicative_mem_pred p (ap : manifest_applicative_pred p) mp := @ApplicativeMemPred ap mp. Lemma mem_topred (pp : pT) : mem (topred pp) = mem pp. Proof. by rewrite /mem; case: pT pp => T1 app1 [mem1 /= ->]. Qed. Lemma topredE x (pp : pT) : topred pp x = (x \in pp). Proof. by rewrite -mem_topred. Qed. Lemma app_predE x p (ap : manifest_applicative_pred p) : ap x = (x \in p). Proof. by case: ap => _ /= ->. Qed. Lemma in_applicative x p (amp : applicative_mem_pred p) : in_mem x amp = p x. Proof. by case: amp => [[_ /= ->]]. Qed. Lemma in_collective x p (msp : manifest_simpl_pred p) : (x \in collective_pred_of_simpl msp) = p x. Proof. by case: msp => _ /= ->. Qed. Lemma in_simpl x p (msp : manifest_simpl_pred p) : in_mem x (Mem [eta fun_of_simpl (msp : simpl_pred T)]) = p x. Proof. by case: msp => _ /= ->. Qed. (* Because of the explicit eta expansion in the left-hand side, this lemma *) (* should only be used in a right-to-left direction. The 8.3 hack allowing *) (* partial right-to-left use does not work with the improved expansion *) (* heuristics in 8.4. *) Lemma unfold_in x p : (x \in ([eta p] : pred T)) = p x. Proof. by []. Qed. Lemma simpl_predE p : SimplPred p =1 p. Proof. by []. Qed. Definition inE := (in_applicative, in_simpl, simpl_predE). (* to be extended *) Lemma mem_simpl sp : mem sp = sp :> pred T. Proof. by []. Qed. Definition memE := mem_simpl. (* could be extended *) Lemma mem_mem (pp : pT) : (mem (mem pp) = mem pp) * (mem [mem pp] = mem pp). Proof. by rewrite -mem_topred. Qed. End simpl_mem. (* Qualifiers and keyed predicates. *) CoInductive qualifier (q : nat) T := Qualifier of predPredType T. Coercion has_quality n T (q : qualifier n T) : pred_class := fun x => let: Qualifier p := q in p x. Implicit Arguments has_quality [T]. Lemma qualifE n T p x : (x \in @Qualifier n T p) = p x. Proof. by []. Qed. Notation "x \is A" := (x \in has_quality 0 A) (at level 70, no associativity, format "'[hv' x '/ ' \is A ']'") : bool_scope. Notation "x \is 'a' A" := (x \in has_quality 1 A) (at level 70, no associativity, format "'[hv' x '/ ' \is 'a' A ']'") : bool_scope. Notation "x \is 'an' A" := (x \in has_quality 2 A) (at level 70, no associativity, format "'[hv' x '/ ' \is 'an' A ']'") : bool_scope. Notation "x \isn't A" := (x \notin has_quality 0 A) (at level 70, no associativity, format "'[hv' x '/ ' \isn't A ']'") : bool_scope. Notation "x \isn't 'a' A" := (x \notin has_quality 1 A) (at level 70, no associativity, format "'[hv' x '/ ' \isn't 'a' A ']'") : bool_scope. Notation "x \isn't 'an' A" := (x \notin has_quality 2 A) (at level 70, no associativity, format "'[hv' x '/ ' \isn't 'an' A ']'") : bool_scope. Notation "[ 'qualify' x | P ]" := (Qualifier 0 (fun x => P%B)) (at level 0, x at level 99, format "'[hv' [ 'qualify' x | '/ ' P ] ']'") : form_scope. Notation "[ 'qualify' x : T | P ]" := (Qualifier 0 (fun x : T => P%B)) (at level 0, x at level 99, only parsing) : form_scope. Notation "[ 'qualify' 'a' x | P ]" := (Qualifier 1 (fun x => P%B)) (at level 0, x at level 99, format "'[hv' [ 'qualify' 'a' x | '/ ' P ] ']'") : form_scope. Notation "[ 'qualify' 'a' x : T | P ]" := (Qualifier 1 (fun x : T => P%B)) (at level 0, x at level 99, only parsing) : form_scope. Notation "[ 'qualify' 'an' x | P ]" := (Qualifier 2 (fun x => P%B)) (at level 0, x at level 99, format "'[hv' [ 'qualify' 'an' x | '/ ' P ] ']'") : form_scope. Notation "[ 'qualify' 'an' x : T | P ]" := (Qualifier 2 (fun x : T => P%B)) (at level 0, x at level 99, only parsing) : form_scope. (* Keyed predicates: support for property-bearing predicate interfaces. *) Section KeyPred. Variable T : Type. CoInductive pred_key (p : predPredType T) := DefaultPredKey. Variable p : predPredType T. Structure keyed_pred (k : pred_key p) := PackKeyedPred {unkey_pred :> pred_class; _ : unkey_pred =i p}. Variable k : pred_key p. Definition KeyedPred := @PackKeyedPred k p (frefl _). Variable k_p : keyed_pred k. Lemma keyed_predE : k_p =i p. Proof. by case: k_p. Qed. (* Instances that strip the mem cast; the first one has "pred_of_mem" as its *) (* projection head value, while the second has "pred_of_simpl". The latter *) (* has the side benefit of preempting accidental misdeclarations. *) (* Note: pred_of_mem is the registered mem >-> pred_class coercion, while *) (* simpl_of_mem; pred_of_simpl is the mem >-> pred >=> Funclass coercion. We *) (* must write down the coercions explicitly as the Canonical head constant *) (* computation does not strip casts !! *) Canonical keyed_mem := @PackKeyedPred k (pred_of_mem (mem k_p)) keyed_predE. Canonical keyed_mem_simpl := @PackKeyedPred k (pred_of_simpl (mem k_p)) keyed_predE. End KeyPred. Notation "x \i 'n' S" := (x \in @unkey_pred _ S _ _) (at level 70, format "'[hv' x '/ ' \i 'n' S ']'") : bool_scope. Section KeyedQualifier. Variables (T : Type) (n : nat) (q : qualifier n T). Structure keyed_qualifier (k : pred_key q) := PackKeyedQualifier {unkey_qualifier; _ : unkey_qualifier = q}. Definition KeyedQualifier k := PackKeyedQualifier k (erefl q). Variables (k : pred_key q) (k_q : keyed_qualifier k). Fact keyed_qualifier_suproof : unkey_qualifier k_q =i q. Proof. by case: k_q => /= _ ->. Qed. Canonical keyed_qualifier_keyed := PackKeyedPred k keyed_qualifier_suproof. End KeyedQualifier. Notation "x \i 's' A" := (x \i n has_quality 0 A) (at level 70, format "'[hv' x '/ ' \i 's' A ']'") : bool_scope. Notation "x \i 's' 'a' A" := (x \i n has_quality 1 A) (at level 70, format "'[hv' x '/ ' \i 's' 'a' A ']'") : bool_scope. Notation "x \i 's' 'an' A" := (x \i n has_quality 2 A) (at level 70, format "'[hv' x '/ ' \i 's' 'an' A ']'") : bool_scope. Module DefaultKeying. Canonical default_keyed_pred T p := KeyedPred (@DefaultPredKey T p). Canonical default_keyed_qualifier T n (q : qualifier n T) := KeyedQualifier (DefaultPredKey q). End DefaultKeying. (* Skolemizing with conditions. *) Lemma all_tag_cond_dep I T (C : pred I) U : (forall x, T x) -> (forall x, C x -> {y : T x & U x y}) -> {f : forall x, T x & forall x, C x -> U x (f x)}. Proof. move=> f0 fP; apply: all_tag (fun x y => C x -> U x y) _ => x. by case Cx: (C x); [case/fP: Cx => y; exists y | exists (f0 x)]. Qed. Lemma all_tag_cond I T (C : pred I) U : T -> (forall x, C x -> {y : T & U x y}) -> {f : I -> T & forall x, C x -> U x (f x)}. Proof. by move=> y0; apply: all_tag_cond_dep. Qed. Lemma all_sig_cond_dep I T (C : pred I) P : (forall x, T x) -> (forall x, C x -> {y : T x | P x y}) -> {f : forall x, T x | forall x, C x -> P x (f x)}. Proof. by move=> f0 /(all_tag_cond_dep f0)[f]; exists f. Qed. Lemma all_sig_cond I T (C : pred I) P : T -> (forall x, C x -> {y : T | P x y}) -> {f : I -> T | forall x, C x -> P x (f x)}. Proof. by move=> y0; apply: all_sig_cond_dep. Qed. Section RelationProperties. (* Caveat: reflexive should not be used to state lemmas, as auto and trivial *) (* will not expand the constant. *) Variable T : Type. Variable R : rel T. Definition total := forall x y, R x y || R y x. Definition transitive := forall y x z, R x y -> R y z -> R x z. Definition symmetric := forall x y, R x y = R y x. Definition antisymmetric := forall x y, R x y && R y x -> x = y. Definition pre_symmetric := forall x y, R x y -> R y x. Lemma symmetric_from_pre : pre_symmetric -> symmetric. Proof. move=> symR x y; apply/idP/idP; exact: symR. Qed. Definition reflexive := forall x, R x x. Definition irreflexive := forall x, R x x = false. Definition left_transitive := forall x y, R x y -> R x =1 R y. Definition right_transitive := forall x y, R x y -> R^~ x =1 R^~ y. Section PER. Hypotheses (symR : symmetric) (trR : transitive). Lemma sym_left_transitive : left_transitive. Proof. by move=> x y Rxy z; apply/idP/idP; apply: trR; rewrite // symR. Qed. Lemma sym_right_transitive : right_transitive. Proof. by move=> x y /sym_left_transitive Rxy z; rewrite !(symR z) Rxy. Qed. End PER. (* We define the equivalence property with prenex quantification so that it *) (* can be localized using the {in ..., ..} form defined below. *) Definition equivalence_rel := forall x y z, R z z * (R x y -> R x z = R y z). Lemma equivalence_relP : equivalence_rel <-> reflexive /\ left_transitive. Proof. split=> [eqiR | [Rxx trR] x y z]; last by split=> [|/trR->]. by split=> [x | x y Rxy z]; [rewrite (eqiR x x x) | rewrite (eqiR x y z)]. Qed. End RelationProperties. Lemma rev_trans T (R : rel T) : transitive R -> transitive (fun x y => R y x). Proof. by move=> trR x y z Ryx Rzy; exact: trR Rzy Ryx. Qed. (* Property localization *) Notation Local "{ 'all1' P }" := (forall x, P x : Prop) (at level 0). Notation Local "{ 'all2' P }" := (forall x y, P x y : Prop) (at level 0). Notation Local "{ 'all3' P }" := (forall x y z, P x y z: Prop) (at level 0). Notation Local ph := (phantom _). Section LocalProperties. Variables T1 T2 T3 : Type. Variables (d1 : mem_pred T1) (d2 : mem_pred T2) (d3 : mem_pred T3). Notation Local ph := (phantom Prop). Definition prop_for (x : T1) P & ph {all1 P} := P x. Lemma forE x P phP : @prop_for x P phP = P x. Proof. by []. Qed. Definition prop_in1 P & ph {all1 P} := forall x, in_mem x d1 -> P x. Definition prop_in11 P & ph {all2 P} := forall x y, in_mem x d1 -> in_mem y d2 -> P x y. Definition prop_in2 P & ph {all2 P} := forall x y, in_mem x d1 -> in_mem y d1 -> P x y. Definition prop_in111 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d2 -> in_mem z d3 -> P x y z. Definition prop_in12 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d2 -> in_mem z d2 -> P x y z. Definition prop_in21 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d1 -> in_mem z d2 -> P x y z. Definition prop_in3 P & ph {all3 P} := forall x y z, in_mem x d1 -> in_mem y d1 -> in_mem z d1 -> P x y z. Variable f : T1 -> T2. Definition prop_on1 Pf P & phantom T3 (Pf f) & ph {all1 P} := forall x, in_mem (f x) d2 -> P x. Definition prop_on2 Pf P & phantom T3 (Pf f) & ph {all2 P} := forall x y, in_mem (f x) d2 -> in_mem (f y) d2 -> P x y. End LocalProperties. Definition inPhantom := Phantom Prop. Definition onPhantom T P (x : T) := Phantom Prop (P x). Definition bijective_in aT rT (d : mem_pred aT) (f : aT -> rT) := exists2 g, prop_in1 d (inPhantom (cancel f g)) & prop_on1 d (Phantom _ (cancel g)) (onPhantom (cancel g) f). Definition bijective_on aT rT (cd : mem_pred rT) (f : aT -> rT) := exists2 g, prop_on1 cd (Phantom _ (cancel f)) (onPhantom (cancel f) g) & prop_in1 cd (inPhantom (cancel g f)). Notation "{ 'for' x , P }" := (prop_for x (inPhantom P)) (at level 0, format "{ 'for' x , P }") : type_scope. Notation "{ 'in' d , P }" := (prop_in1 (mem d) (inPhantom P)) (at level 0, format "{ 'in' d , P }") : type_scope. Notation "{ 'in' d1 & d2 , P }" := (prop_in11 (mem d1) (mem d2) (inPhantom P)) (at level 0, format "{ 'in' d1 & d2 , P }") : type_scope. Notation "{ 'in' d & , P }" := (prop_in2 (mem d) (inPhantom P)) (at level 0, format "{ 'in' d & , P }") : type_scope. Notation "{ 'in' d1 & d2 & d3 , P }" := (prop_in111 (mem d1) (mem d2) (mem d3) (inPhantom P)) (at level 0, format "{ 'in' d1 & d2 & d3 , P }") : type_scope. Notation "{ 'in' d1 & & d3 , P }" := (prop_in21 (mem d1) (mem d3) (inPhantom P)) (at level 0, format "{ 'in' d1 & & d3 , P }") : type_scope. Notation "{ 'in' d1 & d2 & , P }" := (prop_in12 (mem d1) (mem d2) (inPhantom P)) (at level 0, format "{ 'in' d1 & d2 & , P }") : type_scope. Notation "{ 'in' d & & , P }" := (prop_in3 (mem d) (inPhantom P)) (at level 0, format "{ 'in' d & & , P }") : type_scope. Notation "{ 'on' cd , P }" := (prop_on1 (mem cd) (inPhantom P) (inPhantom P)) (at level 0, format "{ 'on' cd , P }") : type_scope. Notation "{ 'on' cd & , P }" := (prop_on2 (mem cd) (inPhantom P) (inPhantom P)) (at level 0, format "{ 'on' cd & , P }") : type_scope. Notation "{ 'on' cd , P & g }" := (prop_on1 (mem cd) (Phantom (_ -> Prop) P) (onPhantom P g)) (at level 0, format "{ 'on' cd , P & g }") : type_scope. Notation "{ 'in' d , 'bijective' f }" := (bijective_in (mem d) f) (at level 0, f at level 8, format "{ 'in' d , 'bijective' f }") : type_scope. Notation "{ 'on' cd , 'bijective' f }" := (bijective_on (mem cd) f) (at level 0, f at level 8, format "{ 'on' cd , 'bijective' f }") : type_scope. (* Weakening and monotonicity lemmas for localized predicates. *) (* Note that using these lemmas in backward reasoning will force expansion of *) (* the predicate definition, as Coq needs to expose the quantifier to apply *) (* these lemmas. We define a few specialized variants to avoid this for some *) (* of the ssrfun predicates. *) Section LocalGlobal. Variables T1 T2 T3 : predArgType. Variables (D1 : pred T1) (D2 : pred T2) (D3 : pred T3). Variables (d1 d1' : mem_pred T1) (d2 d2' : mem_pred T2) (d3 d3' : mem_pred T3). Variables (f f' : T1 -> T2) (g : T2 -> T1) (h : T3). Variables (P1 : T1 -> Prop) (P2 : T1 -> T2 -> Prop). Variable P3 : T1 -> T2 -> T3 -> Prop. Variable Q1 : (T1 -> T2) -> T1 -> Prop. Variable Q1l : (T1 -> T2) -> T3 -> T1 -> Prop. Variable Q2 : (T1 -> T2) -> T1 -> T1 -> Prop. Hypothesis sub1 : sub_mem d1 d1'. Hypothesis sub2 : sub_mem d2 d2'. Hypothesis sub3 : sub_mem d3 d3'. Lemma in1W : {all1 P1} -> {in D1, {all1 P1}}. Proof. by move=> ? ?. Qed. Lemma in2W : {all2 P2} -> {in D1 & D2, {all2 P2}}. Proof. by move=> ? ?. Qed. Lemma in3W : {all3 P3} -> {in D1 & D2 & D3, {all3 P3}}. Proof. by move=> ? ?. Qed. Lemma in1T : {in T1, {all1 P1}} -> {all1 P1}. Proof. by move=> ? ?; auto. Qed. Lemma in2T : {in T1 & T2, {all2 P2}} -> {all2 P2}. Proof. by move=> ? ?; auto. Qed. Lemma in3T : {in T1 & T2 & T3, {all3 P3}} -> {all3 P3}. Proof. by move=> ? ?; auto. Qed. Lemma sub_in1 (Ph : ph {all1 P1}) : prop_in1 d1' Ph -> prop_in1 d1 Ph. Proof. move=> allP x /sub1; exact: allP. Qed. Lemma sub_in11 (Ph : ph {all2 P2}) : prop_in11 d1' d2' Ph -> prop_in11 d1 d2 Ph. Proof. move=> allP x1 x2 /sub1 d1x1 /sub2; exact: allP. Qed. Lemma sub_in111 (Ph : ph {all3 P3}) : prop_in111 d1' d2' d3' Ph -> prop_in111 d1 d2 d3 Ph. Proof. by move=> allP x1 x2 x3 /sub1 d1x1 /sub2 d2x2 /sub3; exact: allP. Qed. Let allQ1 f'' := {all1 Q1 f''}. Let allQ1l f'' h' := {all1 Q1l f'' h'}. Let allQ2 f'' := {all2 Q2 f''}. Lemma on1W : allQ1 f -> {on D2, allQ1 f}. Proof. by move=> ? ?. Qed. Lemma on1lW : allQ1l f h -> {on D2, allQ1l f & h}. Proof. by move=> ? ?. Qed. Lemma on2W : allQ2 f -> {on D2 &, allQ2 f}. Proof. by move=> ? ?. Qed. Lemma on1T : {on T2, allQ1 f} -> allQ1 f. Proof. by move=> ? ?; auto. Qed. Lemma on1lT : {on T2, allQ1l f & h} -> allQ1l f h. Proof. by move=> ? ?; auto. Qed. Lemma on2T : {on T2 &, allQ2 f} -> allQ2 f. Proof. by move=> ? ?; auto. Qed. Lemma subon1 (Phf : ph (allQ1 f)) (Ph : ph (allQ1 f)) : prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph. Proof. by move=> allQ x /sub2; exact: allQ. Qed. Lemma subon1l (Phf : ph (allQ1l f)) (Ph : ph (allQ1l f h)) : prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph. Proof. by move=> allQ x /sub2; exact: allQ. Qed. Lemma subon2 (Phf : ph (allQ2 f)) (Ph : ph (allQ2 f)) : prop_on2 d2' Phf Ph -> prop_on2 d2 Phf Ph. Proof. by move=> allQ x y /sub2=> d2fx /sub2; exact: allQ. Qed. Lemma can_in_inj : {in D1, cancel f g} -> {in D1 &, injective f}. Proof. by move=> fK x y /fK{2}<- /fK{2}<- ->. Qed. Lemma canLR_in x y : {in D1, cancel f g} -> y \in D1 -> x = f y -> g x = y. Proof. by move=> fK D1y ->; rewrite fK. Qed. Lemma canRL_in x y : {in D1, cancel f g} -> x \in D1 -> f x = y -> x = g y. Proof. by move=> fK D1x <-; rewrite fK. Qed. Lemma on_can_inj : {on D2, cancel f & g} -> {on D2 &, injective f}. Proof. by move=> fK x y /fK{2}<- /fK{2}<- ->. Qed. Lemma canLR_on x y : {on D2, cancel f & g} -> f y \in D2 -> x = f y -> g x = y. Proof. by move=> fK D2fy ->; rewrite fK. Qed. Lemma canRL_on x y : {on D2, cancel f & g} -> f x \in D2 -> f x = y -> x = g y. Proof. by move=> fK D2fx <-; rewrite fK. Qed. Lemma inW_bij : bijective f -> {in D1, bijective f}. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma onW_bij : bijective f -> {on D2, bijective f}. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma inT_bij : {in T1, bijective f} -> bijective f. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma onT_bij : {on T2, bijective f} -> bijective f. Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma sub_in_bij (D1' : pred T1) : {subset D1 <= D1'} -> {in D1', bijective f} -> {in D1, bijective f}. Proof. by move=> subD [g' fK g'K]; exists g' => x; move/subD; [exact: fK | exact: g'K]. Qed. Lemma subon_bij (D2' : pred T2) : {subset D2 <= D2'} -> {on D2', bijective f} -> {on D2, bijective f}. Proof. by move=> subD [g' fK g'K]; exists g' => x; move/subD; [exact: fK | exact: g'K]. Qed. End LocalGlobal. Lemma sub_in2 T d d' (P : T -> T -> Prop) : sub_mem d d' -> forall Ph : ph {all2 P}, prop_in2 d' Ph -> prop_in2 d Ph. Proof. by move=> /= sub_dd'; exact: sub_in11. Qed. Lemma sub_in3 T d d' (P : T -> T -> T -> Prop) : sub_mem d d' -> forall Ph : ph {all3 P}, prop_in3 d' Ph -> prop_in3 d Ph. Proof. by move=> /= sub_dd'; exact: sub_in111. Qed. Lemma sub_in12 T1 T d1 d1' d d' (P : T1 -> T -> T -> Prop) : sub_mem d1 d1' -> sub_mem d d' -> forall Ph : ph {all3 P}, prop_in12 d1' d' Ph -> prop_in12 d1 d Ph. Proof. by move=> /= sub1 sub; exact: sub_in111. Qed. Lemma sub_in21 T T3 d d' d3 d3' (P : T -> T -> T3 -> Prop) : sub_mem d d' -> sub_mem d3 d3' -> forall Ph : ph {all3 P}, prop_in21 d' d3' Ph -> prop_in21 d d3 Ph. Proof. by move=> /= sub sub3; exact: sub_in111. Qed. Lemma equivalence_relP_in T (R : rel T) (A : pred T) : {in A & &, equivalence_rel R} <-> {in A, reflexive R} /\ {in A &, forall x y, R x y -> {in A, R x =1 R y}}. Proof. split=> [eqiR | [Rxx trR] x y z *]; last by split=> [|/trR-> //]; exact: Rxx. by split=> [x Ax|x y Ax Ay Rxy z Az]; [rewrite (eqiR x x) | rewrite (eqiR x y)]. Qed. Section MonoHomoMorphismTheory. Variables (aT rT sT : Type) (f : aT -> rT) (g : rT -> aT). Variables (aP : pred aT) (rP : pred rT) (aR : rel aT) (rR : rel rT). Lemma monoW : {mono f : x / aP x >-> rP x} -> {homo f : x / aP x >-> rP x}. Proof. by move=> hf x ax; rewrite hf. Qed. Lemma mono2W : {mono f : x y / aR x y >-> rR x y} -> {homo f : x y / aR x y >-> rR x y}. Proof. by move=> hf x y axy; rewrite hf. Qed. Hypothesis fgK : cancel g f. Lemma homoRL : {homo f : x y / aR x y >-> rR x y} -> forall x y, aR (g x) y -> rR x (f y). Proof. by move=> Hf x y /Hf; rewrite fgK. Qed. Lemma homoLR : {homo f : x y / aR x y >-> rR x y} -> forall x y, aR x (g y) -> rR (f x) y. Proof. by move=> Hf x y /Hf; rewrite fgK. Qed. Lemma homo_mono : {homo f : x y / aR x y >-> rR x y} -> {homo g : x y / rR x y >-> aR x y} -> {mono g : x y / rR x y >-> aR x y}. Proof. move=> mf mg x y; case: (boolP (rR _ _))=> [/mg //|]. by apply: contraNF=> /mf; rewrite !fgK. Qed. Lemma monoLR : {mono f : x y / aR x y >-> rR x y} -> forall x y, rR (f x) y = aR x (g y). Proof. by move=> mf x y; rewrite -{1}[y]fgK mf. Qed. Lemma monoRL : {mono f : x y / aR x y >-> rR x y} -> forall x y, rR x (f y) = aR (g x) y. Proof. by move=> mf x y; rewrite -{1}[x]fgK mf. Qed. Lemma can_mono : {mono f : x y / aR x y >-> rR x y} -> {mono g : x y / rR x y >-> aR x y}. Proof. by move=> mf x y /=; rewrite -mf !fgK. Qed. End MonoHomoMorphismTheory. Section MonoHomoMorphismTheory_in. Variables (aT rT sT : predArgType) (f : aT -> rT) (g : rT -> aT). Variable (aD : pred aT). Variable (aP : pred aT) (rP : pred rT) (aR : rel aT) (rR : rel rT). Notation rD := [pred x | g x \in aD]. Lemma monoW_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in aD &, {homo f : x y / aR x y >-> rR x y}}. Proof. by move=> hf x y hx hy axy; rewrite hf. Qed. Lemma mono2W_in : {in aD, {mono f : x / aP x >-> rP x}} -> {in aD, {homo f : x / aP x >-> rP x}}. Proof. by move=> hf x hx ax; rewrite hf. Qed. Hypothesis fgK_on : {on aD, cancel g & f}. Lemma homoRL_in : {in aD &, {homo f : x y / aR x y >-> rR x y}} -> {in rD & aD, forall x y, aR (g x) y -> rR x (f y)}. Proof. by move=> Hf x y hx hy /Hf; rewrite fgK_on //; apply. Qed. Lemma homoLR_in : {in aD &, {homo f : x y / aR x y >-> rR x y}} -> {in aD & rD, forall x y, aR x (g y) -> rR (f x) y}. Proof. by move=> Hf x y hx hy /Hf; rewrite fgK_on //; apply. Qed. Lemma homo_mono_in : {in aD &, {homo f : x y / aR x y >-> rR x y}} -> {in rD &, {homo g : x y / rR x y >-> aR x y}} -> {in rD &, {mono g : x y / rR x y >-> aR x y}}. Proof. move=> mf mg x y hx hy; case: (boolP (rR _ _))=> [/mg //|]; first exact. by apply: contraNF=> /mf; rewrite !fgK_on //; apply. Qed. Lemma monoLR_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in aD & rD, forall x y, rR (f x) y = aR x (g y)}. Proof. by move=> mf x y hx hy; rewrite -{1}[y]fgK_on // mf. Qed. Lemma monoRL_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in rD & aD, forall x y, rR x (f y) = aR (g x) y}. Proof. by move=> mf x y hx hy; rewrite -{1}[x]fgK_on // mf. Qed. Lemma can_mono_in : {in aD &, {mono f : x y / aR x y >-> rR x y}} -> {in rD &, {mono g : x y / rR x y >-> aR x y}}. Proof. by move=> mf x y hx hy /=; rewrite -mf // !fgK_on. Qed. End MonoHomoMorphismTheory_in. ssreflect-1.5/theories/fintype.v0000644000175000017500000022375012175455021016037 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice. (******************************************************************************) (* The Finite interface describes Types with finitely many elements, *) (* supplying a duplicate-free sequence of all the elements. It is a subclass *) (* of Countable and thus of Choice and Equality. As with Countable, the *) (* interface explicitly includes these somewhat redundant superclasses to *) (* ensure that Canonical instance inference remains consistent. Finiteness *) (* could be stated more simply by bounding the range of the pickle function *) (* supplied by the Countable interface, but this would yield a useless *) (* computational interpretation due to the wasteful Peano integer encodings. *) (* Because the Countable interface is closely tied to the Finite interface *) (* and is not much used on its own, the Countable mixin is included inside *) (* the Finite mixin; this makes it much easier to derive Finite variants of *) (* interfaces, in this file for subFinType, and in the finalg library. *) (* We define the following interfaces and structures: *) (* finType == the packed class type of the Finite interface. *) (* FinType m == the packed class for the Finite mixin m. *) (* Finite.axiom e == every x : T occurs exactly once in e : seq T. *) (* FinMixin ax_e == the Finite mixin for T, encapsulating *) (* ax_e : Finite.axiom e for some e : seq T. *) (* UniqFinMixin uniq_e total_e == an alternative mixin constructor that uses *) (* uniq_e : uniq e and total_e : e =i xpredT. *) (* PcanFinMixin fK == the Finite mixin for T, given f : T -> fT and g with fT *) (* a finType and fK : pcancel f g. *) (* CanFinMixin fK == the Finite mixin for T, given f : T -> fT and g with fT *) (* a finType and fK : cancel f g. *) (* subFinType == the join interface type for subType and finType. *) (* [finType of T for fT] == clone for T of the finType fT. *) (* [finType of T] == clone for T of the finType inferred for T. *) (* [subFinType of T] == a subFinType structure for T, when T already has both *) (* finType and subType structures. *) (* [finMixin of T by <:] == a finType structure for T, when T has a subType *) (* structure over an existing finType. *) (* We define or propagate the finType structure appropriately for all basic *) (* basic types : unit, bool, option, prod, sum, sig and sigT. We also define *) (* a generic type constructor for finite subtypes based on an explicit *) (* enumeration: *) (* seq_sub s == the subType of all x \in s, where s : seq T and T has *) (* a choiceType structure; the seq_sub s type has a *) (* canonical finType structure. *) (* Bounded integers are supported by the following type and operations: *) (* 'I_n, ordinal n == the finite subType of integers i < n, whose *) (* enumeration is {0, ..., n.-1}. 'I_n coerces to nat, *) (* so all the integer arithmetic functions can be used *) (* with 'I_n. *) (* Ordinal lt_i_n == the element of 'I_n with (nat) value i, given *) (* lt_i_n : i < n. *) (* nat_of_ord i == the nat value of i : 'I_n (this function is a *) (* coercion so it is not usually displayed). *) (* ord_enum n == the explicit increasing sequence of the i : 'I_n. *) (* cast_ord eq_n_m i == the element j : 'I_m with the same value as i : 'I_n *) (* given eq_n_m : n = m (indeed, i : nat and j : nat *) (* are convertible). *) (* widen_ord le_n_m i == a j : 'I_m with the same value as i : 'I_n, given *) (* le_n_m : n <= m. *) (* rev_ord i == the complement to n.-1 of i : 'I_n, such that *) (* i + rev_ord i = n.-1. *) (* inord k == the i : 'I_n.+1 with value k (n is inferred from the *) (* context). *) (* sub_ord k == the i : 'I_n.+1 with value n - k (n is inferred from *) (* the context). *) (* ord0 == the i : 'I_n.+1 with value 0 (n is inferred from the *) (* context). *) (* ord_max == the i : 'I_n.+1 with value n (n is inferred from the *) (* context). *) (* bump h k == k.+1 if k >= h, else k (this is a nat function). *) (* unbump h k == k.-1 if k > h, else k (this is a nat function). *) (* lift i j == the j' : 'I_n with value bump i j, where i : 'I_n *) (* and j : 'I_n.-1. *) (* unlift i j == None if i = j, else Some j', where j' : 'I_n.-1 has *) (* value unbump i j, given i, j : 'I_n. *) (* lshift n j == the i : 'I_(m + n) with value j : 'I_m. *) (* rshift m k == the i : 'I_(m + n) with value m + k, k : 'I_n. *) (* unsplit u == either lshift n j or rshift m k, depending on *) (* whether if u : 'I_m + 'I_n is inl j or inr k. *) (* split i == the u : 'I_m + 'I_n such that i = unsplit u; the *) (* type 'I_(m + n) of i determines the split. *) (* Finally, every type T with a finType structure supports the following *) (* operations: *) (* enum A == a duplicate-free list of all the x \in A, where A is a *) (* collective predicate over T. *) (* #|A| == the cardinal of A, i.e., the number of x \in A. *) (* enum_val i == the i'th item of enum A, where i : 'I_(#|A|). *) (* enum_rank x == the i : 'I_(#|T|) such that enum_val i = x. *) (* enum_rank_in Ax0 x == some i : 'I_(#|A|) such that enum_val i = x if *) (* x \in A, given Ax0 : x0 \in A. *) (* A \subset B == all x \in A satisfy x \in B. *) (* A \proper B == all x \in A satisfy x \in B but not the converse. *) (* [disjoint A & B] == no x \in A satisfies x \in B. *) (* image f A == the sequence of f x for all x : T such that x \in A *) (* (where a is an applicative predicate), of length #|P|. *) (* The codomain of F can be any type, but image f A can *) (* only be used as a collective predicate is it is an *) (* eqType. *) (* codom f == a sequence spanning the codomain of f (:= image f T). *) (* [seq F | x : T in A] := image (fun x : T => F) A. *) (* [seq F | x : T] := [seq F | x <- {: T}]. *) (* [seq F | x in A], [seq F | x] == variants without casts. *) (* iinv im_y == some x such that P x holds and f x = y, given *) (* im_y : y \in image f P. *) (* invF inj_f y == the x such that f x = y, for inj_j : injective f with *) (* f : T -> T. *) (* dinjectiveb A f == the restriction of f : T -> R to A is injective *) (* (this is a bolean predicate, R must be an eqType). *) (* injectiveb f == f : T -> R is injective (boolean predicate). *) (* pred0b A == no x : T satisfies x \in A. *) (* [forall x, P] == P (in which x can appear) is true for all values of x; *) (* x must range over a finType. *) (* [exists x, P] == P is true for some value of x. *) (* [forall (x | C), P] := [forall x, C ==> P]. *) (* [forall x in A, P] := [forall (x | x \in A), P]. *) (* [exists (x | C), P] := [exists x, C && P]. *) (* [exists x in A, P] := [exists (x | x \in A), P]. *) (* and typed variants [forall x : T, P], [forall (x : T | C), P], *) (* [exists x : T, P], [exists x : T in A, P], etc. *) (* -> The outer brackets can be omitted when nesting finitary quantifiers, *) (* e.g., [forall i in I, forall j in J, exists a, f i j == a]. *) (* [pick x | P] == Some x, for an x such that P holds, or None if there *) (* is no such x. *) (* [pick x : T] == Some x with x : T, provided T is nonempty, else None. *) (* [pick x in A] == Some x, with x \in A, or None if A is empty. *) (* [pick x in A | P] == Some x, with x \in A s.t. P holds, else None. *) (* [pick x | P & Q] := [pick x | P & Q]. *) (* [pick x in A | P & Q] := [pick x | P & Q]. *) (* and (un)typed variants [pick x : T | P], [pick x : T in A], [pick x], etc. *) (* [arg min_(i < i0 | P) M] == a value of i : T minimizing M : nat, subject *) (* to the condition P (i may appear in P and M), and *) (* provided P holds for i0. *) (* [arg max_(i > i0 | P) M] == a value of i maximizing M subject to P and *) (* provided P holds for i0. *) (* [arg min_(i < i0 in A) M] == an i \in A minimizing M if i0 \in A. *) (* [arg max_(i > i0 in A) M] == an i \in A maximizing M if i0 \in A. *) (* [arg min_(i < i0) M] == an i : T minimizing M, given i0 : T. *) (* [arg max_(i > i0) M] == an i : T maximizing M, given i0 : T. *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Module Finite. Section RawMixin. Variable T : eqType. Definition axiom e := forall x, count (@pred1 T x) e = 1. Lemma uniq_enumP e : uniq e -> e =i T -> axiom e. Proof. by move=> Ue sT x; rewrite count_uniq_mem ?sT. Qed. Record mixin_of := Mixin { mixin_base : Countable.mixin_of T; mixin_enum : seq T; _ : axiom mixin_enum }. End RawMixin. Section Mixins. Variable T : countType. Definition EnumMixin := let: Countable.Pack _ (Countable.Class _ m) _ as cT := T return forall e : seq cT, axiom e -> mixin_of cT in @Mixin (EqType _ _) m. Definition UniqMixin e Ue eT := @EnumMixin e (uniq_enumP Ue eT). Variable n : nat. Definition count_enum := pmap (@pickle_inv T) (iota 0 n). Hypothesis ubT : forall x : T, pickle x < n. Lemma count_enumP : axiom count_enum. Proof. apply: uniq_enumP (pmap_uniq (@pickle_invK T) (iota_uniq _ _)) _ => x. by rewrite mem_pmap -pickleK_inv map_f // mem_iota ubT. Qed. Definition CountMixin := EnumMixin count_enumP. End Mixins. Section ClassDef. Record class_of T := Class { base : Choice.class_of T; mixin : mixin_of (Equality.Pack base T) }. Definition base2 T c := Countable.Class (@base T c) (mixin_base (mixin c)). Local Coercion base : class_of >-> Choice.class_of. Structure type : Type := Pack {sort; _ : class_of sort; _ : Type}. Local Coercion sort : type >-> Sortclass. Variables (T : Type) (cT : type). Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c. Definition clone c of phant_id class c := @Pack T c T. Let xT := let: Pack T _ _ := cT in T. Notation xclass := (class : class_of xT). Definition pack b0 (m0 : mixin_of (EqType T b0)) := fun bT b & phant_id (Choice.class bT) b => fun m & phant_id m0 m => Pack (@Class T b m) T. Definition eqType := @Equality.Pack cT xclass xT. Definition choiceType := @Choice.Pack cT xclass xT. Definition countType := @Countable.Pack cT (base2 xclass) xT. End ClassDef. Module Import Exports. Coercion mixin_base : mixin_of >-> Countable.mixin_of. Coercion base : class_of >-> Choice.class_of. Coercion mixin : class_of >-> mixin_of. Coercion base2 : class_of >-> Countable.class_of. Coercion sort : type >-> Sortclass. Coercion eqType : type >-> Equality.type. Canonical eqType. Coercion choiceType : type >-> Choice.type. Canonical choiceType. Coercion countType : type >-> Countable.type. Canonical countType. Notation finType := type. Notation FinType T m := (@pack T _ m _ _ id _ id). Notation FinMixin := EnumMixin. Notation UniqFinMixin := UniqMixin. Notation "[ 'finType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) (at level 0, format "[ 'finType' 'of' T 'for' cT ]") : form_scope. Notation "[ 'finType' 'of' T ]" := (@clone T _ _ id) (at level 0, format "[ 'finType' 'of' T ]") : form_scope. End Exports. Module Type EnumSig. Parameter enum : forall cT : type, seq cT. Axiom enumDef : enum = fun cT => mixin_enum (class cT). End EnumSig. Module EnumDef : EnumSig. Definition enum cT := mixin_enum (class cT). Definition enumDef := erefl enum. End EnumDef. Notation enum := EnumDef.enum. End Finite. Export Finite.Exports. Canonical finEnum_unlock := Unlockable Finite.EnumDef.enumDef. (* Workaround for the silly syntactic uniformity restriction on coercions; *) (* this avoids a cross-dependency between finset.v and prime.v for the *) (* definition of the \pi(A) notation. *) Definition fin_pred_sort (T : finType) (pT : predType T) := pred_sort pT. Identity Coercion pred_sort_of_fin : fin_pred_sort >-> pred_sort. Definition enum_mem T (mA : mem_pred _) := filter mA (Finite.enum T). Notation enum A := (enum_mem (mem A)). Definition pick (T : finType) (P : pred T) := ohead (enum P). Notation "[ 'pick' x | P ]" := (pick (fun x => P%B)) (at level 0, x ident, format "[ 'pick' x | P ]") : form_scope. Notation "[ 'pick' x : T | P ]" := (pick (fun x : T => P%B)) (at level 0, x ident, only parsing) : form_scope. Definition pick_true T (x : T) := true. Notation "[ 'pick' x : T ]" := [pick x : T | pick_true x] (at level 0, x ident, only parsing). Notation "[ 'pick' x ]" := [pick x : _] (at level 0, x ident, only parsing) : form_scope. Notation "[ 'pic' 'k' x : T ]" := [pick x : T | pick_true _] (at level 0, x ident, format "[ 'pic' 'k' x : T ]") : form_scope. Notation "[ 'pick' x | P & Q ]" := [pick x | P && Q ] (at level 0, x ident, format "[ '[hv ' 'pick' x | P '/ ' & Q ] ']'") : form_scope. Notation "[ 'pick' x : T | P & Q ]" := [pick x : T | P && Q ] (at level 0, x ident, only parsing) : form_scope. Notation "[ 'pick' x 'in' A ]" := [pick x | x \in A] (at level 0, x ident, format "[ 'pick' x 'in' A ]") : form_scope. Notation "[ 'pick' x : T 'in' A ]" := [pick x : T | x \in A] (at level 0, x ident, only parsing) : form_scope. Notation "[ 'pick' x 'in' A | P ]" := [pick x | x \in A & P ] (at level 0, x ident, format "[ '[hv ' 'pick' x 'in' A '/ ' | P ] ']'") : form_scope. Notation "[ 'pick' x : T 'in' A | P ]" := [pick x : T | x \in A & P ] (at level 0, x ident, only parsing) : form_scope. Notation "[ 'pick' x 'in' A | P & Q ]" := [pick x in A | P && Q] (at level 0, x ident, format "[ '[hv ' 'pick' x 'in' A '/ ' | P '/ ' & Q ] ']'") : form_scope. Notation "[ 'pick' x : T 'in' A | P & Q ]" := [pick x : T in A | P && Q] (at level 0, x ident, only parsing) : form_scope. (* We lock the definitions of card and subset to mitigate divergence of the *) (* Coq term comparison algorithm. *) Local Notation card_type := (forall T : finType, mem_pred T -> nat). Local Notation card_def := (fun T mA => size (enum_mem mA)). Module Type CardDefSig. Parameter card : card_type. Axiom cardEdef : card = card_def. End CardDefSig. Module CardDef : CardDefSig. Definition card : card_type := card_def. Definition cardEdef := erefl card. End CardDef. (* Should be Include, but for a silly restriction: can't Include at toplevel! *) Export CardDef. Canonical card_unlock := Unlockable cardEdef. (* A is at level 99 to allow the notation #|G : H| in groups. *) Notation "#| A |" := (card (mem A)) (at level 0, A at level 99, format "#| A |") : nat_scope. Definition pred0b (T : finType) (P : pred T) := #|P| == 0. Prenex Implicits pred0b. Module FiniteQuant. CoInductive quantified := Quantified of bool. Delimit Scope fin_quant_scope with Q. (* Bogus, only used to declare scope. *) Bind Scope fin_quant_scope with quantified. Notation "F ^*" := (Quantified F) (at level 2). Notation "F ^~" := (~~ F) (at level 2). Section Definitions. Variable T : finType. Implicit Types (B : quantified) (x y : T). Definition quant0b Bp := pred0b [pred x : T | let: F^* := Bp x x in F]. (* The first redundant argument protects the notation from Coq's K-term *) (* display kludge; the second protects it from simpl and /=. *) Definition ex B x y := B. (* Binding the predicate value rather than projecting it prevents spurious *) (* unfolding of the boolean connectives by unification. *) Definition all B x y := let: F^* := B in F^~^*. Definition all_in C B x y := let: F^* := B in (C ==> F)^~^*. Definition ex_in C B x y := let: F^* := B in (C && F)^*. End Definitions. Notation "[ x | B ]" := (quant0b (fun x => B x)) (at level 0, x ident). Notation "[ x : T | B ]" := (quant0b (fun x : T => B x)) (at level 0, x ident). Module Exports. Notation ", F" := F^* (at level 200, format ", '/ ' F") : fin_quant_scope. Notation "[ 'forall' x B ]" := [x | all B] (at level 0, x at level 99, B at level 200, format "[ '[hv' 'forall' x B ] ']'") : bool_scope. Notation "[ 'forall' x : T B ]" := [x : T | all B] (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. Notation "[ 'forall' ( x | C ) B ]" := [x | all_in C B] (at level 0, x at level 99, B at level 200, format "[ '[hv' '[' 'forall' ( x '/ ' | C ) ']' B ] ']'") : bool_scope. Notation "[ 'forall' ( x : T | C ) B ]" := [x : T | all_in C B] (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. Notation "[ 'forall' x 'in' A B ]" := [x | all_in (x \in A) B] (at level 0, x at level 99, B at level 200, format "[ '[hv' '[' 'forall' x '/ ' 'in' A ']' B ] ']'") : bool_scope. Notation "[ 'forall' x : T 'in' A B ]" := [x : T | all_in (x \in A) B] (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. Notation ", 'forall' x B" := [x | all B]^* (at level 200, x at level 99, B at level 200, format ", '/ ' 'forall' x B") : fin_quant_scope. Notation ", 'forall' x : T B" := [x : T | all B]^* (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. Notation ", 'forall' ( x | C ) B" := [x | all_in C B]^* (at level 200, x at level 99, B at level 200, format ", '/ ' '[' 'forall' ( x '/ ' | C ) ']' B") : fin_quant_scope. Notation ", 'forall' ( x : T | C ) B" := [x : T | all_in C B]^* (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. Notation ", 'forall' x 'in' A B" := [x | all_in (x \in A) B]^* (at level 200, x at level 99, B at level 200, format ", '/ ' '[' 'forall' x '/ ' 'in' A ']' B") : bool_scope. Notation ", 'forall' x : T 'in' A B" := [x : T | all_in (x \in A) B]^* (at level 200, x at level 99, B at level 200, only parsing) : bool_scope. Notation "[ 'exists' x B ]" := [x | ex B]^~ (at level 0, x at level 99, B at level 200, format "[ '[hv' 'exists' x B ] ']'") : bool_scope. Notation "[ 'exists' x : T B ]" := [x : T | ex B]^~ (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. Notation "[ 'exists' ( x | C ) B ]" := [x | ex_in C B]^~ (at level 0, x at level 99, B at level 200, format "[ '[hv' '[' 'exists' ( x '/ ' | C ) ']' B ] ']'") : bool_scope. Notation "[ 'exists' ( x : T | C ) B ]" := [x : T | ex_in C B]^~ (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. Notation "[ 'exists' x 'in' A B ]" := [x | ex_in (x \in A) B]^~ (at level 0, x at level 99, B at level 200, format "[ '[hv' '[' 'exists' x '/ ' 'in' A ']' B ] ']'") : bool_scope. Notation "[ 'exists' x : T 'in' A B ]" := [x : T | ex_in (x \in A) B]^~ (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. Notation ", 'exists' x B" := [x | ex B]^~^* (at level 200, x at level 99, B at level 200, format ", '/ ' 'exists' x B") : fin_quant_scope. Notation ", 'exists' x : T B" := [x : T | ex B]^~^* (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. Notation ", 'exists' ( x | C ) B" := [x | ex_in C B]^~^* (at level 200, x at level 99, B at level 200, format ", '/ ' '[' 'exists' ( x '/ ' | C ) ']' B") : fin_quant_scope. Notation ", 'exists' ( x : T | C ) B" := [x : T | ex_in C B]^~^* (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. Notation ", 'exists' x 'in' A B" := [x | ex_in (x \in A) B]^~^* (at level 200, x at level 99, B at level 200, format ", '/ ' '[' 'exists' x '/ ' 'in' A ']' B") : bool_scope. Notation ", 'exists' x : T 'in' A B" := [x : T | ex_in (x \in A) B]^~^* (at level 200, x at level 99, B at level 200, only parsing) : bool_scope. End Exports. End FiniteQuant. Export FiniteQuant.Exports. Definition disjoint T (A B : mem_pred _) := @pred0b T (predI A B). Notation "[ 'disjoint' A & B ]" := (disjoint (mem A) (mem B)) (at level 0, format "'[hv' [ 'disjoint' '/ ' A '/' & B ] ']'") : bool_scope. Notation Local subset_type := (forall (T : finType) (A B : mem_pred T), bool). Notation Local subset_def := (fun T A B => pred0b (predD A B)). Module Type SubsetDefSig. Parameter subset : subset_type. Axiom subsetEdef : subset = subset_def. End SubsetDefSig. Module Export SubsetDef : SubsetDefSig. Definition subset : subset_type := subset_def. Definition subsetEdef := erefl subset. End SubsetDef. Canonical subset_unlock := Unlockable subsetEdef. Notation "A \subset B" := (subset (mem A) (mem B)) (at level 70, no associativity) : bool_scope. Definition proper T A B := @subset T A B && ~~ subset B A. Notation "A \proper B" := (proper (mem A) (mem B)) (at level 70, no associativity) : bool_scope. (* image, xinv, inv, and ordinal operations will be defined later. *) Section OpsTheory. Variable T : finType. Implicit Types A B C P Q : pred T. Implicit Types x y : T. Implicit Type s : seq T. Lemma enumP : Finite.axiom (Finite.enum T). Proof. by rewrite unlock; case T => ? [? []]. Qed. Section EnumPick. Variable P : pred T. Lemma enumT : enum T = Finite.enum T. Proof. exact: filter_predT. Qed. Lemma mem_enum A : enum A =i A. Proof. by move=> x; rewrite mem_filter andbC -has_pred1 has_count enumP. Qed. Lemma enum_uniq : uniq (enum P). Proof. by apply/filter_uniq/count_mem_uniq => x; rewrite enumP -enumT mem_enum. Qed. Lemma enum0 : enum pred0 = Nil T. Proof. exact: filter_pred0. Qed. Lemma enum1 x : enum (pred1 x) = [:: x]. Proof. rewrite [enum _](all_pred1P x _ _); first by rewrite -count_filter enumP. by apply/allP=> y; rewrite mem_enum. Qed. CoInductive pick_spec : option T -> Type := | Pick x of P x : pick_spec (Some x) | Nopick of P =1 xpred0 : pick_spec None. Lemma pickP : pick_spec (pick P). Proof. rewrite /pick; case: (enum _) (mem_enum P) => [|x s] Pxs /=. by right; exact: fsym. by left; rewrite -[P _]Pxs mem_head. Qed. End EnumPick. Lemma eq_enum P Q : P =i Q -> enum P = enum Q. Proof. move=> eqPQ; exact: eq_filter. Qed. Lemma eq_pick P Q : P =1 Q -> pick P = pick Q. Proof. by move=> eqPQ; rewrite /pick (eq_enum eqPQ). Qed. Lemma cardE A : #|A| = size (enum A). Proof. by rewrite unlock. Qed. Lemma eq_card A B : A =i B -> #|A| = #|B|. Proof. by move=>eqAB; rewrite !cardE (eq_enum eqAB). Qed. Lemma eq_card_trans A B n : #|A| = n -> B =i A -> #|B| = n. Proof. move <-; exact: eq_card. Qed. Lemma card0 : #|@pred0 T| = 0. Proof. by rewrite cardE enum0. Qed. Lemma cardT : #|T| = size (enum T). Proof. by rewrite cardE. Qed. Lemma card1 x : #|pred1 x| = 1. Proof. by rewrite cardE enum1. Qed. Lemma eq_card0 A : A =i pred0 -> #|A| = 0. Proof. exact: eq_card_trans card0. Qed. Lemma eq_cardT A : A =i predT -> #|A| = size (enum T). Proof. exact: eq_card_trans cardT. Qed. Lemma eq_card1 x A : A =i pred1 x -> #|A| = 1. Proof. exact: eq_card_trans (card1 x). Qed. Lemma cardUI A B : #|[predU A & B]| + #|[predI A & B]| = #|A| + #|B|. Proof. by rewrite !cardE -!count_filter count_predUI. Qed. Lemma cardID B A : #|[predI A & B]| + #|[predD A & B]| = #|A|. Proof. rewrite -cardUI addnC [#|predI _ _|]eq_card0 => [|x] /=. by apply: eq_card => x; rewrite !inE andbC -andb_orl orbN. by rewrite !inE -!andbA andbC andbA andbN. Qed. Lemma cardC A : #|A| + #|[predC A]| = #|T|. Proof. by rewrite !cardE -!count_filter count_predC. Qed. Lemma cardU1 x A : #|[predU1 x & A]| = (x \notin A) + #|A|. Proof. case Ax: (x \in A). by apply: eq_card => y; rewrite inE /=; case: eqP => // ->. rewrite /= -(card1 x) -cardUI addnC. rewrite [#|predI _ _|]eq_card0 => [|y /=]; first exact: eq_card. by rewrite !inE; case: eqP => // ->. Qed. Lemma card2 x y : #|pred2 x y| = (x != y).+1. Proof. by rewrite cardU1 card1 addn1. Qed. Lemma cardC1 x : #|predC1 x| = #|T|.-1. Proof. by rewrite -(cardC (pred1 x)) card1. Qed. Lemma cardD1 x A : #|A| = (x \in A) + #|[predD1 A & x]|. Proof. case Ax: (x \in A); last first. by apply: eq_card => y; rewrite !inE /=; case: eqP => // ->. rewrite /= -(card1 x) -cardUI addnC /=. rewrite [#|predI _ _|]eq_card0 => [|y]; last by rewrite !inE; case: eqP. by apply: eq_card => y; rewrite !inE; case: eqP => // ->. Qed. Lemma max_card A : #|A| <= #|T|. Proof. by rewrite -(cardC A) leq_addr. Qed. Lemma card_size s : #|s| <= size s. Proof. elim: s => [|x s IHs] /=; first by rewrite card0. rewrite cardU1 /=; case: (~~ _) => //; exact: leqW. Qed. Lemma card_uniqP s : reflect (#|s| = size s) (uniq s). Proof. elim: s => [|x s IHs]; first by left; exact card0. rewrite cardU1 /= /addn; case: {+}(x \in s) => /=. by right=> card_Ssz; have:= card_size s; rewrite card_Ssz ltnn. by apply: (iffP IHs) => [<-| [<-]]. Qed. Lemma card0_eq A : #|A| = 0 -> A =i pred0. Proof. by move=> A0 x; apply/idP => Ax; rewrite (cardD1 x) Ax in A0. Qed. Lemma pred0P P : reflect (P =1 pred0) (pred0b P). Proof. apply: (iffP eqP); [exact: card0_eq | exact: eq_card0]. Qed. Lemma pred0Pn P : reflect (exists x, P x) (~~ pred0b P). Proof. case: (pickP P) => [x Px | P0]. by rewrite (introN (pred0P P)) => [|P0]; [left; exists x | rewrite P0 in Px]. by rewrite -lt0n eq_card0 //; right=> [[x]]; rewrite P0. Qed. Lemma card_gt0P A : reflect (exists i, i \in A) (#|A| > 0). Proof. rewrite lt0n; exact: pred0Pn. Qed. Lemma subsetE A B : (A \subset B) = pred0b [predD A & B]. Proof. by rewrite unlock. Qed. Lemma subsetP A B : reflect {subset A <= B} (A \subset B). Proof. rewrite unlock; apply: (iffP (pred0P _)) => [AB0 x | sAB x /=]. by apply/implyP; apply/idPn; rewrite negb_imply andbC [_ && _]AB0. by rewrite andbC -negb_imply; apply/negbF/implyP; exact: sAB. Qed. Lemma subsetPn A B : reflect (exists2 x, x \in A & x \notin B) (~~ (A \subset B)). Proof. rewrite unlock; apply: (iffP (pred0Pn _)) => [[x] | [x Ax nBx]]. by case/andP; exists x. by exists x; rewrite /= nBx. Qed. Lemma subset_leq_card A B : A \subset B -> #|A| <= #|B|. Proof. move=> sAB. rewrite -(cardID A B) [#|predI _ _|](@eq_card _ A) ?leq_addr //= => x. rewrite !inE andbC; case Ax: (x \in A) => //; exact: subsetP Ax. Qed. Lemma subxx_hint (mA : mem_pred T) : subset mA mA. Proof. by case: mA => A; have:= introT (subsetP A A); rewrite !unlock => ->. Qed. Hint Resolve subxx_hint. (* The parametrization by predType makes it easier to apply subxx. *) Lemma subxx (pT : predType T) (pA : pT) : pA \subset pA. Proof. by []. Qed. Lemma eq_subset A1 A2 : A1 =i A2 -> subset (mem A1) =1 subset (mem A2). Proof. move=> eqA12 [B]; rewrite !unlock; congr (_ == 0). by apply: eq_card => x; rewrite inE /= eqA12. Qed. Lemma eq_subset_r B1 B2 : B1 =i B2 -> (@subset T)^~ (mem B1) =1 (@subset T)^~ (mem B2). Proof. move=> eqB12 [A]; rewrite !unlock; congr (_ == 0). by apply: eq_card => x; rewrite !inE /= eqB12. Qed. Lemma eq_subxx A B : A =i B -> A \subset B. Proof. by move/eq_subset->. Qed. Lemma subset_predT A : A \subset T. Proof. by apply/subsetP. Qed. Lemma predT_subset A : T \subset A -> forall x, x \in A. Proof. move/subsetP=> allA x; exact: allA. Qed. Lemma subset_pred1 A x : (pred1 x \subset A) = (x \in A). Proof. by apply/subsetP/idP=> [-> // | Ax y /eqP-> //]; exact: eqxx. Qed. Lemma subset_eqP A B : reflect (A =i B) ((A \subset B) && (B \subset A)). Proof. apply: (iffP andP) => [[sAB sBA] x| eqAB]; last by rewrite !eq_subxx. by apply/idP/idP; apply: subsetP. Qed. Lemma subset_cardP A B : #|A| = #|B| -> reflect (A =i B) (A \subset B). Proof. move=> eqcAB; case: (subsetP A B) (subset_eqP A B) => //= sAB. case: (subsetP B A) => [//|[]] x Bx; apply/idPn => Ax. case/idP: (ltnn #|A|); rewrite {2}eqcAB (cardD1 x B) Bx /=. apply: subset_leq_card; apply/subsetP=> y Ay; rewrite inE /= andbC. by rewrite sAB //; apply/eqP => eqyx; rewrite -eqyx Ay in Ax. Qed. Lemma subset_leqif_card A B : A \subset B -> #|A| <= #|B| ?= iff (B \subset A). Proof. move=> sAB; split; [exact: subset_leq_card | apply/eqP/idP]. by move/subset_cardP=> sABP; rewrite (eq_subset_r (sABP sAB)). by move=> sBA; apply: eq_card; apply/subset_eqP; rewrite sAB. Qed. Lemma subset_trans A B C : A \subset B -> B \subset C -> A \subset C. Proof. by move/subsetP=> sAB /subsetP=> sBC; apply/subsetP=> x /sAB; exact: sBC. Qed. Lemma subset_all s A : (s \subset A) = all (mem A) s. Proof. by exact (sameP (subsetP _ _) allP). Qed. Lemma properE A B : A \proper B = (A \subset B) && ~~(B \subset A). Proof. by []. Qed. Lemma properP A B : reflect (A \subset B /\ (exists2 x, x \in B & x \notin A)) (A \proper B). Proof. by rewrite properE; apply: (iffP andP) => [] [-> /subsetPn]. Qed. Lemma proper_sub A B : A \proper B -> A \subset B. Proof. by case/andP. Qed. Lemma proper_subn A B : A \proper B -> ~~ (B \subset A). Proof. by case/andP. Qed. Lemma proper_trans A B C : A \proper B -> B \proper C -> A \proper C. Proof. case/properP=> sAB [x Bx nAx] /properP[sBC [y Cy nBy]]. rewrite properE (subset_trans sAB) //=; apply/subsetPn; exists y => //. by apply: contra nBy; exact: subsetP. Qed. Lemma proper_sub_trans A B C : A \proper B -> B \subset C -> A \proper C. Proof. case/properP=> sAB [x Bx nAx] sBC; rewrite properE (subset_trans sAB) //. by apply/subsetPn; exists x; rewrite ?(subsetP _ _ sBC). Qed. Lemma sub_proper_trans A B C : A \subset B -> B \proper C -> A \proper C. Proof. move=> sAB /properP[sBC [x Cx nBx]]; rewrite properE (subset_trans sAB) //. by apply/subsetPn; exists x => //; apply: contra nBx; exact: subsetP. Qed. Lemma proper_card A B : A \proper B -> #|A| < #|B|. Proof. by case/andP=> sAB nsBA; rewrite ltn_neqAle !(subset_leqif_card sAB) andbT. Qed. Lemma proper_irrefl A : ~~ (A \proper A). Proof. by rewrite properE subxx. Qed. Lemma properxx A : (A \proper A) = false. Proof. by rewrite properE subxx. Qed. Lemma eq_proper A B : A =i B -> proper (mem A) =1 proper (mem B). Proof. move=> eAB [C]; congr (_ && _); first exact: (eq_subset eAB). by rewrite (eq_subset_r eAB). Qed. Lemma eq_proper_r A B : A =i B -> (@proper T)^~ (mem A) =1 (@proper T)^~ (mem B). Proof. move=> eAB [C]; congr (_ && _); first exact: (eq_subset_r eAB). by rewrite (eq_subset eAB). Qed. Lemma disjoint_sym A B : [disjoint A & B] = [disjoint B & A]. Proof. by congr (_ == 0); apply: eq_card => x; exact: andbC. Qed. Lemma eq_disjoint A1 A2 : A1 =i A2 -> disjoint (mem A1) =1 disjoint (mem A2). Proof. by move=> eqA12 [B]; congr (_ == 0); apply: eq_card => x; rewrite !inE eqA12. Qed. Lemma eq_disjoint_r B1 B2 : B1 =i B2 -> (@disjoint T)^~ (mem B1) =1 (@disjoint T)^~ (mem B2). Proof. by move=> eqB12 [A]; congr (_ == 0); apply: eq_card => x; rewrite !inE eqB12. Qed. Lemma subset_disjoint A B : (A \subset B) = [disjoint A & [predC B]]. Proof. by rewrite disjoint_sym unlock. Qed. Lemma disjoint_subset A B : [disjoint A & B] = (A \subset [predC B]). Proof. by rewrite subset_disjoint; apply: eq_disjoint_r => x; rewrite !inE /= negbK. Qed. Lemma disjoint_trans A B C : A \subset B -> [disjoint B & C] -> [disjoint A & C]. Proof. by rewrite 2!disjoint_subset; exact: subset_trans. Qed. Lemma disjoint0 A : [disjoint pred0 & A]. Proof. exact/pred0P. Qed. Lemma eq_disjoint0 A B : A =i pred0 -> [disjoint A & B]. Proof. by move/eq_disjoint->; exact: disjoint0. Qed. Lemma disjoint1 x A : [disjoint pred1 x & A] = (x \notin A). Proof. apply/negbRL/(sameP (pred0Pn _)). apply: introP => [Ax | notAx [_ /andP[/eqP->]]]; last exact: negP. by exists x; rewrite !inE eqxx. Qed. Lemma eq_disjoint1 x A B : A =i pred1 x -> [disjoint A & B] = (x \notin B). Proof. by move/eq_disjoint->; exact: disjoint1. Qed. Lemma disjointU A B C : [disjoint predU A B & C] = [disjoint A & C] && [disjoint B & C]. Proof. case: [disjoint A & C] / (pred0P (xpredI A C)) => [A0 | nA0] /=. by congr (_ == 0); apply: eq_card => x; rewrite [x \in _]andb_orl A0. apply/pred0P=> nABC; case: nA0 => x; apply/idPn=> /=; move/(_ x): nABC. by rewrite [_ x]andb_orl; case/norP. Qed. Lemma disjointU1 x A B : [disjoint predU1 x A & B] = (x \notin B) && [disjoint A & B]. Proof. by rewrite disjointU disjoint1. Qed. Lemma disjoint_cons x s B : [disjoint x :: s & B] = (x \notin B) && [disjoint s & B]. Proof. exact: disjointU1. Qed. Lemma disjoint_has s A : [disjoint s & A] = ~~ has (mem A) s. Proof. rewrite -(@eq_has _ (mem (enum A))) => [|x]; last exact: mem_enum. rewrite has_sym has_filter -filter_predI -has_filter has_count -eqn0Ngt. by rewrite count_filter /disjoint /pred0b unlock. Qed. Lemma disjoint_cat s1 s2 A : [disjoint s1 ++ s2 & A] = [disjoint s1 & A] && [disjoint s2 & A]. Proof. by rewrite !disjoint_has has_cat negb_or. Qed. End OpsTheory. Hint Resolve subxx_hint. Implicit Arguments pred0P [T P]. Implicit Arguments pred0Pn [T P]. Implicit Arguments subsetP [T A B]. Implicit Arguments subsetPn [T A B]. Implicit Arguments subset_eqP [T A B]. Implicit Arguments card_uniqP [T s]. Implicit Arguments properP [T A B]. Prenex Implicits pred0P pred0Pn subsetP subsetPn subset_eqP card_uniqP. (**********************************************************************) (* *) (* Boolean quantifiers for finType *) (* *) (**********************************************************************) Section Quantifiers. Variables (T : finType) (rT : T -> eqType). Implicit Type (D P : pred T) (f : forall x, rT x). Lemma forallP P : reflect (forall x, P x) [forall x, P x]. Proof. by apply: (iffP pred0P) => /= P_ x; rewrite /= ?P_ ?(negbFE (P_ x)). Qed. Lemma eqfunP f1 f2 : reflect (forall x, f1 x = f2 x) [forall x, f1 x == f2 x]. Proof. by apply: (iffP (forallP _)) => eq_f12 x; apply/eqP/eq_f12. Qed. Lemma forall_inP D P : reflect (forall x, D x -> P x) [forall (x | D x), P x]. Proof. by apply: (iffP (forallP _)) => /= P_ x /=; apply/implyP; apply: P_. Qed. Lemma eqfun_inP D f1 f2 : reflect {in D, forall x, f1 x = f2 x} [forall (x | x \in D), f1 x == f2 x]. Proof. by apply: (iffP (forall_inP _ _)) => eq_f12 x Dx; apply/eqP/eq_f12. Qed. Lemma existsP P : reflect (exists x, P x) [exists x, P x]. Proof. by apply: (iffP pred0Pn) => [] [x]; exists x. Qed. Lemma exists_eqP f1 f2 : reflect (exists x, f1 x = f2 x) [exists x, f1 x == f2 x]. Proof. by apply: (iffP (existsP _)) => [] [x /eqP]; exists x. Qed. Lemma exists_inP D P : reflect (exists2 x, D x & P x) [exists (x | D x), P x]. Proof. by apply: (iffP pred0Pn) => [[x /andP[]] | [x]]; exists x => //; apply/andP. Qed. Lemma exists_eq_inP D f1 f2 : reflect (exists2 x, D x & f1 x = f2 x) [exists (x | D x), f1 x == f2 x]. Proof. by apply: (iffP (exists_inP _ _)) => [] [x Dx /eqP]; exists x. Qed. Lemma eq_existsb P1 P2 : P1 =1 P2 -> [exists x, P1 x] = [exists x, P2 x]. Proof. by move=> eqP12; congr (_ != 0); apply: eq_card. Qed. Lemma eq_existsb_in D P1 P2 : (forall x, D x -> P1 x = P2 x) -> [exists (x | D x), P1 x] = [exists (x | D x), P2 x]. Proof. by move=> eqP12; apply: eq_existsb => x; apply: andb_id2l => /eqP12. Qed. Lemma eq_forallb P1 P2 : P1 =1 P2 -> [forall x, P1 x] = [forall x, P2 x]. Proof. by move=> eqP12; apply/negb_inj/eq_existsb=> /= x; rewrite eqP12. Qed. Lemma eq_forallb_in D P1 P2 : (forall x, D x -> P1 x = P2 x) -> [forall (x | D x), P1 x] = [forall (x | D x), P2 x]. Proof. by move=> eqP12; apply: eq_forallb => i; case Di: (D i); rewrite // eqP12. Qed. Lemma negb_forall P : ~~ [forall x, P x] = [exists x, ~~ P x]. Proof. by []. Qed. Lemma negb_forall_in D P : ~~ [forall (x | D x), P x] = [exists (x | D x), ~~ P x]. Proof. by apply: eq_existsb => x; rewrite negb_imply. Qed. Lemma negb_exists P : ~~ [exists x, P x] = [forall x, ~~ P x]. Proof. by apply/negbLR/esym/eq_existsb=> x; apply: negbK. Qed. Lemma negb_exists_in D P : ~~ [exists (x | D x), P x] = [forall (x | D x), ~~ P x]. Proof. by rewrite negb_exists; apply/eq_forallb => x; rewrite [~~ _]fun_if. Qed. End Quantifiers. Implicit Arguments forallP [T P]. Implicit Arguments eqfunP [T rT f1 f2]. Implicit Arguments forall_inP [T D P]. Implicit Arguments eqfun_inP [T rT D f1 f2]. Implicit Arguments existsP [T P]. Implicit Arguments exists_inP [T D P]. Section Extrema. Variables (I : finType) (i0 : I) (P : pred I) (F : I -> nat). Let arg_pred ord := [pred i | P i & [forall (j | P j), ord (F i) (F j)]]. Definition arg_min := odflt i0 (pick (arg_pred leq)). Definition arg_max := odflt i0 (pick (arg_pred geq)). CoInductive extremum_spec (ord : rel nat) : I -> Type := ExtremumSpec i of P i & (forall j, P j -> ord (F i) (F j)) : extremum_spec ord i. Hypothesis Pi0 : P i0. Let FP n := [exists (i | P i), F i == n]. Let FP_F i : P i -> FP (F i). Proof. by move=> Pi; apply/existsP; exists i; rewrite Pi /=. Qed. Let exFP : exists n, FP n. Proof. by exists (F i0); exact: FP_F. Qed. Lemma arg_minP : extremum_spec leq arg_min. Proof. rewrite /arg_min; case: pickP => [i /andP[Pi /forallP/= min_i] | no_i]. by split=> // j; apply/implyP. case/ex_minnP: exFP => n ex_i min_i; case/pred0P: ex_i => i /=. apply: contraFF (no_i i) => /andP[Pi /eqP def_n]; rewrite /= Pi. by apply/forall_inP=> j Pj; rewrite def_n min_i ?FP_F. Qed. Lemma arg_maxP : extremum_spec geq arg_max. Proof. rewrite /arg_max; case: pickP => [i /andP[Pi /forall_inP/= max_i] | no_i]. by split=> // j; apply/implyP. have (n): FP n -> n <= foldr maxn 0 (map F (enum P)). case/existsP=> i; rewrite -[P i]mem_enum andbC /= => /andP[/eqP <-]. elim: (enum P) => //= j e IHe; rewrite leq_max orbC !inE. by case/predU1P=> [-> | /IHe-> //]; rewrite leqnn orbT. case/ex_maxnP=> // n ex_i max_i; case/pred0P: ex_i => i /=. apply: contraFF (no_i i) => /andP[Pi def_n]; rewrite /= Pi. by apply/forall_inP=> j Pj; rewrite (eqP def_n) max_i ?FP_F. Qed. End Extrema. Notation "[ 'arg' 'min_' ( i < i0 | P ) F ]" := (arg_min i0 (fun i => P%B) (fun i => F)) (at level 0, i, i0 at level 10, format "[ 'arg' 'min_' ( i < i0 | P ) F ]") : form_scope. Notation "[ 'arg' 'min_' ( i < i0 'in' A ) F ]" := [arg min_(i < i0 | i \in A) F] (at level 0, i, i0 at level 10, format "[ 'arg' 'min_' ( i < i0 'in' A ) F ]") : form_scope. Notation "[ 'arg' 'min_' ( i < i0 ) F ]" := [arg min_(i < i0 | true) F] (at level 0, i, i0 at level 10, format "[ 'arg' 'min_' ( i < i0 ) F ]") : form_scope. Notation "[ 'arg' 'max_' ( i > i0 | P ) F ]" := (arg_max i0 (fun i => P%B) (fun i => F)) (at level 0, i, i0 at level 10, format "[ 'arg' 'max_' ( i > i0 | P ) F ]") : form_scope. Notation "[ 'arg' 'max_' ( i > i0 'in' A ) F ]" := [arg max_(i > i0 | i \in A) F] (at level 0, i, i0 at level 10, format "[ 'arg' 'max_' ( i > i0 'in' A ) F ]") : form_scope. Notation "[ 'arg' 'max_' ( i > i0 ) F ]" := [arg max_(i > i0 | true) F] (at level 0, i, i0 at level 10, format "[ 'arg' 'max_' ( i > i0 ) F ]") : form_scope. (**********************************************************************) (* *) (* Boolean injectivity test for functions with a finType domain *) (* *) (**********************************************************************) Section Injectiveb. Variables (aT : finType) (rT : eqType) (f : aT -> rT). Implicit Type D : pred aT. Definition dinjectiveb D := uniq (map f (enum D)). Definition injectiveb := dinjectiveb aT. Lemma dinjectivePn D : reflect (exists2 x, x \in D & exists2 y, y \in [predD1 D & x] & f x = f y) (~~ dinjectiveb D). Proof. apply: (iffP idP) => [injf | [x Dx [y Dxy eqfxy]]]; last first. move: Dx; rewrite -(mem_enum D) => /rot_to[i E defE]. rewrite /dinjectiveb -(rot_uniq i) -map_rot defE /=; apply/nandP; left. rewrite inE /= -(mem_enum D) -(mem_rot i) defE inE in Dxy. rewrite andb_orr andbC andbN in Dxy. by rewrite eqfxy map_f //; case/andP: Dxy. pose p := [pred x in D | [exists (y | y \in [predD1 D & x]), f x == f y]]. case: (pickP p) => [x /= /andP[Dx /exists_inP[y Dxy /eqP eqfxy]] | no_p]. by exists x; last exists y. rewrite /dinjectiveb map_inj_in_uniq ?enum_uniq // in injf => x y Dx Dy eqfxy. apply: contraNeq (negbT (no_p x)) => ne_xy /=; rewrite -mem_enum Dx. by apply/existsP; exists y; rewrite /= !inE eq_sym ne_xy -mem_enum Dy eqfxy /=. Qed. Lemma dinjectiveP D : reflect {in D &, injective f} (dinjectiveb D). Proof. rewrite -[dinjectiveb D]negbK. case: dinjectivePn=> [noinjf | injf]; constructor. case: noinjf => x Dx [y /andP[neqxy /= Dy] eqfxy] injf. by case/eqP: neqxy; exact: injf. move=> x y Dx Dy /= eqfxy; apply/eqP; apply/idPn=> nxy; case: injf. by exists x => //; exists y => //=; rewrite inE /= eq_sym nxy. Qed. Lemma injectivePn : reflect (exists x, exists2 y, x != y & f x = f y) (~~ injectiveb). Proof. apply: (iffP (dinjectivePn _)) => [[x _ [y nxy eqfxy]] | [x [y nxy eqfxy]]]; by exists x => //; exists y => //; rewrite inE /= andbT eq_sym in nxy *. Qed. Lemma injectiveP : reflect (injective f) injectiveb. Proof. apply: (iffP (dinjectiveP _)) => injf x y => [|_ _]; exact: injf. Qed. End Injectiveb. Definition image_mem T T' f mA : seq T' := map f (@enum_mem T mA). Notation image f A := (image_mem f (mem A)). Notation "[ 'seq' F | x 'in' A ]" := (image (fun x => F) A) (at level 0, F at level 99, x ident, format "'[hv' [ 'seq' F '/ ' | x 'in' A ] ']'") : seq_scope. Notation "[ 'seq' F | x : T 'in' A ]" := (image (fun x : T => F) A) (at level 0, F at level 99, x ident, only parsing) : seq_scope. Notation "[ 'seq' F | x : T ]" := [seq F | x : T in sort_of_simpl_pred (@pred_of_argType T)] (at level 0, F at level 99, x ident, format "'[hv' [ 'seq' F '/ ' | x : T ] ']'") : seq_scope. Notation "[ 'seq' F , x ]" := [seq F | x : _ ] (at level 0, F at level 99, x ident, only parsing) : seq_scope. Definition codom T T' f := @image_mem T T' f (mem T). Section Image. Variable T : finType. Implicit Type A : pred T. Section SizeImage. Variables (T' : Type) (f : T -> T'). Lemma size_image A : size (image f A) = #|A|. Proof. by rewrite size_map -cardE. Qed. Lemma size_codom : size (codom f) = #|T|. Proof. exact: size_image. Qed. Lemma codomE : codom f = map f (enum T). Proof. by []. Qed. End SizeImage. Variables (T' : eqType) (f : T -> T'). Lemma imageP A y : reflect (exists2 x, x \in A & y = f x) (y \in image f A). Proof. by apply: (iffP mapP) => [] [x Ax y_fx]; exists x; rewrite // mem_enum in Ax *. Qed. Lemma codomP y : reflect (exists x, y = f x) (y \in codom f). Proof. by apply: (iffP (imageP _ y)) => [][x]; exists x. Qed. Remark iinv_proof A y : y \in image f A -> {x | x \in A & f x = y}. Proof. move=> fy; pose b x := A x && (f x == y). case: (pickP b) => [x /andP[Ax /eqP] | nfy]; first by exists x. by case/negP: fy => /imageP[x Ax fx_y]; case/andP: (nfy x); rewrite fx_y. Qed. Definition iinv A y fAy := s2val (@iinv_proof A y fAy). Lemma f_iinv A y fAy : f (@iinv A y fAy) = y. Proof. exact: s2valP' (iinv_proof fAy). Qed. Lemma mem_iinv A y fAy : @iinv A y fAy \in A. Proof. exact: s2valP (iinv_proof fAy). Qed. Lemma in_iinv_f A : {in A &, injective f} -> forall x fAfx, x \in A -> @iinv A (f x) fAfx = x. Proof. move=> injf x fAfx Ax; apply: injf => //; [exact: mem_iinv | exact: f_iinv]. Qed. Lemma preim_iinv A B y fAy : preim f B (@iinv A y fAy) = B y. Proof. by rewrite /= f_iinv. Qed. Lemma image_f A x : x \in A -> f x \in image f A. Proof. by move=> Ax; apply/imageP; exists x. Qed. Lemma codom_f x : f x \in codom f. Proof. by exact: image_f. Qed. Lemma image_codom A : {subset image f A <= codom f}. Proof. by move=> _ /imageP[x _ ->]; exact: codom_f. Qed. Lemma image_pred0 : image f pred0 =i pred0. Proof. by move=> x; rewrite /image_mem /= enum0. Qed. Section Injective. Hypothesis injf : injective f. Lemma mem_image A x : (f x \in image f A) = (x \in A). Proof. by rewrite mem_map ?mem_enum. Qed. Lemma pre_image A : [preim f of image f A] =i A. Proof. by move=> x; rewrite inE /= mem_image. Qed. Lemma image_iinv A y (fTy : y \in codom f) : (y \in image f A) = (iinv fTy \in A). Proof. by rewrite -mem_image ?f_iinv. Qed. Lemma iinv_f x fTfx : @iinv T (f x) fTfx = x. Proof. by apply: in_iinv_f; first exact: in2W. Qed. Lemma image_pre (B : pred T') : image f [preim f of B] =i [predI B & codom f]. Proof. by move=> y; rewrite /image_mem -filter_map /= mem_filter -enumT. Qed. Lemma bij_on_codom (x0 : T) : {on [pred y in codom f], bijective f}. Proof. pose g y := iinv (valP (insigd (codom_f x0) y)). by exists g => [x fAfx | y fAy]; first apply: injf; rewrite f_iinv insubdK. Qed. Lemma bij_on_image A (x0 : T) : {on [pred y in image f A], bijective f}. Proof. exact: subon_bij (@image_codom A) (bij_on_codom x0). Qed. End Injective. Fixpoint preim_seq s := if s is y :: s' then (if pick (preim f (pred1 y)) is Some x then cons x else id) (preim_seq s') else [::]. Lemma map_preim (s : seq T') : {subset s <= codom f} -> map f (preim_seq s) = s. Proof. elim: s => //= y s IHs; case: pickP => [x /eqP fx_y | nfTy] fTs. by rewrite /= fx_y IHs // => z s_z; apply: fTs; exact: predU1r. by case/imageP: (fTs y (mem_head y s)) => x _ fx_y; case/eqP: (nfTy x). Qed. End Image. Prenex Implicits codom iinv. Implicit Arguments imageP [T T' f A y]. Implicit Arguments codomP [T T' f y]. Lemma flatten_imageP (aT : finType) (rT : eqType) A (P : pred aT) (y : rT) : reflect (exists2 x, x \in P & y \in A x) (y \in flatten [seq A x | x in P]). Proof. by apply: (iffP flatten_mapP) => [][x Px]; exists x; rewrite ?mem_enum in Px *. Qed. Implicit Arguments flatten_imageP [aT rT A P y]. Section CardFunImage. Variables (T T' : finType) (f : T -> T'). Implicit Type A : pred T. Lemma leq_image_card A : #|image f A| <= #|A|. Proof. by rewrite (cardE A) -(size_map f) card_size. Qed. Lemma card_in_image A : {in A &, injective f} -> #|image f A| = #|A|. Proof. move=> injf; rewrite (cardE A) -(size_map f); apply/card_uniqP. rewrite map_inj_in_uniq ?enum_uniq // => x y; rewrite !mem_enum; exact: injf. Qed. Lemma image_injP A : reflect {in A &, injective f} (#|image f A| == #|A|). Proof. apply: (iffP eqP) => [eqfA |]; last exact: card_in_image. by apply/dinjectiveP; apply/card_uniqP; rewrite size_map -cardE. Qed. Hypothesis injf : injective f. Lemma card_image A : #|image f A| = #|A|. Proof. apply: card_in_image; exact: in2W. Qed. Lemma card_codom : #|codom f| = #|T|. Proof. exact: card_image. Qed. Lemma card_preim (B : pred T') : #|[preim f of B]| = #|[predI codom f & B]|. Proof. rewrite -card_image /=; apply: eq_card => y. by rewrite [y \in _]image_pre !inE andbC. Qed. Hypothesis card_range : #|T| = #|T'|. Lemma inj_card_onto y : y \in codom f. Proof. by move: y; apply/subset_cardP; rewrite ?card_codom ?subset_predT. Qed. Lemma inj_card_bij : bijective f. Proof. by exists (fun y => iinv (inj_card_onto y)) => y; rewrite ?iinv_f ?f_iinv. Qed. End CardFunImage. Implicit Arguments image_injP [T T' f A]. Section FinCancel. Variables (T : finType) (f g : T -> T). Section Inv. Hypothesis injf : injective f. Lemma injF_onto y : y \in codom f. Proof. exact: inj_card_onto. Qed. Definition invF y := iinv (injF_onto y). Lemma invF_f : cancel f invF. Proof. by move=> x; exact: iinv_f. Qed. Lemma f_invF : cancel invF f. Proof. by move=> y; exact: f_iinv. Qed. Lemma injF_bij : bijective f. Proof. exact: inj_card_bij. Qed. End Inv. Hypothesis fK : cancel f g. Lemma canF_sym : cancel g f. Proof. exact/(bij_can_sym (injF_bij (can_inj fK))). Qed. Lemma canF_LR x y : x = g y -> f x = y. Proof. exact: canLR canF_sym. Qed. Lemma canF_RL x y : g x = y -> x = f y. Proof. exact: canRL canF_sym. Qed. Lemma canF_eq x y : (f x == y) = (x == g y). Proof. exact: (can2_eq fK canF_sym). Qed. Lemma canF_invF : g =1 invF (can_inj fK). Proof. by move=> y; apply: (canLR fK); rewrite f_invF. Qed. End FinCancel. Section EqImage. Variables (T : finType) (T' : Type). Lemma eq_image (A B : pred T) (f g : T -> T') : A =i B -> f =1 g -> image f A = image g B. Proof. by move=> eqAB eqfg; rewrite /image_mem (eq_enum eqAB) (eq_map eqfg). Qed. Lemma eq_codom (f g : T -> T') : f =1 g -> codom f = codom g. Proof. exact: eq_image. Qed. Lemma eq_invF f g injf injg : f =1 g -> @invF T f injf =1 @invF T g injg. Proof. by move=> eq_fg x; apply: (canLR (invF_f injf)); rewrite eq_fg f_invF. Qed. End EqImage. (* Standard finTypes *) Section SeqFinType. Variables (T : choiceType) (s : seq T). Record seq_sub : Type := SeqSub {ssval : T; ssvalP : in_mem ssval (@mem T _ s)}. Canonical seq_sub_subType := Eval hnf in [subType for ssval]. Definition seq_sub_eqMixin := Eval hnf in [eqMixin of seq_sub by <:]. Canonical seq_sub_eqType := Eval hnf in EqType seq_sub seq_sub_eqMixin. Definition seq_sub_choiceMixin := [choiceMixin of seq_sub by <:]. Canonical seq_sub_choiceType := Eval hnf in ChoiceType seq_sub seq_sub_choiceMixin. Definition seq_sub_enum : seq seq_sub := undup (pmap insub s). Lemma mem_seq_sub_enum x : x \in seq_sub_enum. Proof. by rewrite mem_undup mem_pmap -valK map_f ?ssvalP. Qed. Lemma val_seq_sub_enum : uniq s -> map val seq_sub_enum = s. Proof. move=> Us; rewrite /seq_sub_enum undup_id ?pmap_sub_uniq //. rewrite (pmap_filter (@insubK _ _ _)); apply/all_filterP. by apply/allP => x; rewrite isSome_insub. Qed. Definition seq_sub_pickle x := index x seq_sub_enum. Definition seq_sub_unpickle n := nth None (map some seq_sub_enum) n. Lemma seq_sub_pickleK : pcancel seq_sub_pickle seq_sub_unpickle. Proof. rewrite /seq_sub_unpickle => x. by rewrite (nth_map x) ?nth_index ?index_mem ?mem_seq_sub_enum. Qed. Definition seq_sub_countMixin := CountMixin seq_sub_pickleK. Canonical seq_sub_countType := Eval hnf in CountType seq_sub seq_sub_countMixin. Canonical seq_sub_subCountType := Eval hnf in [subCountType of seq_sub]. Definition seq_sub_finMixin := Eval hnf in UniqFinMixin (undup_uniq _) mem_seq_sub_enum. Canonical seq_sub_finType := Eval hnf in FinType seq_sub seq_sub_finMixin. Lemma card_seq_sub : uniq s -> #|{:seq_sub}| = size s. Proof. by move=> Us; rewrite cardE enumT -(size_map val) unlock val_seq_sub_enum. Qed. End SeqFinType. Lemma unit_enumP : Finite.axiom [::tt]. Proof. by case. Qed. Definition unit_finMixin := Eval hnf in FinMixin unit_enumP. Canonical unit_finType := Eval hnf in FinType unit unit_finMixin. Lemma card_unit : #|{: unit}| = 1. Proof. by rewrite cardT enumT unlock. Qed. Lemma bool_enumP : Finite.axiom [:: true; false]. Proof. by case. Qed. Definition bool_finMixin := Eval hnf in FinMixin bool_enumP. Canonical bool_finType := Eval hnf in FinType bool bool_finMixin. Lemma card_bool : #|{: bool}| = 2. Proof. by rewrite cardT enumT unlock. Qed. Local Notation enumF T := (Finite.enum T). Section OptionFinType. Variable T : finType. Definition option_enum := None :: map some (enumF T). Lemma option_enumP : Finite.axiom option_enum. Proof. by case=> [x|]; rewrite /= count_map (count_pred0, enumP). Qed. Definition option_finMixin := Eval hnf in FinMixin option_enumP. Canonical option_finType := Eval hnf in FinType (option T) option_finMixin. Lemma card_option : #|{: option T}| = #|T|.+1. Proof. by rewrite !cardT !enumT {1}unlock /= !size_map. Qed. End OptionFinType. Section TransferFinType. Variables (eT : countType) (fT : finType) (f : eT -> fT). Lemma pcan_enumP g : pcancel f g -> Finite.axiom (undup (pmap g (enumF fT))). Proof. move=> fK x; rewrite count_uniq_mem ?undup_uniq // mem_undup. by rewrite mem_pmap -fK map_f // -enumT mem_enum. Qed. Definition PcanFinMixin g fK := FinMixin (@pcan_enumP g fK). Definition CanFinMixin g (fK : cancel f g) := PcanFinMixin (can_pcan fK). End TransferFinType. Section SubFinType. Variables (T : choiceType) (P : pred T). Import Finite. Structure subFinType := SubFinType { subFin_sort :> subType P; _ : mixin_of (sub_eqType subFin_sort) }. Definition pack_subFinType U := fun cT b m & phant_id (class cT) (@Class U b m) => fun sT m' & phant_id m' m => @SubFinType sT m'. Implicit Type sT : subFinType. Definition subFin_mixin sT := let: SubFinType _ m := sT return mixin_of (sub_eqType sT) in m. Coercion subFinType_subCountType sT := @SubCountType _ _ sT (subFin_mixin sT). Canonical subFinType_subCountType. Coercion subFinType_finType sT := Pack (@Class sT (sub_choiceClass sT) (subFin_mixin sT)) sT. Canonical subFinType_finType. Lemma codom_val sT x : (x \in codom (val : sT -> T)) = P x. Proof. by apply/codomP/idP=> [[u ->]|Px]; last exists (Sub x Px); rewrite ?valP ?SubK. Qed. End SubFinType. (* This assumes that T has both finType and subCountType structures. *) Notation "[ 'subFinType' 'of' T ]" := (@pack_subFinType _ _ T _ _ _ id _ _ id) (at level 0, format "[ 'subFinType' 'of' T ]") : form_scope. Canonical seq_sub_subFinType T s := Eval hnf in [subFinType of @seq_sub T s]. Section FinTypeForSub. Variables (T : finType) (P : pred T) (sT : subCountType P). Definition sub_enum : seq sT := pmap insub (enumF T). Lemma mem_sub_enum u : u \in sub_enum. Proof. by rewrite mem_pmap_sub -enumT mem_enum. Qed. Lemma sub_enum_uniq : uniq sub_enum. Proof. by rewrite pmap_sub_uniq // -enumT enum_uniq. Qed. Lemma val_sub_enum : map val sub_enum = enum P. Proof. rewrite pmap_filter; last exact: insubK. by apply: eq_filter => x; exact: isSome_insub. Qed. (* We can't declare a canonical structure here because we've already *) (* stated that subType_sort and FinType.sort unify via to the *) (* subType_finType structure. *) Definition SubFinMixin := UniqFinMixin sub_enum_uniq mem_sub_enum. Definition SubFinMixin_for (eT : eqType) of phant eT := eq_rect _ Finite.mixin_of SubFinMixin eT. Variable sfT : subFinType P. Lemma card_sub : #|sfT| = #|[pred x | P x]|. Proof. by rewrite -(eq_card (codom_val sfT)) (card_image val_inj). Qed. Lemma eq_card_sub (A : pred sfT) : A =i predT -> #|A| = #|[pred x | P x]|. Proof. exact: eq_card_trans card_sub. Qed. End FinTypeForSub. (* This assumes that T has a subCountType structure over a type that *) (* has a finType structure. *) Notation "[ 'finMixin' 'of' T 'by' <: ]" := (SubFinMixin_for (Phant T) (erefl _)) (at level 0, format "[ 'finMixin' 'of' T 'by' <: ]") : form_scope. (* Regression for the subFinType stack Record myb : Type := MyB {myv : bool; _ : ~~ myv}. Canonical myb_sub := Eval hnf in [subType for myv]. Definition myb_eqm := Eval hnf in [eqMixin of myb by <:]. Canonical myb_eq := Eval hnf in EqType myb myb_eqm. Definition myb_chm := [choiceMixin of myb by <:]. Canonical myb_ch := Eval hnf in ChoiceType myb myb_chm. Definition myb_cntm := [countMixin of myb by <:]. Canonical myb_cnt := Eval hnf in CountType myb myb_cntm. Canonical myb_scnt := Eval hnf in [subCountType of myb]. Definition myb_finm := [finMixin of myb by <:]. Canonical myb_fin := Eval hnf in FinType myb myb_finm. Canonical myb_sfin := Eval hnf in [subFinType of myb]. Print Canonical Projections. Print myb_finm. Print myb_cntm. *) Section CardSig. Variables (T : finType) (P : pred T). Definition sig_finMixin := [finMixin of {x | P x} by <:]. Canonical sig_finType := Eval hnf in FinType {x | P x} sig_finMixin. Canonical sig_subFinType := Eval hnf in [subFinType of {x | P x}]. Lemma card_sig : #|{: {x | P x}}| = #|[pred x | P x]|. Proof. exact: card_sub. Qed. End CardSig. (**********************************************************************) (* *) (* Ordinal finType : {0, ... , n-1} *) (* *) (**********************************************************************) Section OrdinalSub. Variable n : nat. Inductive ordinal : predArgType := Ordinal m of m < n. Coercion nat_of_ord i := let: Ordinal m _ := i in m. Canonical ordinal_subType := [subType for nat_of_ord]. Definition ordinal_eqMixin := Eval hnf in [eqMixin of ordinal by <:]. Canonical ordinal_eqType := Eval hnf in EqType ordinal ordinal_eqMixin. Definition ordinal_choiceMixin := [choiceMixin of ordinal by <:]. Canonical ordinal_choiceType := Eval hnf in ChoiceType ordinal ordinal_choiceMixin. Definition ordinal_countMixin := [countMixin of ordinal by <:]. Canonical ordinal_countType := Eval hnf in CountType ordinal ordinal_countMixin. Canonical ordinal_subCountType := [subCountType of ordinal]. Lemma ltn_ord (i : ordinal) : i < n. Proof. exact: valP i. Qed. Lemma ord_inj : injective nat_of_ord. Proof. exact: val_inj. Qed. Definition ord_enum : seq ordinal := pmap insub (iota 0 n). Lemma val_ord_enum : map val ord_enum = iota 0 n. Proof. rewrite pmap_filter; last exact: insubK. by apply/all_filterP; apply/allP=> i; rewrite mem_iota isSome_insub. Qed. Lemma ord_enum_uniq : uniq ord_enum. Proof. by rewrite pmap_sub_uniq ?iota_uniq. Qed. Lemma mem_ord_enum i : i \in ord_enum. Proof. by rewrite -(mem_map ord_inj) val_ord_enum mem_iota ltn_ord. Qed. Definition ordinal_finMixin := Eval hnf in UniqFinMixin ord_enum_uniq mem_ord_enum. Canonical ordinal_finType := Eval hnf in FinType ordinal ordinal_finMixin. Canonical ordinal_subFinType := Eval hnf in [subFinType of ordinal]. End OrdinalSub. Notation "''I_' n" := (ordinal n) (at level 8, n at level 2, format "''I_' n"). Hint Resolve ltn_ord. Section OrdinalEnum. Variable n : nat. Lemma val_enum_ord : map val (enum 'I_n) = iota 0 n. Proof. by rewrite enumT unlock val_ord_enum. Qed. Lemma size_enum_ord : size (enum 'I_n) = n. Proof. by rewrite -(size_map val) val_enum_ord size_iota. Qed. Lemma card_ord : #|'I_n| = n. Proof. by rewrite cardE size_enum_ord. Qed. Lemma nth_enum_ord i0 m : m < n -> nth i0 (enum 'I_n) m = m :> nat. Proof. by move=> ?; rewrite -(nth_map _ 0) (size_enum_ord, val_enum_ord) // nth_iota. Qed. Lemma nth_ord_enum (i0 i : 'I_n) : nth i0 (enum 'I_n) i = i. Proof. apply: val_inj; exact: nth_enum_ord. Qed. Lemma index_enum_ord (i : 'I_n) : index i (enum 'I_n) = i. Proof. by rewrite -{1}(nth_ord_enum i i) index_uniq ?(enum_uniq, size_enum_ord). Qed. End OrdinalEnum. Lemma widen_ord_proof n m (i : 'I_n) : n <= m -> i < m. Proof. exact: leq_trans. Qed. Definition widen_ord n m le_n_m i := Ordinal (@widen_ord_proof n m i le_n_m). Lemma cast_ord_proof n m (i : 'I_n) : n = m -> i < m. Proof. by move <-. Qed. Definition cast_ord n m eq_n_m i := Ordinal (@cast_ord_proof n m i eq_n_m). Lemma cast_ord_id n eq_n i : cast_ord eq_n i = i :> 'I_n. Proof. exact: val_inj. Qed. Lemma cast_ord_comp n1 n2 n3 eq_n2 eq_n3 i : @cast_ord n2 n3 eq_n3 (@cast_ord n1 n2 eq_n2 i) = cast_ord (etrans eq_n2 eq_n3) i. Proof. exact: val_inj. Qed. Lemma cast_ordK n1 n2 eq_n : cancel (@cast_ord n1 n2 eq_n) (cast_ord (esym eq_n)). Proof. by move=> i; exact: val_inj. Qed. Lemma cast_ordKV n1 n2 eq_n : cancel (cast_ord (esym eq_n)) (@cast_ord n1 n2 eq_n). Proof. by move=> i; exact: val_inj. Qed. Lemma cast_ord_inj n1 n2 eq_n : injective (@cast_ord n1 n2 eq_n). Proof. exact: can_inj (cast_ordK eq_n). Qed. Lemma rev_ord_proof n (i : 'I_n) : n - i.+1 < n. Proof. by case: n i => [|n] [i lt_i_n] //; rewrite ltnS subSS leq_subr. Qed. Definition rev_ord n i := Ordinal (@rev_ord_proof n i). Lemma rev_ordK n : involutive (@rev_ord n). Proof. by case: n => [|n] [i lti] //; apply: val_inj; rewrite /= !subSS subKn. Qed. Lemma rev_ord_inj {n} : injective (@rev_ord n). Proof. exact: inv_inj (@rev_ordK n). Qed. (* bijection between any finType T and the Ordinal finType of its cardinal *) Section EnumRank. Variable T : finType. Implicit Type A : pred T. Lemma enum_rank_subproof x0 A : x0 \in A -> 0 < #|A|. Proof. by move=> Ax0; rewrite (cardD1 x0) Ax0. Qed. Definition enum_rank_in x0 A (Ax0 : x0 \in A) x := insubd (Ordinal (@enum_rank_subproof x0 [eta A] Ax0)) (index x (enum A)). Definition enum_rank x := @enum_rank_in x T (erefl true) x. Lemma enum_default A : 'I_(#|A|) -> T. Proof. by rewrite cardE; case: (enum A) => [|//] []. Qed. Definition enum_val A i := nth (@enum_default [eta A] i) (enum A) i. Prenex Implicits enum_val. Lemma enum_valP A i : @enum_val A i \in A. Proof. by rewrite -mem_enum mem_nth -?cardE. Qed. Lemma enum_val_nth A x i : @enum_val A i = nth x (enum A) i. Proof. by apply: set_nth_default; rewrite cardE in i *; exact: ltn_ord. Qed. Lemma nth_image T' y0 (f : T -> T') A (i : 'I_#|A|) : nth y0 (image f A) i = f (enum_val i). Proof. by rewrite -(nth_map _ y0) // -cardE. Qed. Lemma nth_codom T' y0 (f : T -> T') (i : 'I_#|T|) : nth y0 (codom f) i = f (enum_val i). Proof. exact: nth_image. Qed. Lemma nth_enum_rank_in x00 x0 A Ax0 : {in A, cancel (@enum_rank_in x0 A Ax0) (nth x00 (enum A))}. Proof. move=> x Ax; rewrite /= insubdK ?nth_index ?mem_enum //. by rewrite cardE [_ \in _]index_mem mem_enum. Qed. Lemma nth_enum_rank x0 : cancel enum_rank (nth x0 (enum T)). Proof. by move=> x; exact: nth_enum_rank_in. Qed. Lemma enum_rankK_in x0 A Ax0 : {in A, cancel (@enum_rank_in x0 A Ax0) enum_val}. Proof. by move=> x; exact: nth_enum_rank_in. Qed. Lemma enum_rankK : cancel enum_rank enum_val. Proof. by move=> x; exact: enum_rankK_in. Qed. Lemma enum_valK_in x0 A Ax0 : cancel enum_val (@enum_rank_in x0 A Ax0). Proof. move=> x; apply: ord_inj; rewrite insubdK; last first. by rewrite cardE [_ \in _]index_mem mem_nth // -cardE. by rewrite index_uniq ?enum_uniq // -cardE. Qed. Lemma enum_valK : cancel enum_val enum_rank. Proof. by move=> x; exact: enum_valK_in. Qed. Lemma enum_rank_inj : injective enum_rank. Proof. exact: can_inj enum_rankK. Qed. Lemma enum_val_inj A : injective (@enum_val A). Proof. by move=> i; exact: can_inj (enum_valK_in (enum_valP i)) (i). Qed. Lemma enum_val_bij_in x0 A : x0 \in A -> {on A, bijective (@enum_val A)}. Proof. move=> Ax0; exists (enum_rank_in Ax0) => [i _|]; last exact: enum_rankK_in. exact: enum_valK_in. Qed. Lemma enum_rank_bij : bijective enum_rank. Proof. by move: enum_rankK enum_valK; exists (@enum_val T). Qed. Lemma enum_val_bij : bijective (@enum_val T). Proof. by move: enum_rankK enum_valK; exists enum_rank. Qed. (* Due to the limitations of the Coq unification patterns, P can only be *) (* inferred from the premise of this lemma, not its conclusion. As a result *) (* this lemma will only be usable in forward chaining style. *) Lemma fin_all_exists U (P : forall x : T, U x -> Prop) : (forall x, exists u, P x u) -> (exists u, forall x, P x (u x)). Proof. move=> ex_u; pose Q m x := enum_rank x < m -> {ux | P x ux}. suffices: forall m, m <= #|T| -> exists w : forall x, Q m x, True. case/(_ #|T|)=> // w _; pose u x := sval (w x (ltn_ord _)). by exists u => x; rewrite {}/u; case: (w x _). elim=> [|m IHm] ltmX; first by have w x: Q 0 x by []; exists w. have{IHm} [w _] := IHm (ltnW ltmX); pose i := Ordinal ltmX. have [u Pu] := ex_u (enum_val i); suffices w' x: Q m.+1 x by exists w'. rewrite /Q ltnS leq_eqVlt (val_eqE _ i); case: eqP => [def_i _ | _ /w //]. by rewrite -def_i enum_rankK in u Pu; exists u. Qed. Lemma fin_all_exists2 U (P Q : forall x : T, U x -> Prop) : (forall x, exists2 u, P x u & Q x u) -> (exists2 u, forall x, P x (u x) & forall x, Q x (u x)). Proof. move=> ex_u; have (x): exists u, P x u /\ Q x u by have [u] := ex_u x; exists u. by case/fin_all_exists=> u /all_and2[]; exists u. Qed. End EnumRank. Implicit Arguments enum_val_inj [[T] [A] x1 x2]. Implicit Arguments enum_rank_inj [[T] x1 x2]. Prenex Implicits enum_val enum_rank. Lemma enum_rank_ord n i : enum_rank i = cast_ord (esym (card_ord n)) i. Proof. by apply: val_inj; rewrite insubdK ?index_enum_ord // card_ord [_ \in _]ltn_ord. Qed. Lemma enum_val_ord n i : enum_val i = cast_ord (card_ord n) i. Proof. by apply: canLR (@enum_rankK _) _; apply: val_inj; rewrite enum_rank_ord. Qed. (* The integer bump / unbump operations. *) Definition bump h i := (h <= i) + i. Definition unbump h i := i - (h < i). Lemma bumpK h : cancel (bump h) (unbump h). Proof. rewrite /bump /unbump => i. have [le_hi | lt_ih] := leqP h i; first by rewrite ltnS le_hi subn1. by rewrite ltnNge ltnW ?subn0. Qed. Lemma neq_bump h i : h != bump h i. Proof. rewrite /bump eqn_leq; have [le_hi | lt_ih] := leqP h i. by rewrite ltnNge le_hi andbF. by rewrite leqNgt lt_ih. Qed. Lemma unbumpKcond h i : bump h (unbump h i) = (i == h) + i. Proof. rewrite /bump /unbump leqNgt -subSKn. case: (ltngtP i h) => /= [-> | ltih | ->] //; last by rewrite ltnn. by rewrite subn1 /= leqNgt !(ltn_predK ltih, ltih, add1n). Qed. Lemma unbumpK h : {in predC1 h, cancel (unbump h) (bump h)}. Proof. by move=> i; move/negbTE=> neq_h_i; rewrite unbumpKcond neq_h_i. Qed. Lemma bump_addl h i k : bump (k + h) (k + i) = k + bump h i. Proof. by rewrite /bump leq_add2l addnCA. Qed. Lemma bumpS h i : bump h.+1 i.+1 = (bump h i).+1. Proof. exact: addnS. Qed. Lemma unbump_addl h i k : unbump (k + h) (k + i) = k + unbump h i. Proof. apply: (can_inj (bumpK (k + h))). by rewrite bump_addl !unbumpKcond eqn_add2l addnCA. Qed. Lemma unbumpS h i : unbump h.+1 i.+1 = (unbump h i).+1. Proof. exact: unbump_addl 1. Qed. Lemma leq_bump h i j : (i <= bump h j) = (unbump h i <= j). Proof. rewrite /bump leq_subLR. case: (leqP i h) (leqP h j) => [le_i_h | lt_h_i] [le_h_j | lt_j_h] //. by rewrite leqW (leq_trans le_i_h). by rewrite !(leqNgt i) ltnW (leq_trans _ lt_h_i). Qed. Lemma leq_bump2 h i j : (bump h i <= bump h j) = (i <= j). Proof. by rewrite leq_bump bumpK. Qed. Lemma bumpC h1 h2 i : bump h1 (bump h2 i) = bump (bump h1 h2) (bump (unbump h2 h1) i). Proof. rewrite {1 5}/bump -leq_bump addnCA; congr (_ + (_ + _)). rewrite 2!leq_bump /unbump /bump; case: (leqP h1 h2) => [le_h12 | lt_h21]. by rewrite subn0 ltnS le_h12 subn1. by rewrite subn1 (ltn_predK lt_h21) (leqNgt h1) lt_h21 subn0. Qed. (* The lift operations on ordinals; to avoid a messy dependent type, *) (* unlift is a partial operation (returns an option). *) Lemma lift_subproof n h (i : 'I_n.-1) : bump h i < n. Proof. by case: n i => [[]|n] //= i; rewrite -addnS (leq_add (leq_b1 _)). Qed. Definition lift n (h : 'I_n) (i : 'I_n.-1) := Ordinal (lift_subproof h i). Lemma unlift_subproof n (h : 'I_n) (u : {j | j != h}) : unbump h (val u) < n.-1. Proof. case: n h u => [|n h] [] //= j ne_jh. rewrite -(leq_bump2 h.+1) bumpS unbumpK // /bump. case: (ltngtP n h) => [|_|eq_nh]; rewrite ?(leqNgt _ h) ?ltn_ord //. by rewrite ltn_neqAle [j <= _](valP j) {2}eq_nh andbT. Qed. Definition unlift n (h i : 'I_n) := omap (fun u : {j | j != h} => Ordinal (unlift_subproof u)) (insub i). CoInductive unlift_spec n h i : option 'I_n.-1 -> Type := | UnliftSome j of i = lift h j : unlift_spec h i (Some j) | UnliftNone of i = h : unlift_spec h i None. Lemma unliftP n (h i : 'I_n) : unlift_spec h i (unlift h i). Proof. rewrite /unlift; case: insubP => [u nhi | ] def_i /=; constructor. by apply: val_inj; rewrite /= def_i unbumpK. by rewrite negbK in def_i; exact/eqP. Qed. Lemma neq_lift n (h : 'I_n) i : h != lift h i. Proof. exact: neq_bump. Qed. Lemma unlift_none n (h : 'I_n) : unlift h h = None. Proof. by case: unliftP => // j Dh; case/eqP: (neq_lift h j). Qed. Lemma unlift_some n (h i : 'I_n) : h != i -> {j | i = lift h j & unlift h i = Some j}. Proof. rewrite eq_sym => /eqP neq_ih. by case Dui: (unlift h i) / (unliftP h i) => [j Dh|//]; exists j. Qed. Lemma lift_inj n (h : 'I_n) : injective (lift h). Proof. move=> i1 i2; move/eqP; rewrite [_ == _](can_eq (@bumpK _)) => eq_i12. exact/eqP. Qed. Lemma liftK n (h : 'I_n) : pcancel (lift h) (unlift h). Proof. by move=> i; case: (unlift_some (neq_lift h i)) => j; move/lift_inj->. Qed. (* Shifting and splitting indices, for cutting and pasting arrays *) Lemma lshift_subproof m n (i : 'I_m) : i < m + n. Proof. by apply: leq_trans (valP i) _; exact: leq_addr. Qed. Lemma rshift_subproof m n (i : 'I_n) : m + i < m + n. Proof. by rewrite ltn_add2l. Qed. Definition lshift m n (i : 'I_m) := Ordinal (lshift_subproof n i). Definition rshift m n (i : 'I_n) := Ordinal (rshift_subproof m i). Lemma split_subproof m n (i : 'I_(m + n)) : i >= m -> i - m < n. Proof. by move/subSn <-; rewrite leq_subLR. Qed. Definition split m n (i : 'I_(m + n)) : 'I_m + 'I_n := match ltnP (i) m with | LtnNotGeq lt_i_m => inl _ (Ordinal lt_i_m) | GeqNotLtn ge_i_m => inr _ (Ordinal (split_subproof ge_i_m)) end. CoInductive split_spec m n (i : 'I_(m + n)) : 'I_m + 'I_n -> bool -> Type := | SplitLo (j : 'I_m) of i = j :> nat : split_spec i (inl _ j) true | SplitHi (k : 'I_n) of i = m + k :> nat : split_spec i (inr _ k) false. Lemma splitP m n (i : 'I_(m + n)) : split_spec i (split i) (i < m). Proof. rewrite /split {-3}/leq. by case: (@ltnP i m) => cmp_i_m //=; constructor; rewrite ?subnKC. Qed. Definition unsplit m n (jk : 'I_m + 'I_n) := match jk with inl j => lshift n j | inr k => rshift m k end. Lemma ltn_unsplit m n (jk : 'I_m + 'I_n) : (unsplit jk < m) = jk. Proof. by case: jk => [j|k]; rewrite /= ?ltn_ord // ltnNge leq_addr. Qed. Lemma splitK m n : cancel (@split m n) (@unsplit m n). Proof. by move=> i; apply: val_inj; case: splitP. Qed. Lemma unsplitK m n : cancel (@unsplit m n) (@split m n). Proof. move=> jk; have:= ltn_unsplit jk. by do [case: splitP; case: jk => //= i j] => [|/addnI] => /ord_inj->. Qed. Section OrdinalPos. Variable n' : nat. Local Notation n := n'.+1. Definition ord0 := Ordinal (ltn0Sn n'). Definition ord_max := Ordinal (ltnSn n'). Lemma leq_ord (i : 'I_n) : i <= n'. Proof. exact: valP i. Qed. Lemma sub_ord_proof m : n' - m < n. Proof. by rewrite ltnS leq_subr. Qed. Definition sub_ord m := Ordinal (sub_ord_proof m). Lemma sub_ordK (i : 'I_n) : n' - (n' - i) = i. Proof. by rewrite subKn ?leq_ord. Qed. Definition inord m : 'I_n := insubd ord0 m. Lemma inordK m : m < n -> inord m = m :> nat. Proof. by move=> lt_m; rewrite val_insubd lt_m. Qed. Lemma inord_val (i : 'I_n) : inord i = i. Proof. by rewrite /inord /insubd valK. Qed. Lemma enum_ordS : enum 'I_n = ord0 :: map (lift ord0) (enum 'I_n'). Proof. apply: (inj_map val_inj); rewrite val_enum_ord /= -map_comp. by rewrite (map_comp (addn 1)) val_enum_ord -iota_addl. Qed. Lemma lift_max (i : 'I_n') : lift ord_max i = i :> nat. Proof. by rewrite /= /bump leqNgt ltn_ord. Qed. Lemma lift0 (i : 'I_n') : lift ord0 i = i.+1 :> nat. Proof. by []. Qed. End OrdinalPos. Implicit Arguments ord0 [[n']]. Implicit Arguments ord_max [[n']]. Implicit Arguments inord [[n']]. Implicit Arguments sub_ord [[n']]. (* Product of two fintypes which is a fintype *) Section ProdFinType. Variable T1 T2 : finType. Definition prod_enum := [seq (x1, x2) | x1 <- enum T1, x2 <- enum T2]. Lemma predX_prod_enum (A1 : pred T1) (A2 : pred T2) : count [predX A1 & A2] prod_enum = #|A1| * #|A2|. Proof. rewrite !cardE -!count_filter -!enumT /prod_enum. elim: (enum T1) => //= x1 s1 IHs; rewrite count_cat {}IHs count_map /preim /=. by case: (x1 \in A1); rewrite ?count_pred0. Qed. Lemma prod_enumP : Finite.axiom prod_enum. Proof. by case=> x1 x2; rewrite (predX_prod_enum (pred1 x1) (pred1 x2)) !card1. Qed. Definition prod_finMixin := Eval hnf in FinMixin prod_enumP. Canonical prod_finType := Eval hnf in FinType (T1 * T2) prod_finMixin. Lemma cardX (A1 : pred T1) (A2 : pred T2) : #|[predX A1 & A2]| = #|A1| * #|A2|. Proof. by rewrite -predX_prod_enum unlock -count_filter unlock. Qed. Lemma card_prod : #|{: T1 * T2}| = #|T1| * #|T2|. Proof. by rewrite -cardX; apply: eq_card; case. Qed. Lemma eq_card_prod (A : pred (T1 * T2)) : A =i predT -> #|A| = #|T1| * #|T2|. Proof. exact: eq_card_trans card_prod. Qed. End ProdFinType. Section TagFinType. Variables (I : finType) (T_ : I -> finType). Definition tag_enum := flatten [seq [seq Tagged T_ x | x <- enumF (T_ i)] | i <- enumF I]. Lemma tag_enumP : Finite.axiom tag_enum. Proof. case=> i x; rewrite -(enumP i) /tag_enum -enumT. elim: (enum I) => //= j e IHe. rewrite count_cat count_map {}IHe; congr (_ + _). rewrite count_filter -cardE /=; case: eqP => [-> | ne_j_i]. by apply: (@eq_card1 _ x) => y; rewrite -topredE /= tagged_asE ?eqxx. by apply: eq_card0 => y. Qed. Definition tag_finMixin := Eval hnf in FinMixin tag_enumP. Canonical tag_finType := Eval hnf in FinType {i : I & T_ i} tag_finMixin. Lemma card_tagged : #|{: {i : I & T_ i}}| = sumn (map (fun i => #|T_ i|) (enum I)). Proof. rewrite cardE !enumT {1}unlock size_flatten /shape -map_comp. by congr (sumn _); apply: eq_map => i; rewrite /= size_map -enumT -cardE. Qed. End TagFinType. Section SumFinType. Variables T1 T2 : finType. Definition sum_enum := [seq inl _ x | x <- enumF T1] ++ [seq inr _ y | y <- enumF T2]. Lemma sum_enum_uniq : uniq sum_enum. Proof. rewrite cat_uniq -!enumT !(enum_uniq, map_inj_uniq); try by move=> ? ? []. by rewrite andbT; apply/hasP=> [[_ /mapP[x _ ->] /mapP[]]]. Qed. Lemma mem_sum_enum u : u \in sum_enum. Proof. by case: u => x; rewrite mem_cat -!enumT map_f ?mem_enum ?orbT. Qed. Definition sum_finMixin := Eval hnf in UniqFinMixin sum_enum_uniq mem_sum_enum. Canonical sum_finType := Eval hnf in FinType (T1 + T2) sum_finMixin. Lemma card_sum : #|{: T1 + T2}| = #|T1| + #|T2|. Proof. by rewrite !cardT !enumT {1}unlock size_cat !size_map. Qed. End SumFinType. ssreflect-1.5/theories/ssrnat.v0000644000175000017500000016217212175455021015673 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import ssreflect ssrfun ssrbool eqtype. Require Import BinNat. Require BinPos Ndec. Require Export Ring. (******************************************************************************) (* A version of arithmetic on nat (natural numbers) that is better suited to *) (* small scale reflection than the Coq Arith library. It contains an *) (* extensive equational theory (including, e.g., the AGM inequality), as well *) (* as support for the ring tactic, and congruence tactics. *) (* The following operations and notations are provided: *) (* *) (* successor and predecessor *) (* n.+1, n.+2, n.+3, n.+4 and n.-1, n.-2 *) (* this frees the names "S" and "pred" *) (* *) (* basic arithmetic *) (* m + n, m - n, m * n *) (* Important: m - n denotes TRUNCATED substraction: m - n = 0 if m <= n. *) (* The definitions use the nosimpl tag to prevent undesirable computation *) (* computation during simplification, but remain compatible with the ones *) (* provided in the Coq.Init.Peano prelude. *) (* For computation, a module NatTrec rebinds all arithmetic notations *) (* to less convenient but also less inefficient tail-recursive functions; *) (* the auxiliary functions used by these versions are flagged with %Nrec. *) (* Also, there is support for input and output of large nat values. *) (* Num 3 082 241 inputs the number 3082241 *) (* [Num of n] outputs the value n *) (* There are coercions num >-> BinNat.N >-> nat; ssrnat rebinds the scope *) (* delimter for BinNat.N to %num, as it uses the shorter %N for its own *) (* notations (Peano notations are flagged with %coq_nat). *) (* *) (* doubling, halving, and parity *) (* n.*2, n./2, odd n, uphalf n, with uphalf n = n.+1./2 *) (* bool coerces to nat so we can write, e.g., n = odd n + n./2.*2. *) (* *) (* iteration *) (* iter n f x0 == f ( .. (f x0)) *) (* iteri n g x0 == g n.-1 (g ... (g 0 x0)) *) (* iterop n op x x0 == op x (... op x x) (n x's) or x0 if n = 0 *) (* *) (* exponentiation, factorial *) (* m ^ n, n`! *) (* m ^ 1 is convertible to m, and m ^ 2 to m * m *) (* *) (* comparison *) (* m <= n, m < n, m >= n, m > n, m == n, m <= n <= p, etc., *) (* comparisons are BOOLEAN operators, and m == n is the generic eqType *) (* operation. *) (* Most compatibility lemmas are stated as boolean equalities; this keeps *) (* the size of the library down. All the inequalities refer to the same *) (* constant "leq"; in particular m < n is identical to m.+1 <= n. *) (* *) (* conditionally strict inequality `leqif' *) (* m <= n ?= iff condition == (m <= n) and ((m == n) = condition) *) (* This is actually a pair of boolean equalities, so rewriting with an *) (* `leqif' lemma can affect several kinds of comparison. The transitivity *) (* lemma for leqif aggregates the conditions, allowing for arguments of *) (* the form ``m <= n <= p <= m, so equality holds throughout''. *) (* *) (* maximum and minimum *) (* maxn m n, minn m n *) (* Note that maxn m n = m + (m - n), due to the truncating subtraction. *) (* Absolute difference (linear distance) between nats is defined in the int *) (* library (in the int.IntDist sublibrary), with the syntax `|m - n|. The *) (* '-' in this notation is the signed integer difference. *) (* *) (* countable choice *) (* ex_minn : forall P : pred nat, (exists n, P n) -> nat *) (* This returns the smallest n such that P n holds. *) (* ex_maxn : forall (P : pred nat) m, *) (* (exists n, P n) -> (forall n, P n -> n <= m) -> nat *) (* This returns the largest n such that P n holds (given an explicit upper *) (* bound). *) (* *) (* This file adds the following suffix conventions to those documented in *) (* ssrbool.v and eqtype.v: *) (* A (infix) -- conjunction, as in *) (* ltn_neqAle : (m < n) = (m != n) && (m <= n). *) (* B -- subtraction, as in subBn : (m - n) - p = m - (n + p). *) (* D -- addition, as in mulnDl : (m + n) * p = m * p + n * p. *) (* M -- multiplication, as in expnMn : (m * n) ^ p = m ^ p * n ^ p. *) (* p (prefix) -- positive, as in *) (* eqn_pmul2l : m > 0 -> (m * n1 == m * n2) = (n1 == n2). *) (* P -- greater than 1, as in *) (* ltn_Pmull : 1 < n -> 0 < m -> m < n * m. *) (* S -- successor, as in addSn : n.+1 + m = (n + m).+1. *) (* V (infix) -- disjunction, as in *) (* leq_eqVlt : (m <= n) = (m == n) || (m < n). *) (* X - exponentiation, as in lognX : logn p (m ^ n) = logn p m * n in *) (* file prime.v (the suffix is not used in ths file). *) (* Suffixes that abreviate operations (D, B, M and X) are used to abbreviate *) (* second-rank operations in equational lemma names that describe left-hand *) (* sides (e.g., mulnDl); they are not used to abbreviate the main operation *) (* of relational lemmas (e.g., leq_add2l). *) (* For the asymmetrical exponentiation operator expn (m ^ n) a right suffix *) (* indicates an operation on the exponent, e.g., expnM : m ^ (n1 * n2) = ...; *) (* a trailing "n" is used to indicate the left operand, e.g., *) (* expnMn : (m1 * m2) ^ n = ... The operands of other operators a selected *) (* using the l/r suffixes. *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. (* Declare legacy Arith operators in new scope. *) Delimit Scope coq_nat_scope with coq_nat. Notation "m + n" := (plus m n) : coq_nat_scope. Notation "m - n" := (minus m n) : coq_nat_scope. Notation "m * n" := (mult m n) : coq_nat_scope. Notation "m <= n" := (le m n) : coq_nat_scope. Notation "m < n" := (lt m n) : coq_nat_scope. Notation "m >= n" := (ge m n) : coq_nat_scope. Notation "m > n" := (gt m n) : coq_nat_scope. (* Rebind scope delimiters, reserving a scope for the "recursive", *) (* i.e., unprotected version of operators. *) Delimit Scope N_scope with num. Delimit Scope nat_scope with N. Delimit Scope nat_rec_scope with Nrec. (* Postfix notation for the successor and predecessor functions. *) (* SSreflect uses "pred" for the generic predicate type, and S as *) (* a local bound variable. *) Notation succn := Datatypes.S. Notation predn := Peano.pred. Notation "n .+1" := (succn n) (at level 2, left associativity, format "n .+1") : nat_scope. Notation "n .+2" := n.+1.+1 (at level 2, left associativity, format "n .+2") : nat_scope. Notation "n .+3" := n.+2.+1 (at level 2, left associativity, format "n .+3") : nat_scope. Notation "n .+4" := n.+2.+2 (at level 2, left associativity, format "n .+4") : nat_scope. Notation "n .-1" := (predn n) (at level 2, left associativity, format "n .-1") : nat_scope. Notation "n .-2" := n.-1.-1 (at level 2, left associativity, format "n .-2") : nat_scope. Lemma succnK : cancel succn predn. Proof. by []. Qed. Lemma succn_inj : injective succn. Proof. by move=> n m []. Qed. (* Predeclare postfix doubling/halving operators. *) Reserved Notation "n .*2" (at level 2, format "n .*2"). Reserved Notation "n ./2" (at level 2, format "n ./2"). (* Canonical comparison and eqType for nat. *) Fixpoint eqn m n {struct m} := match m, n with | 0, 0 => true | m'.+1, n'.+1 => eqn m' n' | _, _ => false end. Lemma eqnP : Equality.axiom eqn. Proof. move=> n m; apply: (iffP idP) => [|<-]; last by elim n. by elim: n m => [|n IHn] [|m] //= /IHn->. Qed. Canonical nat_eqMixin := EqMixin eqnP. Canonical nat_eqType := Eval hnf in EqType nat nat_eqMixin. Implicit Arguments eqnP [x y]. Prenex Implicits eqnP. Lemma eqnE : eqn = eq_op. Proof. by []. Qed. Lemma eqSS m n : (m.+1 == n.+1) = (m == n). Proof. by []. Qed. Lemma nat_irrelevance (x y : nat) (E E' : x = y) : E = E'. Proof. exact: eq_irrelevance. Qed. (* Protected addition, with a more systematic set of lemmas. *) Definition addn_rec := plus. Notation "m + n" := (addn_rec m n) : nat_rec_scope. Definition addn := nosimpl addn_rec. Notation "m + n" := (addn m n) : nat_scope. Lemma addnE : addn = addn_rec. Proof. by []. Qed. Lemma plusE : plus = addn. Proof. by []. Qed. Lemma add0n : left_id 0 addn. Proof. by []. Qed. Lemma addSn m n : m.+1 + n = (m + n).+1. Proof. by []. Qed. Lemma add1n n : 1 + n = n.+1. Proof. by []. Qed. Lemma addn0 : right_id 0 addn. Proof. by move=> n; apply/eqP; elim: n. Qed. Lemma addnS m n : m + n.+1 = (m + n).+1. Proof. by elim: m. Qed. Lemma addSnnS m n : m.+1 + n = m + n.+1. Proof. by rewrite addnS. Qed. Lemma addnCA : left_commutative addn. Proof. by move=> m n p; elim: m => //= m; rewrite addnS => <-. Qed. Lemma addnC : commutative addn. Proof. by move=> m n; rewrite -{1}[n]addn0 addnCA addn0. Qed. Lemma addn1 n : n + 1 = n.+1. Proof. by rewrite addnC. Qed. Lemma addnA : associative addn. Proof. by move=> m n p; rewrite (addnC n) addnCA addnC. Qed. Lemma addnAC : right_commutative addn. Proof. by move=> m n p; rewrite -!addnA (addnC n). Qed. Lemma addnACA : interchange addn addn. Proof. by move=> m n p q; rewrite -!addnA (addnCA n). Qed. Lemma addn_eq0 m n : (m + n == 0) = (m == 0) && (n == 0). Proof. by case: m; case: n. Qed. Lemma eqn_add2l p m n : (p + m == p + n) = (m == n). Proof. by elim: p. Qed. Lemma eqn_add2r p m n : (m + p == n + p) = (m == n). Proof. by rewrite -!(addnC p) eqn_add2l. Qed. Lemma addnI : right_injective addn. Proof. by move=> p m n Heq; apply: eqP; rewrite -(eqn_add2l p) Heq eqxx. Qed. Lemma addIn : left_injective addn. Proof. move=> p m n; rewrite -!(addnC p); apply addnI. Qed. Lemma addn2 m : m + 2 = m.+2. Proof. by rewrite addnC. Qed. Lemma add2n m : 2 + m = m.+2. Proof. by []. Qed. Lemma addn3 m : m + 3 = m.+3. Proof. by rewrite addnC. Qed. Lemma add3n m : 3 + m = m.+3. Proof. by []. Qed. Lemma addn4 m : m + 4 = m.+4. Proof. by rewrite addnC. Qed. Lemma add4n m : 4 + m = m.+4. Proof. by []. Qed. (* Protected, structurally decreasing substraction, and basic lemmas. *) (* Further properties depend on ordering conditions. *) Definition subn_rec := minus. Notation "m - n" := (subn_rec m n) : nat_rec_scope. Definition subn := nosimpl subn_rec. Notation "m - n" := (subn m n) : nat_scope. Lemma subnE : subn = subn_rec. Proof. by []. Qed. Lemma minusE : minus = subn. Proof. by []. Qed. Lemma sub0n : left_zero 0 subn. Proof. by []. Qed. Lemma subn0 : right_id 0 subn. Proof. by case. Qed. Lemma subnn : self_inverse 0 subn. Proof. by elim. Qed. Lemma subSS n m : m.+1 - n.+1 = m - n. Proof. by []. Qed. Lemma subn1 n : n - 1 = n.-1. Proof. by case: n => [|[]]. Qed. Lemma subn2 n : (n - 2)%N = n.-2. Proof. by case: n => [|[|[]]]. Qed. Lemma subnDl p m n : (p + m) - (p + n) = m - n. Proof. by elim: p. Qed. Lemma subnDr p m n : (m + p) - (n + p) = m - n. Proof. by rewrite -!(addnC p) subnDl. Qed. Lemma addKn n : cancel (addn n) (subn^~ n). Proof. by move=> m; rewrite /= -{2}[n]addn0 subnDl subn0. Qed. Lemma addnK n : cancel (addn^~ n) (subn^~ n). Proof. by move=> m; rewrite /= (addnC m) addKn. Qed. Lemma subSnn n : n.+1 - n = 1. Proof. exact (addnK n 1). Qed. Lemma subnDA m n p : n - (m + p) = (n - m) - p. Proof. by elim: m n => [|m IHm] [|n]; try exact (IHm n). Qed. Lemma subnAC : right_commutative subn. Proof. by move=> m n p; rewrite -!subnDA addnC. Qed. Lemma subnS m n : m - n.+1 = (m - n).-1. Proof. by rewrite -addn1 subnDA subn1. Qed. Lemma subSKn m n : (m.+1 - n).-1 = m - n. Proof. by rewrite -subnS. Qed. (* Integer ordering, and its interaction with the other operations. *) Definition leq m n := m - n == 0. Notation "m <= n" := (leq m n) : nat_scope. Notation "m < n" := (m.+1 <= n) : nat_scope. Notation "m >= n" := (n <= m) (only parsing) : nat_scope. Notation "m > n" := (n < m) (only parsing) : nat_scope. (* For sorting, etc. *) Definition geq := [rel m n | m >= n]. Definition ltn := [rel m n | m < n]. Definition gtn := [rel m n | m > n]. Notation "m <= n <= p" := ((m <= n) && (n <= p)) : nat_scope. Notation "m < n <= p" := ((m < n) && (n <= p)) : nat_scope. Notation "m <= n < p" := ((m <= n) && (n < p)) : nat_scope. Notation "m < n < p" := ((m < n) && (n < p)) : nat_scope. Lemma ltnS m n : (m < n.+1) = (m <= n). Proof. by []. Qed. Lemma leq0n n : 0 <= n. Proof. by []. Qed. Lemma ltn0Sn n : 0 < n.+1. Proof. by []. Qed. Lemma ltn0 n : n < 0 = false. Proof. by []. Qed. Lemma leqnn n : n <= n. Proof. by elim: n. Qed. Hint Resolve leqnn. Lemma ltnSn n : n < n.+1. Proof. by []. Qed. Lemma eq_leq m n : m = n -> m <= n. Proof. by move->. Qed. Lemma leqnSn n : n <= n.+1. Proof. by elim: n. Qed. Hint Resolve leqnSn. Lemma leq_pred n : n.-1 <= n. Proof. by case: n => /=. Qed. Lemma leqSpred n : n <= n.-1.+1. Proof. by case: n => /=. Qed. Lemma ltn_predK m n : m < n -> n.-1.+1 = n. Proof. by case: n. Qed. Lemma prednK n : 0 < n -> n.-1.+1 = n. Proof. exact: ltn_predK. Qed. Lemma leqNgt m n : (m <= n) = ~~ (n < m). Proof. by elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. Lemma ltnNge m n : (m < n) = ~~ (n <= m). Proof. by rewrite leqNgt. Qed. Lemma ltnn n : n < n = false. Proof. by rewrite ltnNge leqnn. Qed. Lemma leqn0 n : (n <= 0) = (n == 0). Proof. by case: n. Qed. Lemma lt0n n : (0 < n) = (n != 0). Proof. by case: n. Qed. Lemma lt0n_neq0 n : 0 < n -> n != 0. Proof. by case: n. Qed. Lemma eqn0Ngt n : (n == 0) = ~~ (n > 0). Proof. by case: n. Qed. Lemma neq0_lt0n n : (n == 0) = false -> 0 < n. Proof. by case: n. Qed. Hint Resolve lt0n_neq0 neq0_lt0n. Lemma eqn_leq m n : (m == n) = (m <= n <= m). Proof. elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. Lemma anti_leq : antisymmetric leq. Proof. by move=> m n; rewrite -eqn_leq => /eqP. Qed. Lemma neq_ltn m n : (m != n) = (m < n) || (n < m). Proof. by rewrite eqn_leq negb_and orbC -!ltnNge. Qed. Lemma gtn_eqF m n : m < n -> n == m = false. Proof. by rewrite eqn_leq (leqNgt n) => ->. Qed. Lemma ltn_eqF m n : m < n -> m == n = false. Proof. by move/gtn_eqF; rewrite eq_sym. Qed. Lemma leq_eqVlt m n : (m <= n) = (m == n) || (m < n). Proof. elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. Lemma ltn_neqAle m n : (m < n) = (m != n) && (m <= n). Proof. by rewrite ltnNge leq_eqVlt negb_or -leqNgt eq_sym. Qed. Lemma leq_trans n m p : m <= n -> n <= p -> m <= p. Proof. by elim: n m p => [|i IHn] [|m] [|p] //; exact: IHn m p. Qed. Lemma leq_ltn_trans n m p : m <= n -> n < p -> m < p. Proof. move=> Hmn; exact: leq_trans. Qed. Lemma ltnW m n : m < n -> m <= n. Proof. exact: leq_trans. Qed. Hint Resolve ltnW. Lemma leqW m n : m <= n -> m <= n.+1. Proof. by move=> le_mn; exact: ltnW. Qed. Lemma ltn_trans n m p : m < n -> n < p -> m < p. Proof. by move=> lt_mn /ltnW; exact: leq_trans. Qed. Lemma leq_total m n : (m <= n) || (m >= n). Proof. by rewrite -implyNb -ltnNge; apply/implyP; exact: ltnW. Qed. (* Link to the legacy comparison predicates. *) Lemma leP m n : reflect (m <= n)%coq_nat (m <= n). Proof. apply: (iffP idP); last by elim: n / => // n _ /leq_trans->. elim: n => [|n IHn]; first by case: m. by rewrite leq_eqVlt ltnS => /predU1P[<- // | /IHn]; right. Qed. Implicit Arguments leP [m n]. Lemma le_irrelevance m n le_mn1 le_mn2 : le_mn1 = le_mn2 :> (m <= n)%coq_nat. Proof. elim: {n}n.+1 {-1}n (erefl n.+1) => // n IHn _ [<-] in le_mn1 le_mn2 *. pose def_n2 := erefl n; transitivity (eq_ind _ _ le_mn2 _ def_n2) => //. move def_n1: {1 4 5 7}n le_mn1 le_mn2 def_n2 => n1 le_mn1. case: n1 / le_mn1 def_n1 => [|n1 le_mn1] def_n1 [|n2 le_mn2] def_n2. - by rewrite [def_n2]eq_axiomK. - by move/leP: (le_mn2); rewrite -{1}def_n2 ltnn. - by move/leP: (le_mn1); rewrite {1}def_n2 ltnn. case: def_n2 (def_n2) => ->{n2} def_n2 in le_mn2 *. by rewrite [def_n2]eq_axiomK /=; congr le_S; exact: IHn. Qed. Lemma ltP m n : reflect (m < n)%coq_nat (m < n). Proof. exact leP. Qed. Implicit Arguments ltP [m n]. Lemma lt_irrelevance m n lt_mn1 lt_mn2 : lt_mn1 = lt_mn2 :> (m < n)%coq_nat. Proof. exact: (@le_irrelevance m.+1). Qed. (* Comparison predicates. *) CoInductive leq_xor_gtn m n : bool -> bool -> Set := | LeqNotGtn of m <= n : leq_xor_gtn m n true false | GtnNotLeq of n < m : leq_xor_gtn m n false true. Lemma leqP m n : leq_xor_gtn m n (m <= n) (n < m). Proof. by rewrite ltnNge; case le_mn: (m <= n); constructor; rewrite // ltnNge le_mn. Qed. CoInductive ltn_xor_geq m n : bool -> bool -> Set := | LtnNotGeq of m < n : ltn_xor_geq m n false true | GeqNotLtn of n <= m : ltn_xor_geq m n true false. Lemma ltnP m n : ltn_xor_geq m n (n <= m) (m < n). Proof. by rewrite -(ltnS n); case: leqP; constructor. Qed. CoInductive eqn0_xor_gt0 n : bool -> bool -> Set := | Eq0NotPos of n = 0 : eqn0_xor_gt0 n true false | PosNotEq0 of n > 0 : eqn0_xor_gt0 n false true. Lemma posnP n : eqn0_xor_gt0 n (n == 0) (0 < n). Proof. by case: n; constructor. Qed. CoInductive compare_nat m n : bool -> bool -> bool -> Set := | CompareNatLt of m < n : compare_nat m n true false false | CompareNatGt of m > n : compare_nat m n false true false | CompareNatEq of m = n : compare_nat m n false false true. Lemma ltngtP m n : compare_nat m n (m < n) (n < m) (m == n). Proof. rewrite ltn_neqAle eqn_leq; case: ltnP; first by constructor. by rewrite leq_eqVlt orbC; case: leqP; constructor; first exact/eqnP. Qed. (* Monotonicity lemmas *) Lemma leq_add2l p m n : (p + m <= p + n) = (m <= n). Proof. by elim: p. Qed. Lemma ltn_add2l p m n : (p + m < p + n) = (m < n). Proof. by rewrite -addnS; exact: leq_add2l. Qed. Lemma leq_add2r p m n : (m + p <= n + p) = (m <= n). Proof. by rewrite -!(addnC p); exact: leq_add2l. Qed. Lemma ltn_add2r p m n : (m + p < n + p) = (m < n). Proof. exact: leq_add2r p m.+1 n. Qed. Lemma leq_add m1 m2 n1 n2 : m1 <= n1 -> m2 <= n2 -> m1 + m2 <= n1 + n2. Proof. by move=> le_mn1 le_mn2; rewrite (@leq_trans (m1 + n2)) ?leq_add2l ?leq_add2r. Qed. Lemma leq_addr m n : n <= n + m. Proof. by rewrite -{1}[n]addn0 leq_add2l. Qed. Lemma leq_addl m n : n <= m + n. Proof. by rewrite addnC leq_addr. Qed. Lemma ltn_addr m n p : m < n -> m < n + p. Proof. by move/leq_trans=> -> //; exact: leq_addr. Qed. Lemma ltn_addl m n p : m < n -> m < p + n. Proof. by move/leq_trans=> -> //; exact: leq_addl. Qed. Lemma addn_gt0 m n : (0 < m + n) = (0 < m) || (0 < n). Proof. by rewrite !lt0n -negb_and addn_eq0. Qed. Lemma subn_gt0 m n : (0 < n - m) = (m < n). Proof. by elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. Lemma subn_eq0 m n : (m - n == 0) = (m <= n). Proof. by []. Qed. Lemma leq_subLR m n p : (m - n <= p) = (m <= n + p). Proof. by rewrite -subn_eq0 -subnDA. Qed. Lemma leq_subr m n : n - m <= n. Proof. by rewrite leq_subLR leq_addl. Qed. Lemma subnKC m n : m <= n -> m + (n - m) = n. Proof. by elim: m n => [|m IHm] [|n] // /(IHm n) {2}<-. Qed. Lemma subnK m n : m <= n -> (n - m) + m = n. Proof. by rewrite addnC; exact: subnKC. Qed. Lemma addnBA m n p : p <= n -> m + (n - p) = m + n - p. Proof. by move=> le_pn; rewrite -{2}(subnK le_pn) addnA addnK. Qed. Lemma subnBA m n p : p <= n -> m - (n - p) = m + p - n. Proof. by move=> le_pn; rewrite -{2}(subnK le_pn) subnDr. Qed. Lemma subKn m n : m <= n -> n - (n - m) = m. Proof. by move/subnBA->; rewrite addKn. Qed. Lemma subSn m n : m <= n -> n.+1 - m = (n - m).+1. Proof. by rewrite -add1n => /addnBA <-. Qed. Lemma subnSK m n : m < n -> (n - m.+1).+1 = n - m. Proof. by move/subSn. Qed. Lemma leq_sub2r p m n : m <= n -> m - p <= n - p. Proof. by move=> le_mn; rewrite leq_subLR (leq_trans le_mn) // -leq_subLR. Qed. Lemma leq_sub2l p m n : m <= n -> p - n <= p - m. Proof. rewrite -(leq_add2r (p - m)) leq_subLR. by apply: leq_trans; rewrite -leq_subLR. Qed. Lemma leq_sub m1 m2 n1 n2 : m1 <= m2 -> n2 <= n1 -> m1 - n1 <= m2 - n2. Proof. by move/(leq_sub2r n1)=> le_m12 /(leq_sub2l m2); apply: leq_trans. Qed. Lemma ltn_sub2r p m n : p < n -> m < n -> m - p < n - p. Proof. by move/subnSK <-; exact: (@leq_sub2r p.+1). Qed. Lemma ltn_sub2l p m n : m < p -> m < n -> p - n < p - m. Proof. by move/subnSK <-; exact: leq_sub2l. Qed. Lemma ltn_subRL m n p : (n < p - m) = (m + n < p). Proof. by rewrite !ltnNge leq_subLR. Qed. (* Eliminating the idiom for structurally decreasing compare and subtract. *) Lemma subn_if_gt T m n F (E : T) : (if m.+1 - n is m'.+1 then F m' else E) = (if n <= m then F (m - n) else E). Proof. by case: leqP => [le_nm | /eqnP-> //]; rewrite -{1}(subnK le_nm) -addSn addnK. Qed. (* Max and min. *) Definition maxn m n := if m < n then n else m. Definition minn m n := if m < n then m else n. Lemma max0n : left_id 0 maxn. Proof. by case. Qed. Lemma maxn0 : right_id 0 maxn. Proof. by []. Qed. Lemma maxnC : commutative maxn. Proof. by move=> m n; rewrite /maxn; case ltngtP. Qed. Lemma maxnE m n : maxn m n = m + (n - m). Proof. by rewrite /maxn addnC; case: leqP => [/eqnP-> | /ltnW/subnK]. Qed. Lemma maxnAC : right_commutative maxn. Proof. by move=> m n p; rewrite !maxnE -!addnA !subnDA -!maxnE maxnC. Qed. Lemma maxnA : associative maxn. Proof. by move=> m n p; rewrite !(maxnC m) maxnAC. Qed. Lemma maxnCA : left_commutative maxn. Proof. by move=> m n p; rewrite !maxnA (maxnC m). Qed. Lemma maxnACA : interchange maxn maxn. Proof. by move=> m n p q; rewrite -!maxnA (maxnCA n). Qed. Lemma maxn_idPl {m n} : reflect (maxn m n = m) (m >= n). Proof. by rewrite -subn_eq0 -(eqn_add2l m) addn0 -maxnE; apply: eqP. Qed. Lemma maxn_idPr {m n} : reflect (maxn m n = n) (m <= n). Proof. by rewrite maxnC; apply: maxn_idPl. Qed. Lemma maxnn : idempotent maxn. Proof. by move=> n; apply/maxn_idPl. Qed. Lemma leq_max m n1 n2 : (m <= maxn n1 n2) = (m <= n1) || (m <= n2). Proof. without loss le_n21: n1 n2 / n2 <= n1. by case/orP: (leq_total n2 n1) => le_n12; last rewrite maxnC orbC; apply. by rewrite (maxn_idPl le_n21) orb_idr // => /leq_trans->. Qed. Lemma leq_maxl m n : m <= maxn m n. Proof. by rewrite leq_max leqnn. Qed. Lemma leq_maxr m n : n <= maxn m n. Proof. by rewrite maxnC leq_maxl. Qed. Lemma gtn_max m n1 n2 : (m > maxn n1 n2) = (m > n1) && (m > n2). Proof. by rewrite !ltnNge leq_max negb_or. Qed. Lemma geq_max m n1 n2 : (m >= maxn n1 n2) = (m >= n1) && (m >= n2). Proof. by rewrite -ltnS gtn_max. Qed. Lemma maxnSS m n : maxn m.+1 n.+1 = (maxn m n).+1. Proof. by rewrite !maxnE. Qed. Lemma addn_maxl : left_distributive addn maxn. Proof. by move=> m1 m2 n; rewrite !maxnE subnDr addnAC. Qed. Lemma addn_maxr : right_distributive addn maxn. Proof. by move=> m n1 n2; rewrite !(addnC m) addn_maxl. Qed. Lemma min0n : left_zero 0 minn. Proof. by case. Qed. Lemma minn0 : right_zero 0 minn. Proof. by []. Qed. Lemma minnC : commutative minn. Proof. by move=> m n; rewrite /minn; case ltngtP. Qed. Lemma addn_min_max m n : minn m n + maxn m n = m + n. Proof. by rewrite /minn /maxn; case: ltngtP => // [_|->] //; exact: addnC. Qed. Lemma minnE m n : minn m n = m - (m - n). Proof. by rewrite -(subnDl n) -maxnE -addn_min_max addnK minnC. Qed. Lemma minnAC : right_commutative minn. Proof. by move=> m n p; rewrite !minnE -subnDA subnAC -maxnE maxnC maxnE subnAC subnDA. Qed. Lemma minnA : associative minn. Proof. by move=> m n p; rewrite minnC minnAC (minnC n). Qed. Lemma minnCA : left_commutative minn. Proof. by move=> m n p; rewrite !minnA (minnC n). Qed. Lemma minnACA : interchange minn minn. Proof. by move=> m n p q; rewrite -!minnA (minnCA n). Qed. Lemma minn_idPl {m n} : reflect (minn m n = m) (m <= n). Proof. rewrite (sameP maxn_idPr eqP) -(eqn_add2l m) eq_sym -addn_min_max eqn_add2r. exact: eqP. Qed. Lemma minn_idPr {m n} : reflect (minn m n = n) (m >= n). Proof. by rewrite minnC; apply: minn_idPl. Qed. Lemma minnn : idempotent minn. Proof. by move=> n; apply/minn_idPl. Qed. Lemma leq_min m n1 n2 : (m <= minn n1 n2) = (m <= n1) && (m <= n2). Proof. wlog le_n21: n1 n2 / n2 <= n1. by case/orP: (leq_total n2 n1) => ?; last rewrite minnC andbC; auto. by rewrite /minn ltnNge le_n21 /= andbC; case: leqP => // /leq_trans->. Qed. Lemma gtn_min m n1 n2 : (m > minn n1 n2) = (m > n1) || (m > n2). Proof. by rewrite !ltnNge leq_min negb_and. Qed. Lemma geq_min m n1 n2 : (m >= minn n1 n2) = (m >= n1) || (m >= n2). Proof. by rewrite -ltnS gtn_min. Qed. Lemma geq_minl m n : minn m n <= m. Proof. by rewrite geq_min leqnn. Qed. Lemma geq_minr m n : minn m n <= n. Proof. by rewrite minnC geq_minl. Qed. Lemma addn_minr : right_distributive addn minn. Proof. by move=> m1 m2 n; rewrite !minnE subnDl addnBA ?leq_subr. Qed. Lemma addn_minl : left_distributive addn minn. Proof. by move=> m1 m2 n; rewrite -!(addnC n) addn_minr. Qed. Lemma minnSS m n : minn m.+1 n.+1 = (minn m n).+1. Proof. by rewrite -(addn_minr 1). Qed. (* Quasi-cancellation (really, absorption) lemmas *) Lemma maxnK m n : minn (maxn m n) m = m. Proof. exact/minn_idPr/leq_maxl. Qed. Lemma maxKn m n : minn n (maxn m n) = n. Proof. exact/minn_idPl/leq_maxr. Qed. Lemma minnK m n : maxn (minn m n) m = m. Proof. exact/maxn_idPr/geq_minl. Qed. Lemma minKn m n : maxn n (minn m n) = n. Proof. exact/maxn_idPl/geq_minr. Qed. (* Distributivity. *) Lemma maxn_minl : left_distributive maxn minn. Proof. move=> m1 m2 n; wlog le_m21: m1 m2 / m2 <= m1. move=> IH; case/orP: (leq_total m2 m1) => /IH //. by rewrite minnC [in R in _ = R]minnC. rewrite (minn_idPr le_m21); apply/esym/minn_idPr. by rewrite geq_max leq_maxr leq_max le_m21. Qed. Lemma maxn_minr : right_distributive maxn minn. Proof. by move=> m n1 n2; rewrite !(maxnC m) maxn_minl. Qed. Lemma minn_maxl : left_distributive minn maxn. Proof. by move=> m1 m2 n; rewrite maxn_minr !maxn_minl -minnA maxnn (maxnC _ n) !maxnK. Qed. Lemma minn_maxr : right_distributive minn maxn. Proof. by move=> m n1 n2; rewrite !(minnC m) minn_maxl. Qed. (* Getting a concrete value from an abstract existence proof. *) Section ExMinn. Variable P : pred nat. Hypothesis exP : exists n, P n. Inductive acc_nat i : Prop := AccNat0 of P i | AccNatS of acc_nat i.+1. Lemma find_ex_minn : {m | P m & forall n, P n -> n >= m}. Proof. have: forall n, P n -> n >= 0 by []. have: acc_nat 0. case exP => n; rewrite -(addn0 n); elim: n 0 => [|n IHn] j; first by left. rewrite addSnnS; right; exact: IHn. move: 0; fix 2 => m IHm m_lb; case Pm: (P m); first by exists m. apply: find_ex_minn m.+1 _ _ => [|n Pn]; first by case: IHm; rewrite ?Pm. by rewrite ltn_neqAle m_lb //; case: eqP Pm => // -> /idP[]. Qed. Definition ex_minn := s2val find_ex_minn. Inductive ex_minn_spec : nat -> Type := ExMinnSpec m of P m & (forall n, P n -> n >= m) : ex_minn_spec m. Lemma ex_minnP : ex_minn_spec ex_minn. Proof. by rewrite /ex_minn; case: find_ex_minn. Qed. End ExMinn. Section ExMaxn. Variables (P : pred nat) (m : nat). Hypotheses (exP : exists i, P i) (ubP : forall i, P i -> i <= m). Lemma ex_maxn_subproof : exists i, P (m - i). Proof. by case: exP => i Pi; exists (m - i); rewrite subKn ?ubP. Qed. Definition ex_maxn := m - ex_minn ex_maxn_subproof. CoInductive ex_maxn_spec : nat -> Type := ExMaxnSpec i of P i & (forall j, P j -> j <= i) : ex_maxn_spec i. Lemma ex_maxnP : ex_maxn_spec ex_maxn. Proof. rewrite /ex_maxn; case: ex_minnP => i Pmi min_i; split=> // j Pj. have le_i_mj: i <= m - j by rewrite min_i // subKn // ubP. rewrite -subn_eq0 subnBA ?(leq_trans le_i_mj) ?leq_subr //. by rewrite addnC -subnBA ?ubP. Qed. End ExMaxn. Lemma eq_ex_minn P Q exP exQ : P =1 Q -> @ex_minn P exP = @ex_minn Q exQ. Proof. move=> eqPQ; case: ex_minnP => m1 Pm1 m1_lb; case: ex_minnP => m2 Pm2 m2_lb. by apply/eqP; rewrite eqn_leq m1_lb (m2_lb, eqPQ) // -eqPQ. Qed. Lemma eq_ex_maxn (P Q : pred nat) m n exP ubP exQ ubQ : P =1 Q -> @ex_maxn P m exP ubP = @ex_maxn Q n exQ ubQ. Proof. move=> eqPQ; case: ex_maxnP => i Pi max_i; case: ex_maxnP => j Pj max_j. by apply/eqP; rewrite eqn_leq max_i ?eqPQ // max_j -?eqPQ. Qed. Section Iteration. Variable T : Type. Implicit Types m n : nat. Implicit Types x y : T. Definition iter n f x := let fix loop m := if m is i.+1 then f (loop i) else x in loop n. Definition iteri n f x := let fix loop m := if m is i.+1 then f i (loop i) else x in loop n. Definition iterop n op x := let f i y := if i is 0 then x else op x y in iteri n f. Lemma iterSr n f x : iter n.+1 f x = iter n f (f x). Proof. by elim: n => //= n <-. Qed. Lemma iterS n f x : iter n.+1 f x = f (iter n f x). Proof. by []. Qed. Lemma iter_add n m f x : iter (n + m) f x = iter n f (iter m f x). Proof. by elim: n => //= n ->. Qed. Lemma iteriS n f x : iteri n.+1 f x = f n (iteri n f x). Proof. by []. Qed. Lemma iteropS idx n op x : iterop n.+1 op x idx = iter n (op x) x. Proof. by elim: n => //= n ->. Qed. Lemma eq_iter f f' : f =1 f' -> forall n, iter n f =1 iter n f'. Proof. by move=> eq_f n x; elim: n => //= n ->; rewrite eq_f. Qed. Lemma eq_iteri f f' : f =2 f' -> forall n, iteri n f =1 iteri n f'. Proof. by move=> eq_f n x; elim: n => //= n ->; rewrite eq_f. Qed. Lemma eq_iterop n op op' : op =2 op' -> iterop n op =2 iterop n op'. Proof. by move=> eq_op x; apply: eq_iteri; case. Qed. End Iteration. Lemma iter_succn m n : iter n succn m = m + n. Proof. by elim: n => //= n ->. Qed. Lemma iter_succn_0 n : iter n succn 0 = n. Proof. exact: iter_succn. Qed. Lemma iter_predn m n : iter n predn m = m - n. Proof. by elim: n m => /= [|n IHn] m; rewrite ?subn0 // IHn subnS. Qed. (* Multiplication. *) Definition muln_rec := mult. Notation "m * n" := (muln_rec m n) : nat_rec_scope. Definition muln := nosimpl muln_rec. Notation "m * n" := (muln m n) : nat_scope. Lemma multE : mult = muln. Proof. by []. Qed. Lemma mulnE : muln = muln_rec. Proof. by []. Qed. Lemma mul0n : left_zero 0 muln. Proof. by []. Qed. Lemma muln0 : right_zero 0 muln. Proof. by elim. Qed. Lemma mul1n : left_id 1 muln. Proof. exact: addn0. Qed. Lemma mulSn m n : m.+1 * n = n + m * n. Proof. by []. Qed. Lemma mulSnr m n : m.+1 * n = m * n + n. Proof. exact: addnC. Qed. Lemma mulnS m n : m * n.+1 = m + m * n. Proof. by elim: m => // m; rewrite !mulSn !addSn addnCA => ->. Qed. Lemma mulnSr m n : m * n.+1 = m * n + m. Proof. by rewrite addnC mulnS. Qed. Lemma iter_addn m n p : iter n (addn m) p = m * n + p. Proof. by elim: n => /= [|n ->]; rewrite ?muln0 // mulnS addnA. Qed. Lemma iter_addn_0 m n : iter n (addn m) 0 = m * n. Proof. by rewrite iter_addn addn0. Qed. Lemma muln1 : right_id 1 muln. Proof. by move=> n; rewrite mulnSr muln0. Qed. Lemma mulnC : commutative muln. Proof. by move=> m n; elim: m => [|m]; rewrite (muln0, mulnS) // mulSn => ->. Qed. Lemma mulnDl : left_distributive muln addn. Proof. by move=> m1 m2 n; elim: m1 => //= m1 IHm; rewrite -addnA -IHm. Qed. Lemma mulnDr : right_distributive muln addn. Proof. by move=> m n1 n2; rewrite !(mulnC m) mulnDl. Qed. Lemma mulnBl : left_distributive muln subn. Proof. move=> m n [|p]; first by rewrite !muln0. by elim: m n => // [m IHm] [|n] //; rewrite mulSn subnDl -IHm. Qed. Lemma mulnBr : right_distributive muln subn. Proof. by move=> m n p; rewrite !(mulnC m) mulnBl. Qed. Lemma mulnA : associative muln. Proof. by move=> m n p; elim: m => //= m; rewrite mulSn mulnDl => ->. Qed. Lemma mulnCA : left_commutative muln. Proof. by move=> m n1 n2; rewrite !mulnA (mulnC m). Qed. Lemma mulnAC : right_commutative muln. Proof. by move=> m n p; rewrite -!mulnA (mulnC n). Qed. Lemma mulnACA : interchange muln muln. Proof. by move=> m n p q; rewrite -!mulnA (mulnCA n). Qed. Lemma muln_eq0 m n : (m * n == 0) = (m == 0) || (n == 0). Proof. by case: m n => // m [|n] //=; rewrite muln0. Qed. Lemma muln_eq1 m n : (m * n == 1) = (m == 1) && (n == 1). Proof. by case: m n => [|[|m]] [|[|n]] //; rewrite muln0. Qed. Lemma muln_gt0 m n : (0 < m * n) = (0 < m) && (0 < n). Proof. by case: m n => // m [|n] //=; rewrite muln0. Qed. Lemma leq_pmull m n : n > 0 -> m <= n * m. Proof. by move/prednK <-; exact: leq_addr. Qed. Lemma leq_pmulr m n : n > 0 -> m <= m * n. Proof. by move/leq_pmull; rewrite mulnC. Qed. Lemma leq_mul2l m n1 n2 : (m * n1 <= m * n2) = (m == 0) || (n1 <= n2). Proof. by rewrite {1}/leq -mulnBr muln_eq0. Qed. Lemma leq_mul2r m n1 n2 : (n1 * m <= n2 * m) = (m == 0) || (n1 <= n2). Proof. by rewrite -!(mulnC m) leq_mul2l. Qed. Lemma leq_mul m1 m2 n1 n2 : m1 <= n1 -> m2 <= n2 -> m1 * m2 <= n1 * n2. Proof. move=> le_mn1 le_mn2; apply (@leq_trans (m1 * n2)). by rewrite leq_mul2l le_mn2 orbT. by rewrite leq_mul2r le_mn1 orbT. Qed. Lemma eqn_mul2l m n1 n2 : (m * n1 == m * n2) = (m == 0) || (n1 == n2). Proof. by rewrite eqn_leq !leq_mul2l -orb_andr -eqn_leq. Qed. Lemma eqn_mul2r m n1 n2 : (n1 * m == n2 * m) = (m == 0) || (n1 == n2). Proof. by rewrite eqn_leq !leq_mul2r -orb_andr -eqn_leq. Qed. Lemma leq_pmul2l m n1 n2 : 0 < m -> (m * n1 <= m * n2) = (n1 <= n2). Proof. by move/prednK=> <-; rewrite leq_mul2l. Qed. Implicit Arguments leq_pmul2l [m n1 n2]. Lemma leq_pmul2r m n1 n2 : 0 < m -> (n1 * m <= n2 * m) = (n1 <= n2). Proof. by move/prednK <-; rewrite leq_mul2r. Qed. Implicit Arguments leq_pmul2r [m n1 n2]. Lemma eqn_pmul2l m n1 n2 : 0 < m -> (m * n1 == m * n2) = (n1 == n2). Proof. by move/prednK <-; rewrite eqn_mul2l. Qed. Implicit Arguments eqn_pmul2l [m n1 n2]. Lemma eqn_pmul2r m n1 n2 : 0 < m -> (n1 * m == n2 * m) = (n1 == n2). Proof. by move/prednK <-; rewrite eqn_mul2r. Qed. Implicit Arguments eqn_pmul2r [m n1 n2]. Lemma ltn_mul2l m n1 n2 : (m * n1 < m * n2) = (0 < m) && (n1 < n2). Proof. by rewrite lt0n !ltnNge leq_mul2l negb_or. Qed. Lemma ltn_mul2r m n1 n2 : (n1 * m < n2 * m) = (0 < m) && (n1 < n2). Proof. by rewrite lt0n !ltnNge leq_mul2r negb_or. Qed. Lemma ltn_pmul2l m n1 n2 : 0 < m -> (m * n1 < m * n2) = (n1 < n2). Proof. by move/prednK <-; rewrite ltn_mul2l. Qed. Implicit Arguments ltn_pmul2l [m n1 n2]. Lemma ltn_pmul2r m n1 n2 : 0 < m -> (n1 * m < n2 * m) = (n1 < n2). Proof. by move/prednK <-; rewrite ltn_mul2r. Qed. Implicit Arguments ltn_pmul2r [m n1 n2]. Lemma ltn_Pmull m n : 1 < n -> 0 < m -> m < n * m. Proof. by move=> lt1n m_gt0; rewrite -{1}[m]mul1n ltn_pmul2r. Qed. Lemma ltn_Pmulr m n : 1 < n -> 0 < m -> m < m * n. Proof. by move=> lt1n m_gt0; rewrite mulnC ltn_Pmull. Qed. Lemma ltn_mul m1 m2 n1 n2 : m1 < n1 -> m2 < n2 -> m1 * m2 < n1 * n2. Proof. move=> lt_mn1 lt_mn2; apply (@leq_ltn_trans (m1 * n2)). by rewrite leq_mul2l orbC ltnW. by rewrite ltn_pmul2r // (leq_trans _ lt_mn2). Qed. Lemma maxn_mulr : right_distributive muln maxn. Proof. by case=> // m n1 n2; rewrite /maxn (fun_if (muln _)) ltn_pmul2l. Qed. Lemma maxn_mull : left_distributive muln maxn. Proof. by move=> m1 m2 n; rewrite -!(mulnC n) maxn_mulr. Qed. Lemma minn_mulr : right_distributive muln minn. Proof. by case=> // m n1 n2; rewrite /minn (fun_if (muln _)) ltn_pmul2l. Qed. Lemma minn_mull : left_distributive muln minn. Proof. by move=> m1 m2 n; rewrite -!(mulnC n) minn_mulr. Qed. (* Exponentiation. *) Definition expn_rec m n := iterop n muln m 1. Notation "m ^ n" := (expn_rec m n) : nat_rec_scope. Definition expn := nosimpl expn_rec. Notation "m ^ n" := (expn m n) : nat_scope. Lemma expnE : expn = expn_rec. Proof. by []. Qed. Lemma expn0 m : m ^ 0 = 1. Proof. by []. Qed. Lemma expn1 m : m ^ 1 = m. Proof. by []. Qed. Lemma expnS m n : m ^ n.+1 = m * m ^ n. Proof. by case: n; rewrite ?muln1. Qed. Lemma expnSr m n : m ^ n.+1 = m ^ n * m. Proof. by rewrite mulnC expnS. Qed. Lemma iter_muln m n p : iter n (muln m) p = m ^ n * p. Proof. by elim: n => /= [|n ->]; rewrite ?mul1n // expnS mulnA. Qed. Lemma iter_muln_1 m n : iter n (muln m) 1 = m ^ n. Proof. by rewrite iter_muln muln1. Qed. Lemma exp0n n : 0 < n -> 0 ^ n = 0. Proof. by case: n => [|[]]. Qed. Lemma exp1n n : 1 ^ n = 1. Proof. by elim: n => // n; rewrite expnS mul1n. Qed. Lemma expnD m n1 n2 : m ^ (n1 + n2) = m ^ n1 * m ^ n2. Proof. by elim: n1 => [|n1 IHn]; rewrite !(mul1n, expnS) // IHn mulnA. Qed. Lemma expnMn m1 m2 n : (m1 * m2) ^ n = m1 ^ n * m2 ^ n. Proof. by elim: n => // n IHn; rewrite !expnS IHn -!mulnA (mulnCA m2). Qed. Lemma expnM m n1 n2 : m ^ (n1 * n2) = (m ^ n1) ^ n2. Proof. elim: n1 => [|n1 IHn]; first by rewrite exp1n. by rewrite expnD expnS expnMn IHn. Qed. Lemma expnAC m n1 n2 : (m ^ n1) ^ n2 = (m ^ n2) ^ n1. Proof. by rewrite -!expnM mulnC. Qed. Lemma expn_gt0 m n : (0 < m ^ n) = (0 < m) || (n == 0). Proof. by case: m => [|m]; elim: n => //= n IHn; rewrite expnS // addn_gt0 IHn. Qed. Lemma expn_eq0 m e : (m ^ e == 0) = (m == 0) && (e > 0). Proof. by rewrite !eqn0Ngt expn_gt0 negb_or -lt0n. Qed. Lemma ltn_expl m n : 1 < m -> n < m ^ n. Proof. move=> m_gt1; elim: n => //= n; rewrite -(leq_pmul2l (ltnW m_gt1)) expnS. by apply: leq_trans; exact: ltn_Pmull. Qed. Lemma leq_exp2l m n1 n2 : 1 < m -> (m ^ n1 <= m ^ n2) = (n1 <= n2). Proof. move=> m_gt1; elim: n1 n2 => [|n1 IHn] [|n2] //; last 1 first. - by rewrite !expnS leq_pmul2l ?IHn // ltnW. - by rewrite expn_gt0 ltnW. by rewrite leqNgt (leq_trans m_gt1) // expnS leq_pmulr // expn_gt0 ltnW. Qed. Lemma ltn_exp2l m n1 n2 : 1 < m -> (m ^ n1 < m ^ n2) = (n1 < n2). Proof. by move=> m_gt1; rewrite !ltnNge leq_exp2l. Qed. Lemma eqn_exp2l m n1 n2 : 1 < m -> (m ^ n1 == m ^ n2) = (n1 == n2). Proof. by move=> m_gt1; rewrite !eqn_leq !leq_exp2l. Qed. Lemma expnI m : 1 < m -> injective (expn m). Proof. by move=> m_gt1 e1 e2 /eqP; rewrite eqn_exp2l // => /eqP. Qed. Lemma leq_pexp2l m n1 n2 : 0 < m -> n1 <= n2 -> m ^ n1 <= m ^ n2. Proof. by case: m => [|[|m]] // _; [rewrite !exp1n | rewrite leq_exp2l]. Qed. Lemma ltn_pexp2l m n1 n2 : 0 < m -> m ^ n1 < m ^ n2 -> n1 < n2. Proof. by case: m => [|[|m]] // _; [rewrite !exp1n | rewrite ltn_exp2l]. Qed. Lemma ltn_exp2r m n e : e > 0 -> (m ^ e < n ^ e) = (m < n). Proof. move=> e_gt0; apply/idP/idP=> [|ltmn]. rewrite !ltnNge; apply: contra => lemn. by elim: e {e_gt0} => // e IHe; rewrite !expnS leq_mul. by elim: e e_gt0 => // [[|e] IHe] _; rewrite ?expn1 // ltn_mul // IHe. Qed. Lemma leq_exp2r m n e : e > 0 -> (m ^ e <= n ^ e) = (m <= n). Proof. by move=> e_gt0; rewrite leqNgt ltn_exp2r // -leqNgt. Qed. Lemma eqn_exp2r m n e : e > 0 -> (m ^ e == n ^ e) = (m == n). Proof. by move=> e_gt0; rewrite !eqn_leq !leq_exp2r. Qed. Lemma expIn e : e > 0 -> injective (expn^~ e). Proof. by move=> e_gt1 m n /eqP; rewrite eqn_exp2r // => /eqP. Qed. (* Factorial. *) Fixpoint fact_rec n := if n is n'.+1 then n * fact_rec n' else 1. Definition factorial := nosimpl fact_rec. Notation "n `!" := (factorial n) (at level 2, format "n `!") : nat_scope. Lemma factE : factorial = fact_rec. Proof. by []. Qed. Lemma fact0 : 0`! = 1. Proof. by []. Qed. Lemma factS n : (n.+1)`! = n.+1 * n`!. Proof. by []. Qed. Lemma fact_gt0 n : n`! > 0. Proof. by elim: n => //= n IHn; rewrite muln_gt0. Qed. (* Parity and bits. *) Coercion nat_of_bool (b : bool) := if b then 1 else 0. Lemma leq_b1 (b : bool) : b <= 1. Proof. by case: b. Qed. Lemma addn_negb (b : bool) : ~~ b + b = 1. Proof. by case: b. Qed. Lemma eqb0 (b : bool) : (b == 0 :> nat) = ~~ b. Proof. by case: b. Qed. Lemma eqb1 (b : bool) : (b == 1 :> nat) = b. Proof. by case: b. Qed. Lemma lt0b (b : bool) : (b > 0) = b. Proof. by case: b. Qed. Lemma sub1b (b : bool) : 1 - b = ~~ b. Proof. by case: b. Qed. Lemma mulnb (b1 b2 : bool) : b1 * b2 = b1 && b2. Proof. by case: b1; case: b2. Qed. Lemma mulnbl (b : bool) n : b * n = (if b then n else 0). Proof. by case: b; rewrite ?mul1n. Qed. Lemma mulnbr (b : bool) n : n * b = (if b then n else 0). Proof. by rewrite mulnC mulnbl. Qed. Fixpoint odd n := if n is n'.+1 then ~~ odd n' else false. Lemma oddb (b : bool) : odd b = b. Proof. by case: b. Qed. Lemma odd_add m n : odd (m + n) = odd m (+) odd n. Proof. by elim: m => [|m IHn] //=; rewrite -addTb IHn addbA addTb. Qed. Lemma odd_sub m n : n <= m -> odd (m - n) = odd m (+) odd n. Proof. by move=> le_nm; apply: (@canRL bool) (addbK _) _; rewrite -odd_add subnK. Qed. Lemma odd_opp i m : odd m = false -> i < m -> odd (m - i) = odd i. Proof. by move=> oddm lt_im; rewrite (odd_sub (ltnW lt_im)) oddm. Qed. Lemma odd_mul m n : odd (m * n) = odd m && odd n. Proof. by elim: m => //= m IHm; rewrite odd_add -addTb andb_addl -IHm. Qed. Lemma odd_exp m n : odd (m ^ n) = (n == 0) || odd m. Proof. by elim: n => // n IHn; rewrite expnS odd_mul {}IHn orbC; case odd. Qed. (* Doubling. *) Fixpoint double_rec n := if n is n'.+1 then n'.*2%Nrec.+2 else 0 where "n .*2" := (double_rec n) : nat_rec_scope. Definition double := nosimpl double_rec. Notation "n .*2" := (double n) : nat_scope. Lemma doubleE : double = double_rec. Proof. by []. Qed. Lemma double0 : 0.*2 = 0. Proof. by []. Qed. Lemma doubleS n : n.+1.*2 = n.*2.+2. Proof. by []. Qed. Lemma addnn n : n + n = n.*2. Proof. by apply: eqP; elim: n => // n IHn; rewrite addnS. Qed. Lemma mul2n m : 2 * m = m.*2. Proof. by rewrite mulSn mul1n addnn. Qed. Lemma muln2 m : m * 2 = m.*2. Proof. by rewrite mulnC mul2n. Qed. Lemma doubleD m n : (m + n).*2 = m.*2 + n.*2. Proof. by rewrite -!addnn -!addnA (addnCA n). Qed. Lemma doubleB m n : (m - n).*2 = m.*2 - n.*2. Proof. elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. Lemma leq_double m n : (m.*2 <= n.*2) = (m <= n). Proof. by rewrite /leq -doubleB; case (m - n). Qed. Lemma ltn_double m n : (m.*2 < n.*2) = (m < n). Proof. by rewrite 2!ltnNge leq_double. Qed. Lemma ltn_Sdouble m n : (m.*2.+1 < n.*2) = (m < n). Proof. by rewrite -doubleS leq_double. Qed. Lemma leq_Sdouble m n : (m.*2 <= n.*2.+1) = (m <= n). Proof. by rewrite leqNgt ltn_Sdouble -leqNgt. Qed. Lemma odd_double n : odd n.*2 = false. Proof. by rewrite -addnn odd_add addbb. Qed. Lemma double_gt0 n : (0 < n.*2) = (0 < n). Proof. by case: n. Qed. Lemma double_eq0 n : (n.*2 == 0) = (n == 0). Proof. by case: n. Qed. Lemma doubleMl m n : (m * n).*2 = m.*2 * n. Proof. by rewrite -!mul2n mulnA. Qed. Lemma doubleMr m n : (m * n).*2 = m * n.*2. Proof. by rewrite -!muln2 mulnA. Qed. (* Halving. *) Fixpoint half (n : nat) : nat := if n is n'.+1 then uphalf n' else n with uphalf (n : nat) : nat := if n is n'.+1 then n'./2.+1 else n where "n ./2" := (half n) : nat_scope. Lemma doubleK : cancel double half. Proof. by elim=> //= n ->. Qed. Definition half_double := doubleK. Definition double_inj := can_inj doubleK. Lemma uphalf_double n : uphalf n.*2 = n. Proof. by elim: n => //= n ->. Qed. Lemma uphalf_half n : uphalf n = odd n + n./2. Proof. by elim: n => //= n ->; rewrite addnA addn_negb. Qed. Lemma odd_double_half n : odd n + n./2.*2 = n. Proof. by elim: n => //= n {3}<-; rewrite uphalf_half doubleD; case (odd n). Qed. Lemma half_bit_double n (b : bool) : (b + n.*2)./2 = n. Proof. by case: b; rewrite /= (half_double, uphalf_double). Qed. Lemma halfD m n : (m + n)./2 = (odd m && odd n) + (m./2 + n./2). Proof. rewrite -{1}[n]odd_double_half addnCA -{1}[m]odd_double_half -addnA -doubleD. by do 2!case: odd; rewrite /= ?add0n ?half_double ?uphalf_double. Qed. Lemma half_leq m n : m <= n -> m./2 <= n./2. Proof. by move/subnK <-; rewrite halfD addnA leq_addl. Qed. Lemma half_gt0 n : (0 < n./2) = (1 < n). Proof. by case: n => [|[]]. Qed. Lemma odd_geq m n : odd n -> (m <= n) = (m./2.*2 <= n). Proof. move=> odd_n; rewrite -{1}[m]odd_double_half -[n]odd_double_half odd_n. by case: (odd m); rewrite // leq_Sdouble ltnS leq_double. Qed. Lemma odd_ltn m n : odd n -> (n < m) = (n < m./2.*2). Proof. by move=> odd_n; rewrite !ltnNge odd_geq. Qed. Lemma odd_gt0 n : odd n -> n > 0. Proof. by case: n. Qed. Lemma odd_gt2 n : odd n -> n > 1 -> n > 2. Proof. by move=> odd_n n_gt1; rewrite odd_geq. Qed. (* Squares and square identities. *) Lemma mulnn m : m * m = m ^ 2. Proof. by rewrite !expnS muln1. Qed. Lemma sqrnD m n : (m + n) ^ 2 = m ^ 2 + n ^ 2 + 2 * (m * n). Proof. rewrite -!mulnn mul2n mulnDr !mulnDl (mulnC n) -!addnA. by congr (_ + _); rewrite addnA addnn addnC. Qed. Lemma sqrn_sub m n : n <= m -> (m - n) ^ 2 = m ^ 2 + n ^ 2 - 2 * (m * n). Proof. move/subnK=> def_m; rewrite -{2}def_m sqrnD -addnA addnAC. by rewrite -2!addnA addnn -mul2n -mulnDr -mulnDl def_m addnK. Qed. Lemma sqrnD_sub m n : n <= m -> (m + n) ^ 2 - 4 * (m * n) = (m - n) ^ 2. Proof. move=> le_nm; rewrite -[4]/(2 * 2) -mulnA mul2n -addnn subnDA. by rewrite sqrnD addnK sqrn_sub. Qed. Lemma subn_sqr m n : m ^ 2 - n ^ 2 = (m - n) * (m + n). Proof. by rewrite mulnBl !mulnDr addnC (mulnC m) subnDl !mulnn. Qed. Lemma ltn_sqr m n : (m ^ 2 < n ^ 2) = (m < n). Proof. by rewrite ltn_exp2r. Qed. Lemma leq_sqr m n : (m ^ 2 <= n ^ 2) = (m <= n). Proof. by rewrite leq_exp2r. Qed. Lemma sqrn_gt0 n : (0 < n ^ 2) = (0 < n). Proof. exact: (ltn_sqr 0). Qed. Lemma eqn_sqr m n : (m ^ 2 == n ^ 2) = (m == n). Proof. by rewrite eqn_exp2r. Qed. Lemma sqrn_inj : injective (expn ^~ 2). Proof. exact: expIn. Qed. (* Almost strict inequality: an inequality that is strict unless some *) (* specific condition holds, such as the Cauchy-Schwartz or the AGM *) (* inequality (we only prove the order-2 AGM here; the general one *) (* requires sequences). *) (* We formalize the concept as a rewrite multirule, that can be used *) (* both to rewrite the non-strict inequality to true, and the equality *) (* to the specific condition (for strict inequalities use the ltn_neqAle *) (* lemma); in addition, the conditional equality also coerces to a *) (* non-strict one. *) Definition leqif m n C := ((m <= n) * ((m == n) = C))%type. Notation "m <= n ?= 'iff' C" := (leqif m n C) : nat_scope. Coercion leq_of_leqif m n C (H : m <= n ?= iff C) := H.1 : m <= n. Lemma leqifP m n C : reflect (m <= n ?= iff C) (if C then m == n else m < n). Proof. rewrite ltn_neqAle; apply: (iffP idP) => [|lte]; last by rewrite !lte; case C. by case C => [/eqP-> | /andP[/negPf]]; split=> //; exact: eqxx. Qed. Lemma leqif_refl m C : reflect (m <= m ?= iff C) C. Proof. by apply: (iffP idP) => [-> | <-] //; split; rewrite ?eqxx. Qed. Lemma leqif_trans m1 m2 m3 C12 C23 : m1 <= m2 ?= iff C12 -> m2 <= m3 ?= iff C23 -> m1 <= m3 ?= iff C12 && C23. Proof. move=> ltm12 ltm23; apply/leqifP; rewrite -ltm12. case eqm12: (m1 == m2). by rewrite (eqP eqm12) ltn_neqAle !ltm23 andbT; case C23. by rewrite (@leq_trans m2) ?ltm23 // ltn_neqAle eqm12 ltm12. Qed. Lemma mono_leqif f : {mono f : m n / m <= n} -> forall m n C, (f m <= f n ?= iff C) = (m <= n ?= iff C). Proof. by move=> f_mono m n C; rewrite /leqif !eqn_leq !f_mono. Qed. Lemma leqif_geq m n : m <= n -> m <= n ?= iff (m >= n). Proof. by move=> lemn; split=> //; rewrite eqn_leq lemn. Qed. Lemma leqif_eq m n : m <= n -> m <= n ?= iff (m == n). Proof. by []. Qed. Lemma geq_leqif a b C : a <= b ?= iff C -> (b <= a) = C. Proof. by case=> le_ab; rewrite eqn_leq le_ab. Qed. Lemma ltn_leqif a b C : a <= b ?= iff C -> (a < b) = ~~ C. Proof. by move=> le_ab; rewrite ltnNge (geq_leqif le_ab). Qed. Lemma leqif_add m1 n1 C1 m2 n2 C2 : m1 <= n1 ?= iff C1 -> m2 <= n2 ?= iff C2 -> m1 + m2 <= n1 + n2 ?= iff C1 && C2. Proof. rewrite -(mono_leqif (leq_add2r m2)) -(mono_leqif (leq_add2l n1) m2). exact: leqif_trans. Qed. Lemma leqif_mul m1 n1 C1 m2 n2 C2 : m1 <= n1 ?= iff C1 -> m2 <= n2 ?= iff C2 -> m1 * m2 <= n1 * n2 ?= iff (n1 * n2 == 0) || (C1 && C2). Proof. move=> le1 le2; case: posnP => [n12_0 | ]. rewrite n12_0; move/eqP: n12_0 {le1 le2}le1.1 le2.1; rewrite muln_eq0. by case/orP=> /eqP->; case: m1 m2 => [|m1] [|m2] // _ _; rewrite ?muln0; exact/leqif_refl. rewrite muln_gt0 => /andP[n1_gt0 n2_gt0]. have [m2_0 | m2_gt0] := posnP m2. apply/leqifP; rewrite -le2 andbC eq_sym eqn_leq leqNgt m2_0 muln0. by rewrite muln_gt0 n1_gt0 n2_gt0. have mono_n1 := leq_pmul2l n1_gt0; have mono_m2 := leq_pmul2r m2_gt0. rewrite -(mono_leqif mono_m2) in le1; rewrite -(mono_leqif mono_n1) in le2. exact: leqif_trans le1 le2. Qed. Lemma nat_Cauchy m n : 2 * (m * n) <= m ^ 2 + n ^ 2 ?= iff (m == n). Proof. wlog le_nm: m n / n <= m. by case: (leqP m n); auto; rewrite eq_sym addnC (mulnC m); auto. apply/leqifP; case: ifP => [/eqP-> | ne_mn]; first by rewrite mulnn addnn mul2n. by rewrite -subn_gt0 -sqrn_sub // sqrn_gt0 subn_gt0 ltn_neqAle eq_sym ne_mn. Qed. Lemma nat_AGM2 m n : 4 * (m * n) <= (m + n) ^ 2 ?= iff (m == n). Proof. rewrite -[4]/(2 * 2) -mulnA mul2n -addnn sqrnD; apply/leqifP. by rewrite ltn_add2r eqn_add2r ltn_neqAle !nat_Cauchy; case: ifP => ->. Qed. (* Support for larger integers. The normal definitions of +, - and even *) (* IO are unsuitable for Peano integers larger than 2000 or so because *) (* they are not tail-recursive. We provide a workaround module, along *) (* with a rewrite multirule to change the tailrec operators to the *) (* normal ones. We handle IO via the NatBin module, but provide our *) (* own (more efficient) conversion functions. *) Module NatTrec. (* Usage: *) (* Import NatTrec. *) (* in section definining functions, rebinds all *) (* non-tail recursive operators. *) (* rewrite !trecE. *) (* in the correctness proof, restores operators *) Fixpoint add m n := if m is m'.+1 then m' + n.+1 else n where "n + m" := (add n m) : nat_scope. Fixpoint add_mul m n s := if m is m'.+1 then add_mul m' n (n + s) else s. Definition mul m n := if m is m'.+1 then add_mul m' n n else 0. Notation "n * m" := (mul n m) : nat_scope. Fixpoint mul_exp m n p := if n is n'.+1 then mul_exp m n' (m * p) else p. Definition exp m n := if n is n'.+1 then mul_exp m n' m else 1. Notation "n ^ m" := (exp n m) : nat_scope. Notation Local oddn := odd. Fixpoint odd n := if n is n'.+2 then odd n' else eqn n 1. Notation Local doublen := double. Definition double n := if n is n'.+1 then n' + n.+1 else 0. Notation "n .*2" := (double n) : nat_scope. Lemma addE : add =2 addn. Proof. by elim=> //= n IHn m; rewrite IHn addSnnS. Qed. Lemma doubleE : double =1 doublen. Proof. by case=> // n; rewrite -addnn -addE. Qed. Lemma add_mulE n m s : add_mul n m s = addn (muln n m) s. Proof. by elim: n => //= n IHn in m s *; rewrite IHn addE addnCA addnA. Qed. Lemma mulE : mul =2 muln. Proof. by case=> //= n m; rewrite add_mulE addnC. Qed. Lemma mul_expE m n p : mul_exp m n p = muln (expn m n) p. Proof. by elim: n => [|n IHn] in p *; rewrite ?mul1n //= expnS IHn mulE mulnCA mulnA. Qed. Lemma expE : exp =2 expn. Proof. by move=> m [|n] //=; rewrite mul_expE expnS mulnC. Qed. Lemma oddE : odd =1 oddn. Proof. move=> n; rewrite -{1}[n]odd_double_half addnC. by elim: n./2 => //=; case (oddn n). Qed. Definition trecE := (addE, (doubleE, oddE), (mulE, add_mulE, (expE, mul_expE))). End NatTrec. Notation natTrecE := NatTrec.trecE. Lemma eq_binP : Equality.axiom Ndec.Neqb. Proof. move=> p q; apply: (iffP idP) => [|<-]; last by case: p => //; elim. by case: q; case: p => //; elim=> [p IHp|p IHp|] [q|q|] //=; case/IHp=> ->. Qed. Canonical bin_nat_eqMixin := EqMixin eq_binP. Canonical bin_nat_eqType := Eval hnf in EqType N bin_nat_eqMixin. Section NumberInterpretation. Import BinPos. Section Trec. Import NatTrec. Fixpoint nat_of_pos p0 := match p0 with | xO p => (nat_of_pos p).*2 | xI p => (nat_of_pos p).*2.+1 | xH => 1 end. End Trec. Coercion Local nat_of_pos : positive >-> nat. Coercion nat_of_bin b := if b is Npos p then p : nat else 0. Fixpoint pos_of_nat n0 m0 := match n0, m0 with | n.+1, m.+2 => pos_of_nat n m | n.+1, 1 => xO (pos_of_nat n n) | n.+1, 0 => xI (pos_of_nat n n) | 0, _ => xH end. Definition bin_of_nat n0 := if n0 is n.+1 then Npos (pos_of_nat n n) else 0%num. Lemma bin_of_natK : cancel bin_of_nat nat_of_bin. Proof. have sub2nn n : n.*2 - n = n by rewrite -addnn addKn. case=> //= n; rewrite -{3}[n]sub2nn. by elim: n {2 4}n => // m IHm [|[|n]] //=; rewrite IHm // natTrecE sub2nn. Qed. Lemma nat_of_binK : cancel nat_of_bin bin_of_nat. Proof. case=> //=; elim=> //= p; case: (nat_of_pos p) => //= n [<-]. by rewrite natTrecE !addnS {2}addnn; elim: {1 3}n. by rewrite natTrecE addnS /= addnS {2}addnn; elim: {1 3}n. Qed. Lemma nat_of_succ_gt0 p : Psucc p = p.+1 :> nat. Proof. by elim: p => //= p ->; rewrite !natTrecE. Qed. Lemma nat_of_addn_gt0 p q : (p + q)%positive = p + q :> nat. Proof. apply: fst (Pplus_carry p q = (p + q).+1 :> nat) _. elim: p q => [p IHp|p IHp|] [q|q|] //=; rewrite !natTrecE //; by rewrite ?IHp ?nat_of_succ_gt0 ?(doubleS, doubleD, addn1, addnS). Qed. Lemma nat_of_add_bin b1 b2 : (b1 + b2)%num = b1 + b2 :> nat. Proof. case: b1 b2 => [|p] [|q] //=; exact: nat_of_addn_gt0. Qed. Lemma nat_of_mul_bin b1 b2 : (b1 * b2)%num = b1 * b2 :> nat. Proof. case: b1 b2 => [|p] [|q] //=; elim: p => [p IHp|p IHp|] /=; by rewrite ?(mul1n, nat_of_addn_gt0, mulSn) //= !natTrecE IHp doubleMl. Qed. Lemma nat_of_exp_bin n (b : N) : n ^ b = pow_N 1 muln n b. Proof. case: b => [|p] /=; first exact: expn0. by elim: p => //= p <-; rewrite natTrecE mulnn -expnM muln2 ?expnS. Qed. End NumberInterpretation. (* Big(ger) nat IO; usage: *) (* Num 1 072 399 *) (* to create large numbers for test cases *) (* Eval compute in [Num of some expression] *) (* to display the resut of an expression that *) (* returns a larger integer. *) Record number : Type := Num {bin_of_number :> N}. Definition extend_number (nn : number) m := Num (nn * 1000 + bin_of_nat m). Coercion extend_number : number >-> Funclass. Canonical number_subType := [newType for bin_of_number]. Definition number_eqMixin := Eval hnf in [eqMixin of number by <:]. Canonical number_eqType := Eval hnf in EqType number number_eqMixin. Notation "[ 'Num' 'of' e ]" := (Num (bin_of_nat e)) (at level 0, format "[ 'Num' 'of' e ]") : nat_scope. (* Interface to ring/ring_simplify tactics *) Lemma nat_semi_ring : semi_ring_theory 0 1 addn muln (@eq _). Proof. exact: mk_srt add0n addnC addnA mul1n mul0n mulnC mulnA mulnDl. Qed. Lemma nat_semi_morph : semi_morph 0 1 addn muln (@eq _) 0%num 1%num Nplus Nmult pred1 nat_of_bin. Proof. by move: nat_of_add_bin nat_of_mul_bin; split=> //= m n; move/eqP->. Qed. Lemma nat_power_theory : power_theory 1 muln (@eq _) nat_of_bin expn. Proof. split; exact: nat_of_exp_bin. Qed. (* Interface to the ring tactic machinery. *) Fixpoint pop_succn e := if e is e'.+1 then fun n => pop_succn e' n.+1 else id. Ltac pop_succn e := eval lazy beta iota delta [pop_succn] in (pop_succn e 1). Ltac nat_litteral e := match pop_succn e with | ?n.+1 => constr: (bin_of_nat n) | _ => NotConstant end. Ltac succn_to_add := match goal with | |- context G [?e.+1] => let x := fresh "NatLit0" in match pop_succn e with | ?n.+1 => pose x := n.+1; let G' := context G [x] in change G' | _ ?e' ?n => pose x := n; let G' := context G [x + e'] in change G' end; succn_to_add; rewrite {}/x | _ => idtac end. Add Ring nat_ring_ssr : nat_semi_ring (morphism nat_semi_morph, constants [nat_litteral], preprocess [succn_to_add], power_tac nat_power_theory [nat_litteral]). (* A congruence tactic, similar to the boolean one, along with an .+1/+ *) (* normalization tactic. *) Ltac nat_norm := succn_to_add; rewrite ?add0n ?addn0 -?addnA ?(addSn, addnS, add0n, addn0). Ltac nat_congr := first [ apply: (congr1 succn _) | apply: (congr1 predn _) | apply: (congr1 (addn _) _) | apply: (congr1 (subn _) _) | apply: (congr1 (addn^~ _) _) | match goal with |- (?X1 + ?X2 = ?X3) => symmetry; rewrite -1?(addnC X1) -?(addnCA X1); apply: (congr1 (addn X1) _); symmetry end ]. ssreflect-1.5/theories/ssreflect.mllib0000644000175000017500000000002612175453667017210 0ustar garesgaresSsrmatching Ssreflect ssreflect-1.5/theories/ssrfun.v0000644000175000017500000012157312175455021015701 0ustar garesgares(* (c) Copyright Microsoft Corporation and Inria. *) (* You may distribute this file under the terms of the CeCILL-B license *) Require Import ssreflect. (******************************************************************************) (* This file contains the basic definitions and notations for working with *) (* functions. The definitions provide for: *) (* *) (* - Pair projections: *) (* p.1 == first element of a pair *) (* p.2 == second element of a pair *) (* These notations also apply to p : P /\ Q, via an and >-> pair coercion. *) (* *) (* - Simplifying functions, beta-reduced by /= and simpl: *) (* [fun : T => E] == constant function from type T that returns E *) (* [fun x => E] == unary function *) (* [fun x : T => E] == unary function with explicit domain type *) (* [fun x y => E] == binary function *) (* [fun x y : T => E] == binary function with common domain type *) (* [fun (x : T) y => E] \ *) (* [fun (x : xT) (y : yT) => E] | == binary function with (some) explicit, *) (* [fun x (y : T) => E] / independent domain types for each argument *) (* *) (* - Partial functions using option type: *) (* oapp f d ox == if ox is Some x returns f x, d otherwise *) (* odflt d ox == if ox is Some x returns x, d otherwise *) (* obind f ox == if ox is Some x returns f x, None otherwise *) (* omap f ox == if ox is Some x returns Some (f x), None otherwise *) (* *) (* - Singleton types: *) (* all_equal_to x0 == x0 is the only value in its type, so any such value *) (* can be rewritten to x0. *) (* *) (* - A generic wrapper type: *) (* wrapped T == the inductive type with values Wrap x for x : T. *) (* unwrap w == the projection of w : wrapped T on T. *) (* wrap x == the canonical injection of x : T into wrapped T; it is *) (* equivalent to Wrap x, but is declared as a (default) *) (* Canonical Structure, which lets the Coq HO unification *) (* automatically expand x into unwrap (wrap x). The delta *) (* reduction of wrap x to Wrap can be exploited to *) (* introduce controlled nondeterminism in Canonical *) (* Structure inference, as in the implementation of *) (* the mxdirect predicate in matrix.v. *) (* *) (* - Sigma types: *) (* tag w == the i of w : {i : I & T i}. *) (* tagged w == the T i component of w : {i : I & T i}. *) (* Tagged T x == the {i : I & T i} with component x : T i. *) (* tag2 w == the i of w : {i : I & T i & U i}. *) (* tagged2 w == the T i component of w : {i : I & T i & U i}. *) (* tagged2' w == the U i component of w : {i : I & T i & U i}. *) (* Tagged2 T U x y == the {i : I & T i} with components x : T i and y : U i. *) (* sval u == the x of u : {x : T | P x}. *) (* s2val u == the x of u : {x : T | P x & Q x}. *) (* The properties of sval u, s2val u are given by lemmas svalP, s2valP, and *) (* s2valP'. We provide coercions sigT2 >-> sigT and sig2 >-> sig >-> sigT. *) (* A suite of lemmas (all_sig, ...) let us skolemize sig, sig2, sigT, sigT2 *) (* and pair, e.g., *) (* have /all_sig[f fP] (x : T): {y : U | P y} by ... *) (* yields an f : T -> U such that fP : forall x, P (f x). *) (* - Identity functions: *) (* id == NOTATION for the explicit identity function fun x => x. *) (* @id T == notation for the explicit identity at type T. *) (* idfun == an expression with a head constant, convertible to id; *) (* idfun x simplifies to x. *) (* @idfun T == the expression above, specialized to type T. *) (* phant_id x y == the function type phantom _ x -> phantom _ y. *) (* *** In addition to their casual use in functional programming, identity *) (* functions are often used to trigger static unification as part of the *) (* construction of dependent Records and Structures. For example, if we need *) (* a structure sT over a type T, we take as arguments T, sT, and a "dummy" *) (* function T -> sort sT: *) (* Definition foo T sT & T -> sort sT := ... *) (* We can avoid specifying sT directly by calling foo (@id T), or specify *) (* the call completely while still ensuring the consistency of T and sT, by *) (* calling @foo T sT idfun. The phant_id type allows us to extend this trick *) (* to non-Type canonical projections. It also allows us to sidestep *) (* dependent type constraints when building explicit records, e.g., given *) (* Record r := R {x; y : T(x)}. *) (* if we need to build an r from a given y0 while inferring some x0, such *) (* that y0 : T(x0), we pose *) (* Definition mk_r .. y .. (x := ...) y' & phant_id y y' := R x y'. *) (* Calling @mk_r .. y0 .. id will cause Coq to use y' := y0, while checking *) (* the dependent type constraint y0 : T(x0). *) (* *) (* - Extensional equality for functions and relations (i.e. functions of two *) (* arguments): *) (* f1 =1 f2 == f1 x is equal to f2 x for all x. *) (* f1 =1 f2 :> A == ... and f2 is explicitly typed. *) (* f1 =2 f2 == f1 x y is equal to f2 x y for all x y. *) (* f1 =2 f2 :> A == ... and f2 is explicitly typed. *) (* *) (* - Composition for total and partial functions: *) (* f^~ y == function f with second argument specialised to y, *) (* i.e., fun x => f x y *) (* CAVEAT: conditional (non-maximal) implicit arguments *) (* of f are NOT inserted in this context *) (* @^~ x == application at x, i.e., fun f => f x *) (* [eta f] == the explicit eta-expansion of f, i.e., fun x => f x *) (* CAVEAT: conditional (non-maximal) implicit arguments *) (* of f are NOT inserted in this context. *) (* fun=> v := the constant function fun _ => v. *) (* f1 \o f2 == composition of f1 and f2. *) (* Note: (f1 \o f2) x simplifies to f1 (f2 x). *) (* f1 \; f2 == categorical composition of f1 and f2. This expands to *) (* to f2 \o f1 and (f1 \; f2) x simplifies to f2 (f1 x). *) (* pcomp f1 f2 == composition of partial functions f1 and f2. *) (* *) (* - Reserved notation for various arithmetic and algebraic operations: *) (* e.[a1, ..., a_n] evaluation (e.g., polynomials). *) (* e`_i indexing (number list, integer pi-part). *) (* x^-1 inverse (group, field). *) (* x *+ n, x *- n integer multiplier (modules and rings). *) (* x ^+ n, x ^- n integer exponent (groups and rings). *) (* x *: A, A :* x external product (scaling/module product in rings, *) (* left/right cosets in groups). *) (* A :&: B intersection (of sets, groups, subspaces, ...). *) (* A :|: B, a |: B union, union with a singleton (of sets). *) (* A :\: B, A :\ b relative complement (of sets, subspaces, ...). *) (* <>, <[a]> generated group/subspace, generated cycle/line. *) (* 'C[x], 'C_A[x] point centralisers (in groups and F-algebras). *) (* 'C(A), 'C_B(A) centralisers (in groups and matrix and F_algebras). *) (* 'Z(A) centers (in groups and matrix and F-algebras). *) (* m %/ d, m %% d Euclidean division and remainder (nat, polynomials). *) (* d %| m Euclidean divisibility (nat, polynomial). *) (* m = n %[mod d] equality mod d (also defined for <>, ==, and !=). *) (* e^`(n) nth formal derivative (groups, polynomials). *) (* e^`() simple formal derivative (polynomials only). *) (* `|x| norm, absolute value, distance (rings, int, nat). *) (* x <= y ?= iff C x is less than y, and equal iff C holds (nat, rings). *) (* x <= y :> T, etc cast comparison (rings, all comparison operators). *) (* [rec a1, ..., an] standard shorthand for hidden recursor (see prime.v). *) (* The interpretation of these notations is not defined here, but the *) (* declarations help maintain consistency across the library. *) (* *) (* - Properties of functions: *) (* injective f <-> f is injective. *) (* cancel f g <-> g is a left inverse of f / f is a right inverse of g. *) (* pcancel f g <-> g is a left inverse of f where g is partial. *) (* ocancel f g <-> g is a left inverse of f where f is partial. *) (* bijective f <-> f is bijective (has a left and right inverse). *) (* involutive f <-> f is involutive. *) (* *) (* - Properties for operations. *) (* left_id e op <-> e is a left identity for op (e op x = x). *) (* right_id e op <-> e is a right identity for op (x op e = x). *) (* left_inverse e inv op <-> inv is a left inverse for op wrt identity e, *) (* i.e., (inv x) op x = e. *) (* right_inverse e inv op <-> inv is a right inverse for op wrt identity e *) (* i.e., x op (i x) = e. *) (* self_inverse e op <-> each x is its own op-inverse (x op x = e). *) (* idempotent op <-> op is idempotent for op (x op x = x). *) (* associative op <-> op is associative, i.e., *) (* x op (y op z) = (x op y) op z. *) (* commutative op <-> op is commutative (x op y = y op x). *) (* left_commutative op <-> op is left commutative, i.e., *) (* x op (y op z) = y op (x op z). *) (* right_commutative op <-> op is right commutative, i.e., *) (* (x op y) op z = (x op z) op y. *) (* left_zero z op <-> z is a left zero for op (z op x = z). *) (* right_zero z op <-> z is a right zero for op (x op z = z). *) (* left_distributive op1 op2 <-> op1 distributes over op2 to the left: *) (* (x op2 y) op1 z = (x op1 z) op2 (y op1 z). *) (* right_distributive op1 op2 <-> op distributes over add to the right: *) (* x op1 (y op2 z) = (x op1 z) op2 (x op1 z). *) (* interchange op1 op2 <-> op1 and op2 satisfy an interchange law: *) (* (x op2 y) op1 (z op2 t) = (x op1 z) op2 (y op1 t). *) (* Note that interchange op op is a commutativity property. *) (* left_injective op <-> op is injective in its left argument: *) (* x op y = z op y -> x = z. *) (* right_injective op <-> op is injective in its right argument: *) (* x op y = x op z -> y = z. *) (* left_loop inv op <-> op, inv obey the inverse loop left axiom: *) (* (inv x) op (x op y) = y for all x, y, i.e., *) (* op (inv x) is always a left inverse of op x *) (* rev_left_loop inv op <-> op, inv obey the inverse loop reverse left *) (* axiom: x op ((inv x) op y) = y, for all x, y. *) (* right_loop inv op <-> op, inv obey the inverse loop right axiom: *) (* (x op y) op (inv y) = x for all x, y. *) (* rev_right_loop inv op <-> op, inv obey the inverse loop reverse right *) (* axiom: (x op y) op (inv y) = x for all x, y. *) (* Note that familiar "cancellation" identities like x + y - y = x or *) (* x - y + x = x are respectively instances of right_loop and rev_right_loop *) (* The corresponding lemmas will use the K and NK/VK suffixes, respectively. *) (* *) (* - Morphisms for functions and relations: *) (* {morph f : x / a >-> r} <-> f is a morphism with respect to functions *) (* (fun x => a) and (fun x => r); if r == R[x], *) (* this states that f a = R[f x] for all x. *) (* {morph f : x / a} <-> f is a morphism with respect to the *) (* function expression (fun x => a). This is *) (* shorthand for {morph f : x / a >-> a}; note *) (* that the two instances of a are often *) (* interpreted at different types. *) (* {morph f : x y / a >-> r} <-> f is a morphism with respect to functions *) (* (fun x y => a) and (fun x y => r). *) (* {morph f : x y / a} <-> f is a morphism with respect to the *) (* function expression (fun x y => a). *) (* {homo f : x / a >-> r} <-> f is a homomorphism with respect to the *) (* predicates (fun x => a) and (fun x => r); *) (* if r == R[x], this states that a -> R[f x] *) (* for all x. *) (* {homo f : x / a} <-> f is a homomorphism with respect to the *) (* predicate expression (fun x => a). *) (* {homo f : x y / a >-> r} <-> f is a homomorphism with respect to the *) (* relations (fun x y => a) and (fun x y => r). *) (* {homo f : x y / a} <-> f is a homomorphism with respect to the *) (* relation expression (fun x y => a). *) (* {mono f : x / a >-> r} <-> f is monotone with respect to projectors *) (* (fun x => a) and (fun x => r); if r == R[x], *) (* this states that R[f x] = a for all x. *) (* {mono f : x / a} <-> f is monotone with respect to the projector *) (* expression (fun x => a). *) (* {mono f : x y / a >-> r} <-> f is monotone with respect to relators *) (* (fun x y => a) and (fun x y => r). *) (* {mono f : x y / a} <-> f is monotone with respect to the relator *) (* expression (fun x y => a). *) (* *) (* The file also contains some basic lemmas for the above concepts. *) (* Lemmas relative to cancellation laws use some abbreviated suffixes: *) (* K - a cancellation rule like esymK : cancel (@esym T x y) (@esym T y x). *) (* LR - a lemma moving an operation from the left hand side of a relation to *) (* the right hand side, like canLR: cancel g f -> x = g y -> f x = y. *) (* RL - a lemma moving an operation from the right to the left, e.g., canRL. *) (* Beware that the LR and RL orientations refer to an "apply" (back chaining) *) (* usage; when using the same lemmas with "have" or "move" (forward chaining) *) (* the directions will be reversed!. *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Delimit Scope fun_scope with FUN. Open Scope fun_scope. (* Notations for argument transpose *) Notation "f ^~ y" := (fun x => f x y) (at level 10, y at level 8, no associativity, format "f ^~ y") : fun_scope. Notation "@^~ x" := (fun f => f x) (at level 10, x at level 8, no associativity, format "@^~ x") : fun_scope. Delimit Scope pair_scope with PAIR. Open Scope pair_scope. (* Notations for pair/conjunction projections *) Notation "p .1" := (fst p) (at level 2, left associativity, format "p .1") : pair_scope. Notation "p .2" := (snd p) (at level 2, left associativity, format "p .2") : pair_scope. Coercion pair_of_and P Q (PandQ : P /\ Q) := (proj1 PandQ, proj2 PandQ). Definition all_pair I T U (w : forall i : I, T i * U i) := (fun i => (w i).1, fun i => (w i).2). (* Reserved notation for evaluation *) Reserved Notation "e .[ x ]" (at level 2, left associativity, format "e .[ x ]"). Reserved Notation "e .[ x1 , x2 , .. , xn ]" (at level 2, left associativity, format "e '[ ' .[ x1 , '/' x2 , '/' .. , '/' xn ] ']'"). (* Reserved notation for subscripting and superscripting *) Reserved Notation "s `_ i" (at level 3, i at level 2, left associativity, format "s `_ i"). Reserved Notation "x ^-1" (at level 3, left associativity, format "x ^-1"). (* Reserved notation for integer multipliers and exponents *) Reserved Notation "x *+ n" (at level 40, left associativity). Reserved Notation "x *- n" (at level 40, left associativity). Reserved Notation "x ^+ n" (at level 29, left associativity). Reserved Notation "x ^- n" (at level 29, left associativity). (* Reserved notation for external multiplication. *) Reserved Notation "x *: A" (at level 40). Reserved Notation "A :* x" (at level 40). (* Reserved notation for set-theretic operations. *) Reserved Notation "A :&: B" (at level 48, left associativity). Reserved Notation "A :|: B" (at level 52, left associativity). Reserved Notation "a |: A" (at level 52, left associativity). Reserved Notation "A :\: B" (at level 50, left associativity). Reserved Notation "A :\ b" (at level 50, left associativity). (* Reserved notation for generated structures *) Reserved Notation "<< A >>" (at level 0, format "<< A >>"). Reserved Notation "<[ a ] >" (at level 0, format "<[ a ] >"). (* Reserved notation for centralisers and centers. *) Reserved Notation "''C' [ x ]" (at level 8, format "''C' [ x ]"). Reserved Notation "''C_' A [ x ]" (at level 8, A at level 2, format "''C_' A [ x ]"). Reserved Notation "''C' ( A )" (at level 8, format "''C' ( A )"). Reserved Notation "''C_' B ( A )" (at level 8, B at level 2, format "''C_' B ( A )"). Reserved Notation "''Z' ( A )" (at level 8, format "''Z' ( A )"). (* Compatibility with group action centraliser notation. *) Reserved Notation "''C_' ( A ) [ x ]" (at level 8, only parsing). Reserved Notation "''C_' ( B ) ( A )" (at level 8, only parsing). (* Reserved notation for Euclidean division and divisibility. *) Reserved Notation "m %/ d" (at level 40, no associativity). Reserved Notation "m %% d" (at level 40, no associativity). Reserved Notation "m %| d" (at level 70, no associativity). Reserved Notation "m = n %[mod d ]" (at level 70, n at next level, format "'[hv ' m '/' = n '/' %[mod d ] ']'"). Reserved Notation "m == n %[mod d ]" (at level 70, n at next level, format "'[hv ' m '/' == n '/' %[mod d ] ']'"). Reserved Notation "m <> n %[mod d ]" (at level 70, n at next level, format "'[hv ' m '/' <> n '/' %[mod d ] ']'"). Reserved Notation "m != n %[mod d ]" (at level 70, n at next level, format "'[hv ' m '/' != n '/' %[mod d ] ']'"). (* Reserved notation for derivatives. *) Reserved Notation "a ^` ()" (at level 8, format "a ^` ()"). Reserved Notation "a ^` ( n )" (at level 8, format "a ^` ( n )"). (* Reserved notation for absolute value. *) Reserved Notation "`| x |" (at level 0, x at level 99, format "`| x |"). (* Reserved notation for conditional comparison *) Reserved Notation "x <= y ?= 'iff' c" (at level 70, y, c at next level, format "x '[hv' <= y '/' ?= 'iff' c ']'"). (* Reserved notation for cast comparison. *) Reserved Notation "x <= y :> T" (at level 70, y at next level). Reserved Notation "x >= y :> T" (at level 70, y at next level, only parsing). Reserved Notation "x < y :> T" (at level 70, y at next level). Reserved Notation "x > y :> T" (at level 70, y at next level, only parsing). Reserved Notation "x <= y ?= 'iff' c :> T" (at level 70, y, c at next level, format "x '[hv' <= y '/' ?= 'iff' c :> T ']'"). (* Complements on the option type constructor, used below to *) (* encode partial functions. *) Module Option. Definition apply aT rT (f : aT -> rT) x u := if u is Some y then f y else x. Definition default T := apply (fun x : T => x). Definition bind aT rT (f : aT -> option rT) := apply f None. Definition map aT rT (f : aT -> rT) := bind (fun x => Some (f x)). End Option. Notation oapp := Option.apply. Notation odflt := Option.default. Notation obind := Option.bind. Notation omap := Option.map. Notation some := (@Some _) (only parsing). (* Shorthand for some basic equality lemmas. *) Notation erefl := refl_equal. Notation ecast i T e x := (let: erefl in _ = i := e return T in x). Definition esym := sym_eq. Definition nesym := sym_not_eq. Definition etrans := trans_eq. Definition congr1 := f_equal. Definition congr2 := f_equal2. (* Force at least one implicit when used as a view. *) Prenex Implicits esym nesym. (* A predicate for singleton types. *) Definition all_equal_to T (x0 : T) := forall x, unkeyed x = x0. Lemma unitE : all_equal_to tt. Proof. by case. Qed. (* A generic wrapper type *) Structure wrapped T := Wrap {unwrap : T}. Canonical wrap T x := @Wrap T x. Prenex Implicits unwrap wrap Wrap. (* Syntax for defining auxiliary recursive function. *) (* Usage: *) (* Section FooDefinition. *) (* Variables (g1 : T1) (g2 : T2). (globals) *) (* Fixoint foo_auxiliary (a3 : T3) ... := *) (* body, using [rec e3, ...] for recursive calls *) (* where "[ 'rec' a3 , a4 , ... ]" := foo_auxiliary. *) (* Definition foo x y .. := [rec e1, ...]. *) (* + proofs about foo *) (* End FooDefinition. *) Reserved Notation "[ 'rec' a0 ]" (at level 0, format "[ 'rec' a0 ]"). Reserved Notation "[ 'rec' a0 , a1 ]" (at level 0, format "[ 'rec' a0 , a1 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 ]"). Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 , a9 ]" (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 , a9 ]"). (* Definitions and notation for explicit functions with simplification, *) (* i.e., which simpl and /= beta expand (this is complementary to nosimpl). *) Section SimplFun. Variables aT rT : Type. CoInductive simpl_fun := SimplFun of aT -> rT. Definition fun_of_simpl f := fun x => let: SimplFun lam := f in lam x. Coercion fun_of_simpl : simpl_fun >-> Funclass. End SimplFun. Notation "[ 'fun' : T => E ]" := (SimplFun (fun _ : T => E)) (at level 0, format "'[hv' [ 'fun' : T => '/ ' E ] ']'") : fun_scope. Notation "[ 'fun' x => E ]" := (SimplFun (fun x => E)) (at level 0, x ident, format "'[hv' [ 'fun' x => '/ ' E ] ']'") : fun_scope. Notation "[ 'fun' x : T => E ]" := (SimplFun (fun x : T => E)) (at level 0, x ident, only parsing) : fun_scope. Notation "[ 'fun' x y => E ]" := (fun x => [fun y => E]) (at level 0, x ident, y ident, format "'[hv' [ 'fun' x y => '/ ' E ] ']'") : fun_scope. Notation "[ 'fun' x y : T => E ]" := (fun x : T => [fun y : T => E]) (at level 0, x ident, y ident, only parsing) : fun_scope. Notation "[ 'fun' ( x : T ) y => E ]" := (fun x : T => [fun y => E]) (at level 0, x ident, y ident, only parsing) : fun_scope. Notation "[ 'fun' x ( y : T ) => E ]" := (fun x => [fun y : T => E]) (at level 0, x ident, y ident, only parsing) : fun_scope. Notation "[ 'fun' ( x : xT ) ( y : yT ) => E ]" := (fun x : xT => [fun y : yT => E]) (at level 0, x ident, y ident, only parsing) : fun_scope. (* For delta functions in eqtype.v. *) Definition SimplFunDelta aT rT (f : aT -> aT -> rT) := [fun z => f z z]. (* Extensional equality, for unary and binary functions, including syntactic *) (* sugar. *) Section ExtensionalEquality. Variables A B C : Type. Definition eqfun (f g : B -> A) : Prop := forall x, f x = g x. Definition eqrel (r s : C -> B -> A) : Prop := forall x y, r x y = s x y. Lemma frefl f : eqfun f f. Proof. by []. Qed. Lemma fsym f g : eqfun f g -> eqfun g f. Proof. by move=> eq_fg x. Qed. Lemma ftrans f g h : eqfun f g -> eqfun g h -> eqfun f h. Proof. by move=> eq_fg eq_gh x; rewrite eq_fg. Qed. Lemma rrefl r : eqrel r r. Proof. by []. Qed. End ExtensionalEquality. Typeclasses Opaque eqfun. Typeclasses Opaque eqrel. Hint Resolve frefl rrefl. Notation "f1 =1 f2" := (eqfun f1 f2) (at level 70, no associativity) : fun_scope. Notation "f1 =1 f2 :> A" := (f1 =1 (f2 : A)) (at level 70, f2 at next level, A at level 90) : fun_scope. Notation "f1 =2 f2" := (eqrel f1 f2) (at level 70, no associativity) : fun_scope. Notation "f1 =2 f2 :> A" := (f1 =2 (f2 : A)) (at level 70, f2 at next level, A, B at level 90) : fun_scope. Section Composition. Variables A B C : Type. Definition funcomp u (f : B -> A) (g : C -> B) x := let: tt := u in f (g x). Definition catcomp u g f := funcomp u f g. Local Notation comp := (funcomp tt). Definition pcomp (f : B -> option A) (g : C -> option B) x := obind f (g x). Lemma eq_comp f f' g g' : f =1 f' -> g =1 g' -> comp f g =1 comp f' g'. Proof. by move=> eq_ff' eq_gg' x; rewrite /= eq_gg' eq_ff'. Qed. End Composition. Notation comp := (funcomp tt). Notation "@ 'comp'" := (fun A B C => @funcomp A B C tt). Notation "f1 \o f2" := (comp f1 f2) (at level 50, format "f1 \o '/ ' f2") : fun_scope. Notation "f1 \; f2" := (catcomp tt f1 f2) (at level 60, right associativity, format "f1 \; '/ ' f2") : fun_scope. Notation "[ 'eta' f ]" := (fun x => f x) (at level 0, format "[ 'eta' f ]") : fun_scope. Notation "'fun' => E" := (fun _ => E) (at level 200, only parsing) : fun_scope. Notation id := (fun x => x). Notation "@ 'id' T" := (fun x : T => x) (at level 10, T at level 8, only parsing) : fun_scope. Definition id_head T u x : T := let: tt := u in x. Definition explicit_id_key := tt. Notation idfun := (id_head tt). Notation "@ 'idfun' T " := (@id_head T explicit_id_key) (at level 10, T at level 8, format "@ 'idfun' T") : fun_scope. Definition phant_id T1 T2 v1 v2 := phantom T1 v1 -> phantom T2 v2. (* Strong sigma types. *) Section Tag. Variables (I : Type) (i : I) (T_ U_ : I -> Type). Definition tag := projS1. Definition tagged : forall w, T_(tag w) := @projS2 I [eta T_]. Definition Tagged x := @existS I [eta T_] i x. Definition tag2 (w : @sigT2 I T_ U_) := let: existT2 i _ _ := w in i. Definition tagged2 w : T_(tag2 w) := let: existT2 _ x _ := w in x. Definition tagged2' w : U_(tag2 w) := let: existT2 _ _ y := w in y. Definition Tagged2 x y := @existS2 I [eta T_] [eta U_] i x y. End Tag. Implicit Arguments Tagged [I i]. Implicit Arguments Tagged2 [I i]. Prenex Implicits tag tagged Tagged tag2 tagged2 tagged2' Tagged2. Coercion tag_of_tag2 I T_ U_ (w : @sigT2 I T_ U_) := Tagged (fun i => T_ i * U_ i)%type (tagged2 w, tagged2' w). Lemma all_tag I T U : (forall x : I, {y : T x & U x y}) -> {f : forall x, T x & forall x, U x (f x)}. Proof. by move=> fP; exists (fun x => tag (fP x)) => x; case: (fP x). Qed. Lemma all_tag2 I T U V : (forall i : I, {y : T i & U i y & V i y}) -> {f : forall i, T i & forall i, U i (f i) & forall i, V i (f i)}. Proof. by case/all_tag=> f /all_pair[]; exists f. Qed. (* Refinement types. *) (* Prenex Implicits and renaming. *) Notation sval := (@proj1_sig _ _). Notation "@ 'sval'" := (@proj1_sig) (at level 10, format "@ 'sval'"). Section Sig. Variables (T : Type) (P Q : T -> Prop). Lemma svalP (u : sig P) : P (sval u). Proof. by case: u. Qed. Definition s2val (u : sig2 P Q) := let: exist2 x _ _ := u in x. Lemma s2valP u : P (s2val u). Proof. by case: u. Qed. Lemma s2valP' u : Q (s2val u). Proof. by case: u. Qed. End Sig. Prenex Implicits svalP s2val s2valP s2valP'. Coercion tag_of_sig I P (u : @sig I P) := Tagged P (svalP u). Coercion sig_of_sig2 I P Q (u : @sig2 I P Q) := exist (fun i => P i /\ Q i) (s2val u) (conj (s2valP u) (s2valP' u)). Lemma all_sig I T P : (forall x : I, {y : T x | P x y}) -> {f : forall x, T x | forall x, P x (f x)}. Proof. by case/all_tag=> f; exists f. Qed. Lemma all_sig2 I T P Q : (forall x : I, {y : T x | P x y & Q x y}) -> {f : forall x, T x | forall x, P x (f x) & forall x, Q x (f x)}. Proof. by case/all_sig=> f /all_pair[]; exists f. Qed. Section Morphism. Variables (aT rT sT : Type) (f : aT -> rT). (* Morphism property for unary and binary functions *) Definition morphism_1 aF rF := forall x, f (aF x) = rF (f x). Definition morphism_2 aOp rOp := forall x y, f (aOp x y) = rOp (f x) (f y). (* Homomorphism property for unary and binary relations *) Definition homomorphism_1 (aP rP : _ -> Prop) := forall x, aP x -> rP (f x). Definition homomorphism_2 (aR rR : _ -> _ -> Prop) := forall x y, aR x y -> rR (f x) (f y). (* Stability property for unary and binary relations *) Definition monomorphism_1 (aP rP : _ -> sT) := forall x, rP (f x) = aP x. Definition monomorphism_2 (aR rR : _ -> _ -> sT) := forall x y, rR (f x) (f y) = aR x y. End Morphism. Notation "{ 'morph' f : x / a >-> r }" := (morphism_1 f (fun x => a) (fun x => r)) (at level 0, f at level 99, x ident, format "{ 'morph' f : x / a >-> r }") : type_scope. Notation "{ 'morph' f : x / a }" := (morphism_1 f (fun x => a) (fun x => a)) (at level 0, f at level 99, x ident, format "{ 'morph' f : x / a }") : type_scope. Notation "{ 'morph' f : x y / a >-> r }" := (morphism_2 f (fun x y => a) (fun x y => r)) (at level 0, f at level 99, x ident, y ident, format "{ 'morph' f : x y / a >-> r }") : type_scope. Notation "{ 'morph' f : x y / a }" := (morphism_2 f (fun x y => a) (fun x y => a)) (at level 0, f at level 99, x ident, y ident, format "{ 'morph' f : x y / a }") : type_scope. Notation "{ 'homo' f : x / a >-> r }" := (homomorphism_1 f (fun x => a) (fun x => r)) (at level 0, f at level 99, x ident, format "{ 'homo' f : x / a >-> r }") : type_scope. Notation "{ 'homo' f : x / a }" := (homomorphism_1 f (fun x => a) (fun x => a)) (at level 0, f at level 99, x ident, format "{ 'homo' f : x / a }") : type_scope. Notation "{ 'homo' f : x y / a >-> r }" := (homomorphism_2 f (fun x y => a) (fun x y => r)) (at level 0, f at level 99, x ident, y ident, format "{ 'homo' f : x y / a >-> r }") : type_scope. Notation "{ 'homo' f : x y / a }" := (homomorphism_2 f (fun x y => a) (fun x y => a)) (at level 0, f at level 99, x ident, y ident, format "{ 'homo' f : x y / a }") : type_scope. Notation "{ 'homo' f : x y /~ a }" := (homomorphism_2 f (fun y x => a) (fun x y => a)) (at level 0, f at level 99, x ident, y ident, format "{ 'homo' f : x y /~ a }") : type_scope. Notation "{ 'mono' f : x / a >-> r }" := (monomorphism_1 f (fun x => a) (fun x => r)) (at level 0, f at level 99, x ident, format "{ 'mono' f : x / a >-> r }") : type_scope. Notation "{ 'mono' f : x / a }" := (monomorphism_1 f (fun x => a) (fun x => a)) (at level 0, f at level 99, x ident, format "{ 'mono' f : x / a }") : type_scope. Notation "{ 'mono' f : x y / a >-> r }" := (monomorphism_2 f (fun x y => a) (fun x y => r)) (at level 0, f at level 99, x ident, y ident, format "{ 'mono' f : x y / a >-> r }") : type_scope. Notation "{ 'mono' f : x y / a }" := (monomorphism_2 f (fun x y => a) (fun x y => a)) (at level 0, f at level 99, x ident, y ident, format "{ 'mono' f : x y / a }") : type_scope. Notation "{ 'mono' f : x y /~ a }" := (monomorphism_2 f (fun y x => a) (fun x y => a)) (at level 0, f at level 99, x ident, y ident, format "{ 'mono' f : x y /~ a }") : type_scope. (* In an intuitionistic setting, we have two degrees of injectivity. The *) (* weaker one gives only simplification, and the strong one provides a left *) (* inverse (we show in `fintype' that they coincide for finite types). *) (* We also define an intermediate version where the left inverse is only a *) (* partial function. *) Section Injections. (* rT must come first so we can use @ to mitigate the Coq 1st order *) (* unification bug (e..g., Coq can't infer rT from a "cancel" lemma). *) Variables (rT aT : Type) (f : aT -> rT). Definition injective := forall x1 x2, f x1 = f x2 -> x1 = x2. Definition cancel g := forall x, g (f x) = x. Definition pcancel g := forall x, g (f x) = Some x. Definition ocancel (g : aT -> option rT) h := forall x, oapp h x (g x) = x. Lemma can_pcan g : cancel g -> pcancel (fun y => Some (g y)). Proof. by move=> fK x; congr (Some _). Qed. Lemma pcan_inj g : pcancel g -> injective. Proof. by move=> fK x y /(congr1 g); rewrite !fK => [[]]. Qed. Lemma can_inj g : cancel g -> injective. Proof. by move/can_pcan; exact: pcan_inj. Qed. Lemma canLR g x y : cancel g -> x = f y -> g x = y. Proof. by move=> fK ->. Qed. Lemma canRL g x y : cancel g -> f x = y -> x = g y. Proof. by move=> fK <-. Qed. End Injections. Lemma Some_inj {T} : injective (@Some T). Proof. by move=> x y []. Qed. (* cancellation lemmas for dependent type casts. *) Lemma esymK T x y : cancel (@esym T x y) (@esym T y x). Proof. by case: y /. Qed. Lemma etrans_id T x y (eqxy : x = y :> T) : etrans (erefl x) eqxy = eqxy. Proof. by case: y / eqxy. Qed. Section InjectionsTheory. Variables (A B C : Type) (f g : B -> A) (h : C -> B). Lemma inj_id : injective (@id A). Proof. by []. Qed. Lemma inj_can_sym f' : cancel f f' -> injective f' -> cancel f' f. Proof. move=> fK injf' x; exact: injf'. Qed. Lemma inj_comp : injective f -> injective h -> injective (f \o h). Proof. move=> injf injh x y /injf; exact: injh. Qed. Lemma can_comp f' h' : cancel f f' -> cancel h h' -> cancel (f \o h) (h' \o f'). Proof. by move=> fK hK x; rewrite /= fK hK. Qed. Lemma pcan_pcomp f' h' : pcancel f f' -> pcancel h h' -> pcancel (f \o h) (pcomp h' f'). Proof. by move=> fK hK x; rewrite /pcomp fK /= hK. Qed. Lemma eq_inj : injective f -> f =1 g -> injective g. Proof. by move=> injf eqfg x y; rewrite -2!eqfg; exact: injf. Qed. Lemma eq_can f' g' : cancel f f' -> f =1 g -> f' =1 g' -> cancel g g'. Proof. by move=> fK eqfg eqfg' x; rewrite -eqfg -eqfg'. Qed. Lemma inj_can_eq f' : cancel f f' -> injective f' -> cancel g f' -> f =1 g. Proof. by move=> fK injf' gK x; apply: injf'; rewrite fK. Qed. End InjectionsTheory. Section Bijections. Variables (A B : Type) (f : B -> A). CoInductive bijective : Prop := Bijective g of cancel f g & cancel g f. Hypothesis bijf : bijective. Lemma bij_inj : injective f. Proof. by case: bijf => g fK _; apply: can_inj fK. Qed. Lemma bij_can_sym f' : cancel f' f <-> cancel f f'. Proof. split=> fK; first exact: inj_can_sym fK bij_inj. by case: bijf => h _ hK x; rewrite -[x]hK fK. Qed. Lemma bij_can_eq f' f'' : cancel f f' -> cancel f f'' -> f' =1 f''. Proof. by move=> fK fK'; apply: (inj_can_eq _ bij_inj); apply/bij_can_sym. Qed. End Bijections. Section BijectionsTheory. Variables (A B C : Type) (f : B -> A) (h : C -> B). Lemma eq_bij : bijective f -> forall g, f =1 g -> bijective g. Proof. by case=> f' fK f'K g eqfg; exists f'; eapply eq_can; eauto. Qed. Lemma bij_comp : bijective f -> bijective h -> bijective (f \o h). Proof. by move=> [f' fK f'K] [h' hK h'K]; exists (h' \o f'); apply: can_comp; auto. Qed. Lemma bij_can_bij : bijective f -> forall f', cancel f f' -> bijective f'. Proof. by move=> bijf; exists f; first by apply/(bij_can_sym bijf). Qed. End BijectionsTheory. Section Involutions. Variables (A : Type) (f : A -> A). Definition involutive := cancel f f. Hypothesis Hf : involutive. Lemma inv_inj : injective f. Proof. exact: can_inj Hf. Qed. Lemma inv_bij : bijective f. Proof. by exists f. Qed. End Involutions. Section OperationProperties. Variables S T R : Type. Section SopTisR. Implicit Type op : S -> T -> R. Definition left_inverse e inv op := forall x, op (inv x) x = e. Definition right_inverse e inv op := forall x, op x (inv x) = e. Definition left_injective op := forall x, injective (op^~ x). Definition right_injective op := forall y, injective (op y). End SopTisR. Section SopTisS. Implicit Type op : S -> T -> S. Definition right_id e op := forall x, op x e = x. Definition left_zero z op := forall x, op z x = z. Definition right_commutative op := forall x y z, op (op x y) z = op (op x z) y. Definition left_distributive op add := forall x y z, op (add x y) z = add (op x z) (op y z). Definition right_loop inv op := forall y, cancel (op^~ y) (op^~ (inv y)). Definition rev_right_loop inv op := forall y, cancel (op^~ (inv y)) (op^~ y). End SopTisS. Section SopTisT. Implicit Type op : S -> T -> T. Definition left_id e op := forall x, op e x = x. Definition right_zero z op := forall x, op x z = z. Definition left_commutative op := forall x y z, op x (op y z) = op y (op x z). Definition right_distributive op add := forall x y z, op x (add y z) = add (op x y) (op x z). Definition left_loop inv op := forall x, cancel (op x) (op (inv x)). Definition rev_left_loop inv op := forall x, cancel (op (inv x)) (op x). End SopTisT. Section SopSisT. Implicit Type op : S -> S -> T. Definition self_inverse e op := forall x, op x x = e. Definition commutative op := forall x y, op x y = op y x. End SopSisT. Section SopSisS. Implicit Type op : S -> S -> S. Definition idempotent op := forall x, op x x = x. Definition associative op := forall x y z, op x (op y z) = op (op x y) z. Definition interchange op1 op2 := forall x y z t, op1 (op2 x y) (op2 z t) = op2 (op1 x z) (op1 y t). End SopSisS. End OperationProperties. ssreflect-1.5/Makefile0000644000175000017500000000053212175453667014015 0ustar garesgaresCOQMAKEFILE := Makefile.coq COQMAKE := +$(MAKE) -f $(COQMAKEFILE) ifneq "$(COQBIN)" "" COQBIN := $(COQBIN)/ endif all: $(COQMAKEFILE) mkdir -p bin $(COQMAKE) all $(COQMAKEFILE): Make $(COQBIN)coq_makefile -f Make > $(COQMAKEFILE) install: $(COQMAKE) install clean: -$(COQMAKE) clean rm -f $(COQMAKEFILE) bin .PHONY: clean all install ssreflect-1.5/MANIFEST0000644000175000017500000000047512175455021013476 0ustar garesgaresANNOUNCE AUTHORS CeCILL-B INSTALL Make Makefile pg-ssr.el README src/ssreflect.ml4 src/ssrmatching.ml4 src/ssrmatching.mli theories/choice.v theories/eqtype.v theories/fintype.v theories/seq.v theories/ssrbool.v theories/ssreflect.mllib theories/ssreflect.v theories/ssrfun.v theories/ssrmatching.v theories/ssrnat.v