Seems to work but takes a million years to type check

This commit is contained in:
Yiyun Liu 2025-02-27 21:33:25 -05:00
parent 6c11f5560d
commit 7503dea251
5 changed files with 580 additions and 572 deletions

View file

@ -16,7 +16,7 @@ uninstall: $(COQMAKEFILE)
$(MAKE) -f $(COQMAKEFILE) uninstall $(MAKE) -f $(COQMAKEFILE) uninstall
theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v : syntax.sig theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v : syntax.sig
autosubst -f -v ge813 -s coq -o theories/Autosubst2/syntax.v syntax.sig autosubst -f -v ge813 -s ucoq -o theories/Autosubst2/syntax.v syntax.sig
.PHONY: clean FORCE .PHONY: clean FORCE

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,213 @@
(** * Autosubst Header for Unnamed Syntax
Version: December 11, 2019.
*)
(* Adrian:
I changed this library a bit to work better with my generated code.
1. I use nat directly instead of defining fin to be nat and using Some/None as S/O
2. I removed the "s, sigma" notation for scons because it interacts with dependent function types "forall x, A"*)
Require Import core.
Require Import Setoid Morphisms Relation_Definitions.
Definition ap {X Y} (f : X -> Y) {x y : X} (p : x = y) : f x = f y :=
match p with eq_refl => eq_refl end.
Definition apc {X Y} {f g : X -> Y} {x y : X} (p : f = g) (q : x = y) : f x = g y :=
match q with eq_refl => match p with eq_refl => eq_refl end end.
(** ** Primitives of the Sigma Calculus. *)
Definition shift := S.
Definition var_zero := 0.
Definition id {X} := @Datatypes.id X.
Definition scons {X: Type} (x : X) (xi : nat -> X) :=
fun n => match n with
| 0 => x
| S n => xi n
end.
#[ export ]
Hint Opaque scons : rewrite.
(** ** Type Class Instances for Notation
Required to make notation work. *)
(** *** Type classes for renamings. *)
Class Ren1 (X1 : Type) (Y Z : Type) :=
ren1 : X1 -> Y -> Z.
Class Ren2 (X1 X2 : Type) (Y Z : Type) :=
ren2 : X1 -> X2 -> Y -> Z.
Class Ren3 (X1 X2 X3 : Type) (Y Z : Type) :=
ren3 : X1 -> X2 -> X3 -> Y -> Z.
Class Ren4 (X1 X2 X3 X4 : Type) (Y Z : Type) :=
ren4 : X1 -> X2 -> X3 -> X4 -> Y -> Z.
Class Ren5 (X1 X2 X3 X4 X5 : Type) (Y Z : Type) :=
ren5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z.
Module RenNotations.
Notation "s ⟨ xi1 ⟩" := (ren1 xi1 s) (at level 7, left associativity, format "s ⟨ xi1 ⟩") : subst_scope.
Notation "s ⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ⟩") : subst_scope.
Notation "s ⟨ xi1 ; xi2 ; xi3 ⟩" := (ren3 xi1 xi2 xi3 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ⟩") : subst_scope.
Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩" := (ren4 xi1 xi2 xi3 xi4 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩") : subst_scope.
Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩" := (ren5 xi1 xi2 xi3 xi4 xi5 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩") : subst_scope.
Notation "⟨ xi ⟩" := (ren1 xi) (at level 1, left associativity, format "⟨ xi ⟩") : fscope.
Notation "⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2) (at level 1, left associativity, format "⟨ xi1 ; xi2 ⟩") : fscope.
End RenNotations.
(** *** Type Classes for Substiution *)
Class Subst1 (X1 : Type) (Y Z: Type) :=
subst1 : X1 -> Y -> Z.
Class Subst2 (X1 X2 : Type) (Y Z: Type) :=
subst2 : X1 -> X2 -> Y -> Z.
Class Subst3 (X1 X2 X3 : Type) (Y Z: Type) :=
subst3 : X1 -> X2 -> X3 -> Y -> Z.
Class Subst4 (X1 X2 X3 X4: Type) (Y Z: Type) :=
subst4 : X1 -> X2 -> X3 -> X4 -> Y -> Z.
Class Subst5 (X1 X2 X3 X4 X5 : Type) (Y Z: Type) :=
subst5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z.
Module SubstNotations.
Notation "s [ sigma ]" := (subst1 sigma s) (at level 7, left associativity, format "s '/' [ sigma ]") : subst_scope.
Notation "s [ sigma ; tau ]" := (subst2 sigma tau s) (at level 7, left associativity, format "s '/' [ sigma ; '/' tau ]") : subst_scope.
End SubstNotations.
(** *** Type Class for Variables *)
Class Var X Y :=
ids : X -> Y.
#[export] Instance idsRen : Var nat nat := id.
(** ** Proofs for the substitution primitives. *)
Arguments funcomp {X Y Z} (g)%fscope (f)%fscope.
Module CombineNotations.
Notation "f >> g" := (funcomp g f) (at level 50) : fscope.
Notation "s .: sigma" := (scons s sigma) (at level 55, sigma at next level, right associativity) : subst_scope.
#[ global ]
Open Scope fscope.
#[ global ]
Open Scope subst_scope.
End CombineNotations.
Import CombineNotations.
(** A generic lifting of a renaming. *)
Definition up_ren (xi : nat -> nat) :=
0 .: (xi >> S).
(** A generic proof that lifting of renamings composes. *)
Lemma up_ren_ren (xi: nat -> nat) (zeta : nat -> nat) (rho: nat -> nat) (E: forall x, (xi >> zeta) x = rho x) :
forall x, (up_ren xi >> up_ren zeta) x = up_ren rho x.
Proof.
intros [|x].
- reflexivity.
- unfold up_ren. cbn. unfold funcomp. f_equal. apply E.
Qed.
(** Eta laws. *)
Lemma scons_eta' {T} (f : nat -> T) :
pointwise_relation _ eq (f var_zero .: (funcomp f shift)) f.
Proof. intros x. destruct x; reflexivity. Qed.
Lemma scons_eta_id' :
pointwise_relation _ eq (var_zero .: shift) id.
Proof. intros x. destruct x; reflexivity. Qed.
Lemma scons_comp' (T: Type) {U} (s: T) (sigma: nat -> T) (tau: T -> U) :
pointwise_relation _ eq (funcomp tau (s .: sigma)) ((tau s) .: (funcomp tau sigma)).
Proof. intros x. destruct x; reflexivity. Qed.
(* Morphism for Setoid Rewriting. The only morphism that can be defined statically. *)
#[export] Instance scons_morphism {X: Type} :
Proper (eq ==> pointwise_relation _ eq ==> pointwise_relation _ eq) (@scons X).
Proof.
intros ? t -> sigma tau H.
intros [|x].
cbn. reflexivity.
apply H.
Qed.
#[export] Instance scons_morphism2 {X: Type} :
Proper (eq ==> pointwise_relation _ eq ==> eq ==> eq) (@scons X).
Proof.
intros ? t -> sigma tau H ? x ->.
destruct x as [|x].
cbn. reflexivity.
apply H.
Qed.
(** ** Generic lifting of an allfv predicate *)
Definition up_allfv (p: nat -> Prop) : nat -> Prop := scons True p.
(** ** Notations for unscoped syntax *)
Module UnscopedNotations.
Include RenNotations.
Include SubstNotations.
Include CombineNotations.
(* Notation "s , sigma" := (scons s sigma) (at level 60, format "s , sigma", right associativity) : subst_scope. *)
Notation "s '..'" := (scons s ids) (at level 1, format "s ..") : subst_scope.
Notation "" := (shift) : subst_scope.
#[global]
Open Scope fscope.
#[global]
Open Scope subst_scope.
End UnscopedNotations.
(** ** Tactics for unscoped syntax *)
(** Automatically does a case analysis on a natural number, useful for proofs with context renamings/context morphisms. *)
Tactic Notation "auto_case" tactic(t) := (match goal with
| [|- forall (i : nat), _] => intros []; t
end).
(** Generic fsimpl tactic: simplifies the above primitives in a goal. *)
Ltac fsimpl :=
repeat match goal with
| [|- context[id >> ?f]] => change (id >> f) with f (* AsimplCompIdL *)
| [|- context[?f >> id]] => change (f >> id) with f (* AsimplCompIdR *)
| [|- context [id ?s]] => change (id s) with s
| [|- context[(?f >> ?g) >> ?h]] => change ((f >> g) >> h) with (f >> (g >> h))
| [|- context[(?v .: ?g) var_zero]] => change ((v .: g) var_zero) with v
| [|- context[(?v .: ?g) 0]] => change ((v .: g) 0) with v
| [|- context[(?v .: ?g) (S ?n)]] => change ((v .: g) (S n)) with (g n)
| [|- context[?f >> (?x .: ?g)]] => change (f >> (x .: g)) with g (* f should evaluate to shift *)
| [|- context[var_zero]] => change var_zero with 0
| [|- context[?x2 .: (funcomp ?f shift)]] => change (scons x2 (funcomp f shift)) with (scons (f var_zero) (funcomp f shift)); setoid_rewrite (@scons_eta' _ _ f)
| [|- context[?f var_zero .: ?g]] => change (scons (f var_zero) g) with (scons (f var_zero) (funcomp f shift)); rewrite scons_eta'
| [|- _ = ?h (?f ?s)] => change (h (f s)) with ((f >> h) s)
| [|- ?h (?f ?s) = _] => change (h (f s)) with ((f >> h) s)
(* DONE had to put an underscore as the last argument to scons. This might be an argument against unfolding funcomp *)
| [|- context[funcomp _ (scons _ _)]] => setoid_rewrite scons_comp'; eta_reduce
| [|- context[scons var_zero shift]] => setoid_rewrite scons_eta_id'; eta_reduce
end.

View file

@ -4,6 +4,7 @@ Import Ltac2.Notations.
Import Ltac2.Control. Import Ltac2.Control.
From Hammer Require Import Tactics. From Hammer Require Import Tactics.
Definition renaming_ok {n m} (Γ : fin n -> PTm n) (Δ : fin m -> PTm m) (ξ : fin m -> fin n) := 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). forall (i : fin m), ren_PTm ξ (Δ i) = Γ (ξ i).

View file

@ -1,13 +1,64 @@
From Equations Require Import Equations. From Equations Require Import Equations.
Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax.
common typing preservation admissible fp_red structural soundness. Derive NoConfusion for nat PTag BTag PTm.
Require Import algorithmic.
From stdpp Require Import relations (rtc(..), nsteps(..)).
Require Import ssreflect ssrbool. Require Import ssreflect ssrbool.
From Hammer Require Import Tactics. From Hammer Require Import Tactics.
Inductive algo_dom {n} : PTm n -> PTm n -> Prop := Definition ishf (a : PTm) :=
match a with
| PPair _ _ => true
| PAbs _ => true
| PUniv _ => true
| PBind _ _ _ => true
| PNat => true
| PSuc _ => true
| PZero => true
| _ => false
end.
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.
Module HRed.
Inductive R : PTm -> PTm -> Prop :=
(****************** Beta ***********************)
| AppAbs a b :
R (PApp (PAbs a) b) (subst_PTm (scons b VarPTm) a)
| ProjPair p a b :
R (PProj p (PPair a b)) (if p is PL then a else b)
| IndZero P b c :
R (PInd P PZero b c) b
| IndSuc P a b c :
R (PInd P (PSuc a) b c) (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c)
(*************** Congruence ********************)
| AppCong a0 a1 b :
R a0 a1 ->
R (PApp a0 b) (PApp a1 b)
| ProjCong p a0 a1 :
R a0 a1 ->
R (PProj p a0) (PProj p a1)
| IndCong P a0 a1 b c :
R a0 a1 ->
R (PInd P a0 b c) (PInd P a1 b c).
Definition nf a := forall b, ~ R a b.
End HRed.
Inductive algo_dom : PTm -> PTm -> Prop :=
| A_AbsAbs a b : | A_AbsAbs a b :
algo_dom_r a b -> algo_dom_r a b ->
(* --------------------- *) (* --------------------- *)
@ -74,7 +125,7 @@ Inductive algo_dom {n} : PTm n -> PTm n -> Prop :=
(* ------------------------- *) (* ------------------------- *)
algo_dom (PApp u0 a0) (PApp u1 a1) algo_dom (PApp u0 a0) (PApp u1 a1)
with algo_dom_r {n} : PTm n -> PTm n -> Prop := with algo_dom_r : PTm -> PTm -> Prop :=
| A_NfNf a b : | A_NfNf a b :
algo_dom a b -> algo_dom a b ->
algo_dom_r a b algo_dom_r a b
@ -92,67 +143,26 @@ with algo_dom_r {n} : PTm n -> PTm n -> Prop :=
(* ----------------------- *) (* ----------------------- *)
algo_dom_r a b. algo_dom_r a b.
Derive Signature for algo_dom algo_dom_r. Lemma algo_dom_hf_hne (a b : PTm) :
Derive NoConfusion for PTm.
Next Obligation.
Admitted.
Next Obligation.
Admitted.
Derive Dependent Inversion adom_inv with (forall n (a b : PTm n), algo_dom a b) Sort Prop.
Lemma algo_dom_hf_hne n (a b : PTm n) :
algo_dom a b -> algo_dom a b ->
(ishf a \/ ishne a) /\ (ishf b \/ ishne b). (ishf a \/ ishne a) /\ (ishf b \/ ishne b).
Proof. Proof.
induction 1 =>//=; hauto lq:on. induction 1 =>//=; hauto lq:on.
Qed. Qed.
Lemma hf_no_hred n (a b : PTm n) : Lemma hf_no_hred (a b : PTm) :
ishf a -> ishf a ->
HRed.R a b -> HRed.R a b ->
False. False.
Proof. hauto l:on inv:HRed.R. Qed. Proof. hauto l:on inv:HRed.R. Qed.
Lemma hne_no_hred n (a b : PTm n) : Lemma hne_no_hred (a b : PTm) :
ishne a -> ishne a ->
HRed.R a b -> HRed.R a b ->
False. False.
Proof. elim : a b => //=; hauto l:on inv:HRed.R. Qed. Proof. elim : a b => //=; hauto l:on inv:HRed.R. Qed.
Definition fin_beq {n} (i j : fin n) : bool. Derive Signature for algo_dom.
Proof.
induction n.
- by exfalso.
- refine (match i , j with
| None, None => true
| Some i, Some j => IHn i j
| _, _ => false
end).
Defined.
Lemma fin_eq_dec {n} (i j : fin n) :
Bool.reflect (i = j) (fin_beq i j).
Proof.
revert i j. induction n.
- destruct i.
- destruct i; destruct j.
+ specialize (IHn f f0).
inversion IHn; subst.
simpl. rewrite -H.
apply ReflectT.
reflexivity.
simpl. rewrite -H.
apply ReflectF.
injection. tauto.
+ by apply ReflectF.
+ by apply ReflectF.
+ by apply ReflectT.
Defined.
Scheme Equality for PTag.
Scheme Equality for BTag.
(* Fixpoint PTm_eqb {n} (a b : PTm n) := *) (* Fixpoint PTm_eqb {n} (a b : PTm n) := *)
(* match a, b with *) (* match a, b with *)
(* | VarPTm i, VarPTm j => fin_eq i j *) (* | VarPTm i, VarPTm j => fin_eq i j *)
@ -171,7 +181,7 @@ Scheme Equality for BTag.
(* destruct IHa1. *) (* destruct IHa1. *)
(* destruct a1. *) (* destruct a1. *)
Fixpoint hred {n} (a : PTm n) : option (PTm n) := Fixpoint hred (a : PTm) : option (PTm) :=
match a with match a with
| VarPTm i => None | VarPTm i => None
| PAbs a => None | PAbs a => None
@ -204,31 +214,31 @@ Fixpoint hred {n} (a : PTm n) : option (PTm n) :=
end end
end. end.
Lemma hred_complete n (a b : PTm n) : Lemma hred_complete (a b : PTm) :
HRed.R a b -> hred a = Some b. HRed.R a b -> hred a = Some b.
Proof. Proof.
induction 1; hauto lq:on rew:off inv:HRed.R b:on. induction 1; hauto lq:on rew:off inv:HRed.R b:on.
Qed. Qed.
Lemma hred_sound n (a b : PTm n): Lemma hred_sound (a b : PTm):
hred a = Some b -> HRed.R a b. hred a = Some b -> HRed.R a b.
Proof. Proof.
elim : a b; hauto q:on dep:on ctrs:HRed.R. elim : a b; hauto q:on dep:on ctrs:HRed.R.
Qed. Qed.
Lemma hred_deter n (a b0 b1 : PTm n) : Lemma hred_deter (a b0 b1 : PTm) :
HRed.R a b0 -> HRed.R a b1 -> b0 = b1. HRed.R a b0 -> HRed.R a b1 -> b0 = b1.
Proof. Proof.
move /hred_complete => + /hred_complete. congruence. move /hred_complete => + /hred_complete. congruence.
Qed. Qed.
Definition hred_fancy n (a : PTm n) : Definition hred_fancy (a : PTm) :
relations.nf HRed.R a + {x | HRed.R a x}. HRed.nf a + {x | HRed.R a x}.
Proof. Proof.
destruct (hred a) as [a'|] eqn:eq . destruct (hred a) as [a'|] eqn:eq .
- right. exists a'. hauto q:on use:hred_sound. - right. exists a'. hauto q:on use:hred_sound.
- left. - left.
move => [a' h]. move => a' h.
move /hred_complete in h. move /hred_complete in h.
congruence. congruence.
Defined. Defined.
@ -241,6 +251,8 @@ Ltac check_equal_triv :=
| _ => idtac | _ => idtac
end. end.
Scheme Equality for nat. Scheme Equality for PTag.
Scheme Equality for BTag. Scheme Equality for PTm.
(* Program Fixpoint check_equal {n} (a b : PTm n) (h : algo_dom a b) {struct h} : bool := *) (* Program Fixpoint check_equal {n} (a b : PTm n) (h : algo_dom a b) {struct h} : bool := *)
(* match a, b with *) (* match a, b with *)
(* | VarPTm i, VarPTm j => fin_beq i j *) (* | VarPTm i, VarPTm j => fin_beq i j *)
@ -255,9 +267,9 @@ Ltac check_equal_triv :=
(* Next Obligation. *) (* Next Obligation. *)
(* simpl. *) (* simpl. *)
Equations check_equal {n} (a b : PTm n) (h : algo_dom a b) : Equations check_equal (a b : PTm) (h : algo_dom a b) :
bool by struct h := bool by struct h :=
check_equal (VarPTm i) (VarPTm j) h := fin_beq i j; check_equal (VarPTm i) (VarPTm j) h := nat_beq i j;
check_equal (PAbs a) (PAbs b) h := check_equal_r a b ltac:(check_equal_triv); check_equal (PAbs a) (PAbs b) h := check_equal_r a b ltac:(check_equal_triv);
check_equal (PAbs a) b h := check_equal_r a (PApp (ren_PTm shift b) (VarPTm var_zero)) ltac:(check_equal_triv); check_equal (PAbs a) b h := check_equal_r a (PApp (ren_PTm shift b) (VarPTm var_zero)) ltac:(check_equal_triv);
check_equal a (PAbs b) h := check_equal_r (PApp (ren_PTm shift a) (VarPTm var_zero)) b ltac:(check_equal_triv); check_equal a (PAbs b) h := check_equal_r (PApp (ren_PTm shift a) (VarPTm var_zero)) b ltac:(check_equal_triv);
@ -274,26 +286,24 @@ Equations check_equal {n} (a b : PTm n) (h : algo_dom a b) :
check_equal (PSuc a) (PSuc b) h := check_equal_r a b ltac:(check_equal_triv); check_equal (PSuc a) (PSuc b) h := check_equal_r a b ltac:(check_equal_triv);
check_equal (PUniv i) (PUniv j) _ := Nat.eqb i j; check_equal (PUniv i) (PUniv j) _ := Nat.eqb i j;
check_equal a b h := false; check_equal a b h := false;
with check_equal_r {n} (a b : PTm n) (h : algo_dom_r a b) : with check_equal_r (a b : PTm) (h : algo_dom_r a b) :
bool by struct h := bool by struct h :=
check_equal_r a b h with hred_fancy _ a => check_equal_r a b h with hred_fancy a =>
{ check_equal_r a b h (inr a') := check_equal_r (proj1_sig a') b _; { check_equal_r a b h (inr a') := check_equal_r (proj1_sig a') b _;
check_equal_r a b h (inl _) with hred_fancy _ b => check_equal_r a b h (inl _) with hred_fancy b =>
{ check_equal_r a b h (inl _) (inl _) := check_equal a b _; { check_equal_r a b h (inl _) (inl _) := check_equal a b _;
check_equal_r a b h (inl _) (inr b') := check_equal_r a (proj1_sig b') _}} . check_equal_r a b h (inl _) (inr b') := check_equal_r a (proj1_sig b') _}} .
Next Obligation. Next Obligation.
move => /= ih ihr n a nfa b nfb. inversion h; subst=>//=.
inversion 1; subst=>//=.
exfalso. sfirstorder. exfalso. sfirstorder.
exfalso. sfirstorder. exfalso. sfirstorder.
Defined. Defined.
Next Obligation. Next Obligation.
simpl. simpl.
move => /= ih ihr n a nfa b [b' hb']. inversion h; subst =>//=.
inversion 1; subst =>//=.
exfalso. hauto lq:on use:algo_dom_hf_hne, hf_no_hred, hne_no_hred. exfalso. hauto lq:on use:algo_dom_hf_hne, hf_no_hred, hne_no_hred.
exfalso. sfirstorder. exfalso. sfirstorder.
have ? : b' = b'0 by eauto using hred_deter. have ? : b' = b'0 by eauto using hred_deter.
@ -302,8 +312,7 @@ Next Obligation.
Defined. Defined.
Next Obligation. Next Obligation.
simpl => ih ihr n a [a' ha'] b. inversion h; subst => //=.
inversion 1; subst => //=.
exfalso. hauto lq:on use:algo_dom_hf_hne, hf_no_hred, hne_no_hred. exfalso. hauto lq:on use:algo_dom_hf_hne, hf_no_hred, hne_no_hred.
suff ? : a'0 = a' by subst; assumption. suff ? : a'0 = a' by subst; assumption.
by eauto using hred_deter. by eauto using hred_deter.