(**************************************************************************) (* ___ *) (* ||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/notation/functions/apply_2.ma". include "ground/arith/pnat_le_plus.ma". include "ground/relocation/pstream_eq.ma". include "ground/relocation/rtmap_istot.ma". (* RELOCATION P-STREAM ******************************************************) rec definition apply (i: pnat) on i: gr_map → pnat. * #p #f cases i -i [ @p | #i lapply (apply i f) -apply -i -f #i @(i+p) ] defined. interpretation "functional application (nstream)" 'Apply f i = (apply i f). (* Properties on at (specific) ************************************************) lemma at_O1: ∀i2,f. @❪𝟏, i2⨮f❫ ≘ i2. #i2 elim i2 -i2 /2 width=5 by gr_pat_refl, gr_pat_next/ qed. lemma at_S1: ∀p,f,i1,i2. @❪i1, f❫ ≘ i2 → @❪↑i1, p⨮f❫ ≘ i2+p. #p elim p -p /3 width=7 by gr_pat_push, gr_pat_next/ qed. lemma at_total: ∀i1,f. @❪i1, f❫ ≘ f@❨i1❩. #i1 elim i1 -i1 [ * // | #i #IH * /3 width=1 by at_S1/ ] qed. lemma at_istot: ∀f. 𝐓❪f❫. /2 width=2 by ex_intro/ qed. lemma at_plus2: ∀f,i1,i,p,q. @❪i1, p⨮f❫ ≘ i → @❪i1, (p+q)⨮f❫ ≘ i+q. #f #i1 #i #p #q #H elim q -q /2 width=5 by gr_pat_next/ qed. (* Inversion lemmas on at (specific) ******************************************) lemma at_inv_O1: ∀f,p,i2. @❪𝟏, p⨮f❫ ≘ i2 → p = i2. #f #p elim p -p /2 width=6 by gr_pat_inv_unit_push/ #p #IH #i2 #H elim (gr_pat_inv_next … H) -H [|*: // ] #j2 #Hj * -i2 /3 width=1 by eq_f/ qed-. lemma at_inv_S1: ∀f,p,j1,i2. @❪↑j1, p⨮f❫ ≘ i2 → ∃∃j2. @❪j1, f❫ ≘ j2 & j2+p = i2. #f #p elim p -p /2 width=5 by gr_pat_inv_succ_push/ #p #IH #j1 #i2 #H elim (gr_pat_inv_next … H) -H [|*: // ] #j2 #Hj * -i2 elim (IH … Hj) -IH -Hj #i2 #Hi * -j2 /2 width=3 by ex2_intro/ qed-. lemma at_inv_total: ∀f,i1,i2. @❪i1, f❫ ≘ i2 → f@❨i1❩ = i2. /2 width=6 by fr2_nat_mono/ qed-. (* Forward lemmas on at (specific) *******************************************) lemma at_increasing_plus: ∀f,p,i1,i2. @❪i1, p⨮f❫ ≘ i2 → i1 + p ≤ ↑i2. #f #p * [ #i2 #H <(at_inv_O1 … H) -i2 // | #i1 #i2 #H elim (at_inv_S1 … H) -H #j1 #Ht * -i2 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 ****************************************************) theorem apply_inj: ∀f,i1,i2,j. f@❨i1❩ = j → f@❨i2❩ = j → i1 = i2. /2 width=4 by gr_pat_inj/ qed-. corec theorem nstream_eq_inv_ext: ∀f1,f2. (∀i. f1@❨i❩ = f2@❨i❩) → f1 ≗ f2. * #p1 #f1 * #p2 #f2 #Hf @stream_eq_cons [ @(Hf (𝟏)) | @nstream_eq_inv_ext -nstream_eq_inv_ext #i lapply (Hf (𝟏)) >apply_O1 >apply_O1 #H destruct lapply (Hf (↑i)) >apply_S1 >apply_S1 #H /3 width=2 by eq_inv_pplus_bi_dx, eq_inv_psucc_bi/ ] qed-.