Add nostuck antisubstitution
This commit is contained in:
parent
7fb60e1c2f
commit
c4a13daa54
2 changed files with 69 additions and 6 deletions
71
cosn.v
71
cosn.v
|
@ -9,6 +9,7 @@ Require Import Psatz.
|
||||||
From stdpp Require Import relations (rtc (..), rtc_once, rtc_r, sn).
|
From stdpp Require Import relations (rtc (..), rtc_once, rtc_r, sn).
|
||||||
From Hammer Require Import Tactics.
|
From Hammer Require Import Tactics.
|
||||||
Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common fp_red.
|
Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common fp_red.
|
||||||
|
Require Import Btauto.
|
||||||
|
|
||||||
Fixpoint nostuck (a : PTm) :=
|
Fixpoint nostuck (a : PTm) :=
|
||||||
match a with
|
match a with
|
||||||
|
@ -20,18 +21,19 @@ Fixpoint nostuck (a : PTm) :=
|
||||||
| PProj _ a => (ishf a ==> ispair a) && nostuck a
|
| PProj _ a => (ishf a ==> ispair a) && nostuck a
|
||||||
| PZero => true
|
| PZero => true
|
||||||
| PSuc a => nostuck a
|
| PSuc a => nostuck a
|
||||||
| PInd P a b c => (ishf a ==> iszero a || issuc a) && nostuck b && nostuck c
|
| PInd P a b c => nostuck P && (ishf a ==> iszero a || issuc a) && nostuck b && nostuck c
|
||||||
| PNat => true
|
| PNat => true
|
||||||
| PUniv _ => true
|
| PUniv _ => true
|
||||||
end.
|
end.
|
||||||
|
|
||||||
|
|
||||||
CoInductive safe a : Prop :=
|
CoInductive safe a : Prop :=
|
||||||
safe_intro {safe_ok : nostuck a; safe_red : forall b,RRed.R a b -> safe b}.
|
safe_intro :
|
||||||
|
nostuck a ->
|
||||||
|
(forall b,RRed.R a b -> safe b) ->
|
||||||
|
safe a.
|
||||||
|
|
||||||
Arguments safe_intro {a}.
|
Arguments safe_intro {a}.
|
||||||
Arguments safe_ok {a}.
|
|
||||||
Arguments safe_red {a}.
|
|
||||||
|
|
||||||
Lemma safe_coind P : (forall a, P a -> nostuck a /\ (forall b, RRed.R a b -> P b)) -> forall a, P a -> safe a.
|
Lemma safe_coind P : (forall a, P a -> nostuck a /\ (forall b, RRed.R a b -> P b)) -> forall a, P a -> safe a.
|
||||||
move => h.
|
move => h.
|
||||||
|
@ -42,3 +44,64 @@ Lemma safe_coind P : (forall a, P a -> nostuck a /\ (forall b, RRed.R a b -> P
|
||||||
apply ha0.
|
apply ha0.
|
||||||
move => b hb. apply ha1 in hb. apply ih. apply hb.
|
move => b hb. apply ha1 in hb. apply ih. apply hb.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
Lemma safe_app_inv0 : forall a b, safe (PApp a b) -> safe a.
|
||||||
|
Proof.
|
||||||
|
suff : forall a, (exists b, safe (PApp a b)) -> safe a by firstorder.
|
||||||
|
apply safe_coind.
|
||||||
|
sauto lqb:on.
|
||||||
|
Qed.
|
||||||
|
|
||||||
|
Lemma safe_app_inv1 : forall a b, safe (PApp a b) -> safe b.
|
||||||
|
Proof.
|
||||||
|
suff : forall b, (exists a, safe (PApp a b)) -> safe b by firstorder.
|
||||||
|
apply safe_coind.
|
||||||
|
sauto lqb:on.
|
||||||
|
Qed.
|
||||||
|
|
||||||
|
Lemma safe_abs_inv : forall a, safe (PAbs a) -> safe a.
|
||||||
|
Proof.
|
||||||
|
apply safe_coind.
|
||||||
|
sauto lqb:on.
|
||||||
|
Qed.
|
||||||
|
|
||||||
|
Lemma nostuck_antisubstitution : forall ρ a, nostuck (subst_PTm ρ a) -> nostuck a.
|
||||||
|
Proof.
|
||||||
|
suff : forall (ρ : nat -> PTm) (a : PTm), nostuck (subst_PTm ρ a) ==> nostuck a by sauto lqb:on.
|
||||||
|
move => /[swap]. elim => //=.
|
||||||
|
- move => *. rewrite !Bool.implb_orb /is_true. btauto.
|
||||||
|
- move => b ihb a iha ρ.
|
||||||
|
move /(_ ρ) : ihb. apply /implyP.
|
||||||
|
move /(_ ρ) : iha. apply /implyP.
|
||||||
|
case : b => //= *; rewrite /is_true !Bool.implb_orb; btauto.
|
||||||
|
- move => a iha b ihb ρ.
|
||||||
|
move /(_ ρ) : ihb. apply /implyP.
|
||||||
|
move /(_ ρ) : iha. apply /implyP.
|
||||||
|
rewrite /is_true !Bool.implb_orb; btauto.
|
||||||
|
- move => p u hu ρ.
|
||||||
|
move /(_ ρ) : hu. apply /implyP.
|
||||||
|
case : u => //= *; rewrite /is_true !Bool.implb_orb; btauto.
|
||||||
|
- move => _ a iha b ihb ρ.
|
||||||
|
move /(_ (up_PTm_PTm ρ)) : ihb. apply /implyP.
|
||||||
|
move /(_ ρ) : iha. apply /implyP.
|
||||||
|
rewrite /is_true !Bool.implb_orb; btauto.
|
||||||
|
- move => P ihP a iha b ihb c ihc ρ.
|
||||||
|
move /(_ (up_PTm_PTm ρ)) : ihP. apply /implyP.
|
||||||
|
move /(_ ρ) : iha. apply /implyP.
|
||||||
|
move /(_ ρ) : ihb. apply /implyP.
|
||||||
|
move /(_ (up_PTm_PTm (up_PTm_PTm ρ))) : ihc. apply /implyP.
|
||||||
|
case : a => //= *; rewrite /is_true !Bool.implb_orb; btauto.
|
||||||
|
Qed.
|
||||||
|
|
||||||
|
Lemma safe_antisubstitution : forall ρ a, safe (subst_PTm ρ a) -> safe a.
|
||||||
|
Proof.
|
||||||
|
suff : forall a, (exists ρ, safe (subst_PTm ρ a)) -> safe a by firstorder.
|
||||||
|
apply safe_coind.
|
||||||
|
move => a [ρ ha].
|
||||||
|
split.
|
||||||
|
have {}ha : nostuck (subst_PTm ρ a) by hauto lq:on inv:safe lq:on.
|
||||||
|
by eauto using nostuck_antisubstitution.
|
||||||
|
move => b hr. exists ρ.
|
||||||
|
inversion ha as [ha0 ha1].
|
||||||
|
hauto lq:on use:RRed.substing.
|
||||||
|
Qed.
|
||||||
|
|
|
@ -1270,7 +1270,7 @@ Module Type NoForbid.
|
||||||
|
|
||||||
Axiom P_PairInv : forall (a b : PTm), P (PPair a b) -> P a /\ P b.
|
Axiom P_PairInv : forall (a b : PTm), P (PPair a b) -> P a /\ P b.
|
||||||
Axiom P_ProjInv : forall p (a : PTm), P (PProj p a) -> P a.
|
Axiom P_ProjInv : forall p (a : PTm), P (PProj p a) -> P a.
|
||||||
Axiom P_renaming : forall (ξ : nat -> nat) a , P (ren_PTm ξ a) <-> P a.
|
Axiom P_renaming : forall (ξ : nat -> nat) a , P (ren_PTm ξ a) -> P a.
|
||||||
End NoForbid.
|
End NoForbid.
|
||||||
|
|
||||||
Module Type NoForbid_FactSig (M : NoForbid).
|
Module Type NoForbid_FactSig (M : NoForbid).
|
||||||
|
@ -1342,7 +1342,7 @@ Module SN_NoForbid <: NoForbid.
|
||||||
induction h; sauto lq:on rew:off.
|
induction h; sauto lq:on rew:off.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma P_renaming : forall (ξ : nat -> nat) a , P (ren_PTm ξ a) <-> P a.
|
Lemma P_renaming : forall (ξ : nat -> nat) a , P (ren_PTm ξ a) -> P a.
|
||||||
Proof. hauto lq:on use:sn_antirenaming, sn_renaming. Qed.
|
Proof. hauto lq:on use:sn_antirenaming, sn_renaming. Qed.
|
||||||
|
|
||||||
Lemma P_ProjBind : forall p p' (A : PTm) B, ~ P (PProj p (PBind p' A B)).
|
Lemma P_ProjBind : forall p p' (A : PTm) B, ~ P (PProj p (PBind p' A B)).
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue