package coq
Formal proof management system
Install
Dune Dependency
Authors
Maintainers
Sources
coq-8.14.0.tar.gz
sha256=b1501d686c21836302191ae30f610cca57fb309214c126518ca009363ad2cd3c
doc/src/coq-core.tactics/btermdn.ml.html
Source file btermdn.ml
1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70 71 72 73 74 75 76 77 78 79 80 81 82 83 84 85 86 87 88 89 90 91 92 93 94 95 96 97 98 99 100 101 102 103 104 105 106 107 108 109 110 111 112 113 114 115 116 117 118 119 120 121 122 123 124 125 126 127 128 129 130 131 132 133 134 135 136 137 138 139 140 141 142 143 144 145 146 147 148 149 150 151 152 153 154 155 156 157 158 159 160 161 162 163 164 165 166 167 168 169 170 171 172 173 174 175 176 177 178 179 180 181 182 183 184 185 186 187 188 189 190 191 192 193 194 195 196 197
(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * Copyright INRIA, CNRS and contributors *) (* <O___,, * (see version control and CREDITS file for authors & dates) *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (* * (see LICENSE file for the text of the license) *) (************************************************************************) open Util open Constr open EConstr open Names open Pattern (* Discrimination nets with bounded depth. See the module dn.ml for further explanations. Eduardo (5/8/97). *) let dnet_depth = ref 8 type term_label = | GRLabel of GlobRef.t | ProdLabel | LambdaLabel | SortLabel let compare_term_label t1 t2 = match t1, t2 with | GRLabel gr1, GRLabel gr2 -> GlobRef.CanOrd.compare gr1 gr2 | _ -> pervasives_compare t1 t2 (** OK *) type 'res lookup_res = 'res Dn.lookup_res = Label of 'res | Nothing | Everything let decomp_pat p = let rec decrec acc = function | PApp (f,args) -> decrec (Array.to_list args @ acc) f | PProj (p, c) -> let hole = PMeta None in let params = List.make (Projection.npars p) hole in (PRef (GlobRef.ConstRef (Projection.constant p)), params @ c :: acc) | c -> (c,acc) in decrec [] p let decomp sigma t = let rec decrec acc c = match EConstr.kind sigma c with | App (f,l) -> decrec (Array.fold_right (fun a l -> a::l) l acc) f | Proj (p, c) -> (* Hack: fake evar to generate [Everything] in the functions below *) let hole = mkEvar (Evar.unsafe_of_int (-1), []) in let params = List.make (Projection.npars p) hole in (mkConst (Projection.constant p), params @ c :: acc) | Cast (c1,_,_) -> decrec acc c1 | _ -> (c,acc) in decrec [] t let evaluable_constant c env = (* This is a hack to work around a broken Print Module implementation, see bug #2668. *) if Environ.mem_constant c env then Environ.evaluable_constant c env else true let constr_val_discr env sigma t = let open GlobRef in let c, l = decomp sigma t in match EConstr.kind sigma c with | Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l) | Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l) | Var id -> Label(GRLabel (VarRef id),l) | Const (c, _) -> if evaluable_constant c env then Everything else Label(GRLabel (ConstRef c),l) | _ -> Nothing let constr_pat_discr env t = if not (Patternops.occur_meta_pattern t) then None else let open GlobRef in match decomp_pat t with | PRef ((IndRef _) as ref), args | PRef ((ConstructRef _ ) as ref), args -> Some (GRLabel ref,args) | PRef ((VarRef v) as ref), args -> Some(GRLabel ref,args) | PRef ((ConstRef c) as ref), args -> if evaluable_constant c env then None else Some (GRLabel ref, args) | _ -> None let constr_val_discr_st env sigma ts t = let c, l = decomp sigma t in let open GlobRef in match EConstr.kind sigma c with | Const (c,u) -> if evaluable_constant c env && TransparentState.is_transparent_constant ts c then Everything else Label(GRLabel (ConstRef c),l) | Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l) | Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l) | Var id -> if Environ.evaluable_named id env && TransparentState.is_transparent_variable ts id then Everything else Label(GRLabel (VarRef id),l) | Prod (n, d, c) -> Label(ProdLabel, [d; c]) | Lambda (n, d, c) -> if List.is_empty l then Label(LambdaLabel, [d; c] @ l) else Everything | Sort _ -> Label(SortLabel, []) | Evar _ -> Everything | Rel _ | Meta _ | Cast _ | LetIn _ | App _ | Case _ | Fix _ | CoFix _ | Proj _ | Int _ | Float _ | Array _ -> Nothing let constr_pat_discr_st env ts t = let open GlobRef in match decomp_pat t with | PRef ((IndRef _) as ref), args | PRef ((ConstructRef _ ) as ref), args -> Some (GRLabel ref,args) | PRef ((VarRef v) as ref), args -> if Environ.evaluable_named v env && (TransparentState.is_transparent_variable ts v) then None else Some(GRLabel ref,args) | PRef ((ConstRef c) as ref), args -> if evaluable_constant c env && TransparentState.is_transparent_constant ts c then None else Some (GRLabel ref, args) | PVar v, args when not (TransparentState.is_transparent_variable ts v) -> Some(GRLabel (VarRef v),args) | PProd (_, d, c), [] -> Some (ProdLabel, [d ; c]) | PLambda (_, d, c), [] -> Some (LambdaLabel, [d ; c]) | PSort s, [] -> Some (SortLabel, []) | _ -> None let bounded_constr_pat_discr_st env st (t,depth) = if Int.equal depth 0 then None else match constr_pat_discr_st env st t with | None -> None | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l) let bounded_constr_val_discr_st env sigma st (t,depth) = if Int.equal depth 0 then Nothing else match constr_val_discr_st env sigma st t with | Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l) | Nothing -> Nothing | Everything -> Everything let bounded_constr_pat_discr env (t,depth) = if Int.equal depth 0 then None else match constr_pat_discr env t with | None -> None | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l) let bounded_constr_val_discr env sigma (t,depth) = if Int.equal depth 0 then Nothing else match constr_val_discr env sigma t with | Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l) | Nothing -> Nothing | Everything -> Everything module Make = functor (Z : Map.OrderedType) -> struct module Y = struct type t = term_label let compare = compare_term_label end module Dn = Dn.Make(Y)(Z) type t = Dn.t type pattern = Dn.pattern let pattern env st pat = match st with | None -> Dn.pattern (bounded_constr_pat_discr env) (pat, !dnet_depth) | Some st -> Dn.pattern (bounded_constr_pat_discr_st env st) (pat, !dnet_depth) let empty = Dn.empty let add = Dn.add let rmv = Dn.rmv let lookup env sigma = function | None -> (fun dn t -> Dn.lookup dn (bounded_constr_val_discr env sigma) (t,!dnet_depth)) | Some st -> (fun dn t -> Dn.lookup dn (bounded_constr_val_discr_st env sigma st) (t,!dnet_depth)) end
sectionYPositions = computeSectionYPositions($el), 10)"
x-init="setTimeout(() => sectionYPositions = computeSectionYPositions($el), 10)"
>