From: Ferruccio Guidi Date: Sat, 7 Sep 2013 17:07:12 +0000 (+0000) Subject: support for nat-labeled reflexive and transitive closure added to lambdadelta X-Git-Tag: make_still_working~1109 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=b08fbe87149ae7c2e14dfd7485dbf50d6e7dbfe6;p=helm.git support for nat-labeled reflexive and transitive closure added to lambdadelta --- diff --git a/matita/matita/contribs/lambdadelta/ground_2/lstar.ma b/matita/matita/contribs/lambdadelta/ground_2/lstar.ma new file mode 100644 index 000000000..ae707f2b9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground_2/lstar.ma @@ -0,0 +1,20 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "arithmetics/lstar.ma". + +(* PROPERTIES OF NAT-LABELED REFLEXIVE AND TRANSITIVE CLOSURE ***************) + +definition llstar: ∀A:Type[0]. ∀B. (A→relation B) → nat → (A→relation B) ≝ + λA,B,R,l,a. lstar … (R a) l. diff --git a/matita/matita/lib/arithmetics/lstar.ma b/matita/matita/lib/arithmetics/lstar.ma index ddd4113e3..e3dcd55d6 100644 --- a/matita/matita/lib/arithmetics/lstar.ma +++ b/matita/matita/lib/arithmetics/lstar.ma @@ -48,6 +48,11 @@ lemma lstar_step: ∀B,R,b1,b2. R b1 b2 → lstar B R 1 b1 b2. /2 width=3/ qed. +lemma lstar_dx: ∀B,R,l,b1,b. lstar B R l b1 b → ∀b2. R b b2 → + lstar B R (l+1) b1 b2. +#B #R #l #b1 #b #H @(lstar_ind_l … l b1 H) -l -b1 /2 width=1/ /3 width=3/ +qed. + lemma lstar_inv_O: ∀B,R,l,b1,b2. lstar B R l b1 b2 → 0 = l → b1 = b2. #B #R #l #b1 #b2 * -l -b1 -b2 // #b1 #b #_ #l #b2 #_