Require Import Coq.Vectors.Vector.
Local Notation vec := t.
From PostTheorem Require Export FOL.Syntax.Core.
Set Default Proof Using "Type".
Local Set Implicit Arguments.
Local Unset Strict Implicit.
Section Subst.
Context {Σ_funcs : funcs_signature}.
Context {Σ_preds : preds_signature}.
Context {ops : operators}.
Lemma subst_term_ext (t : term) sigma tau :
(forall n, sigma n = tau n) -> t`[sigma] = t`[tau].
Proof.
intros H. induction t; cbn.
- now apply H.
- f_equal. now apply map_ext_in.
Qed.
Lemma subst_term_id (t : term) sigma :
(forall n, sigma n = var n) -> t`[sigma] = t.
Proof.
intros H. induction t; cbn.
- now apply H.
- f_equal. now erewrite map_ext_in, map_id.
Qed.
Lemma subst_term_var (t : term) :
t`[var] = t.
Proof.
now apply subst_term_id.
Qed.
Lemma subst_term_comp (t : term) sigma tau :
t`[sigma]`[tau] = t`[sigma >> subst_term tau].
Proof.
induction t; cbn.
- reflexivity.
- f_equal. rewrite map_map. now apply map_ext_in.
Qed.
Lemma subst_term_shift (t : term) s :
t`[↑]`[s..] = t.
Proof.
rewrite subst_term_comp. apply subst_term_id. now intros [|].
Qed.
Lemma up_term (t : term) xi :
t`[↑]`[up xi] = t`[xi]`[↑].
Proof.
rewrite !subst_term_comp. apply subst_term_ext. reflexivity.
Qed.
Lemma up_ext sigma tau :
(forall n, sigma n = tau n) -> forall n, up sigma n = up tau n.
Proof.
destruct n; cbn; trivial.
unfold funcomp. now rewrite H.
Qed.
Lemma up_var sigma :
(forall n, sigma n = var n) -> forall n, up sigma n = var n.
Proof.
destruct n; cbn; trivial.
unfold funcomp. now rewrite H.
Qed.
Lemma up_funcomp sigma tau :
forall n, (up sigma >> subst_term (up tau)) n = up (sigma >> subst_term tau) n.
Proof.
intros [|]; cbn; trivial.
setoid_rewrite subst_term_comp.
apply subst_term_ext. now intros [|].
Qed.
Lemma subst_ext {ff : falsity_flag} (phi : form) sigma tau :
(forall n, sigma n = tau n) -> phi[sigma] = phi[tau].
Proof.
induction phi in sigma, tau |- *; cbn; intros H.
- reflexivity.
- f_equal. apply map_ext. intros s. now apply subst_term_ext.
- now erewrite IHphi1, IHphi2.
- erewrite IHphi; trivial. now apply up_ext.
Qed.
Lemma subst_id {ff : falsity_flag} (phi : form) sigma :
(forall n, sigma n = var n) -> phi[sigma] = phi.
Proof.
induction phi in sigma |- *; cbn; intros H.
- reflexivity.
- f_equal. erewrite map_ext; try apply map_id. intros s. now apply subst_term_id.
- now erewrite IHphi1, IHphi2.
- erewrite IHphi; trivial. now apply up_var.
Qed.
Lemma subst_var {ff : falsity_flag} (phi : form) :
phi[var] = phi.
Proof.
now apply subst_id.
Qed.
Lemma subst_comp {ff : falsity_flag} (phi : form) sigma tau :
phi[sigma][tau] = phi[sigma >> subst_term tau].
Proof.
induction phi in sigma, tau |- *; cbn.
- reflexivity.
- f_equal. rewrite map_map. apply map_ext. intros s. apply subst_term_comp.
- now rewrite IHphi1, IHphi2.
- rewrite IHphi. f_equal. now apply subst_ext, up_funcomp.
Qed.
Lemma subst_shift {ff : falsity_flag} (phi : form) s :
phi[↑][s..] = phi.
Proof.
rewrite subst_comp. apply subst_id. now intros [|].
Qed.
Lemma up_form {ff : falsity_flag} xi psi :
psi[↑][up xi] = psi[xi][↑].
Proof.
rewrite !subst_comp. apply subst_ext. reflexivity.
Qed.
Lemma up_decompose {ff : falsity_flag} sigma phi :
phi[up (S >> sigma)][(sigma 0)..] = phi[sigma].
Proof.
rewrite subst_comp. apply subst_ext.
intros [].
- reflexivity.
- apply subst_term_shift.
Qed.
End Subst.