X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Fmr2_at.ma;h=d2a2499a3af4dcc026edd9f30a79fb7766b388af;hp=77aa3917cfd4138784d4d89d4c2b2fac71752af9;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hpb=3b7b8afcb429a60d716d5226a5b6ab0d003228b1 diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_at.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_at.ma index 77aa3917c..d2a2499a3 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_at.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_at.ma @@ -20,9 +20,9 @@ include "ground_2/relocation/mr2.ma". inductive at: mr2 → relation nat ≝ | at_nil: ∀i. at (◊) i i | at_lt : ∀cs,l,m,i1,i2. i1 < l → - at cs i1 i2 → at ({l, m};cs) i1 i2 + at cs i1 i2 → at (❨l, m❩;cs) i1 i2 | at_ge : ∀cs,l,m,i1,i2. l ≤ i1 → - at cs (i1 + m) i2 → at ({l, m};cs) i1 i2 + at cs (i1 + m) i2 → at (❨l, m❩;cs) i1 i2 . interpretation "application (multiple relocation with pairs)" @@ -30,7 +30,7 @@ interpretation "application (multiple relocation with pairs)" (* Basic inversion lemmas ***************************************************) -fact at_inv_nil_aux: ∀cs,i1,i2. @⦃i1, cs⦄ ≘ i2 → cs = ◊ → i1 = i2. +fact at_inv_nil_aux: ∀cs,i1,i2. @❪i1, cs❫ ≘ i2 → cs = ◊ → i1 = i2. #cs #i1 #i2 * -cs -i1 -i2 [ // | #cs #l #m #i1 #i2 #_ #_ #H destruct @@ -38,13 +38,13 @@ fact at_inv_nil_aux: ∀cs,i1,i2. @⦃i1, cs⦄ ≘ i2 → cs = ◊ → i1 = i2. ] qed-. -lemma at_inv_nil: ∀i1,i2. @⦃i1, ◊⦄ ≘ i2 → i1 = i2. +lemma at_inv_nil: ∀i1,i2. @❪i1, ◊❫ ≘ i2 → i1 = i2. /2 width=3 by at_inv_nil_aux/ qed-. -fact at_inv_cons_aux: ∀cs,i1,i2. @⦃i1, cs⦄ ≘ i2 → - ∀l,m,cs0. cs = {l, m};cs0 → - i1 < l ∧ @⦃i1, cs0⦄ ≘ i2 ∨ - l ≤ i1 ∧ @⦃i1 + m, cs0⦄ ≘ i2. +fact at_inv_cons_aux: ∀cs,i1,i2. @❪i1, cs❫ ≘ i2 → + ∀l,m,cs0. cs = ❨l, m❩;cs0 → + i1 < l ∧ @❪i1, cs0❫ ≘ i2 ∨ + l ≤ i1 ∧ @❪i1 + m, cs0❫ ≘ i2. #cs #i1 #i2 * -cs -i1 -i2 [ #i #l #m #cs #H destruct | #cs1 #l1 #m1 #i1 #i2 #Hil1 #Hi12 #l2 #m2 #cs2 #H destruct /3 width=1 by or_introl, conj/ @@ -52,20 +52,20 @@ fact at_inv_cons_aux: ∀cs,i1,i2. @⦃i1, cs⦄ ≘ i2 → ] qed-. -lemma at_inv_cons: ∀cs,l,m,i1,i2. @⦃i1, {l, m};cs⦄ ≘ i2 → - i1 < l ∧ @⦃i1, cs⦄ ≘ i2 ∨ - l ≤ i1 ∧ @⦃i1 + m, cs⦄ ≘ i2. +lemma at_inv_cons: ∀cs,l,m,i1,i2. @❪i1, ❨l, m❩;cs❫ ≘ i2 → + i1 < l ∧ @❪i1, cs❫ ≘ i2 ∨ + l ≤ i1 ∧ @❪i1 + m, cs❫ ≘ i2. /2 width=3 by at_inv_cons_aux/ qed-. -lemma at_inv_cons_lt: ∀cs,l,m,i1,i2. @⦃i1, {l, m};cs⦄ ≘ i2 → - i1 < l → @⦃i1, cs⦄ ≘ i2. +lemma at_inv_cons_lt: ∀cs,l,m,i1,i2. @❪i1, ❨l, m❩;cs❫ ≘ i2 → + i1 < l → @❪i1, cs❫ ≘ i2. #cs #l #m #i1 #m2 #H elim (at_inv_cons … H) -H * // #Hli1 #_ #Hi1l elim (lt_le_false … Hi1l Hli1) qed-. -lemma at_inv_cons_ge: ∀cs,l,m,i1,i2. @⦃i1, {l, m};cs⦄ ≘ i2 → - l ≤ i1 → @⦃i1 + m, cs⦄ ≘ i2. +lemma at_inv_cons_ge: ∀cs,l,m,i1,i2. @❪i1, ❨l, m❩;cs❫ ≘ i2 → + l ≤ i1 → @❪i1 + m, cs❫ ≘ i2. #cs #l #m #i1 #m2 #H elim (at_inv_cons … H) -H * // #Hi1l #_ #Hli1 elim (lt_le_false … Hi1l Hli1) @@ -73,7 +73,7 @@ qed-. (* Main properties **********************************************************) -theorem at_mono: ∀cs,i,i1. @⦃i, cs⦄ ≘ i1 → ∀i2. @⦃i, cs⦄ ≘ i2 → i1 = i2. +theorem at_mono: ∀cs,i,i1. @❪i, cs❫ ≘ i1 → ∀i2. @❪i, cs❫ ≘ i2 → i1 = i2. #cs #i #i1 #H elim H -cs -i -i1 [ #i #x #H <(at_inv_nil … H) -x // | #cs #l #m #i #i1 #Hil #_ #IHi1 #x #H