X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_sand.ma;h=fcf82d2cc2afebd2862a51336cebd009d5cc98e6;hb=629687db8a55432e95c82f0c79e3f51c023e65a6;hp=86ad267c6958f1da814bfe8ade93203efb4668ee;hpb=5832735b721c0bd8567c8f0be761a9136363a2a6;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sand.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sand.ma index 86ad267c6..fcf82d2cc 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sand.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sand.ma @@ -73,6 +73,39 @@ qed-. (* Basic properties *********************************************************) +corec lemma sand_eq_repl_back1: ∀f2,f. eq_repl_back … (λf1. f1 ⋒ f2 ≡ f). +#f2 #f #f1 * -f1 -f2 -f +#f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #x #Hx +try cases (eq_inv_px … Hx … H1) try cases (eq_inv_nx … Hx … H1) -g1 +/3 width=7 by sand_pp, sand_np, sand_pn, sand_nn/ +qed-. + +lemma sand_eq_repl_fwd1: ∀f2,f. eq_repl_fwd … (λf1. f1 ⋒ f2 ≡ f). +#f2 #f @eq_repl_sym /2 width=3 by sand_eq_repl_back1/ +qed-. + +corec lemma sand_eq_repl_back2: ∀f1,f. eq_repl_back … (λf2. f1 ⋒ f2 ≡ f). +#f1 #f #f2 * -f1 -f2 -f +#f1 #f2 #f #g1 #g2 #g #Hf #H #H2 #H0 #x #Hx +try cases (eq_inv_px … Hx … H2) try cases (eq_inv_nx … Hx … H2) -g2 +/3 width=7 by sand_pp, sand_np, sand_pn, sand_nn/ +qed-. + +lemma sand_eq_repl_fwd2: ∀f1,f. eq_repl_fwd … (λf2. f1 ⋒ f2 ≡ f). +#f1 #f @eq_repl_sym /2 width=3 by sand_eq_repl_back2/ +qed-. + +corec lemma sand_eq_repl_back3: ∀f1,f2. eq_repl_back … (λf. f1 ⋒ f2 ≡ f). +#f1 #f2 #f * -f1 -f2 -f +#f1 #f2 #f #g1 #g2 #g #Hf #H #H2 #H0 #x #Hx +try cases (eq_inv_px … Hx … H0) try cases (eq_inv_nx … Hx … H0) -g +/3 width=7 by sand_pp, sand_np, sand_pn, sand_nn/ +qed-. + +lemma sand_eq_repl_fwd3: ∀f1,f2. eq_repl_fwd … (λf. f1 ⋒ f2 ≡ f). +#f1 #f2 @eq_repl_sym /2 width=3 by sand_eq_repl_back3/ +qed-. + corec lemma sand_refl: ∀f. f ⋒ f ≡ f. #f cases (pn_split f) * #g #H [ @(sand_pp … H H H) | @(sand_nn … H H H) ] -H //