X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Ftrace_isid.ma;h=6dce9e0a2fd8db60fa006cfbcf4d26fab5ad5291;hb=a8cd6cc85182245df447a21caf16b6503fa4b3e5;hp=d87487c897fa098b1b59e56c107ce1c2fda70c75;hpb=46815bb7af06b235ead2fd67a4aee2d294b51928;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma index d87487c89..6dce9e0a2 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma @@ -12,15 +12,19 @@ (* *) (**************************************************************************) -include "ground_2/notation/relations/isid_1.ma". +include "ground_2/notation/relations/isidentity_1.ma". include "ground_2/relocation/trace_after.ma". +include "ground_2/relocation/trace_sle.ma". (* RELOCATION TRACE *********************************************************) definition isid: predicate trace ≝ λcs. ∥cs∥ = |cs|. interpretation "test for identity (trace)" - 'IsId cs = (isid cs). + 'IsIdentity cs = (isid cs). + +definition t_reflexive: ∀S:Type[0]. predicate (trace → relation S) ≝ + λS,R. ∀a. ∃∃t. 𝐈⦃t⦄ & R t a a. (* Basic properties *********************************************************) @@ -98,3 +102,13 @@ lemma after_inv_isid3: ∀t1,t2,t. t1 ⊚ t2 ≡ t → 𝐈⦃t⦄ → 𝐈⦃t1 | #t1 #t2 #t #_ #_ #H elim (isid_inv_false … H) ] qed-. + +(* Forward on inclusion *****************************************************) + +lemma sle_isid1_fwd: ∀t1,t2. t1 ⊆ t2 → 𝐈⦃t1⦄ → t1 = t2. +#t1 #t2 #H elim H -t1 -t2 // +[ #t1 #t2 #_ #IH #H lapply (isid_inv_true … H) -H + #HT1 @eq_f2 // @IH @HT1 (**) (* full auto fails *) +| #t1 #t2 #b #_ #_ #H elim (isid_inv_false … H) +] +qed-.