From 742e21da086654af82f308027250d00b50d67f52 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Thu, 2 Dec 2021 17:31:43 +0100 Subject: [PATCH] update in ground + advances on total redlocation --- .../lambdadelta/ground/relocation/tr_pap.ma | 41 ++++++++++--------- .../ground/relocation/tr_pap_pat.ma | 34 +++++++++++++++ 2 files changed, 56 insertions(+), 19 deletions(-) create mode 100644 matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_pat.ma diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap.ma b/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap.ma index 7acf280d2..67231de44 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap.ma @@ -13,7 +13,8 @@ (**************************************************************************) include "ground/notation/functions/apply_2.ma". -include "ground/relocation/tr_pat.ma". +include "ground/arith/pnat_plus.ma". +include "ground/relocation/tr_map.ma". (* POSITIVE APPLICATION FOR TOTAL RELOCATION MAPS ***************************) @@ -30,36 +31,37 @@ interpretation "functional positive application (total relocation maps)" 'Apply f i = (tr_pap i f). -(* Constructions with pr_pat ***********************************************) +(* Basic constructions ******************************************************) -(*** at_total *) -lemma tr_pat_total: ∀i1,f. @❨i1,𝐭❨f❩❩ ≘ f@❨i1❩. -#i1 elim i1 -i1 -[ * // | #i #IH * /3 width=1 by pr_pat_succ_sn/ ] -qed. +(*** apply_O1 *) +lemma tr_pap_unit (f): + ∀p. p = (p⨮f)@❨𝟏❩. +// qed. -(* Inversions with pr_pat ***************************************************) +(*** apply_S1 *) +lemma tr_pap_succ (f): + ∀p,i. f@❨i❩+p = (p⨮f)@❨↑i❩. +// qed. +(* +(*** apply_S2 *) +lemma tr_pap_next (f): + ∀i. ↑(f@❨i❩) = (↑f)@❨i❩. +* #p #f * // +qed. -lemma at_inv_total: ∀f,i1,i2. @❨i1, f❩ ≘ i2 → f@❨i1❩ = i2. -/2 width=6 by fr2_nat_mono/ qed-. -(* Basic properties *********************************************************) -lemma apply_O1: ∀p,f. (p⨮f)@❨𝟏❩ = p. -// qed. +(*** apply_eq_repl *) +lemma apply_eq_repl (i): + ∀f1,f2. f1 ≗ f2 → f1@❨i❩ = f2@❨i❩. -lemma apply_S1: ∀p,f,i. (p⨮f)@❨↑i❩ = f@❨i❩+p. -// qed. -lemma apply_eq_repl (i): gr_eq_repl … (λf1,f2. f1@❨i❩ = f2@❨i❩). +(i): pr_eq_repl … (λf1,f2. f1@❨i❩ = f2@❨i❩). #i elim i -i [2: #i #IH ] * #p1 #f1 * #p2 #f2 #H elim (eq_inv_seq_aux … H) -H #Hp #Hf // >apply_S1 >apply_S1 /3 width=1 by eq_f2/ qed. -lemma apply_S2: ∀f,i. (↑f)@❨i❩ = ↑(f@❨i❩). -* #p #f * // -qed. (* Main inversion lemmas ****************************************************) @@ -86,3 +88,4 @@ include "ground/relocation/rtmap_istot.ma". lemma at_istot: ∀f. 𝐓❨f❩. /2 width=2 by ex_intro/ qed. *) +*) \ No newline at end of file diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_pat.ma b/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_pat.ma new file mode 100644 index 000000000..90dd95a77 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_pat.ma @@ -0,0 +1,34 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground/relocation/pr_pat_pat.ma". +include "ground/relocation/tr_pat.ma". +include "ground/relocation/tr_pap.ma". + +(* POSITIVE APPLICATION FOR TOTAL RELOCATION MAPS ***************************) + +(* Constructions with pr_pat ***********************************************) + +(*** at_total *) +lemma tr_pat_total: ∀i1,f. @❨i1,𝐭❨f❩❩ ≘ f@❨i1❩. +#i1 elim i1 -i1 +[ * // | #i #IH * /3 width=1 by tr_pat_succ_sn/ ] +qed. + +(* Inversions with pr_pat ***************************************************) + +(*** at_inv_total *) +lemma tr_pat_inv_total (f): + ∀i1,i2. @❨i1, 𝐭❨f❩❩ ≘ i2 → f@❨i1❩ = i2. +/2 width=6 by pr_pat_mono/ qed-. -- 2.39.2