Skip to content

Commit

Permalink
change mcltt to mctt
Browse files Browse the repository at this point in the history
  • Loading branch information
HuStmpHrrr committed Nov 23, 2024
1 parent 295ba57 commit 4a412da
Show file tree
Hide file tree
Showing 5 changed files with 44 additions and 44 deletions.
16 changes: 8 additions & 8 deletions theories/Core/Completeness/EqualityCases.v
Original file line number Diff line number Diff line change
Expand Up @@ -30,7 +30,7 @@ Proof.
per_univ_elem_econstructor; mauto 3; try solve_refl.
Qed.
#[export]
Hint Resolve rel_exp_eq_cong : mcltt.
Hint Resolve rel_exp_eq_cong : mctt.

Lemma valid_exp_eq : forall {Γ i A M1 M2},
{{ Γ ⊨ A : Type@i }} ->
Expand All @@ -40,7 +40,7 @@ Lemma valid_exp_eq : forall {Γ i A M1 M2},
Proof. mauto. Qed.

#[export]
Hint Resolve valid_exp_eq : mcltt.
Hint Resolve valid_exp_eq : mctt.

Lemma rel_exp_refl_cong : forall {Γ i A A' M M'},
{{ Γ ⊨ A ≈ A' : Type@i }} ->
Expand All @@ -67,7 +67,7 @@ Proof.
symmetry; mauto 3.
Qed.
#[export]
Hint Resolve rel_exp_refl_cong : mcltt.
Hint Resolve rel_exp_refl_cong : mctt.

Lemma rel_exp_eq_sub : forall {Γ σ Δ i A M1 M2},
{{ Γ ⊨s σ : Δ }} ->
Expand Down Expand Up @@ -97,7 +97,7 @@ Proof.
Qed.

#[export]
Hint Resolve rel_exp_eq_sub : mcltt.
Hint Resolve rel_exp_eq_sub : mctt.

Lemma rel_exp_refl_sub : forall {Γ σ Δ i A M},
{{ Γ ⊨s σ : Δ }} ->
Expand Down Expand Up @@ -126,7 +126,7 @@ Proof.
Qed.

#[export]
Hint Resolve rel_exp_refl_sub : mcltt.
Hint Resolve rel_exp_refl_sub : mctt.

Lemma rel_exp_eqrec_sub : forall {Γ σ Δ i A M1 M2 j B BR N},
{{ Γ ⊨s σ : Δ }} ->
Expand Down Expand Up @@ -208,7 +208,7 @@ Proof.
It might be better to do some optimization first. *)
Admitted.
#[export]
Hint Resolve rel_exp_eqrec_sub : mcltt.
Hint Resolve rel_exp_eqrec_sub : mctt.

Lemma rel_exp_eqrec_cong : forall {Γ i A A' M1 M1' M2 M2' j B B' BR BR' N N'},
{{ Γ ⊨ A : Type@i }} ->
Expand All @@ -226,7 +226,7 @@ Lemma rel_exp_eqrec_cong : forall {Γ i A A' M1 M1' M2 M2' j B B' BR BR' N N'},
Proof.
Admitted.
#[export]
Hint Resolve rel_exp_eqrec_cong : mcltt.
Hint Resolve rel_exp_eqrec_cong : mctt.

Lemma rel_exp_eqrec_beta : forall {Γ i A M j B BR},
{{ Γ ⊨ A : Type@i }} ->
Expand All @@ -239,4 +239,4 @@ Lemma rel_exp_eqrec_beta : forall {Γ i A M j B BR},
Proof.
Admitted.
#[export]
Hint Resolve rel_exp_eqrec_beta : mcltt.
Hint Resolve rel_exp_eqrec_beta : mctt.
4 changes: 2 additions & 2 deletions theories/Core/Semantic/PER/Definitions.v
Original file line number Diff line number Diff line change
Expand Up @@ -82,7 +82,7 @@ Variant per_eq (point_rel : relation domain) m1 m2' : relation domain :=
{{ Dom ⇑ a n ≈ ⇑ a' n' ∈ per_eq point_rel m1 m2' }} }
.
#[export]
Hint Constructors per_eq : mcltt.
Hint Constructors per_eq : mctt.

Variant per_ne : relation domain :=
| per_ne_neut :
Expand Down Expand Up @@ -319,7 +319,7 @@ Variant cons_per_ctx_env tail_rel (head_rel : forall {ρ ρ'} (equiv_ρ_ρ' : {{
{{ Dom ^(ρ 0) ≈ ^(ρ' 0) ∈ head_rel equiv_ρ_drop_ρ'_drop }} ->
{{ Dom ρ ≈ ρ' ∈ cons_per_ctx_env tail_rel (@head_rel) }} }.
#[export]
Hint Constructors cons_per_ctx_env : mcltt.
Hint Constructors cons_per_ctx_env : mctt.

Inductive per_ctx_env : relation env -> ctx -> ctx -> Prop :=
| per_ctx_env_nil :
Expand Down
4 changes: 2 additions & 2 deletions theories/Core/Semantic/PER/Lemmas.v
Original file line number Diff line number Diff line change
Expand Up @@ -184,7 +184,7 @@ Proof with mautosolve 3.
Qed.

#[export]
Hint Resolve per_eq_sym : mcltt.
Hint Resolve per_eq_sym : mctt.

Lemma per_eq_trans : forall point_rel m1 m2 n n' n'',
PER point_rel ->
Expand All @@ -198,7 +198,7 @@ Proof with mautosolve 3.
Qed.

#[export]
Hint Resolve per_eq_trans : mcltt.
Hint Resolve per_eq_trans : mctt.

#[export]
Instance per_eq_PER {point_rel m1 m2} {Hpoint : PER point_rel} : PER (per_eq point_rel m1 m2).
Expand Down
6 changes: 3 additions & 3 deletions theories/Core/Soundness/LogicalRelation/CoreLemmas.v
Original file line number Diff line number Diff line change
Expand Up @@ -679,7 +679,7 @@ Proof.
Qed.

#[local]
Hint Resolve glu_nat_resp_per_nat : mcltt.
Hint Resolve glu_nat_resp_per_nat : mctt.

#[local]
Ltac resp_per_IH :=
Expand Down Expand Up @@ -965,7 +965,7 @@ Proof.
destruct_rel_mod_eval.
simplify_evals.
deepexec H1 ltac:(fun H => pose proof H).
autorewrite with mcltt in *.
autorewrite with mctt in *.
mauto 3.

- simpl_glu_rel.
Expand Down Expand Up @@ -1043,7 +1043,7 @@ Proof.
split; [mauto 3 |].
intros.
saturate_weakening_escape.
autorewrite with mcltt.
autorewrite with mctt.
mauto 3.
- simpl_glu_rel.
econstructor; repeat split; mauto 3;
Expand Down
58 changes: 29 additions & 29 deletions theories/Core/Syntactic/Presup.v
Original file line number Diff line number Diff line change
Expand Up @@ -33,7 +33,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_natrec_cong_right : mcltt.
Hint Resolve presup_exp_eq_natrec_cong_right : mctt.

Lemma presup_exp_eq_natrec_sub_left : forall {Γ σ Δ i A MZ MS M},
{{ ⊢ Γ }} ->
Expand All @@ -58,7 +58,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_natrec_sub_left : mcltt.
Hint Resolve presup_exp_eq_natrec_sub_left : mctt.

Lemma presup_exp_eq_natrec_sub_right : forall {Γ σ Δ i A MZ MS M},
{{ ⊢ Γ }} ->
Expand Down Expand Up @@ -112,7 +112,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_natrec_sub_right : mcltt.
Hint Resolve presup_exp_eq_natrec_sub_right : mctt.

Lemma presup_exp_eq_beta_succ_right : forall {Γ i A MZ MS M},
{{ ⊢ Γ, ℕ }} ->
Expand Down Expand Up @@ -156,7 +156,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_beta_succ_right : mcltt.
Hint Resolve presup_exp_eq_beta_succ_right : mctt.

Lemma presup_exp_eq_fn_cong_right : forall {Γ i A A' j B M'},
{{ ⊢ Γ }} ->
Expand All @@ -179,7 +179,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_fn_cong_right : mcltt.
Hint Resolve presup_exp_eq_fn_cong_right : mctt.

Lemma presup_exp_eq_fn_sub_right : forall {Γ σ Δ i A j B M},
{{ ⊢ Γ }} ->
Expand All @@ -204,7 +204,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_fn_sub_right : mcltt.
Hint Resolve presup_exp_eq_fn_sub_right : mctt.

Lemma presup_exp_eq_app_cong_right : forall {Γ i A B M' N N'},
{{ ⊢ Γ }} ->
Expand All @@ -227,7 +227,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_app_cong_right : mcltt.
Hint Resolve presup_exp_eq_app_cong_right : mctt.

Lemma presup_exp_eq_app_sub_left : forall {Γ σ Δ i A B M N},
{{ ⊢ Γ }} ->
Expand Down Expand Up @@ -258,7 +258,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_app_sub_left : mcltt.
Hint Resolve presup_exp_eq_app_sub_left : mctt.

Lemma presup_exp_eq_app_sub_right : forall {Γ σ Δ i A B M N},
{{ ⊢ Γ }} ->
Expand Down Expand Up @@ -287,7 +287,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_app_sub_right : mcltt.
Hint Resolve presup_exp_eq_app_sub_right : mctt.

Lemma presup_exp_eq_pi_eta_right : forall {Γ i A B M},
{{ ⊢ Γ }} ->
Expand All @@ -312,7 +312,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_pi_eta_right : mcltt.
Hint Resolve presup_exp_eq_pi_eta_right : mctt.

Lemma presup_exp_eq_prop_eq_var0 : forall {Γ i A},
{{ Γ ⊢ A : Type@i }} ->
Expand All @@ -326,7 +326,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_var0 : mcltt.
Hint Resolve presup_exp_eq_prop_eq_var0 : mctt.

Lemma presup_exp_eq_prop_eq_var1 : forall {Γ i A},
{{ Γ ⊢ A : Type@i }} ->
Expand All @@ -340,7 +340,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_var1 : mcltt.
Hint Resolve presup_exp_eq_prop_eq_var1 : mctt.

Lemma presup_exp_eq_prop_eq_wf : forall {Γ i A},
{{ Γ ⊢ A : Type@i }} ->
Expand All @@ -356,7 +356,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_wf : mcltt.
Hint Resolve presup_exp_eq_prop_eq_wf : mctt.

Lemma presup_exp_eq_prop_eq_sub_helper2 : forall {Γ σ Δ i A M1 M2},
{{ Γ ⊢s σ : Δ }} ->
Expand All @@ -377,7 +377,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_sub_helper2 : mcltt.
Hint Resolve presup_exp_eq_prop_eq_sub_helper2 : mctt.

Lemma presup_exp_eq_prop_eq_typ_sub2_eq : forall {Γ σ Δ i A M1 M2},
{{ Γ ⊢s σ : Δ }} ->
Expand Down Expand Up @@ -407,7 +407,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_typ_sub2_eq : mcltt.
Hint Resolve presup_exp_eq_prop_eq_typ_sub2_eq : mctt.

Lemma presup_exp_eq_prop_eq_sub_helper3 : forall {Γ σ Δ i A M1 M2 N},
{{ Γ ⊢s σ : Δ }} ->
Expand All @@ -426,7 +426,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_sub_helper3 : mcltt.
Hint Resolve presup_exp_eq_prop_eq_sub_helper3 : mctt.

Lemma presup_exp_eq_prop_eq_id_sub_helper2 : forall {Γ i A M1 M2},
{{ Γ ⊢ A : Type@i }} ->
Expand All @@ -447,7 +447,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_id_sub_helper2 : mcltt.
Hint Resolve presup_exp_eq_prop_eq_id_sub_helper2 : mctt.

Lemma presup_exp_eq_prop_eq_typ_id_sub2_eq : forall {Γ i A M1 M2},
{{ Γ ⊢ A : Type@i }} ->
Expand All @@ -472,7 +472,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_typ_id_sub2_eq : mcltt.
Hint Resolve presup_exp_eq_prop_eq_typ_id_sub2_eq : mctt.

Lemma presup_exp_eq_prop_eq_id_sub_helper3 : forall {Γ i A M1 M2 N},
{{ Γ ⊢ A : Type@i }} ->
Expand All @@ -490,7 +490,7 @@ Proof.
Qed.

#[export]
Hint Resolve presup_exp_eq_prop_eq_id_sub_helper3 : mcltt.
Hint Resolve presup_exp_eq_prop_eq_id_sub_helper3 : mctt.

Lemma presup_exp_eq_refl_sub_right : forall {Γ σ Δ i A M},
{{ ⊢ Γ }} ->
Expand All @@ -509,7 +509,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_refl_sub_right : mcltt.
Hint Resolve presup_exp_eq_refl_sub_right : mctt.

Lemma presup_exp_eq_eqrec_sub_left : forall {Γ σ Δ i A M1 M2 j B N BR},
{{ ⊢ Γ }} ->
Expand Down Expand Up @@ -554,7 +554,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_eqrec_sub_left : mcltt.
Hint Resolve presup_exp_eq_eqrec_sub_left : mctt.

Lemma presup_exp_eq_eqrec_sub_right : forall {Γ σ Δ i A M1 M2 j B N BR},
{{ ⊢ Γ }} ->
Expand Down Expand Up @@ -881,7 +881,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_eqrec_sub_right : mcltt.
Hint Resolve presup_exp_eq_eqrec_sub_right : mctt.

Lemma presup_exp_eq_refl_cong_right : forall {Γ i A A' M M'},
{{ ⊢ Γ }} ->
Expand All @@ -899,7 +899,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_refl_cong_right : mcltt.
Hint Resolve presup_exp_eq_refl_cong_right : mctt.

Lemma presup_exp_eq_eqrec_cong_right : forall {Γ i A A' M1 M1' M2 M2' j B B' BR BR' N N'},
{{ ⊢ Γ }} ->
Expand Down Expand Up @@ -1029,7 +1029,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_eqrec_cong_right : mcltt.
Hint Resolve presup_exp_eq_eqrec_cong_right : mctt.

Lemma presup_exp_eq_eqrec_beta_right : forall {Γ i A M j B BR},
{{ ⊢ Γ }} ->
Expand Down Expand Up @@ -1137,7 +1137,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_eqrec_beta_right : mcltt.
Hint Resolve presup_exp_eq_eqrec_beta_right : mctt.

Lemma presup_exp_eq_var_0_sub_left : forall {Γ σ Δ i A M},
{{ ⊢ Γ }} ->
Expand All @@ -1154,7 +1154,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_var_0_sub_left : mcltt.
Hint Resolve presup_exp_eq_var_0_sub_left : mctt.

Lemma presup_exp_eq_var_S_sub_left : forall {Γ σ Δ i A M B x},
{{ ⊢ Γ }} ->
Expand All @@ -1173,7 +1173,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_var_S_sub_left : mcltt.
Hint Resolve presup_exp_eq_var_S_sub_left : mctt.

Lemma presup_exp_eq_sub_cong_right : forall {Γ σ σ' Δ i A M M'},
{{ ⊢ Δ }} ->
Expand All @@ -1193,7 +1193,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_sub_cong_right : mcltt.
Hint Resolve presup_exp_eq_sub_cong_right : mctt.

Lemma presup_exp_eq_sub_compose_right : forall {Γ τ Γ' σ Γ'' i A M},
{{ ⊢ Γ }} ->
Expand All @@ -1210,7 +1210,7 @@ Proof.
Qed.

#[local]
Hint Resolve presup_exp_eq_sub_compose_right : mcltt.
Hint Resolve presup_exp_eq_sub_compose_right : mctt.

#[local]
Ltac gen_presup_IH presup_exp_eq presup_sub_eq presup_subtyp H :=
Expand Down

0 comments on commit 4a412da

Please sign in to comment.