ladder-calculus/coq/terms_debruijn.v

131 lines
5.5 KiB
Coq
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

From Coq Require Import Strings.String.
From Coq Require Import Lists.List.
Import ListNotations.
From Coq Require Import Arith.EqNat.
Local Open Scope nat_scope.
Require Import Atom.
Require Import Metatheory.
Require Import FSetNotin.
Inductive type_DeBruijn : Type :=
| ty_id : string -> type_DeBruijn
| ty_fvar : atom -> type_DeBruijn
| ty_bvar : nat -> type_DeBruijn
| ty_univ : type_DeBruijn -> type_DeBruijn
| ty_spec : type_DeBruijn -> type_DeBruijn -> type_DeBruijn
| ty_func : type_DeBruijn -> type_DeBruijn -> type_DeBruijn
| ty_morph : type_DeBruijn -> type_DeBruijn -> type_DeBruijn
| ty_ladder : type_DeBruijn -> type_DeBruijn -> type_DeBruijn
.
Inductive expr_DeBruijn : Type :=
| ex_fvar : atom -> expr_DeBruijn
| ex_bvar : nat -> expr_DeBruijn
| ex_ty_abs : expr_DeBruijn -> expr_DeBruijn
| ex_ty_app : expr_DeBruijn -> type_DeBruijn -> expr_DeBruijn
| ex_abs : type_DeBruijn -> expr_DeBruijn -> expr_DeBruijn
| ex_morph : type_DeBruijn -> expr_DeBruijn -> expr_DeBruijn
| ex_app : expr_DeBruijn -> expr_DeBruijn -> expr_DeBruijn
| ex_let : type_DeBruijn -> expr_DeBruijn -> expr_DeBruijn -> expr_DeBruijn
| ex_ascend : type_DeBruijn -> expr_DeBruijn -> expr_DeBruijn
| ex_descend : type_DeBruijn -> expr_DeBruijn -> expr_DeBruijn
.
(* get the list of all free variables in a type term *)
Fixpoint type_fv (τ : type_DeBruijn) {struct τ} : atoms :=
match τ with
| ty_id s => {}
| ty_fvar α => singleton α
| ty_bvar x => {}
| ty_univ τ => (type_fv τ)
| ty_spec σ τ => (type_fv σ) `union` (type_fv τ)
| ty_func σ τ => (type_fv σ) `union` (type_fv τ)
| ty_morph σ τ => (type_fv σ) `union` (type_fv τ)
| ty_ladder σ τ => (type_fv σ) `union` (type_fv τ)
end.
(* substitute free variable x with type σ in τ *)
Fixpoint subst_type (x:atom) (σ:type_DeBruijn) (τ:type_DeBruijn) {struct τ} : type_DeBruijn :=
match τ with
| ty_id s => ty_id s
| ty_fvar s => if x == s then σ else τ
| ty_bvar y => ty_bvar y
| ty_univ τ => ty_univ (subst_type x σ τ)
| ty_spec τ1 τ2 => ty_spec (subst_type x σ τ1) (subst_type x σ τ2)
| ty_func τ1 τ2 => ty_func (subst_type x σ τ1) (subst_type x σ τ2)
| ty_morph τ1 τ2 => ty_morph (subst_type x σ τ1) (subst_type x σ τ2)
| ty_ladder τ1 τ2 => ty_ladder (subst_type x σ τ1) (subst_type x σ τ2)
end.
(* replace a free variable with a new (dangling) bound variable *)
Fixpoint type_bind_fvar (x:atom) (n:nat) (τ:type_DeBruijn) {struct τ} : type_DeBruijn :=
match τ with
| ty_id s => ty_id s
| ty_fvar s => if x == s then ty_bvar n else τ
| ty_bvar n => ty_bvar n
| ty_univ τ1 => ty_univ (type_bind_fvar x (S n) τ1)
| ty_spec τ1 τ2 => ty_spec (type_bind_fvar x n τ1) (type_bind_fvar x n τ2)
| ty_func τ1 τ2 => ty_func (type_bind_fvar x n τ1) (type_bind_fvar x n τ2)
| ty_morph τ1 τ2 => ty_morph (type_bind_fvar x n τ1) (type_bind_fvar x n τ2)
| ty_ladder τ1 τ2 => ty_ladder (type_bind_fvar x n τ1) (type_bind_fvar x n τ2)
end.
(* replace (dangling) index with another type *)
Fixpoint type_open_rec (k:nat) (σ:type_DeBruijn) (τ:type_DeBruijn) {struct τ} : type_DeBruijn :=
match τ with
| ty_id s => ty_id s
| ty_fvar s => ty_fvar s
| ty_bvar i => if k === i then σ else τ
| ty_univ τ1 => ty_univ (type_open_rec (S k) σ τ1)
| ty_spec τ1 τ2 => ty_spec (type_open_rec k σ τ1) (type_open_rec k σ τ2)
| ty_func τ1 τ2 => ty_func (type_open_rec k σ τ1) (type_open_rec k σ τ2)
| ty_morph τ1 τ2 => ty_morph (type_open_rec k σ τ1) (type_open_rec k σ τ2)
| ty_ladder τ1 τ2 => ty_ladder (type_open_rec k σ τ1) (type_open_rec k σ τ2)
end.
Notation "'[' z '~>' u ']' e" := (subst_type z u e) (at level 68).
Notation "'{' k '~>' σ '}' τ" := (type_open_rec k σ τ) (at level 67).
Definition type_open σ τ := type_open_rec 0 σ τ.
(* is the type locally closed ? *)
Inductive type_lc : type_DeBruijn -> Prop :=
| Tlc_Id : forall s, type_lc (ty_id s)
| Tlc_Var : forall s, type_lc (ty_fvar s)
| Tlc_Univ : forall τ1 L,
(forall x, (x `notin` L) -> type_lc (type_open (ty_fvar x) τ1)) ->
type_lc (ty_univ τ1)
| Tlc_Spec : forall τ1 τ2, type_lc τ1 -> type_lc τ2 -> type_lc (ty_spec τ1 τ2)
| Tlc_Func : forall τ1 τ2, type_lc τ1 -> type_lc τ2 -> type_lc (ty_func τ1 τ2)
| Tlc_Morph : forall τ1 τ2, type_lc τ1 -> type_lc τ2 -> type_lc (ty_morph τ1 τ2)
| Tlc_Ladder : forall τ1 τ2, type_lc τ1 -> type_lc τ2 -> type_lc (ty_ladder τ1 τ2)
.
(* number of abstractions *)
Fixpoint type_debruijn_depth (τ:type_DeBruijn) : nat :=
match τ with
| ty_id s => 0
| ty_fvar s => 0
| ty_bvar x => 0
| ty_func s t => max (type_debruijn_depth s) (type_debruijn_depth t)
| ty_morph s t => max (type_debruijn_depth s) (type_debruijn_depth t)
| ty_univ t => (1 + (type_debruijn_depth t))
| ty_spec s t => ((type_debruijn_depth s) - 1)
| ty_ladder s t => max (type_debruijn_depth s) (type_debruijn_depth t)
end.
(*
Fixpoint type_named2debruijn (τ:type_term) {struct τ} : type_DeBruijn :=
match τ with
| type_id s => ty_id s
| type_var s => ty_fvar (pick fresh )
| type_univ x t => let t':=(type_named2debruijn t) in (ty_univ (type_bind_fvar x 0 t'))
| type_spec s t => ty_spec (type_named2debruijn s) (type_named2debruijn t)
| type_fun s t => ty_func (type_named2debruijn s) (type_named2debruijn t)
| type_morph s t => ty_morph (type_named2debruijn s) (type_named2debruijn t)
| type_ladder s t => ty_ladder (type_named2debruijn s) (type_named2debruijn t)
end.
Coercion type_named2debruijn : type_term >-> type_DeBruijn.
*)