Set Implicit Arguments.
Require Import List Lia.
From Undecidability.HOU Require Import std.std.
From Undecidability.HOU.calculus Require Import prelim terms syntax semantics.

Set Default Proof Using "Type".


Section Typing.

  Context {X: Const}.

  Definition ctx := list type.

  Implicit Types (s t: exp X) (n m: ) ( : ctx)
           (x y: fin) (c d: X) (A B: type)
           ( : fin exp X) ( : fin fin).


  Reserved Notation "Gamma ⊢ s : A" (at level 80, s at level 99).

  Inductive typing Gamma : exp X type Prop :=
  | typingVar i A: nth i = Some A var i : A
  | typingConst c: const c : ctype X c
  | typingLam s A B: A :: s : B s : A B
  | typingApp s t A B: s : A B t : A s t : B
  where "Gamma ⊢ s : A" := (typing s A).

  Hint Constructors typing : core.

  Definition typingRen Delta delta Gamma :=
     x A, nth x = Some A nth ( x) = Some A.

  Notation "Delta ⊫ delta : Gamma" := (typingRen ) (at level 80, at level 99).

  Definition typingSubst Delta sigma Gamma :=
     x A, nth x = Some A x : A.

  Notation "Delta ⊩ sigma : Gamma" := (typingSubst ) (at level 80, at level 99).

  Lemma typingSubst_cons Delta sigma s Gamma A:
     : s : A s .: : A :: .
  Proof.
    intros ??[|]; cbn; eauto; now intros ? [= ].
  Qed.


  Section Preservation.
    Lemma preservation_under_renaming delta Gamma Delta s A:
       s : A : ren_exp s : A.
    Proof.
      induction 1 in , |-*; intros H'; cbn [ren_exp]; eauto.
      econstructor; apply IHtyping.
      intros []; cbn; eauto.
    Qed.


    Lemma preservation_under_substitution sigma Gamma Delta s A:
       s : A : s : A.
    Proof.
      induction 1 in , |-*; intros H'; cbn [subst_exp]; subst; eauto.
      econstructor; apply IHtyping; intros []; cbn; eauto.
      intros; eapply preservation_under_renaming; eauto.
      intros ?; eauto.
    Qed.


    Lemma preservation_under_reduction s s' Gamma A:
      s > s' s : A s' : A.
    Proof.
      induction 1 in , A |-*; intros ; invp @typing; eauto.
      inv . eapply preservation_under_substitution; eauto.
      intros []; cbn; eauto.
      intros ? H; now injection H as .
    Qed.


    Lemma preservation_under_steps s s' Gamma A:
      s >* s' s : A s' : A.
    Proof.
      induction 1 in , A |-*; eauto using preservation_under_reduction.
    Qed.


  End Preservation.

  Lemma weakening_app Gamma Delta s A:
     s : A s : A.
  Proof.
    intros H; replace s with (ren id s) by now asimpl.
    eapply preservation_under_renaming; eauto.
    intros i B H'; unfold id.
    rewrite nth_error_app1; eauto.
    eapply nth_error_Some_lt; eauto.
  Qed.


  Lemma typing_variables Gamma s A:
     s : A x, x vars s x dom .
  Proof.
    intros ? x % vars_varof.
    induction H in x, |-*; inv ; eauto.
    - now domin H.
    - specialize (IHtyping _ ).
      eapply dom_lt_iff in IHtyping.
      cbn in *; intuition.
  Qed.


End Typing.

Notation "Gamma ⊢ s : A" := (typing s A) (at level 80, s at level 99).
Notation "Delta ⊫ delta : Gamma" := (typingRen ) (at level 80, at level 99).
Notation "Delta ⊩ sigma : Gamma" := (typingSubst ) (at level 80, at level 99).

#[export] Hint Constructors typing : core.
#[export] Hint Resolve typing_variables : core.

#[export] Hint Resolve
     preservation_under_renaming
     preservation_under_substitution : core.