Add unscoped syntax

This commit is contained in:
Yiyun Liu 2025-03-02 16:40:39 -05:00
parent 4509a64bf5
commit 657c1325c9
5 changed files with 532 additions and 541 deletions

View file

@ -1,41 +1,52 @@
Require Import Autosubst2.fintype Autosubst2.syntax Autosubst2.core ssreflect.
Require Import Autosubst2.unscoped Autosubst2.syntax Autosubst2.core ssreflect.
From Ltac2 Require Ltac2.
Import Ltac2.Notations.
Import Ltac2.Control.
From Hammer Require Import Tactics.
Definition renaming_ok {n m} (Γ : fin n -> PTm n) (Δ : fin m -> PTm m) (ξ : fin m -> fin n) :=
forall (i : fin m), ren_PTm ξ (Δ i) = Γ (ξ i).
Inductive lookup : nat -> list PTm -> PTm -> Prop :=
| here A Γ : lookup 0 (cons A Γ) (ren_PTm shift A)
| there i Γ A B :
lookup i Γ A ->
lookup (S i) (cons B Γ) (ren_PTm shift A).
Lemma up_injective n m (ξ : fin n -> fin m) :
(forall i j, ξ i = ξ j -> i = j) ->
forall i j, (upRen_PTm_PTm ξ) i = (upRen_PTm_PTm ξ) j -> i = j.
Definition renaming_ok (Γ : list PTm) (Δ : list PTm) (ξ : nat -> nat) :=
forall i A, lookup i Δ A -> lookup (ξ i) Γ (ren_PTm ξ A).
Definition ren_inj (ξ : nat -> nat) := forall i j, ξ i = ξ j -> i = j.
Lemma up_injective (ξ : nat -> nat) :
ren_inj ξ ->
ren_inj (upRen_PTm_PTm ξ).
Proof.
sblast inv:option.
move => h i j.
case : i => //=; case : j => //=.
move => i j. rewrite /funcomp. hauto lq:on rew:off unfold:ren_inj.
Qed.
Local Ltac2 rec solve_anti_ren () :=
let x := Fresh.in_goal (Option.get (Ident.of_string "x")) in
intro $x;
lazy_match! Constr.type (Control.hyp x) with
| fin _ -> _ _ => (ltac1:(case;hauto lq:on rew:off use:up_injective))
| nat -> nat => (ltac1:(case => *//=; qauto l:on use:up_injective unfold:ren_inj))
| _ => solve_anti_ren ()
end.
Local Ltac solve_anti_ren := ltac2:(Control.enter solve_anti_ren).
Lemma ren_injective n m (a b : PTm n) (ξ : fin n -> fin m) :
(forall i j, ξ i = ξ j -> i = j) ->
Lemma ren_injective (a b : PTm) (ξ : nat -> nat) :
ren_inj ξ ->
ren_PTm ξ a = ren_PTm ξ b ->
a = b.
Proof.
move : m ξ b. elim : n / a => //; try solve_anti_ren.
move : ξ b. elim : a => //; try solve_anti_ren.
move => p ihp ξ []//=. hauto lq:on inv:PTm, nat ctrs:- use:up_injective.
Qed.
Inductive HF : Set :=
| H_Pair | H_Abs | H_Univ | H_Bind (p : BTag) | H_Nat | H_Suc | H_Zero | H_Bot.
Definition ishf {n} (a : PTm n) :=
Definition ishf (a : PTm) :=
match a with
| PPair _ _ => true
| PAbs _ => true
@ -47,7 +58,7 @@ Definition ishf {n} (a : PTm n) :=
| _ => false
end.
Definition toHF {n} (a : PTm n) :=
Definition toHF (a : PTm) :=
match a with
| PPair _ _ => H_Pair
| PAbs _ => H_Abs
@ -59,54 +70,53 @@ Definition toHF {n} (a : PTm n) :=
| _ => H_Bot
end.
Fixpoint ishne {n} (a : PTm n) :=
Fixpoint ishne (a : PTm) :=
match a with
| VarPTm _ => true
| PApp a _ => ishne a
| PProj _ a => ishne a
| PBot => true
| PInd _ n _ _ => ishne n
| _ => false
end.
Definition isbind {n} (a : PTm n) := if a is PBind _ _ _ then true else false.
Definition isbind (a : PTm) := if a is PBind _ _ _ then true else false.
Definition isuniv {n} (a : PTm n) := if a is PUniv _ then true else false.
Definition isuniv (a : PTm) := if a is PUniv _ then true else false.
Definition ispair {n} (a : PTm n) :=
Definition ispair (a : PTm) :=
match a with
| PPair _ _ => true
| _ => false
end.
Definition isnat {n} (a : PTm n) := if a is PNat then true else false.
Definition isnat (a : PTm) := if a is PNat then true else false.
Definition iszero {n} (a : PTm n) := if a is PZero then true else false.
Definition iszero (a : PTm) := if a is PZero then true else false.
Definition issuc {n} (a : PTm n) := if a is PSuc _ then true else false.
Definition issuc (a : PTm) := if a is PSuc _ then true else false.
Definition isabs {n} (a : PTm n) :=
Definition isabs (a : PTm) :=
match a with
| PAbs _ => true
| _ => false
end.
Definition ishf_ren n m (a : PTm n) (ξ : fin n -> fin m) :
Definition ishf_ren (a : PTm) (ξ : nat -> nat) :
ishf (ren_PTm ξ a) = ishf a.
Proof. case : a => //=. Qed.
Definition isabs_ren n m (a : PTm n) (ξ : fin n -> fin m) :
Definition isabs_ren (a : PTm) (ξ : nat -> nat) :
isabs (ren_PTm ξ a) = isabs a.
Proof. case : a => //=. Qed.
Definition ispair_ren n m (a : PTm n) (ξ : fin n -> fin m) :
Definition ispair_ren (a : PTm) (ξ : nat -> nat) :
ispair (ren_PTm ξ a) = ispair a.
Proof. case : a => //=. Qed.
Definition ishne_ren n m (a : PTm n) (ξ : fin n -> fin m) :
Definition ishne_ren (a : PTm) (ξ : nat -> nat) :
ishne (ren_PTm ξ a) = ishne a.
Proof. move : m ξ. elim : n / a => //=. Qed.
Proof. move : ξ. elim : a => //=. Qed.
Lemma renaming_shift n m Γ (ρ : fin n -> PTm m) A :
renaming_ok (funcomp (ren_PTm shift) (scons (subst_PTm ρ A) Γ)) Γ shift.
Proof. sfirstorder. Qed.
Lemma renaming_shift Γ (ρ : nat -> PTm) A :
renaming_ok (cons A Γ) Γ shift.
Proof. rewrite /renaming_ok. hauto lq:on ctrs:lookup. Qed.