]> matita.cs.unibo.it Git - helm.git/commitdiff
- lambda: parallel reduction to obtain diamond property
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 1 Dec 2012 17:07:33 +0000 (17:07 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 1 Dec 2012 17:07:33 +0000 (17:07 +0000)
- list:   local "norm" notation removed in favour of "card" core notation

matita/matita/contribs/lambda/delifting_substitution.ma
matita/matita/contribs/lambda/labelled_sequential_computation.ma [new file with mode: 0644]
matita/matita/contribs/lambda/labelled_sequential_reduction.ma
matita/matita/contribs/lambda/lift.ma
matita/matita/contribs/lambda/parallel_reduction.ma [new file with mode: 0644]
matita/matita/contribs/lambda/policy.txt [new file with mode: 0644]
matita/matita/contribs/lambda/preamble.ma
matita/matita/contribs/lambda/redex_pointer.ma
matita/matita/contribs/lambda/redex_pointer_sequence.ma [new file with mode: 0644]
matita/matita/lib/basics/lists/list.ma
matita/matita/lib/basics/relations.ma

index 3d4c2150bca2d6b4fdc88814d568ce7baba251b7..11ce887611a10fef06cf2f1cc6aa61476a89ead6 100644 (file)
@@ -111,8 +111,8 @@ theorem dsubst_lift_ge: ∀h,D,M,d1,d2. d1 + h ≤ d2 →
 ]
 qed.
 
-theorem subst_subst_ge: ∀D1,D2,M,d1,d2. d1 ≤ d2 →
-                        [d2 ⬐ D2] [d1 ⬐ D1] M = [d1 ⬐ [d2 - d1 ⬐ D2] D1] [d2 + 1 ⬐ D2] M.
+theorem dsubst_dsubst_ge: ∀D1,D2,M,d1,d2. d1 ≤ d2 →
+                          [d2 ⬐ D2] [d1 ⬐ D1] M = [d1 ⬐ [d2 - d1 ⬐ D2] D1] [d2 + 1 ⬐ D2] M.
 #D1 #D2 #M elim M -M
 [ #i #d1 #d2 #Hd12 elim (lt_or_eq_or_gt i d1) #Hid1
   [ lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2
@@ -134,9 +134,12 @@ theorem subst_subst_ge: ∀D1,D2,M,d1,d2. d1 ≤ d2 →
 ]
 qed.
 
-theorem subst_subst_lt: ∀D1,D2,M,d1,d2. d2 < d1 →
-                        [d2 ⬐ [d1 - d2 -1 ⬐ D1] D2] [d1 ⬐ D1] M = [d1 - 1 ⬐ D1] [d2 ⬐ D2] M.
+theorem dsubst_dsubst_lt: ∀D1,D2,M,d1,d2. d2 < d1 →
+                          [d2 ⬐ [d1 - d2 -1 ⬐ D1] D2] [d1 ⬐ D1] M = [d1 - 1 ⬐ D1] [d2 ⬐ D2] M.
 #D1 #D2 #M #d1 #d2 #Hd21
 lapply (ltn_to_ltO … Hd21) #Hd1
->subst_subst_ge in ⊢ (???%); /2 width=1/ <plus_minus_m_m //
+>dsubst_dsubst_ge in ⊢ (???%); /2 width=1/ <plus_minus_m_m //
 qed.
+
+definition dsubstable_dx: predicate (relation term) ≝ λR.
+                          ∀D,M1,M2. R M1 M2 → ∀d. R ([d ⬐ D] M1) ([d ⬐ D] M2).
diff --git a/matita/matita/contribs/lambda/labelled_sequential_computation.ma b/matita/matita/contribs/lambda/labelled_sequential_computation.ma
new file mode 100644 (file)
index 0000000..fd096c0
--- /dev/null
@@ -0,0 +1,93 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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 "redex_pointer_sequence.ma".
+include "labelled_sequential_reduction.ma".
+
+(* LABELLED SEQUENTIAL COMPUTATION (MULTISTEP) ******************************)
+
+(* Note: this comes from "star term lsred" *)
+inductive lsreds: rpseq → relation term ≝
+| lsreds_nil : ∀M. lsreds (◊) M M
+| lsreds_cons: ∀p,M1,M. M1 ⇀[p] M →
+               ∀s,M2. lsreds s M M2 → lsreds (p::s) M1 M2
+.
+
+interpretation "labelled sequential computation"
+   'SeqRedStar M s N = (lsreds s M N).
+
+notation "hvbox( M break ⇀* [ term 46 s ] break term 46 N )"
+   non associative with precedence 45
+   for @{ 'SeqRedStar $M $s $N }.
+
+lemma lsred_lsreds: ∀p,M1,M2. M1 ⇀[p] M2 → M1 ⇀*[p::◊] M2.
+/2 width=3/
+qed.
+
+lemma lsreds_inv_nil: ∀s,M1,M2. M1 ⇀*[s] M2 → ◊ = s → M1 = M2.
+#s #M1 #M2 * -s -M1 -M2 //
+#p #M1 #M #_ #s #M2 #_ #H destruct
+qed-.
+
+lemma lsreds_inv_cons: ∀s,M1,M2. M1 ⇀*[s] M2 → ∀q,r. q::r = s →
+                       ∃∃M. M1 ⇀[q] M & M ⇀*[r] M2.
+#s #M1 #M2 * -s -M1 -M2
+[ #M #q #r #H destruct 
+| #p #M1 #M #HM1 #s #M2 #HM2 #q #r #H destruct /2 width=3/
+]
+qed-.
+
+lemma lsreds_inv_lsred: ∀p,M1,M2. M1 ⇀*[p::◊] M2 → M1 ⇀[p] M2.
+#p #M1 #M2 #H
+elim (lsreds_inv_cons … H ???) -H [4: // |2,3: skip ] #M #HM1 #H (**) (* simplify line *)
+<(lsreds_inv_nil … H ?) -H //
+qed-.
+
+(* Note: "|s|" should be unparetesized *)
+lemma lsreds_fwd_mult: ∀s,M1,M2. M1 ⇀*[s] M2 → #{M2} ≤ #{M1} ^ (2 ^ (|s|)).
+#s #M1 #M2 #H elim H -s -M1 -M2 normalize //
+#p #M1 #M #HM1 #s #M2 #_ #IHM2
+lapply (lsred_fwd_mult … HM1) -p #HM1
+@(transitive_le … IHM2) -M2
+/3 width=1 by le_exp1, lt_O_exp, lt_to_le/ (**) (* auto: slow without trace *)
+qed-.
+
+lemma lsreds_lift: ∀s. liftable (lsreds s).
+#s #h #M1 #M2 #H elim H -s -M1 -M2 // /3 width=3/
+qed.
+
+lemma lsreds_inv_lift: ∀s. deliftable (lsreds s).
+#s #h #N1 #N2 #H elim H -s -N1 -N2 /2 width=3/
+#p #N1 #N #HN1 #s #N2 #_ #IHN2 #d #M1 #HMN1
+elim (lsred_inv_lift … HN1 … HMN1) -N1 #M #HM1 #HMN
+elim (IHN2 … HMN) -N /3 width=3/
+qed-.
+
+lemma lsreds_dsubst: ∀s. dsubstable_dx (lsreds s).
+#s #D #M1 #M2 #H elim H -s -M1 -M2 // /3 width=3/
+qed.
+
+theorem lsreds_mono: ∀s. singlevalued … (lsreds s).
+#s #M #N1 #H elim H -s -M -N1
+[ /2 width=3 by lsreds_inv_nil/
+| #p #M #M1 #HM1 #s #N1 #_ #IHMN1 #N2 #H
+  elim (lsreds_inv_cons … H ???) -H [4: // |2,3: skip ] #M2 #HM2 #HMN2 (**) (* simplify line *)
+  lapply (lsred_mono … HM1 … HM2) -M #H destruct /2 width=1/
+]
+qed-.
+
+theorem lsreds_trans: ∀s1,M1,M. M1 ⇀*[s1] M →
+                      ∀s2,M2. M ⇀*[s2] M2 → M1 ⇀*[s1@s2] M2.
+#s1 #M1 #M #H elim H -s1 -M1 -M normalize // /3 width=3/
+qed-.
index c5bf4ba999564949a784eced2a19890315e5c2d6..be86cb2346283a8da55b0ac11f4377844f4ef7eb 100644 (file)
@@ -15,7 +15,7 @@
 include "redex_pointer.ma".
 include "multiplicity.ma".
 
-(* LABELLED SEQUENTIAL REDUCTION (ONE STEP) *********************************)
+(* LABELLED SEQUENTIAL REDUCTION (SINGLE STEP) ******************************)
 
 (* Note: the application "(A B)" is represented by "@B.A" following:
          F. Kamareddine and R.P. Nederpelt: "A useful λ-notation".
@@ -29,14 +29,63 @@ inductive lsred: rpointer → relation term ≝
 .
 
 interpretation "labelled sequential reduction"
-    'LablSeqRed M p N = (lsred p M N).
+   'SeqRed M p N = (lsred p M N).
 
 (* Note: we do not use → since it is reserved by CIC *)
 notation "hvbox( M break ⇀ [ term 46 p ] break term 46 N )"
    non associative with precedence 45
-   for @{ 'LablSeqRed $M $p $N }.
+   for @{ 'SeqRed $M $p $N }.
 
-theorem lsred_fwd_mult: ∀p,M,N. M ⇀[p] N → #{N} < #{M} * #{M}.
+lemma lsred_inv_vref: ∀p,M,N. M ⇀[p] N → ∀i. #i = M → ⊥.
+#p #M #N * -p -M -N
+[ #A #D #i #H destruct
+| #p #A #C #_ #i #H destruct
+| #p #B #D #A #_ #i #H destruct
+| #p #B #A #C #_ #i #H destruct
+]
+qed-.
+
+lemma lsred_inv_beta: ∀p,M,N. M ⇀[p] N → ∀D,C. @D.C = M → ◊ = p →
+                      ∃∃A. 𝛌.A = C & [⬐D] A = N.
+#p #M #N * -p -M -N
+[ #A #D #D0 #C0 #H #_ destruct /2 width=3/
+| #p #A #C #_ #D0 #C0 #H destruct
+| #p #B #D #A #_ #D0 #C0 #_ #H destruct
+| #p #B #A #C #_ #D0 #C0 #_ #H destruct
+]
+qed-.
+
+lemma lsred_inv_abst: ∀p,M,N. M ⇀[p] N → ∀A. 𝛌.A = M →
+                      ∃∃C. A ⇀[p] C & 𝛌.C = N.
+#p #M #N * -p -M -N
+[ #A #D #A0 #H destruct
+| #p #A #C #HAC #A0 #H destruct /2 width=3/
+| #p #B #D #A #_ #A0 #H destruct
+| #p #B #A #C #_ #A0 #H destruct
+]
+qed-.
+
+lemma lsred_inv_appl_sn: ∀p,M,N. M ⇀[p] N → ∀B,A,q. @B.A = M → true::q = p →
+                         ∃∃D. B ⇀[q] D & @D.A = N.
+#p #M #N * -p -M -N
+[ #A #D #B0 #A0 #p0 #_ #H destruct
+| #p #A #C #_ #B0 #D0 #p0 #H destruct
+| #p #B #D #A #HBD #B0 #A0 #p0 #H1 #H2 destruct /2 width=3/
+| #p #B #A #C #_ #B0 #A0 #p0 #_ #H destruct
+]
+qed-.
+
+lemma lsred_inv_appl_dx: ∀p,M,N. M ⇀[p] N → ∀B,A,q. @B.A = M → false::q = p →
+                         ∃∃C. A ⇀[q] C & @B.C = N.
+#p #M #N * -p -M -N
+[ #A #D #B0 #A0 #p0 #_ #H destruct
+| #p #A #C #_ #B0 #D0 #p0 #H destruct
+| #p #B #D #A #_ #B0 #A0 #p0 #_ #H destruct
+| #p #B #A #C #HAC #B0 #A0 #p0 #H1 #H2 destruct /2 width=3/
+]
+qed-.
+
+lemma lsred_fwd_mult: ∀p,M,N. M ⇀[p] N → #{N} < #{M} * #{M}.
 #p #M #N #H elim H -p -M -N
 [ #A #D @(le_to_lt_to_lt … (#{A}*#{D})) //
   normalize /3 width=1 by lt_minus_to_plus_r, lt_times/ (**) (* auto: too slow without trace *) 
@@ -49,3 +98,42 @@ theorem lsred_fwd_mult: ∀p,M,N. M ⇀[p] N → #{N} < #{M} * #{M}.
 @(transitive_le … (#{B}*#{B}+#{A}*#{A})) [ /2 width=1/ ]
 >distributive_times_plus normalize /2 width=1/
 qed-.
+
+lemma lsred_lift: ∀p. liftable (lsred p).
+#p #h #M1 #M2 #H elim H -p -M1 -M2 normalize /2 width=1/
+#A #D #d <dsubst_lift_le //
+qed.
+
+lemma lsred_inv_lift: ∀p. deliftable (lsred p).
+#p #h #N1 #N2 #H elim H -p -N1 -N2
+[ #C #D #d #M1 #H
+  elim (lift_inv_appl … H) -H #B #M #H0 #HM #H destruct
+  elim (lift_inv_abst … HM) -HM #A #H0 #H destruct /3 width=3/
+| #p #C1 #C2 #_ #IHC12 #d #M1 #H
+  elim (lift_inv_abst … H) -H #A1 #H0 #H destruct
+  elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #H destruct (**) (* simplify line *)
+  @(ex2_1_intro … (𝛌.A2)) // /2 width=1/
+| #p #D1 #D2 #C1 #_ #IHD12 #d #M1 #H
+  elim (lift_inv_appl … H) -H #B1 #A #H1 #H2 #H destruct
+  elim (IHD12 ???) -IHD12 [4: // |2,3: skip ] #B2 #HB12 #H destruct (**) (* simplify line *)
+  @(ex2_1_intro … (@B2.A)) // /2 width=1/
+| #p #D1 #C1 #C2 #_ #IHC12 #d #M1 #H
+  elim (lift_inv_appl … H) -H #B #A1 #H1 #H2 #H destruct
+  elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #H destruct (**) (* simplify line *)
+  @(ex2_1_intro … (@B.A2)) // /2 width=1/
+]
+qed-.
+
+lemma lsred_dsubst: ∀p. dsubstable_dx (lsred p).
+#p #D1 #M1 #M2 #H elim H -p -M1 -M2 normalize /2 width=1/
+#A #D2 #d >dsubst_dsubst_ge //
+qed.
+
+theorem lsred_mono: ∀p. singlevalued … (lsred p).
+#p #M #N1 #H elim H -p -M -N1
+[ #A #D #N2 #H elim (lsred_inv_beta … H ????) -H [4,5: // |2,3: skip ] #A0 #H1 #H2 destruct // (**) (* simplify line *)
+| #p #A #C #_ #IHAC #N2 #H elim (lsred_inv_abst … H ??) -H [3: // |2: skip ] #C0 #HAC #H destruct /3 width=1/ (**) (* simplify line *)
+| #p #B #D #A #_ #IHBD #N2 #H elim (lsred_inv_appl_sn … H ?????) -H [5,6: // |2,3,4: skip ] #D0 #HBD #H destruct /3 width=1/ (**) (* simplify line *)
+| #p #B #A #C #_ #IHAC #N2 #H elim (lsred_inv_appl_dx … H ?????) -H [5,6: // |2,3,4: skip ] #C0 #HAC #H destruct /3 width=1/ (**) (* simplify line *)
+]
+qed-.
index 817c0c1d7f8a68dddb49f8d31828723e7d9bec24..8f2199fc3a1d097b322cbe229f4b6007bb534e96 100644 (file)
@@ -239,3 +239,10 @@ elim (le_inv_plus_l … Hd12) -Hd12 #Hd12 #Hh1d2
 lapply (sym_eq term … H) -H >(plus_minus_m_m … Hh1d2) in ⊢ (???%→?); -Hh1d2 #H
 elim (lift_inv_lift_le … Hd12 … H) -H -Hd12 /2 width=3/
 qed-.
+
+definition liftable: predicate (relation term) ≝ λR.
+                     ∀h,M1,M2. R M1 M2 → ∀d. R (↑[d, h] M1) (↑[d, h] M2).
+
+definition deliftable: predicate (relation term) ≝ λR.
+                       ∀h,N1,N2. R N1 N2 → ∀d,M1. ↑[d, h] M1 = N1 →
+                       ∃∃M2. R M1 M2 & ↑[d, h] M2 = N2.
diff --git a/matita/matita/contribs/lambda/parallel_reduction.ma b/matita/matita/contribs/lambda/parallel_reduction.ma
new file mode 100644 (file)
index 0000000..4cb204b
--- /dev/null
@@ -0,0 +1,81 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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 "multiplicity.ma".
+
+(* PARALLEL REDUCTION (SINGLE STEP) *****************************************)
+
+(* Note: the application "(A B)" is represented by "@B.A"
+         as for labelled sequential reduction
+*)
+inductive pred: relation term ≝
+| pred_vref: ∀i. pred (#i) (#i)
+| pred_abst: ∀A,C. pred A C → pred (𝛌.A) (𝛌.C) 
+| pred_appl: ∀B,D,A,C. pred B D → pred A C → pred (@B.A) (@D.C)
+| pred_beta: ∀B,D,A,C. pred B D → pred A C → pred (@B.𝛌.A) ([⬐D]C)
+.
+
+interpretation "parallel reduction"
+    'ParRed M N = (pred M N).
+
+notation "hvbox( M break ⥤ break term 46 N )"
+   non associative with precedence 45
+   for @{ 'ParRed $M $N }.
+
+lemma pred_refl: reflexive … pred.
+#M elim M -M // /2 width=1/
+qed.
+
+lemma pred_inv_vref: ∀M,N. M ⥤ N → ∀i. #i = M → #i = N.
+#M #N * -M -N //
+[ #A #C #_ #i #H destruct
+| #B #D #A #C #_ #_ #i #H destruct
+| #B #D #A #C #_ #_ #i #H destruct
+]
+qed-.
+
+lemma pred_inv_abst: ∀M,N. M ⥤ N → ∀A. 𝛌.A = M →
+                     ∃∃C. A ⥤ C & 𝛌.C = N.
+#M #N * -M -N
+[ #i #A0 #H destruct
+| #A #C #HAC #A0 #H destruct /2 width=3/
+| #B #D #A #C #_ #_ #A0 #H destruct
+| #B #D #A #C #_ #_ #A0 #H destruct
+]
+qed-.
+
+lemma pred_lift: liftable pred.
+#h #M1 #M2 #H elim H -M1 -M2 normalize // /2 width=1/
+#D #D #A #C #_ #_ #IHBD #IHAC #d <dsubst_lift_le // /2 width=1/
+qed.
+
+lemma pred_inv_lift: deliftable pred.
+#h #N1 #N2 #H elim H -N1 -N2 /2 width=3/
+[ #C1 #C2 #_ #IHC12 #d #M1 #H
+  elim (lift_inv_abst … H) -H #A1 #HAC1 #H
+  elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
+  @(ex2_1_intro … (𝛌.A2)) // /2 width=1/
+| #D1 #D2 #C1 #C2 #_ #_ #IHD12 #IHC12 #d #M1 #H
+  elim (lift_inv_appl … H) -H #B1 #A1 #HBD1 #HAC1 #H
+  elim (IHD12 … HBD1) -D1 #B2 #HB12 #HBD2
+  elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
+  @(ex2_1_intro … (@B2.A2)) // /2 width=1/
+| #D1 #D2 #C1 #C2 #_ #_ #IHD12 #IHC12 #d #M1 #H
+  elim (lift_inv_appl … H) -H #B1 #M #HBD1 #HM #H1
+  elim (lift_inv_abst … HM) -HM #A1 #HAC1 #H
+  elim (IHD12 … HBD1) -D1 #B2 #HB12 #HBD2
+  elim (IHC12 … HAC1) -C1 #A2 #HA12 #HAC2 destruct
+  @(ex2_1_intro … ([⬐B2]A2)) /2 width=1/
+]
+qed-.
diff --git a/matita/matita/contribs/lambda/policy.txt b/matita/matita/contribs/lambda/policy.txt
new file mode 100644 (file)
index 0000000..4d6e5bd
--- /dev/null
@@ -0,0 +1,14 @@
+NAMING CONVENTIONS FOR METAVARIABLES
+
+A, B, C, D: term
+H         : transient premise
+IH        : inductive premise
+M, N      : term
+
+a, b      : boolean
+d, e      : variable reference depth 
+h         : relocation height
+i, j      : de Bruijn index
+k         : relocation height
+p, q      : redex pointer
+r, s      : redex pointer sequence
index 3d5efb8f2acf300c1107057d1e5b50d7a5ca0073..f3884b229489f7bbe58cd0bde50e28c7086b6833 100644 (file)
@@ -14,7 +14,7 @@
 
 include "basics/star.ma".
 include "basics/lists/list.ma".
-include "arithmetics/nat.ma".
+include "arithmetics/exp.ma".
 
 include "xoa_notation.ma".
 include "xoa.ma".
index f6e05c93fa740e664bad3e766e73cde027eb232d..6895ea787e5440f07250c2799b8607fed6313a56 100644 (file)
@@ -47,6 +47,6 @@ notation "hvbox(a break ≺ b)"
 interpretation "'less than' on redex pointers"
    'lt p q = (TC rpointer rpprec p q).
 
-(* Note: this is p ≤ q, that *really* is p < q ∨ p = q *)
+(* Note: this is p ≤ q *)
 interpretation "'less or equal to' on redex pointers"
-   'leq x y = (RC rpointer (TC rpointer rpprec) x y).
+   'leq x y = (star rpointer x y).
diff --git a/matita/matita/contribs/lambda/redex_pointer_sequence.ma b/matita/matita/contribs/lambda/redex_pointer_sequence.ma
new file mode 100644 (file)
index 0000000..5b05f47
--- /dev/null
@@ -0,0 +1,21 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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 "redex_pointer.ma".
+
+(* REDEX POINTER SEQUENCE ***************************************************)
+
+(* Policy: pointer sequence metavariables: r, s *)
+
+definition rpseq: Type[0] \def list rpointer.
index a330f6224ef1fdb6c94e356b71955f9bad0129e5..5d000e5c19a81a7e6869da8214d4d65a3d31fc67 100644 (file)
@@ -175,8 +175,7 @@ let rec length (A:Type[0]) (l:list A) on l ≝
     [ nil ⇒ 0
     | cons a tl ⇒ S (length A tl)].
 
-notation "|M|" non associative with precedence 65 for @{'norm $M}.
-interpretation "norm" 'norm l = (length ? l).
+interpretation "list length" 'card l = (length ? l).
 
 lemma length_tail: ∀A,l. length ? (tail A l) = pred (length ? l).
 #A #l elim l // 
index ed5130955886e557a1e11b8c83b09f7ad649cb48..8b746ef6896d4f58bb4efa17728a27b7f8bba0e7 100644 (file)
@@ -48,6 +48,9 @@ definition tight_apart: ∀A.∀eq,ap:relation A.Prop
 definition antisymmetric: ∀A.∀R:relation A.Prop
 ≝ λA.λR.∀x,y:A. R x y → ¬(R y x).
 
+definition singlevalued: ∀A,B. predicate (relation2 A B) ≝ λA,B,R.
+                         ∀a,b1. R a b1 → ∀b2. R a b2 → b1 = b2.
+
 (* Reflexive closure ************)
 
 definition RC: ∀A:Type[0]. relation A → relation A ≝