Require Import Autosubst2.core Autosubst2.unscoped compile Autosubst2.syntax ssreflect. From Hammer Require Import Tactics. Reserved Notation "Γ ⊢ a ∈ A" (at level 70). Reserved Notation "⊢ Γ" (at level 70). Inductive lookup : nat -> list Tm -> Tm -> Prop := | here A Γ : lookup 0 (cons A Γ) (ren_Tm shift A) | there i Γ A B : lookup i Γ A -> lookup (S i) (cons B Γ) (ren_Tm shift A). Lemma lookup_deter i Γ A B : lookup i Γ A -> lookup i Γ B -> A = B. Proof. move => h. move : B. induction h; hauto lq:on inv:lookup. Qed. Lemma here' A Γ U : U = ren_Tm shift A -> lookup 0 (A :: Γ) U. Proof. move => ->. apply here. Qed. Lemma there' i Γ A B U : U = ren_Tm shift A -> lookup i Γ A -> lookup (S i) (cons B Γ) U. Proof. move => ->. apply there. Qed. Derive Inversion lookup_inv with (forall i Γ A, lookup i Γ A). Inductive Wt : list Tm -> Tm -> Tm -> Prop := | T_Var i Γ A : ⊢ Γ -> lookup i Γ A -> Γ ⊢ VarTm i ∈ A | T_Bind Γ i p A B : Γ ⊢ A ∈ Univ i -> cons A Γ ⊢ B ∈ Univ i -> Γ ⊢ TBind p A B ∈ Univ i | T_Abs Γ a A B i : Γ ⊢ TBind TPi A B ∈ (Univ i) -> (cons A Γ) ⊢ a ∈ B -> Γ ⊢ Abs A a ∈ TBind TPi A B | T_App Γ b a A B : Γ ⊢ b ∈ TBind TPi A B -> Γ ⊢ a ∈ A -> Γ ⊢ App b a ∈ subst_Tm (scons a VarTm) B | T_Pair Γ (a b : Tm) A B i : Γ ⊢ TBind TSig A B ∈ (Univ i) -> Γ ⊢ a ∈ A -> Γ ⊢ b ∈ subst_Tm (scons a VarTm) B -> Γ ⊢ Pair a b ∈ TBind TSig A B | T_Proj1 Γ (a : Tm) A B : Γ ⊢ a ∈ TBind TSig A B -> Γ ⊢ Proj PL a ∈ A | T_Proj2 Γ (a : Tm) A B : Γ ⊢ a ∈ TBind TSig A B -> Γ ⊢ Proj PR a ∈ subst_Tm (scons (Proj PL a) VarTm) B | T_Univ Γ i : ⊢ Γ -> Γ ⊢ Univ i ∈ Univ (S i) | T_Conv Γ (a : Tm) A B i : Γ ⊢ a ∈ A -> Γ ⊢ B ∈ Univ i -> Join.R A B -> Γ ⊢ a ∈ B with Wff : list Tm -> Prop := | Wff_Nil : ⊢ nil | Wff_Cons Γ (A : Tm) i : ⊢ Γ -> Γ ⊢ A ∈ Univ i -> (* -------------------------------- *) ⊢ (cons A Γ) where "Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ). Scheme wf_ind := Induction for Wff Sort Prop with wt_ind := Induction for Wt Sort Prop. Combined Scheme wt_mutual from wf_ind, wt_ind. (* Lemma lem : *) (* (forall n (Γ : fin n -> Tm n), ⊢ Γ -> ...) /\ *) (* (forall n Γ (a A : Tm n), Γ ⊢ a ∈ A -> ...) /\ *) (* Proof. apply wt_mutual. ... *)