ladder-calculus/coq/typing/env.v

38 lines
1 KiB
Coq
Raw Normal View History

2024-09-22 15:52:06 +02:00
Require Import Metatheory.
2024-09-21 17:26:26 +02:00
Require Import Atom.
Require Import Environment.
Require Import debruijn.
Notation env := (list (atom * type_DeBruijn)).
2024-09-22 15:52:06 +02:00
(* env is well-formed if each atom is bound at most once
* and every binding is to a locally closed type
*)
Inductive env_wf : env -> Prop :=
| env_wf_empty :
env_wf []
| env_wf_type : forall Γ x τ,
env_wf Γ ->
type_lc τ ->
x `notin` (dom Γ) ->
env_wf ((x,τ)::Γ)
.
Ltac gather_atoms :=
let A := gather_atoms_with (fun x : atoms => x) in
let B := gather_atoms_with (fun x : atom => singleton x) in
let C := gather_atoms_with (fun x : list (atom * type_DeBruijn) => dom x) in
let D := gather_atoms_with (fun x : expr_DeBruijn => expr_fv x) in
constr:(A `union` B `union` C `union` D).
Tactic Notation "pick" "fresh" ident(x) :=
let L := gather_atoms in
(pick fresh x for L).
Tactic Notation
"pick" "fresh" ident(atom_name) "and" "apply" constr(lemma) :=
let L := gather_atoms in
pick fresh atom_name excluding L and apply lemma.