coq: alpha conversion in types

This commit is contained in:
Michael Sippel 2024-08-18 10:47:35 +02:00
parent 3a84dada65
commit 221c017640
Signed by: senvas
GPG key ID: 060F22F65102F95C

View file

@ -30,13 +30,33 @@
* satisfies all properties required of an equivalence relation. * satisfies all properties required of an equivalence relation.
*) *)
Require Import terms. Require Import terms.
Require Import subst.
From Coq Require Import Strings.String. From Coq Require Import Strings.String.
Include Terms. Include Terms.
Include Subst.
Module Equiv. Module Equiv.
(** Alpha conversion in types *)
Reserved Notation "S '-->α' T" (at level 40).
Inductive type_conv_alpha : type_term -> type_term -> Prop :=
| Eq_Alpha : forall x y t,
(type_univ x t) -->α (type_univ y (type_subst x (type_var y) t))
where "S '-->α' T" := (type_conv_alpha S T).
(** Alpha conversion is symmetric *)
Lemma type_alpha_symm :
forall σ τ,
(σ -->α τ) -> (τ -->α σ).
Proof.
(* TODO *)
Admitted.
(** Define all rewrite steps $\label{coq:type-dist}$ *) (** Define all rewrite steps $\label{coq:type-dist}$ *)
Reserved Notation "S '-->distribute-ladder' T" (at level 40). Reserved Notation "S '-->distribute-ladder' T" (at level 40).
@ -123,6 +143,10 @@ Inductive type_eq : type_term -> type_term -> Prop :=
y === z -> y === z ->
x === z x === z
| L_Rename : forall x y,
x -->α y ->
x === y
| L_Distribute : forall x y, | L_Distribute : forall x y,
x -->distribute-ladder y -> x -->distribute-ladder y ->
x === y x === y
@ -147,12 +171,12 @@ Proof.
apply L_Refl. apply L_Refl.
} }
2:{ 3:{
apply L_Condense. apply L_Condense.
apply distribute_inverse. apply distribute_inverse.
apply H. apply H.
} }
2:{ 3:{
apply L_Distribute. apply L_Distribute.
apply condense_inverse. apply condense_inverse.
apply H. apply H.
@ -161,6 +185,10 @@ Proof.
apply L_Trans with (y:=y). apply L_Trans with (y:=y).
apply IHtype_eq2. apply IHtype_eq2.
apply IHtype_eq1. apply IHtype_eq1.
apply type_alpha_symm in H.
apply L_Rename.
apply H.
Qed. Qed.
(** "flat" types do not contain ladders $\label{coq:type-flat}$ *) (** "flat" types do not contain ladders $\label{coq:type-flat}$ *)