File: Fresh.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: forky, sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (34 lines) | stat: -rw-r--r-- 1,451 bytes parent folder | download | duplicates (2)
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
(************************************************************************)
(*         *   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)         *)
(************************************************************************)

Require Import Ltac2.Init.
Require Ltac2.Control.
Require Ltac2.List.

Module Free.

Ltac2 Type t.
(** Type of sets of free variables *)

Ltac2 @ external union : t -> t -> t := "coq-core.plugins.ltac2" "fresh_free_union".

Ltac2 @ external of_ids : ident list -> t := "coq-core.plugins.ltac2" "fresh_free_of_ids".

Ltac2 @ external of_constr : constr -> t := "coq-core.plugins.ltac2" "fresh_free_of_constr".

Ltac2 of_goal () := of_ids (List.map (fun (id, _, _) => id) (Control.hyps ())).

End Free.

Ltac2 @ external fresh : Free.t -> ident -> ident := "coq-core.plugins.ltac2" "fresh_fresh".
(** Generate a fresh identifier with the given base name which is not a
    member of the provided set of free variables. *)

Ltac2 in_goal id := Fresh.fresh (Free.of_goal ()) id.