Prove selected upgrade
This commit is contained in:
parent
a28b7fd808
commit
fcdbee2a91
1 changed files with 139 additions and 6 deletions
|
@ -1,5 +1,5 @@
|
|||
Require Import Autosubst2.syntax
|
||||
Autosubst2.core Autosubst2.unscoped typing ssreflect List.
|
||||
Autosubst2.core Autosubst2.unscoped typing ssreflect List ssrbool.
|
||||
|
||||
From Hammer Require Import Tactics.
|
||||
|
||||
|
@ -39,6 +39,18 @@ End ren.
|
|||
|
||||
Arguments ren_ok {T}.
|
||||
|
||||
Definition bool_leb a b :=
|
||||
match a,b with
|
||||
| true, false => false
|
||||
| _,_ => true
|
||||
end.
|
||||
|
||||
Lemma bool_leb_refl a : bool_leb a a.
|
||||
Proof. case : a => //=. Qed.
|
||||
|
||||
Lemma bool_le_spec a b : Bool.reflect (Bool.le a b) (bool_leb a b).
|
||||
Proof. hauto l:on inv:bool ctrs:Bool.reflect. Qed.
|
||||
|
||||
Module Liftable.
|
||||
Lemma lookup_fixed i Γ ℓ A :
|
||||
lookup i Γ (ℓ, A) ->
|
||||
|
@ -74,11 +86,92 @@ Module Liftable.
|
|||
Proof.
|
||||
elim : a ξ Φ Ψ => //=; hauto l:on use:ren_up unfold:ren_ok.
|
||||
Qed.
|
||||
|
||||
Fixpoint ctx_leq Φ Ψ :=
|
||||
match Φ, Ψ with
|
||||
| nil, nil => true
|
||||
| ℓ::Φ, ℓ'::Ψ => bool_leb ℓ ℓ' && ctx_leq Φ Ψ
|
||||
| _, _ => false
|
||||
end.
|
||||
|
||||
Lemma ctx_leq_lookup Φ Ψ i ℓ1 :
|
||||
ctx_leq Φ Ψ ->
|
||||
lookup i Ψ ℓ1 ->
|
||||
exists ℓ0, lookup i Φ ℓ0 /\
|
||||
Bool.le ℓ0 ℓ1.
|
||||
Proof.
|
||||
move => + h. move : Φ.
|
||||
elim : i Ψ ℓ1 / h.
|
||||
- move => A Γ Ψ h.
|
||||
case : Ψ h => //=.
|
||||
move => B Ψ.
|
||||
case : bool_le_spec => //=.
|
||||
sfirstorder ctrs:lookup.
|
||||
- move => i Γ A B hi ihi Ψ.
|
||||
case : Ψ => //=.
|
||||
move => C Ψ.
|
||||
case : bool_le_spec => //=.
|
||||
hauto lq:on ctrs:lookup.
|
||||
Qed.
|
||||
|
||||
Lemma narrowing Φ Ψ a :
|
||||
ctx_leq Φ Ψ ->
|
||||
liftable Ψ a ->
|
||||
liftable Φ a.
|
||||
Proof.
|
||||
elim : a Φ Ψ => //=.
|
||||
- move => n Φ Ψ.
|
||||
move : ctx_leq_lookup; repeat move/[apply].
|
||||
move => [ℓ1 [h0 h1]].
|
||||
suff : ℓ1 = false by congruence.
|
||||
hauto q:on inv:bool.
|
||||
- sfirstorder.
|
||||
- sfirstorder.
|
||||
- sfirstorder.
|
||||
Qed.
|
||||
|
||||
End Liftable.
|
||||
|
||||
|
||||
Definition lvl_leb a b :=
|
||||
match a, b with
|
||||
| Fixed ℓ0, Fixed ℓ1 => Bool.eqb ℓ0 ℓ1
|
||||
| Floating ℓ0, Floating ℓ1 => bool_leb ℓ0 ℓ1
|
||||
| _,_ => false
|
||||
end.
|
||||
|
||||
Lemma lvl_leb_refl a : lvl_leb a a.
|
||||
Proof. destruct a,b; sfirstorder use:Bool.eqb_reflx, bool_leb_refl. Qed.
|
||||
|
||||
Module Typing.
|
||||
|
||||
Inductive ctx_fleq : list bool -> ctx -> ctx -> Prop :=
|
||||
| FL_Nil : ctx_fleq nil nil nil
|
||||
| FL_ConsFixed Φ Γ Δ A :
|
||||
ctx_fleq Φ Γ Δ ->
|
||||
ctx_fleq (false :: Φ) (A :: Γ) (A :: Δ)
|
||||
| FL_ConsFloat Φ Γ Δ ℓ0 ℓ1 A :
|
||||
ctx_fleq Φ Γ Δ ->
|
||||
lvl_leb ℓ0 ℓ1 ->
|
||||
ctx_fleq (true :: Φ) ((ℓ0, A) :: Γ) ((ℓ1, A) :: Δ).
|
||||
|
||||
Lemma ctx_fleq_fixed Φ Γ Δ :
|
||||
ctx_fleq Φ Γ Δ ->
|
||||
ctx_fleq (fixed_ctx Γ) Γ Δ.
|
||||
Proof.
|
||||
move => h. elim : Φ Γ Δ/ h => //=.
|
||||
- constructor.
|
||||
- move => Φ Γ Δ [ℓ B].
|
||||
move => h0 h1.
|
||||
destruct ℓ => //=; constructor; eauto using lvl_leb_refl.
|
||||
- move => Φ Γ Δ ℓ0 ℓ1 A hctx ihctx h0.
|
||||
destruct ℓ0, ℓ1 => //=.
|
||||
simpl in h0.
|
||||
move : h0. case : Bool.eqb_spec => //= ? _. subst.
|
||||
by constructor.
|
||||
by constructor.
|
||||
Qed.
|
||||
|
||||
Lemma renaming ξ Γ Δ ℓ a A :
|
||||
Δ ⊢ a ;; ℓ ∈ A ->
|
||||
ren_ok ξ Γ Δ ->
|
||||
|
@ -96,11 +189,51 @@ Module Typing.
|
|||
- hauto lq:on ctrs:Wt use:ren_up.
|
||||
Qed.
|
||||
|
||||
(* Liftable: stable *)
|
||||
(* If a well-typed term is liftable, then it should remain well-typed at the same level
|
||||
if the floating variables are lifted *)
|
||||
Lemma ctx_fleq_lookup i Φ (Γ Δ : ctx) A :
|
||||
ctx_fleq Φ Γ Δ ->
|
||||
lookup i Φ false ->
|
||||
lookup i Γ A ->
|
||||
lookup i Δ A.
|
||||
Proof.
|
||||
move => h. move : A i.
|
||||
elim : Φ Γ Δ / h => //=; sauto.
|
||||
Qed.
|
||||
|
||||
(* Upgrade *)
|
||||
(* If Γ ⊢ a ;; ℓ ∈ A, we want to lift up the levels in Γ *)
|
||||
Lemma ctx_fleq_fixed_eq Φ Γ Δ :
|
||||
ctx_fleq Φ Γ Δ ->
|
||||
fixed_ctx Γ = fixed_ctx Δ.
|
||||
Proof.
|
||||
move => h. elim : Φ Γ Δ / h => //=.
|
||||
- hauto lq:on.
|
||||
- move => Φ Γ Δ ℓ0 ℓ1 h ih h0 h1.
|
||||
f_equal => //.
|
||||
move {h0 ih h}.
|
||||
hauto lq:on inv:Level.
|
||||
Qed.
|
||||
|
||||
Lemma upgrading Φ Γ Δ a ℓ A :
|
||||
Γ ⊢ a ;; ℓ ∈ A ->
|
||||
liftable Φ a ->
|
||||
ctx_fleq Φ Γ Δ ->
|
||||
Δ ⊢ a ;; ℓ ∈ A.
|
||||
Proof.
|
||||
move => h. move : Φ Δ.
|
||||
elim : Γ a ℓ A /h.
|
||||
- move => Γ ℓ i ℓ0 A hℓ hi Φ Δ /= h0 h1.
|
||||
econstructor; eauto.
|
||||
eauto using ctx_fleq_lookup.
|
||||
- hauto lq:on ctrs:Wt.
|
||||
- move => Γ ℓ a A B ha iha Φ Δ hla hleq /=.
|
||||
constructor. apply : iha => /=; cycle 1; eauto.
|
||||
by constructor.
|
||||
- move => Γ ℓ ℓ0 a b A B hb ihb ha iha hla Φ Δ /= hlb hctx.
|
||||
apply : T_App; eauto.
|
||||
apply : iha; eauto.
|
||||
sfirstorder use:ctx_fleq_fixed.
|
||||
erewrite <- ctx_fleq_fixed_eq; eauto.
|
||||
- move => Γ ℓ ℓ0 b A B hb ihb Δ /=.
|
||||
move => h0 h1. constructor. apply : ihb; eauto.
|
||||
simpl. by constructor.
|
||||
Qed.
|
||||
|
||||
End Typing.
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue