to view this site correctly, please select a font
with <link to="http://www.unicode.org/">Unicode</link> support.
</body>
+ <body>
+ <ucs-bronze char="03BB"/>
+ </body>
<!-- ===================================================================== -->
</xsl:choose>
</xsl:template>
+<xsl:template match="ld:ucs-bronze">
+ <img
+ alt="[Official bronze sponsor of Unicode Consortium]"
+ title="Official bronze sponsor of Unicode Consortium"
+ src="{$baseurl}images/bronze-{@char}.png"
+ />
+</xsl:template>
+
<xsl:template match="ld:footer">
<xsl:call-template name="rule"/>
<div class="spacer"><br/></div>
--- /dev/null
+include "basic_2/dynamic/cnv_cpce.ma".
+
+(*
+lemma cpce_inv_eta_drops (h) (n) (G) (L) (i):
+ ∀X. ⦃G,L⦄ ⊢ #i ⬌η[h] X →
+ ∀K,W. ⇩*[i] L ≘ K.ⓛW →
+ ∀p,V1,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V1.U →
+ ∀V2. ⦃G,K⦄ ⊢ V1 ⬌η[h] V2 →
+ ∀W2. ⇧*[↑i] V2 ≘ W2 → X = +ⓛW2.ⓐ#0.#↑i.
+*)
+
+
+theorem cpce_mono_cnv (h) (a) (G) (L):
+ ∀T. ⦃G,L⦄ ⊢ T ![h,a] →
+ ∀T1. ⦃G,L⦄ ⊢ T ⬌η[h] T1 → ∀T2. ⦃G,L⦄ ⊢ T ⬌η[h] T2 → T1 = T2.
+#h #a #G #L #T #HT
--- /dev/null
+(* Advanced inversion lemmas ************************************************)
+
+lemma cpt_inv_sort_sn (h) (n) (G) (L) (s):
+ ∀X2. ⦃G,L⦄ ⊢ ⋆s ⬆[h,n] X2 →
+ ∨∨ ∧∧ X2 = ⋆s & n = 0
+ | ∧∧ X2 = ⋆(⫯[h]s) & n =1.
+#h #n #G #L #s #X2 * #c #Hc #H
+elim (cpg_inv_sort1 … H) -H * #H1 #H2 destruct
+/3 width=1 by or_introl, or_intror, conj/
+qed-.
+
+lemma cpt_inv_zero_sn (h) (n) (G) (L):
+ ∀X2. ⦃G,L⦄ ⊢ #0 ⬆[h,n] X2 →
+ ∨∨ ∧∧ X2 = #0 & n = 0
+ | ∃∃K,V1,V2. ⦃G,K⦄ ⊢ V1 ⬆[h,n] V2 & ⇧*[1] V2 ≘ X2 & L = K.ⓓV1
+ | ∃∃m,K,V1,V2. ⦃G,K⦄ ⊢ V1 ⬆[h,m] V2 & ⇧*[1] V2 ≘ X2 & L = K.ⓛV1 & n = ↑m.
+#h #n #G #L #X2 * #c #Hc #H
+elim (cpg_inv_zero1 … H) -H *
+[ #H1 #H2 destruct /3 width=1 by or3_intro0, conj/
+| #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 destruct
+ /4 width=8 by or3_intro1, ex3_3_intro, ex2_intro/
+| #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 destruct
+ elim (ist_inv_plus_SO_dx … H2) -H2 [| // ] #m #Hc #H destruct
+ /4 width=8 by or3_intro2, ex4_4_intro, ex2_intro/
+]
+qed-.
+
+lemma cpt_inv_lref_sn (h) (n) (G) (L) (i):
+ ∀X2. ⦃G,L⦄ ⊢ #↑i ⬆[h,n] X2 →
+ ∨∨ ∧∧ X2 = #(↑i) & n = 0
+ | ∃∃I,K,T. ⦃G,K⦄ ⊢ #i ⬆[h,n] T & ⇧*[1] T ≘ X2 & L = K.ⓘ{I}.
+#h #n #G #L #i #X2 * #c #Hc #H
+elim (cpg_inv_lref1 … H) -H *
+[ #H1 #H2 destruct /3 width=1 by or_introl, conj/
+| #I #K #V2 #HV2 #HVT2 #H destruct
+ /4 width=6 by ex3_3_intro, ex2_intro, or_intror/
+]
+qed-.
+
+lemma cpt_inv_gref_sn (h) (n) (G) (L) (l):
+ ∀X2. ⦃G,L⦄ ⊢ §l ⬆[h,n] X2 → ∧∧ X2 = §l & n = 0.
+#h #n #G #L #l #X2 * #c #Hc #H
+elim (cpg_inv_gref1 … H) -H #H1 #H2 destruct
+/2 width=1 by conj/
+qed-.
+
+
+lemma cpt_inv_sort_sn_iter (h) (n) (G) (L) (s):
+ ∀X2. ⦃G,L⦄ ⊢ ⋆s ⬆[h,n] X2 →
+ ∧∧ X2 = ⋆(((next h)^n) s) & n ≤ 1.
+#h #n #G #L #s #X2 #H
+elim (cpt_inv_sort_sn … H) -H * #H1 #H2 destruct
+/2 width=1 by conj/
+qed-.
--- /dev/null
+(* Properties with t-bound rt-transition for terms **************************)
+
+axiom cpt_total (h) (n) (G) (L) (T):
+ ∃U. ⦃G,L⦄ ⊢ T ⬆[h,n] U.
+
+lemma pippo (h) (G) (L) (T0):
+ ∀T1. ⦃G,L⦄ ⊢ T1 ➡[h] T0 → ∀n,T2. ⦃G,L⦄ ⊢ T0 ⬆[h,n] T2 →
+ ∃∃T. ⦃G,L⦄ ⊢ T1 ⬆[h,n] T & ⦃G,L⦄ ⊢ T ➡[h] T2.
+#h #G #L #T0 #T1 #H
+@(cpr_ind … H) -G -L -T0 -T1
+[ #I #G #L #n #T2 #HT2
+ /2 width=3 by ex2_intro/
+| #G #K #V1 #V0 #W0 #_ #IH #HVW0 #n #T2 #HT2
+ elim (cpt_inv_lifts_sn … HT2 (Ⓣ) … K … HVW0) -W0
+ [| /3 width=1 by drops_refl, drops_drop/ ] #V2 #HVT2 #HV02
+ elim (IH … HV02) -V0 #V0 #HV10 #HV02
+|
+|
+|
+|
+|
+| #p #G #L #V1 #V0 #W1 #W0 #T1 #T0 #_ #_ #_ #IHV #IHW #IHT #n #X #HX
+ elim (cpt_inv_bind_sn … HX) -HX #X0 #T2 #HX #HT02 #H destruct
+ elim (cpt_inv_cast_sn … HX) -HX *
+ [ #W2 #V2 #HW02 #HV02 #H destruct
+ elim (cpt_total h n G (L.ⓛW1) T0) #T3 #HT03
+ elim (IHV … HV02) -V0 #V0 #HV10 #HV02
+ elim (IHW … HW02) -W0 #W0 #HW10 #HW02
+ elim (IHT … HT02) -T0 #T0 #HT10 #HT02
+ @(ex2_intro … (ⓐV0.ⓛ{p}W0.T0))
+ [ /3 width=1 by cpt_appl, cpt_bind/
+ | @(cpm_beta … HV02 HW02)
+
+ | #m #_ #H destruct
+ ]
+
+lemma cpm_cpt_cpr (h) (n) (G) (L):
+ ∀T1,T2. ⦃G,L⦄ ⊢ T1 ➡[n,h] T2 →
+ ∃∃T0. ⦃G,L⦄ ⊢ T1 ⬆[h,n] T0 & ⦃G,L⦄ ⊢ T0 ➡[h] T2.
+#h #n #G #L #T1 #T2 #H
+@(cpm_ind … H) -n -G -L -T1 -T2
+[ #I #G #L /2 width=3 by ex2_intro/
+| #G #L #s /3 width=3 by cpm_sort, ex2_intro/
+| #n #G #K #V1 #V2 #W2 #_ * #V0 #HV10 #HV02 #HVW2
+ elim (lifts_total V0 (𝐔❴1❵)) #W0 #HVW0
+ lapply (cpm_lifts_bi … HV02 (Ⓣ) … (K.ⓓV1) … HVW0 … HVW2) -HVW2
+ [ /3 width=1 by drops_refl, drops_drop/ ] -HV02 #HW02
+ /3 width=3 by cpt_delta, ex2_intro/
+| #n #G #K #V1 #V2 #W2 #_ * #V0 #HV10 #HV02 #HVW2
+ elim (lifts_total V0 (𝐔❴1❵)) #W0 #HVW0
+ lapply (cpm_lifts_bi … HV02 (Ⓣ) … (K.ⓛV1) … HVW0 … HVW2) -HVW2
+ [ /3 width=1 by drops_refl, drops_drop/ ] -HV02 #HW02
+ /3 width=3 by cpt_ell, ex2_intro/
+| #n #I #G #K #T2 #U2 #i #_ * #T0 #HT0 #HT02 #HTU2
+ elim (lifts_total T0 (𝐔❴1❵)) #U0 #HTU0
+ lapply (cpm_lifts_bi … HT02 (Ⓣ) … (K.ⓘ{I}) … HTU0 … HTU2) -HTU2
+ [ /3 width=1 by drops_refl, drops_drop/ ] -HT02 #HU02
+ /3 width=3 by cpt_lref, ex2_intro/
+| #n #p #I #G #L #V1 #V2 #T1 #T2 #HV12 #_ #_ * #T0 #HT10 #HT02
+ /3 width=5 by cpt_bind, cpm_bind, ex2_intro/
+| #n #G #L #V1 #V2 #T1 #T2 #HV12 #_ #_ * #T0 #HT10 #HT02
+ /3 width=5 by cpt_appl, cpm_appl, ex2_intro/
+| #n #G #L #V1 #V2 #T1 #T2 #_ #_ * #V0 #HV10 #HV02 * #T0 #HT10 #HT02
+ /3 width=5 by cpt_cast, cpm_cast, ex2_intro/
+| #n #G #L #V #U1 #T1 #T2 #HTU1 #_ * #T0 #HT10 #HT02
+ elim (cpt_lifts_sn … HT10 (Ⓣ) … (L.ⓓV) … HTU1) -T1
+ [| /3 width=1 by drops_refl, drops_drop/ ] #U0 #HTU0 #HU10
+ /3 width=6 by cpt_bind, cpm_zeta, ex2_intro/
+| #n #G #L #U #T1 #T2 #_ * #T0 #HT10 #HT02
+| #n #G #L #U1 #U2 #T #_ * #U0 #HU10 #HU02
+ /3 width=3 by cpt_ee, ex2_intro/
+| #n #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #_ #_ #_ * #T0 #HT10 #HT02
+ /4 width=7 by cpt_appl, cpt_bind, cpm_beta, ex2_intro/
+| #n #p #G #L #V1 #V2 #V0 #W1 #W2 #T1 #T2 #HV12 #HW12 #_ #_ #_ * #T0 #HT10 #HT02 #HV20
+ /4 width=9 by cpt_appl, cpt_bind, cpm_theta, ex2_intro/
+]
+
+(* Forward lemmas with t-bound rt-transition for terms **********************)
+
+lemma pippo (h) (G) (L) (T0):
+ ∀T1. ⦃G,L⦄ ⊢ T1 ➡[h] T0 →
+ ∀n,T2. ⦃G,L⦄ ⊢ T0 ⬆[h,n] T2 → ⦃G,L⦄ ⊢ T1 ➡[n,h] T2.
+#h #G #L #T0 #T1 #H
+@(cpr_ind … H) -G -L -T0 -T1
+[ #I #G #L #n #T2 #HT2
+ /2 width=1 by cpt_fwd_cpm/
+| #G #K #V1 #V0 #W0 #_ #IH #HVW0 #n #T2 #HT2
+ elim (cpt_inv_lifts_sn … HT2 (Ⓣ) … K … HVW0) -W0
+ [| /3 width=1 by drops_refl, drops_drop/ ] #V2 #HVT2 #HV02
+ lapply (IH … HV02) -V0 #HV12
+ /2 width=3 by cpm_delta/
+|
+|
+|
+|
+|
+| #p #G #L #V1 #V0 #W1 #W0 #T1 #T0 #_ #_ #_ #IHV #IHW #IHT #n #X #HX
+ elim (cpt_inv_bind_sn … HX) -HX #X0 #T2 #HX #HT02 #H destruct
+ elim (cpt_inv_cast_sn … HX) -HX *
+ [ #W2 #V2 #HW02 #HV02 #H destruct
+ elim (cpt_total h n G (L.ⓛW1) T0) #T2 #HT02
+ lapply (IHV … HV02) -V0 #HV12
+ lapply (IHW … HW02) -W0 #HW12
+ lapply (IHT … HT02) -T0 #HT12
+ @(cpm_beta … HV12 HW12) //
+
+ | #m #_ #H destruct
+ ]
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||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 "basic_2/rt_transition/cpt_drops.ma".
+include "basic_2/rt_transition/cpt_fqu.ma".
+
+(* T-BOUND CONTEXT-SENSITIVE PARALLEL T-TRANSITION FOR TERMS ****************)
+
+(* Main properties **********************************************************)
+
+theorem cpt_n_O_trans (h) (n) (G) (L) (T0):
+ ∀T1. ⦃G,L⦄ ⊢ T1 ⬆[h,n] T0 →
+ ∀T2. ⦃G,L⦄ ⊢ T0 ⬆[h,0] T2 → ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2.
+#h #n #G #L #T0 #T1 #H
+@(cpt_ind … H) -H
+[ #I #G #L #X2 #HX2 //
+| #G #L #s #X2 #HX2
+ elim (cpt_inv_sort_sn_iter … HX2) -HX2 #H #_ destruct //
+| #n #G #K #V1 #V0 #W0 #_ #IH #HVW0 #W2 #HW02
+ elim (cpt_inv_lifts_sn … HW02 (Ⓣ) … K … HVW0) -W0
+ [| /3 width=1 by drops_refl, drops_drop/ ] #V2 #HVW2 #HV02
+ /3 width=3 by cpt_delta/
+| #n #G #K #W1 #W0 #V0 #_ #IH #HWV0 #V2 #HV02
+ elim (cpt_inv_lifts_sn … HV02 (Ⓣ) … K … HWV0) -V0
+ [| /3 width=1 by drops_refl, drops_drop/ ] #W2 #HWV2 #HW02
+ /3 width=3 by cpt_ell/
+| #n #I #G #K #T0 #U0 #i #_ #IH #HTU0 #U2 #HU02
+ elim (cpt_inv_lifts_sn … HU02 (Ⓣ) … K … HTU0) -U0
+ [| /3 width=1 by drops_refl, drops_drop/ ] #T2 #HTU2 #HT02
+ /3 width=3 by cpt_lref/
+| #n #p #I #G #L #V1 #V0 #T1 #T0 #_ #_ #IHV #IHT #X2 #HX2
+ elim (cpt_inv_bind_sn … HX2) -HX2 #V2 #T2 #HV02 #HT02 #H destruct
+ @cpt_bind
+ [ /2 width=1 by/
+ | @IHT
+ ]
+| #n #G #L #V1 #V0 #T1 #T0 #_ #_ #IHV #IHT #X2 #HX2
+ elim (cpt_inv_appl_sn … HX2) -HX2 #V2 #T2 #HV02 #HT02 #H destruct
+ /3 width=1 by cpt_appl/
+| #n #G #L #V1 #V0 #T1 #T0 #_ #_ #IHV #IHT #X2 #HX2
+ elim (cpt_inv_cast_sn … HX2) -HX2 *
+ [ #V2 #T2 #HV02 #HT02 #H destruct /3 width=1 by cpt_cast/
+ | #m #_ #H destruct
+ ]
+| #n #G #L #V1 #V0 #T1 #_ #IH #V2 #HV02
+ /3 width=1 by cpt_ee/
+]
+
+
+
+
\ No newline at end of file
(**************************************************************************)
include "static_2/relocation/drops.ma".
+include "static_2/relocation/lifts_lifts.ma".
include "basic_2/rt_conversion/cpce.ma".
(* CONTEXT-SENSITIVE PARALLEL ETA-CONVERSION FOR TERMS **********************)
/5 width=8 by cpce_lref, drops_drop/
]
qed.
+
+(* Inversion lemmas with uniform slicing for local environments *************)
+
+lemma cpce_inv_lref_sn_drops (h) (G) (i) (L):
+ ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
+ ∀I,K. ⇩*[i] L ≘ K.ⓘ{I} →
+ ∨∨ ∧∧ ∀n,p,W,V,U. I = BPair Abst W → ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥ & #i = X2
+ | ∃∃n,p,W,V1,V2,W2,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V1.U & ⦃G,K⦄ ⊢ V1 ⬌η[h] V2
+ & ⇧*[↑i] V2 ≘ W2 & I = BPair Abst W & +ⓛW2.ⓐ#0.#(↑i) = X2.
+#h #G #i elim i -i
+[ #L #X2 #HX2 #I #K #HLK
+ lapply (drops_fwd_isid … HLK ?) -HLK [ // ] #H destruct
+ /2 width=1 by cpce_inv_zero_sn/
+| #i #IH #L0 #X0 #HX0 #J #K #H0
+ elim (drops_inv_succ … H0) -H0 #I #L #HLK #H destruct
+ elim (cpce_inv_lref_sn … HX0) -HX0 #X2 #HX2 #HX20
+ elim (IH … HX2 … HLK) -IH -I -L *
+ [ #HJ #H destruct
+ lapply (lifts_inv_lref1_uni … HX20) -HX20 #H destruct
+ /4 width=7 by or_introl, conj/
+ | #n #p #W #V1 #V2 #W2 #U #HWU #HV12 #HVW2 #H1 #H2 destruct
+ elim (lifts_inv_bind1 … HX20) -HX20 #X2 #X #HWX2 #HX #H destruct
+ elim (lifts_inv_flat1 … HX) -HX #X0 #X1 #H0 #H1 #H destruct
+ lapply (lifts_inv_push_zero_sn … H0) -H0 #H destruct
+ elim (lifts_inv_push_succ_sn … H1) -H1 #j #Hj #H destruct
+ lapply (lifts_inv_lref1_uni … Hj) -Hj #H destruct
+ /4 width=12 by lifts_trans_uni, ex5_7_intro, or_intror/
+ ]
+]
+qed-.
+
+lemma cpce_inv_zero_sn_drops (h) (G) (i) (L):
+ ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
+ ∀I,K. ⇩*[i] L ≘ K.ⓘ{I} →
+ (∀n,p,W,V,U. I = BPair Abst W → ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥) →
+ #i = X2.
+#h #G #i #L #X2 #HX2 #I #K #HLK #HI
+elim (cpce_inv_lref_sn_drops … HX2 … HLK) -L *
+[ #_ #H //
+| #n #p #W #V1 #V2 #W2 #U #HWU #_ #_ #H destruct
+ elim (HI … HWU) -n -p -K -X2 -V1 -V2 -W2 -U -i //
+]
+qed-.
lemma cpt_refl (h) (G) (L): reflexive … (cpt h G L 0).
/3 width=3 by cpg_refl, ex2_intro/ qed.
-(* Advanced properties ******************************************************)
-
-lemma cpt_sort (h) (G) (L):
- ∀n. n ≤ 1 → ∀s. ⦃G,L⦄ ⊢ ⋆s ⬆[h,n] ⋆((next h)^n s).
-#h #G #L * //
-#n #H #s <(le_n_O_to_eq n) /2 width=1 by le_S_S_to_le/
-qed.
+(* Basic inversion lemmas ***************************************************)
+
+lemma cpt_inv_atom_sn (h) (n) (J) (G) (L):
+ ∀X2. ⦃G,L⦄ ⊢ ⓪{J} ⬆[h,n] X2 →
+ ∨∨ ∧∧ X2 = ⓪{J} & n = 0
+ | ∃∃s. X2 = ⋆(⫯[h]s) & J = Sort s & n =1
+ | ∃∃K,V1,V2. ⦃G,K⦄ ⊢ V1 ⬆[h,n] V2 & ⇧*[1] V2 ≘ X2 & L = K.ⓓV1 & J = LRef 0
+ | ∃∃m,K,V1,V2. ⦃G,K⦄ ⊢ V1 ⬆[h,m] V2 & ⇧*[1] V2 ≘ X2 & L = K.ⓛV1 & J = LRef 0 & n = ↑m
+ | ∃∃I,K,T,i. ⦃G,K⦄ ⊢ #i ⬆[h,n] T & ⇧*[1] T ≘ X2 & L = K.ⓘ{I} & J = LRef (↑i).
+#h #n #J #G #L #X2 * #c #Hc #H
+elim (cpg_inv_atom1 … H) -H *
+[ #H1 #H2 destruct /3 width=1 by or5_intro0, conj/
+| #s #H1 #H2 #H3 destruct /3 width=3 by or5_intro1, ex3_intro/
+| #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 #H3 destruct
+ /4 width=6 by or5_intro2, ex4_3_intro, ex2_intro/
+| #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 #H3 destruct
+ elim (ist_inv_plus_SO_dx … H3) -H3 [| // ] #m #Hc #H destruct
+ /4 width=9 by or5_intro3, ex5_4_intro, ex2_intro/
+| #I #K #V2 #i #HV2 #HVT2 #H1 #H2 destruct
+ /4 width=8 by or5_intro4, ex4_4_intro, ex2_intro/
+]
+qed-.
+
+lemma cpt_inv_bind_sn (h) (n) (p) (I) (G) (L) (V1) (T1):
+ ∀X2. ⦃G,L⦄ ⊢ ⓑ{p,I}V1.T1 ⬆[h,n] X2 →
+ ∃∃V2,T2. ⦃G,L⦄ ⊢ V1 ⬆[h,0] V2 & ⦃G,L.ⓑ{I}V1⦄ ⊢ T1 ⬆[h,n] T2
+ & X2 = ⓑ{p,I}V2.T2.
+#h #n #p #I #G #L #V1 #T1 #X2 * #c #Hc #H
+elim (cpg_inv_bind1 … H) -H *
+[ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct
+ elim (ist_inv_max … H2) -H2 #nV #nT #HcV #HcT #H destruct
+ elim (ist_inv_shift … HcV) -HcV #HcV #H destruct
+ /3 width=5 by ex3_2_intro, ex2_intro/
+| #cT #T2 #_ #_ #_ #_ #H destruct
+ elim (ist_inv_plus_10_dx … H)
+]
+qed-.
+
+lemma cpt_inv_appl_sn (h) (n) (G) (L) (V1) (T1):
+ ∀X2. ⦃G,L⦄ ⊢ ⓐV1.T1 ⬆[h,n] X2 →
+ ∃∃V2,T2. ⦃G,L⦄ ⊢ V1 ⬆[h,0] V2 & ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2 & X2 = ⓐV2.T2.
+#h #n #G #L #V1 #T1 #X2 * #c #Hc #H elim (cpg_inv_appl1 … H) -H *
+[ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct
+ elim (ist_inv_max … H2) -H2 #nV #nT #HcV #HcT #H destruct
+ elim (ist_inv_shift … HcV) -HcV #HcV #H destruct
+ /3 width=5 by ex3_2_intro, ex2_intro/
+| #cV #cW #cU #p #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #_ #H destruct
+ elim (ist_inv_plus_10_dx … H)
+| #cV #cW #cU #p #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #_ #_ #H destruct
+ elim (ist_inv_plus_10_dx … H)
+]
+qed-.
+
+lemma cpt_inv_cast_sn (h) (n) (G) (L) (V1) (T1):
+ ∀X2. ⦃G,L⦄ ⊢ ⓝV1.T1 ⬆[h,n] X2 →
+ ∨∨ ∃∃V2,T2. ⦃G,L⦄ ⊢ V1 ⬆[h,n] V2 & ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2 & X2 = ⓝV2.T2
+ | ∃∃m. ⦃G,L⦄ ⊢ V1 ⬆[h,m] X2 & n = ↑m.
+#h #n #G #L #V1 #T1 #X2 * #c #Hc #H elim (cpg_inv_cast1 … H) -H *
+[ #cV #cT #V2 #T2 #HV12 #HT12 #HcVT #H1 #H2 destruct
+ elim (ist_inv_max … H2) -H2 #nV #nT #HcV #HcT #H destruct
+ <idempotent_max
+ /4 width=5 by or_introl, ex3_2_intro, ex2_intro/
+| #cT #_ #H destruct
+ elim (ist_inv_plus_10_dx … H)
+| #cV #H12 #H destruct
+ elim (ist_inv_plus_SO_dx … H) -H [| // ] #m #Hm #H destruct
+ /4 width=3 by ex2_intro, or_intror/
+]
+qed-.
(* *)
(**************************************************************************)
-include "basic_2/rt_transition/cpm_drops.ma".
-include "basic_2/rt_transition/cpt_drops.ma".
+include "basic_2/rt_transition/cpm.ma".
+include "basic_2/rt_transition/cpt_fqu.ma".
(* T-BOUND CONTEXT-SENSITIVE PARALLEL T-TRANSITION FOR TERMS ****************)
-(* Properties with t-bound rt-transition for terms **************************)
-
-lemma cpm_cpt_cpr (h) (n) (G) (L):
- ∀T1,T2. ⦃G,L⦄ ⊢ T1 ➡[n,h] T2 →
- ∃∃T0. ⦃G,L⦄ ⊢ T1 ⬆[h,n] T0 & ⦃G,L⦄ ⊢ T0 ➡[h] T2.
-#h #n #G #L #T1 #T2 #H
-@(cpm_ind … H) -n -G -L -T1 -T2
-[ #I #G #L /2 width=3 by ex2_intro/
-| #G #L #s /3 width=3 by cpm_sort, ex2_intro/
-| #n #G #K #V1 #V2 #W2 #_ * #V0 #HV10 #HV02 #HVW2
- elim (lifts_total V0 (𝐔❴1❵)) #W0 #HVW0
- lapply (cpm_lifts_bi … HV02 (Ⓣ) … (K.ⓓV1) … HVW0 … HVW2) -HVW2
- [ /3 width=1 by drops_refl, drops_drop/ ] -HV02 #HW02
- /3 width=3 by cpt_delta, ex2_intro/
-| #n #G #K #V1 #V2 #W2 #_ * #V0 #HV10 #HV02 #HVW2
- elim (lifts_total V0 (𝐔❴1❵)) #W0 #HVW0
- lapply (cpm_lifts_bi … HV02 (Ⓣ) … (K.ⓛV1) … HVW0 … HVW2) -HVW2
- [ /3 width=1 by drops_refl, drops_drop/ ] -HV02 #HW02
- /3 width=3 by cpt_ell, ex2_intro/
-| #n #I #G #K #T2 #U2 #i #_ * #T0 #HT0 #HT02 #HTU2
- elim (lifts_total T0 (𝐔❴1❵)) #U0 #HTU0
- lapply (cpm_lifts_bi … HT02 (Ⓣ) … (K.ⓘ{I}) … HTU0 … HTU2) -HTU2
- [ /3 width=1 by drops_refl, drops_drop/ ] -HT02 #HU02
- /3 width=3 by cpt_lref, ex2_intro/
-| #n #p #I #G #L #V1 #V2 #T1 #T2 #HV12 #_ #_ * #T0 #HT10 #HT02
- /3 width=5 by cpt_bind, cpm_bind, ex2_intro/
-| #n #G #L #V1 #V2 #T1 #T2 #HV12 #_ #_ * #T0 #HT10 #HT02
- /3 width=5 by cpt_appl, cpm_appl, ex2_intro/
-| #n #G #L #V1 #V2 #T1 #T2 #_ #_ * #V0 #HV10 #HV02 * #T0 #HT10 #HT02
- /3 width=5 by cpt_cast, cpm_cast, ex2_intro/
-| #n #G #L #V #U1 #T1 #T2 #HTU1 #_ * #T0 #HT10 #HT02
- elim (cpt_lifts_sn … HT10 (Ⓣ) … (L.ⓓV) … HTU1) -T1
- [| /3 width=1 by drops_refl, drops_drop/ ] #U0 #HTU0 #HU10
- /3 width=6 by cpt_bind, cpm_zeta, ex2_intro/
-| #n #G #L #U #T1 #T2 #_ * #T0 #HT10 #HT02
-| #n #G #L #U1 #U2 #T #_ * #U0 #HU10 #HU02
- /3 width=3 by cpt_ee, ex2_intro/
-| #n #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #_ #_ #_ * #T0 #HT10 #HT02
- /4 width=7 by cpt_appl, cpt_bind, cpm_beta, ex2_intro/
-| #n #p #G #L #V1 #V2 #V0 #W1 #W2 #T1 #T2 #HV12 #HW12 #_ #_ #_ * #T0 #HT10 #HT02 #HV20
- /4 width=9 by cpt_appl, cpt_bind, cpm_theta, ex2_intro/
-]
-
(* Forward lemmas with t-bound rt-transition for terms **********************)
lemma cpt_fwd_cpm (h) (n) (G) (L):
- ∀T1,T2. ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2 → ⦃G,L⦄ ⊢ T1 ➡[n,h] T2.
\ No newline at end of file
+ ∀T1,T2. ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2 → ⦃G,L⦄ ⊢ T1 ➡[n,h] T2.
+#h #n #G #L #T1 #T2 #H
+@(cpt_ind … H) -n -G -L -T1 -T2
+/3 width=3 by cpm_ee, cpm_cast, cpm_appl, cpm_bind, cpm_lref, cpm_ell, cpm_delta/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||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 "static_2/s_transition/fqu_weight.ma".
+include "basic_2/rt_transition/cpt.ma".
+
+(* T-BOUND CONTEXT-SENSITIVE PARALLEL T-TRANSITION FOR TERMS ****************)
+
+(* Basic eliminators ********************************************************)
+
+lemma cpt_ind (h) (Q:relation5 …):
+ (∀I,G,L. Q 0 G L (⓪{I}) (⓪{I})) →
+ (∀G,L,s. Q 1 G L (⋆s) (⋆(⫯[h]s))) →
+ (∀n,G,K,V1,V2,W2. ⦃G,K⦄ ⊢ V1 ⬆[h,n] V2 → Q n G K V1 V2 →
+ ⇧*[1] V2 ≘ W2 → Q n G (K.ⓓV1) (#0) W2
+ ) → (∀n,G,K,V1,V2,W2. ⦃G,K⦄ ⊢ V1 ⬆[h,n] V2 → Q n G K V1 V2 →
+ ⇧*[1] V2 ≘ W2 → Q (↑n) G (K.ⓛV1) (#0) W2
+ ) → (∀n,I,G,K,T,U,i. ⦃G,K⦄ ⊢ #i ⬆[h,n] T → Q n G K (#i) T →
+ ⇧*[1] T ≘ U → Q n G (K.ⓘ{I}) (#↑i) (U)
+ ) → (∀n,p,I,G,L,V1,V2,T1,T2. ⦃G,L⦄ ⊢ V1 ⬆[h,0] V2 → ⦃G,L.ⓑ{I}V1⦄ ⊢ T1 ⬆[h,n] T2 →
+ Q 0 G L V1 V2 → Q n G (L.ⓑ{I}V1) T1 T2 → Q n G L (ⓑ{p,I}V1.T1) (ⓑ{p,I}V2.T2)
+ ) → (∀n,G,L,V1,V2,T1,T2. ⦃G,L⦄ ⊢ V1 ⬆[h,0] V2 → ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2 →
+ Q 0 G L V1 V2 → Q n G L T1 T2 → Q n G L (ⓐV1.T1) (ⓐV2.T2)
+ ) → (∀n,G,L,V1,V2,T1,T2. ⦃G,L⦄ ⊢ V1 ⬆[h,n] V2 → ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2 →
+ Q n G L V1 V2 → Q n G L T1 T2 → Q n G L (ⓝV1.T1) (ⓝV2.T2)
+ ) → (∀n,G,L,V1,V2,T. ⦃G,L⦄ ⊢ V1 ⬆[h,n] V2 →
+ Q n G L V1 V2 → Q (↑n) G L (ⓝV1.T) V2
+ ) →
+ ∀n,G,L,T1,T2. ⦃G,L⦄ ⊢ T1 ⬆[h,n] T2 → Q n G L T1 T2.
+#h #Q #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #IH8 #IH9 #n #G #L #T1
+generalize in match n; -n
+@(fqu_wf_ind (Ⓣ) … G L T1) -G -L -T1 #G0 #L0 * [| * [| * ]]
+[ #I #IH #n #X2 #HX2 -IH6 -IH7 -IH8 -IH9
+ elim (cpt_inv_atom_sn … HX2) -HX2 *
+ [ #H1 #H2 destruct -IH2 -IH3 -IH4 -IH5 //
+ | #s #H1 #H2 #H3 destruct -IH1 -IH3 -IH4 -IH5 //
+ | #K #V1 #V2 #HV12 #HVX2 #H1 #H2 destruct -IH1 -IH2 -IH4 -IH5 /3 width=4 by/
+ | #m #K #W1 #W2 #HW12 #HWX2 #H1 #H2 #H3 destruct -IH1 -IH2 -IH3 -IH5 /3 width=4 by/
+ | #J #K #T2 #i #HT2 #HTX2 #H1 #H2 destruct -IH1 -IH2 -IH3 -IH4 /3 width=4 by/
+ ]
+| #p #I #V1 #T1 #IH #n #X2 #HX2 -IH1 -IH2 -IH3 -IH4 -IH5 -IH7 -IH8 -IH9
+ elim (cpt_inv_bind_sn … HX2) -HX2 #V2 #T2 #HV12 #HT12 #H destruct
+ /4 width=1 by fqu_bind_dx/
+| #V1 #T1 #IH #n #X2 #HX2 -IH1 -IH2 -IH3 -IH4 -IH5 -IH6 -IH8 -IH9
+ elim (cpt_inv_appl_sn … HX2) -HX2 #V2 #T2 #HV12 #HT12 #H destruct /3 width=1 by/
+| #U1 #T1 #IH #n #X2 #HX2 -IH1 -IH2 -IH3 -IH4 -IH5 -IH6 -IH7
+ elim (cpt_inv_cast_sn … HX2) -HX2 *
+ [ #U2 #T2 #HU12 #HT12 #H destruct -IH9 /3 width=1 by/
+ | #m #Hm #H destruct -IH8 /3 width=1 by/
+ ]
+]
+qed-.
class "cyan"
[ { "rt-transition" * } {
[ { "context-sensitive parallel t-transition" * } {
- [ [ "for terms" ] "cpt" + "( ⦃?,?⦄ ⊢ ? ⬆[?,?] ? )" "cpt_drops" + "cpt_cpm" * ]
+ [ [ "for terms" ] "cpt" + "( ⦃?,?⦄ ⊢ ? ⬆[?,?] ? )" "cpt_drops" + "cpt_fqu" + "cpt_cpm" * ]
}
]
[ { "context-sensitive parallel r-transition" * } {
#n #H destruct //
qed-.
+lemma ist_inv_10: ∀n. 𝐓⦃n,𝟙𝟘⦄ → ⊥.
+#h #H destruct
+qed-.
+
(* Main inversion properties ************************************************)
theorem ist_inj: ∀n1,n2,c. 𝐓⦃n1,c⦄ → 𝐓⦃n2,c⦄ → n1 = n2.
elim (ist_inv_plus … H) -H #n1 #n2 #Hn1 #Hn2 #H destruct
/2 width=3 by ex2_intro/
qed-.
+
+lemma ist_inv_plus_10_dx: ∀n,c. 𝐓⦃n,c+𝟙𝟘⦄ → ⊥.
+#n #c #H
+elim (ist_inv_plus … H) -H #n1 #n2 #_ #H #_
+/2 width=2 by ist_inv_10/
+qed-.
]
qed-.
+lemma lifts_inv_push_zero_sn (f):
+ ∀X. ⇧*[⫯f]#0 ≘ X → #0 = X.
+#f #X #H
+elim (lifts_inv_lref1 … H) -H #i #Hi #H destruct
+lapply (at_inv_ppx … Hi ???) -Hi //
+qed-.
+
+lemma lifts_inv_push_succ_sn (f) (i1):
+ ∀X. ⇧*[⫯f]#(↑i1) ≘ X →
+ ∃∃i2. ⇧*[f]#i1 ≘ #i2 & #(↑i2) = X.
+#f #i1 #X #H
+elim (lifts_inv_lref1 … H) -H #j #Hij #H destruct
+elim (at_inv_npx … Hij) -Hij [|*: // ] #i2 #Hi12 #H destruct
+/3 width=3 by lifts_lref, ex2_intro/
+qed-.
+
(* Inversion lemmas with uniform relocations ********************************)
lemma lifts_inv_lref1_uni: ∀l,Y,i. ⇧*[l] #i ≘ Y → Y = #(l+i).
elim (HR … HU12 … HTU1) -HR -U1 #X #HUX #HTX
<(lifts_inj … HUX … HTU2) -U2 -T2 -f //
qed-.
+
+lemma lifts_trans_uni (T):
+ ∀l1,T1. ⇧*[l1] T1 ≘ T →
+ ∀l2,T2. ⇧*[l2] T ≘ T2 → ⇧*[l1+l2] T1 ≘ T2.
+#T #l1 #T1 #HT1 #l2 #T2 #HT2
+@(lifts_trans … HT1 … HT2) //
+qed-.
--- /dev/null
+(*
+ ||M|| This file is part of HELM, an Hypertextual, Electronic
+ ||A|| Library of Mathematics, developed at the Computer Science
+ ||T|| Department of the University of Bologna, Italy.
+ ||I||
+ ||T||
+ ||A||
+ \ / This file is distributed under the terms of the
+ \ / GNU General Public License Version 2
+ V_______________________________________________________________ *)
+
+include "basics/deqsets.ma".
+include "basics/lists/listb.ma".
+
+(*
+
+record DeqSet : Type[1] ≝ {
+ carr :> Type[0];
+ eqb: carr → carr → bool;
+ eqb_true: ∀x,y. (eqb x y = true) ↔ (x = y)
+}. *)
+
+
+(* list *)
+
+let rec eq_list (A:DeqSet) (l1,l2:list A) on l1 ≝
+ match l1 with
+ [ nil ⇒ match l2 with [nil ⇒ true | _ ⇒ false]
+ | cons a1 tl1 ⇒
+ match l2 with [nil ⇒ false | cons a2 tl2 ⇒ a1 ==a2 ∧ eq_list A tl1 tl2]].
+
+lemma eq_list_true: ∀A:DeqSet.∀l1,l2:list A.
+ eq_list A l1 l2 = true ↔ l1 = l2.
+#A #l1 elim l1
+ [* [% // |#a #tl % normalize #H destruct ]
+ |#a1 #tl1 #Hind *
+ [normalize % #H destruct
+ |#a2 #tl2 normalize %
+ [cases (true_or_false (a1==a2)) #Heq >Heq normalize
+ [#Htl >(\P Heq) >(proj1 … (Hind tl2) Htl) // | #H destruct ]
+ |#H destruct >(\b (refl … )) >(proj2 … (Hind tl2) (refl …)) //
+ ]
+ ]
+ ]
+qed.
+
+definition DeqList ≝ λA:DeqSet.
+ mk_DeqSet (list A) (eq_list A) (eq_list_true A).
+
+unification hint 0 ≔ C;
+ T ≟ carr C,
+ X ≟ DeqList C
+(* ---------------------------------------- *) ⊢
+ list T ≡ carr X.
+
+alias id "eqb" = "cic:/matita/basics/deqsets/eqb#fix:0:0:3".
+alias symbol "hint_decl" (instance 1) = "hint_decl_Type0".
+unification hint 0 ≔ T,a1,a2;
+ X ≟ DeqList T
+(* ---------------------------------------- *) ⊢
+ eq_list T a1 a2 ≡ eqb X a1 a2.
+
+
\ / GNU General Public License Version 2
V_______________________________________________________________ *)
-include "basics/lists/listb.ma".
+include "basics/deqlist.ma".
(****** DeqSet: a set with a decidable equality ******)
#A #B #f #a #b #eqf @memb_filter_l [@(\b eqf)]
@enum_prod_complete //
qed.
+
+(* FinFun *)
+
+definition enum_fun_raw: ∀A,B:DeqSet.list A → list B → list (list (DeqProd A B)) ≝
+ λA,B,lA,lB.foldr A (list (list (DeqProd A B)))
+ (λa.compose ??? (λb. cons ? 〈a,b〉) lB) [[]] lA.
+
+lemma enum_fun_raw_cons: ∀A,B,a,lA,lB.
+ enum_fun_raw A B (a::lA) lB =
+ compose ??? (λb. cons ? 〈a,b〉) lB (enum_fun_raw A B lA lB).
+//
+qed.
+
+definition is_functional ≝ λA,B:DeqSet.λlA:list A.λl: list (DeqProd A B).
+ map ?? (fst A B) l = lA.
+
+definition carr_fun ≝ λA,B:FinSet.
+ DeqSig (DeqList (DeqProd A B)) (is_functional A B (enum A)).
+
+definition carr_fun_l ≝ λA,B:DeqSet.λl.
+ DeqSig (DeqList (DeqProd A B)) (is_functional A B l).
+
+lemma compose_spec1 : ∀A,B,C:DeqSet.∀f:A→B→C.∀a:A.∀b:B.∀lA:list A.∀lB:list B.
+ a ∈ lA = true → b ∈ lB = true → ((f a b) ∈ (compose A B C f lA lB)) = true.
+#A #B #C #f #a #b #lA elim lA
+ [normalize #lB #H destruct
+ |#a1 #tl #Hind #lB #Ha #Hb cases (orb_true_l ?? Ha) #Hcase
+ [>(\P Hcase) normalize @memb_append_l1 @memb_map //
+ |@memb_append_l2 @Hind //
+ ]
+ ]
+qed.
+
+lemma compose_cons: ∀A,B,C.∀f:A→B→C.∀l1,l2,a.
+ compose A B C f (a::l1) l2 =
+ (map ?? (f a) l2)@(compose A B C f l1 l2).
+// qed.
+
+lemma compose_spec2 : ∀A,B,C:DeqSet.∀f:A→B→C.∀c:C.∀lA:list A.∀lB:list B.
+ c ∈ (compose A B C f lA lB) = true →
+ ∃a,b.a ∈ lA = true ∧ b ∈ lB = true ∧ c = f a b.
+#A #B #C #f #c #lA elim lA
+ [normalize #lB #H destruct
+ |#a1 #tl #Hind #lB >compose_cons #Hc cases (memb_append … Hc) #Hcase
+ [lapply(memb_map_to_exists … Hcase) * #b * #Hb #Hf
+ %{a1} %{b} /3/
+ |lapply(Hind ? Hcase) * #a2 * #b * * #Ha #Hb #Hf %{a2} %{b} % // % //
+ @orb_true_r2 //
+ ]
+ ]
+qed.
+
+definition compose2 ≝
+ λA,B:DeqSet.λa:A.λl. compose B (carr_fun_l A B l) (carr_fun_l A B (a::l))
+ (λb,tl. mk_Sig ?? (〈a,b〉::(pi1 … tl)) ?).
+normalize @eq_f @(pi2 … tl)
+qed.
+
+let rec Dfoldr (A:Type[0]) (B:list A → Type[0])
+ (f:∀a:A.∀l.B l → B (a::l)) (b:B [ ]) (l:list A) on l : B l ≝
+ match l with [ nil ⇒ b | cons a l ⇒ f a l (Dfoldr A B f b l)].
+
+definition empty_graph: ∀A,B:DeqSet. carr_fun_l A B [].
+#A #B %{[]} // qed.
+
+definition enum_fun: ∀A,B:DeqSet.∀lA:list A.list B → list (carr_fun_l A B lA) ≝
+ λA,B,lA,lB.Dfoldr A (λl.list (carr_fun_l A B l))
+ (λa,l.compose2 ?? a l lB) [empty_graph A B] lA.
+
+lemma mem_enum_fun: ∀A,B:DeqSet.∀lA,lB.∀x:carr_fun_l A B lA.
+ pi1 … x ∈ map ?? (pi1 … ) (enum_fun A B lA lB) = true →
+ x ∈ enum_fun A B lA lB = true .
+#A #B #lA #lB #x @(memb_map_inj
+ (DeqSig (DeqList (DeqProd A B))
+ (λx0:DeqList (DeqProd A B).is_functional A B lA x0))
+ (DeqList (DeqProd A B)) (pi1 …))
+* #l1 #H1 * #l2 #H2 #Heq lapply H1 lapply H2 >Heq //
+qed.
+
+lemma enum_fun_cons: ∀A,B,a,lA,lB.
+ enum_fun A B (a::lA) lB =
+ compose ??? (λb,tl. mk_Sig ?? (〈a,b〉::(pi1 … tl)) ?) lB (enum_fun A B lA lB).
+//
+qed.
+
+lemma map_map: ∀A,B,C.∀f:A→B.∀g:B→C.∀l.
+ map ?? g (map ?? f l) = map ?? (g ∘ f) l.
+#A #B #C #f #g #l elim l [//]
+#a #tl #Hind normalize @eq_f @Hind
+qed.
+
+lemma map_compose: ∀A,B,C,D.∀f:A→B→C.∀g:C→D.∀l1,l2.
+ map ?? g (compose A B C f l1 l2) = compose A B D (λa,b. g (f a b)) l1 l2.
+#A #B #C #D #f #g #l1 elim l1 [//]
+#a #tl #Hind #l2 >compose_cons >compose_cons <map_append @eq_f2
+ [@map_map |@Hind]
+qed.
+
+definition enum_fun_graphs: ∀A,B,lA,lB.
+ map ?? (pi1 … ) (enum_fun A B lA lB) = enum_fun_raw A B lA lB.
+#A #B #lA elim lA [normalize //]
+#a #tl #Hind #lB >(enum_fun_cons A B a tl lB) >enum_fun_raw_cons >map_compose
+cut (∀lB2. compose B (Σx:DeqList (DeqProd A B).is_functional A B tl x)
+ (DeqList (DeqProd A B))
+ (λa0:B
+ .λb:Σx:DeqList (DeqProd A B).is_functional A B tl x
+ .〈a,a0〉
+ ::pi1 (list (A×B)) (λx:DeqList (DeqProd A B).is_functional A B tl x) b) lB
+ (enum_fun A B tl lB2)
+ =compose B (list (A×B)) (list (A×B)) (λb:B.cons (A×B) 〈a,b〉) lB
+ (enum_fun_raw A B tl lB2))
+ [#lB2 elim lB
+ [normalize //
+ |#b #tlb #Hindb >compose_cons in ⊢ (???%); >compose_cons
+ @eq_f2 [<Hind >map_map // |@Hindb]]]
+#Hcut @Hcut
+qed.
+
+lemma uniqueb_compose: ∀A,B,C:DeqSet.∀f,l1,l2.
+ (∀a1,a2,b1,b2. f a1 b1 = f a2 b2 → a1 = a2 ∧ b1 = b2) →
+ uniqueb ? l1 = true → uniqueb ? l2 = true →
+ uniqueb ? (compose A B C f l1 l2) = true.
+#A #B #C #f #l1 #l2 #Hinj elim l1 //
+#a #tl #Hind #HuA #HuB >compose_cons @uniqueb_append
+ [@(unique_map_inj … HuB) #b1 #b2 #Hb1b2 @(proj2 … (Hinj … Hb1b2))
+ |@Hind // @(andb_true_r … HuA)
+ |#c #Hc lapply(memb_map_to_exists … Hc) * #b * #Hb2 #Hfab % #Hc
+ lapply(compose_spec2 … Hc) * #a1 * #b1 * * #Ha1 #Hb1 <Hfab #H
+ @(absurd (a=a1))
+ [@(proj1 … (Hinj … H))
+ |% #eqa @(absurd … Ha1) % <eqa #H lapply(andb_true_l … HuA) >H
+ normalize #H1 destruct (H1)
+ ]
+ ]
+qed.
+
+lemma enum_fun_unique: ∀A,B:DeqSet.∀lA,lB.
+ uniqueb ? lA = true → uniqueb ? lB = true →
+ uniqueb ? (enum_fun A B lA lB) = true.
+#A #B #lA elim lA
+ [#lB #_ #ulB //
+ |#a #tlA #Hind #lB #uA #uB lapply (enum_fun_cons A B a tlA lB) #H >H
+ @(uniqueb_compose B (carr_fun_l A B tlA) (carr_fun_l A B (a::tlA)))
+ [#b1 #b2 * #l1 #funl1 * #l2 #funl2 #H1 destruct (H1) /2/
+ |//
+ |@(Hind … uB) @(andb_true_r … uA)
+ ]
+ ]
+qed.
+
+lemma enum_fun_complete: ∀A,B:FinSet.∀l1,l2.
+ (∀x:A. memb A x l1 = true) →
+ (∀x:B. memb B x l2 = true) →
+ ∀x:carr_fun_l A B l1. memb ? x (enum_fun A B l1 l2) = true.
+#A #B #l1 #l2 #H1 #H2 * #g #H @mem_enum_fun >enum_fun_graphs
+lapply H -H lapply g -g elim l1
+ [* // #p #tlg normalize #H destruct (H)
+ |#a #tl #Hind #g cases g
+ [normalize in ⊢ (%→?); #H destruct (H)
+ |* #a1 #b #tl1 normalize in ⊢ (%→?); #H
+ cut (is_functional A B tl tl1) [destruct (H) //] #Hfun
+ >(cons_injective_l ????? H)
+ >(enum_fun_raw_cons … ) @(compose_spec1 … (λb. cons ? 〈a,b〉))
+ [@H2 |@Hind @Hfun]
+ ]
+ ]
+qed.
+
+definition FinFun ≝
+λA,B:FinSet.mk_FinSet (carr_fun A B)
+ (enum_fun A B (enum A) (enum B))
+ (enum_fun_unique A B … (enum_unique A) (enum_unique B))
+ (enum_fun_complete A B … (enum_complete A) (enum_complete B)).
+
+(*
+unification hint 0 ≔ C1,C2;
+ T1 ≟ FinSetcarr C1,
+ T2 ≟ FinSetcarr C2,
+ X ≟ FinProd C1 C2
+(* ---------------------------------------- *) ⊢
+ T1×T2 ≡ FinSetcarr X. *)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||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 "basics/lists/list.ma".
-
-axiom alpha : Type[0].
-notation "𝔸" non associative with precedence 90 for @{'alphabet}.
-interpretation "set of names" 'alphabet = alpha.
-
-inductive tp : Type[0] ≝
-| top : tp
-| arr : tp → tp → tp.
-inductive tm : Type[0] ≝
-| var : nat → tm
-| par : 𝔸 → tm
-| abs : tp → tm → tm
-| app : tm → tm → tm.
-
-let rec Nth T n (l:list T) on n ≝
- match l with
- [ nil ⇒ None ?
- | cons hd tl ⇒ match n with
- [ O ⇒ Some ? hd
- | S n0 ⇒ Nth T n0 tl ] ].
-
-inductive judg : list tp → tm → tp → Prop ≝
-| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t
-| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u
-| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u).
-
-definition Env := list (𝔸 × tp).
-
-axiom vclose_env : Env → list tp.
-axiom vclose_tm : Env → tm → tm.
-axiom Lam : 𝔸 → tp → tm → tm.
-definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T.
-definition dom ≝ λG:Env.map ?? (fst ??) G.
-
-definition sctx ≝ 𝔸 × tm.
-axiom swap_tm : 𝔸 → 𝔸 → tm → tm.
-definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M.
-
-axiom in_list : ∀A:Type[0].A → list A → Prop.
-interpretation "list membership" 'mem x l = (in_list ? x l).
-interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)).
-
-axiom in_Env : 𝔸 × tp → Env → Prop.
-notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}.
-interpretation "Env membership" 'lefttriangle x l = (in_Env x l).
-
-let rec FV M ≝ match M with
- [ par X ⇒ [X]
- | app M1 M2 ⇒ FV M1@FV M2
- | abs T M0 ⇒ FV M0
- | _ ⇒ [ ] ].
-
-(* axiom Lookup : 𝔸 → Env → option tp. *)
-
-(* forma alto livello del judgment
- t_abs* : ∀G,T,X,M,U.
- (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) →
- Judg G (Lam X T (M[X])) (arr T U) *)
-
-(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *)
-
-axiom Judg_ind : ∀P:Env → tm → tp → Prop.
- (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) →
- (∀G,M,N,T,U.
- Judg G M (arr T U) → Judg G N T →
- P G M (arr T U) → P G N T → P G (app M N) U) →
- (∀G,T1,T2,X,M1.
- (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) →
- (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) →
- P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) →
- ∀G,M,T.Judg G M T → P G M T.
-
-axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T.
-axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U.
-axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U).
-
-definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2.
-interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2).
-
-axiom daemon : ∀P:Prop.P.
-
-theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T.
-#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?)
-@(Judg_ind … HJ)
-[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub //
-| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2
- [| @IH1 // | @IH2 // ]
-| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH
- [ (* trivial property of Lam *) @daemon
- | (* trivial property of subenv *) @daemon ]
-]
-qed.
-
-(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e
- ricorsione *)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||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 "binding/names.ma".
-
-(* permutations *)
-definition finite_perm : ∀X:Nset.(X → X) → Prop ≝
- λX,f.injective X X f ∧ surjective X X f ∧ ∃l.∀x.x ∉ l → f x = x.
-
-(* maps a permutation to a list of parameters *)
-definition Pi_list : ∀X:Nset.(X → X) → list X → list X ≝
- λX,p,xl.map ?? (λx.p x) xl.
-
-interpretation "permutation of X list" 'middot p x = (Pi_list p x).
-
-definition swap : ∀N:Nset.N → N → N → N ≝
- λN,u,v,x.match (x == u) with
- [true ⇒ v
- |false ⇒ match (x == v) with
- [true ⇒ u
- |false ⇒ x]].
-
-axiom swap_right : ∀N,x,y.swap N x y y = x.
-(*
-#N x y;nnormalize;nrewrite > (p_eqb3 ? y y …);//;
-nlapply (refl ? (y ≟ x));ncases (y ≟ x) in ⊢ (???% → %);nnormalize;//;
-#H1;napply p_eqb1;//;
-nqed.
-*)
-
-axiom swap_left : ∀N,x,y.swap N x y x = y.
-(*
-#N x y;nnormalize;nrewrite > (p_eqb3 ? x x …);//;
-nqed.
-*)
-
-axiom swap_other : ∀N,x,y,z.x ≠ z → y ≠ z → swap N x y z = z.
-(*
-#N x y z H1 H2;nnormalize;nrewrite > (p_eqb4 …);
-##[nrewrite > (p_eqb4 …);//;@;ncases H2;/2/;
-##|@;ncases H1;/2/
-##]
-nqed.
-*)
-
-axiom swap_inv : ∀N,x,y,z.swap N x y (swap N x y z) = z.
-(*
-#N x y z;nlapply (refl ? (x ≟ z));ncases (x ≟ z) in ⊢ (???% → ?);#H1
-##[nrewrite > (p_eqb1 … H1);nrewrite > (swap_left …);//;
-##|nlapply (refl ? (y ≟ z));ncases (y ≟ z) in ⊢ (???% → ?);#H2
- ##[nrewrite > (p_eqb1 … H2);nrewrite > (swap_right …);//;
- ##|nrewrite > (swap_other …) in ⊢ (??(????%)?);/2/;
- nrewrite > (swap_other …);/2/;
- ##]
-##]
-nqed.
-*)
-
-axiom swap_fp : ∀N,x1,x2.finite_perm ? (swap N x1 x2).
-(*
-#N x1 x2;@
-##[@
- ##[nwhd;#xa xb;nnormalize;nlapply (refl ? (xa ≟ x1));
- ncases (xa ≟ x1) in ⊢ (???% → %);#H1
- ##[nrewrite > (p_eqb1 … H1);nlapply (refl ? (xb ≟ x1));
- ncases (xb ≟ x1) in ⊢ (???% → %);#H2
- ##[nrewrite > (p_eqb1 … H2);//
- ##|nlapply (refl ? (xb ≟ x2));
- ncases (xb ≟ x2) in ⊢ (???% → %);#H3
- ##[nnormalize;#H4;nrewrite > H4 in H3;
- #H3;nrewrite > H3 in H2;#H2;ndestruct (H2)
- ##|nnormalize;#H4;nrewrite > H4 in H3;
- nrewrite > (p_eqb3 …);//;#H5;ndestruct (H5)
- ##]
- ##]
- ##|nlapply (refl ? (xa ≟ x2));
- ncases (xa ≟ x2) in ⊢ (???% → %);#H2
- ##[nrewrite > (p_eqb1 … H2);nlapply (refl ? (xb ≟ x1));
- ncases (xb ≟ x1) in ⊢ (???% → %);#H3
- ##[nnormalize;#H4;nrewrite > H4 in H3;
- #H3;nrewrite > (p_eqb1 … H3);@
- ##|nnormalize;nlapply (refl ? (xb ≟ x2));
- ncases (xb ≟ x2) in ⊢ (???% → %);#H4
- ##[nrewrite > (p_eqb1 … H4);//
- ##|nnormalize;#H5;nrewrite > H5 in H3;
- nrewrite > (p_eqb3 …);//;#H6;ndestruct (H6);
- ##]
- ##]
- ##|nnormalize;nlapply (refl ? (xb ≟ x1));
- ncases (xb ≟ x1) in ⊢ (???% → %);#H3
- ##[nnormalize;#H4;nrewrite > H4 in H2;nrewrite > (p_eqb3 …);//;
- #H5;ndestruct (H5)
- ##|nlapply (refl ? (xb ≟ x2));
- ncases (xb ≟ x2) in ⊢ (???% → %);#H4
- ##[nnormalize;#H5;nrewrite > H5 in H1;nrewrite > (p_eqb3 …);//;
- #H6;ndestruct (H6)
- ##|nnormalize;//
- ##]
- ##]
- ##]
- ##]
- ##|nwhd;#z;nnormalize;nlapply (refl ? (z ≟ x1));
- ncases (z ≟ x1) in ⊢ (???% → %);#H1
- ##[nlapply (refl ? (z ≟ x2));
- ncases (z ≟ x2) in ⊢ (???% → %);#H2
- ##[@ z;nrewrite > H1;nrewrite > H2;napply p_eqb1;//
- ##|@ x2;nrewrite > (p_eqb4 …);
- ##[nrewrite > (p_eqb3 …);//;
- nnormalize;napply p_eqb1;//
- ##|nrewrite < (p_eqb1 … H1);@;#H3;nrewrite > H3 in H2;
- nrewrite > (p_eqb3 …);//;#H2;ndestruct (H2)
- ##]
- ##]
- ##|nlapply (refl ? (z ≟ x2));
- ncases (z ≟ x2) in ⊢ (???% → %);#H2
- ##[@ x1;nrewrite > (p_eqb3 …);//;
- napply p_eqb1;nnormalize;//
- ##|@ z;nrewrite > H1;nrewrite > H2;@;
- ##]
- ##]
- ##]
-##|@ [x1;x2];#x0 H1;nrewrite > (swap_other …)
- ##[@
- ##|@;#H2;nrewrite > H2 in H1;*;#H3;napply H3;/2/;
- ##|@;#H2;nrewrite > H2 in H1;*;#H3;napply H3;//;
- ##]
-##]
-nqed.
-*)
-
-axiom inj_swap : ∀N,u,v.injective ?? (swap N u v).
-(*
-#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//;
-nqed.
-*)
-
-axiom surj_swap : ∀N,u,v.surjective ?? (swap N u v).
-(*
-#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//;
-nqed.
-*)
-
-axiom finite_swap : ∀N,u,v.∃l.∀x.x ∉ l → swap N u v x = x.
-(*
-#N u v;ncases (swap_fp N u v);*;#H1 H2 H3;//;
-nqed.
-*)
-
-(* swaps two lists of parameters
-definition Pi_swap_list : ∀xl,xl':list X.X → X ≝
- λxl,xl',x.foldr2 ??? (λu,v,r.swap ? u v r) x xl xl'.
-
-nlemma fp_swap_list :
- ∀xl,xl'.finite_perm ? (Pi_swap_list xl xl').
-#xl xl';@
-##[@;
- ##[ngeneralize in match xl';nelim xl
- ##[nnormalize;//;
- ##|#x0 xl0;#IH xl'';nelim xl''
- ##[nnormalize;//
- ##|#x1 xl1 IH1 y0 y1;nchange in ⊢ (??%% → ?) with (swap ????);
- #H1;nlapply (inj_swap … H1);#H2;
- nlapply (IH … H2);//
- ##]
- ##]
- ##|ngeneralize in match xl';nelim xl
- ##[nnormalize;#_;#z;@z;@
- ##|#x' xl0 IH xl'';nelim xl''
- ##[nnormalize;#z;@z;@
- ##|#x1 xl1 IH1 z;
- nchange in ⊢ (??(λ_.???%)) with (swap ????);
- ncases (surj_swap X x' x1 z);#x2 H1;
- ncases (IH xl1 x2);#x3 H2;@ x3;
- nrewrite < H2;napply H1
- ##]
- ##]
- ##]
-##|ngeneralize in match xl';nelim xl
- ##[#;@ [];#;@
- ##|#x0 xl0 IH xl'';ncases xl''
- ##[@ [];#;@
- ##|#x1 xl1;ncases (IH xl1);#xl2 H1;
- ncases (finite_swap X x0 x1);#xl3 H2;
- @ (xl2@xl3);#x2 H3;
- nchange in ⊢ (??%?) with (swap ????);
- nrewrite > (H1 …);
- ##[nrewrite > (H2 …);//;@;#H4;ncases H3;#H5;napply H5;
- napply in_list_to_in_list_append_r;//
- ##|@;#H4;ncases H3;#H5;napply H5;
- napply in_list_to_in_list_append_l;//
- ##]
- ##]
- ##]
-##]
-nqed.
-
-(* the 'reverse' swap of lists of parameters
- composing Pi_swap_list and Pi_swap_list_r yields the identity function
- (see the Pi_swap_list_inv lemma) *)
-ndefinition Pi_swap_list_r : ∀xl,xl':list X. Pi ≝
- λxl,xl',x.foldl2 ??? (λr,u,v.swap ? u v r ) x xl xl'.
-
-nlemma fp_swap_list_r :
- ∀xl,xl'.finite_perm ? (Pi_swap_list_r xl xl').
-#xl xl';@
-##[@;
- ##[ngeneralize in match xl';nelim xl
- ##[nnormalize;//;
- ##|#x0 xl0;#IH xl'';nelim xl''
- ##[nnormalize;//
- ##|#x1 xl1 IH1 y0 y1;nwhd in ⊢ (??%% → ?);
- #H1;nlapply (IH … H1);#H2;
- napply (inj_swap … H2);
- ##]
- ##]
- ##|ngeneralize in match xl';nelim xl
- ##[nnormalize;#_;#z;@z;@
- ##|#x' xl0 IH xl'';nelim xl''
- ##[nnormalize;#z;@z;@
- ##|#x1 xl1 IH1 z;nwhd in ⊢ (??(λ_.???%));
- ncases (IH xl1 z);#x2 H1;
- ncases (surj_swap X x' x1 x2);#x3 H2;
- @ x3;nrewrite < H2;napply H1;
- ##]
- ##]
- ##]
-##|ngeneralize in match xl';nelim xl
- ##[#;@ [];#;@
- ##|#x0 xl0 IH xl'';ncases xl''
- ##[@ [];#;@
- ##|#x1 xl1;
- ncases (IH xl1);#xl2 H1;
- ncases (finite_swap X x0 x1);#xl3 H2;
- @ (xl2@xl3);#x2 H3;nwhd in ⊢ (??%?);
- nrewrite > (H2 …);
- ##[nrewrite > (H1 …);//;@;#H4;ncases H3;#H5;napply H5;
- napply in_list_to_in_list_append_l;//
- ##|@;#H4;ncases H3;#H5;napply H5;
- napply in_list_to_in_list_append_r;//
- ##]
- ##]
- ##]
-##]
-nqed.
-
-nlemma Pi_swap_list_inv :
- ∀xl1,xl2,x.
- Pi_swap_list xl1 xl2 (Pi_swap_list_r xl1 xl2 x) = x.
-#xl;nelim xl
-##[#;@
-##|#x1 xl1 IH xl';ncases xl'
- ##[#;@
- ##|#x2 xl2;#x;
- nchange in ⊢ (??%?) with
- (swap ??? (Pi_swap_list ??
- (Pi_swap_list_r ?? (swap ????))));
- nrewrite > (IH xl2 ?);napply swap_inv;
- ##]
-##]
-nqed.
-
-nlemma Pi_swap_list_fresh :
- ∀x,xl1,xl2.x ∉ xl1 → x ∉ xl2 → Pi_swap_list xl1 xl2 x = x.
-#x xl1;nelim xl1
-##[#;@
-##|#x3 xl3 IH xl2' H1;ncases xl2'
- ##[#;@
- ##|#x4 xl4 H2;ncut (x ∉ xl3 ∧ x ∉ xl4);
- ##[@
- ##[@;#H3;ncases H1;#H4;napply H4;/2/;
- ##|@;#H3;ncases H2;#H4;napply H4;/2/
- ##]
- ##] *; #H1' H2';
- nchange in ⊢ (??%?) with (swap ????);
- nrewrite > (swap_other …)
- ##[napply IH;//;
- ##|nchange in ⊢ (?(???%)) with (Pi_swap_list ???);
- nrewrite > (IH …);//;@;#H3;ncases H2;#H4;napply H4;//;
- ##|nchange in ⊢ (?(???%)) with (Pi_swap_list ???);
- nrewrite > (IH …);//;@;#H3;ncases H1;#H4;napply H4;//
- ##]
- ##]
-##]
-nqed.
-*)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||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 "basics/lists/list.ma".
-include "basics/deqsets.ma".
-include "binding/names.ma".
-include "binding/fp.ma".
-
-axiom alpha : Nset.
-notation "𝔸" non associative with precedence 90 for @{'alphabet}.
-interpretation "set of names" 'alphabet = alpha.
-
-inductive tp : Type[0] ≝
-| top : tp
-| arr : tp → tp → tp.
-inductive pretm : Type[0] ≝
-| var : nat → pretm
-| par : 𝔸 → pretm
-| abs : tp → pretm → pretm
-| app : pretm → pretm → pretm.
-
-let rec Nth T n (l:list T) on n ≝
- match l with
- [ nil ⇒ None ?
- | cons hd tl ⇒ match n with
- [ O ⇒ Some ? hd
- | S n0 ⇒ Nth T n0 tl ] ].
-
-let rec vclose_tm_aux u x k ≝ match u with
- [ var n ⇒ if (leb k n) then var (S n) else u
- | par x0 ⇒ if (x0 == x) then (var k) else u
- | app v1 v2 ⇒ app (vclose_tm_aux v1 x k) (vclose_tm_aux v2 x k)
- | abs s v ⇒ abs s (vclose_tm_aux v x (S k)) ].
-definition vclose_tm ≝ λu,x.vclose_tm_aux u x O.
-
-definition vopen_var ≝ λn,x,k.match eqb n k with
- [ true ⇒ par x
- | false ⇒ match leb n k with
- [ true ⇒ var n
- | false ⇒ var (pred n) ] ].
-
-let rec vopen_tm_aux u x k ≝ match u with
- [ var n ⇒ vopen_var n x k
- | par x0 ⇒ u
- | app v1 v2 ⇒ app (vopen_tm_aux v1 x k) (vopen_tm_aux v2 x k)
- | abs s v ⇒ abs s (vopen_tm_aux v x (S k)) ].
-definition vopen_tm ≝ λu,x.vopen_tm_aux u x O.
-
-let rec FV u ≝ match u with
- [ par x ⇒ [x]
- | app v1 v2 ⇒ FV v1@FV v2
- | abs s v ⇒ FV v
- | _ ⇒ [ ] ].
-
-definition lam ≝ λx,s,u.abs s (vclose_tm u x).
-
-let rec Pi_map_tm p u on u ≝ match u with
-[ par x ⇒ par (p x)
-| var _ ⇒ u
-| app v1 v2 ⇒ app (Pi_map_tm p v1) (Pi_map_tm p v2)
-| abs s v ⇒ abs s (Pi_map_tm p v) ].
-
-interpretation "permutation of tm" 'middot p x = (Pi_map_tm p x).
-
-notation "hvbox(u⌈x⌉)"
- with precedence 45
- for @{ 'open $u $x }.
-
-(*
-notation "hvbox(u⌈x⌉)"
- with precedence 45
- for @{ 'open $u $x }.
-notation "❴ u ❵ x" non associative with precedence 90 for @{ 'open $u $x }.
-*)
-interpretation "ln term variable open" 'open u x = (vopen_tm u x).
-notation < "hvbox(ν x break . u)"
- with precedence 20
-for @{'nu $x $u }.
-notation > "ν list1 x sep , . term 19 u" with precedence 20
- for ${ fold right @{$u} rec acc @{'nu $x $acc)} }.
-interpretation "ln term variable close" 'nu x u = (vclose_tm u x).
-
-let rec tm_height u ≝ match u with
-[ app v1 v2 ⇒ S (max (tm_height v1) (tm_height v2))
-| abs s v ⇒ S (tm_height v)
-| _ ⇒ O ].
-
-theorem le_n_O_rect_Type0 : ∀n:nat. n ≤ O → ∀P: nat →Type[0]. P O → P n.
-#n (cases n) // #a #abs cases (?:False) /2/ qed.
-
-theorem nat_rect_Type0_1 : ∀n:nat.∀P:nat → Type[0].
-(∀m.(∀p. p < m → P p) → P m) → P n.
-#n #P #H
-cut (∀q:nat. q ≤ n → P q) /2/
-(elim n)
- [#q #HleO (* applica male *)
- @(le_n_O_rect_Type0 ? HleO)
- @H #p #ltpO cases (?:False) /2/ (* 3 *)
- |#p #Hind #q #HleS
- @H #a #lta @Hind @le_S_S_to_le /2/
- ]
-qed.
-
-lemma leb_false_to_lt : ∀n,m. leb n m = false → m < n.
-#n elim n
-[ #m normalize #H destruct(H)
-| #n0 #IH * // #m normalize #H @le_S_S @IH // ]
-qed.
-
-lemma nominal_eta_aux : ∀x,u.x ∉ FV u → ∀k.vclose_tm_aux (vopen_tm_aux u x k) x k = u.
-#x #u elim u
-[ #n #_ #k normalize cases (decidable_eq_nat n k) #Hnk
- [ >Hnk >eqb_n_n normalize >(\b ?) //
- | >(not_eq_to_eqb_false … Hnk) normalize cases (true_or_false (leb n k)) #Hleb
- [ >Hleb normalize >(?:leb k n = false) //
- @lt_to_leb_false @not_eq_to_le_to_lt /2/
- | >Hleb normalize >(?:leb k (pred n) = true) normalize
- [ cases (leb_false_to_lt … Hleb) //
- | @le_to_leb_true cases (leb_false_to_lt … Hleb) normalize /2/ ] ] ]
-| #y normalize #Hy >(\bf ?) // @(not_to_not … Hy) //
-| #s #v #IH normalize #Hv #k >IH // @Hv
-| #v1 #v2 #IH1 #IH2 normalize #Hv1v2 #k
- >IH1 [ >IH2 // | @(not_to_not … Hv1v2) @in_list_to_in_list_append_l ]
- @(not_to_not … Hv1v2) @in_list_to_in_list_append_r ]
-qed.
-
-corollary nominal_eta : ∀x,u.x ∉ FV u → (νx.u⌈x⌉) = u.
-#x #u #Hu @nominal_eta_aux //
-qed.
-
-lemma eq_height_vopen_aux : ∀v,x,k.tm_height (vopen_tm_aux v x k) = tm_height v.
-#v #x elim v
-[ #n #k normalize cases (eqb n k) // cases (leb n k) //
-| #u #k %
-| #s #u #IH #k normalize >IH %
-| #u1 #u2 #IH1 #IH2 #k normalize >IH1 >IH2 % ]
-qed.
-
-corollary eq_height_vopen : ∀v,x.tm_height (v⌈x⌉) = tm_height v.
-#v #x @eq_height_vopen_aux
-qed.
-
-theorem pretm_ind_plus_aux :
- ∀P:pretm → Type[0].
- (∀x:𝔸.P (par x)) →
- (∀n:ℕ.P (var n)) →
- (∀v1,v2. P v1 → P v2 → P (app v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) →
- ∀n,u.tm_height u ≤ n → P u.
-#P #Hpar #Hvar #Happ #C #Hlam #n change with ((λn.?) n); @(nat_rect_Type0_1 n ??)
-#m cases m
-[ #_ * /2/
- [ normalize #s #v #Hfalse cases (?:False) cases (not_le_Sn_O (tm_height v)) /2/
- | #v1 #v2 whd in ⊢ (?%?→?); #Hfalse cases (?:False) cases (not_le_Sn_O (max ??))
- [ #H @H @Hfalse|*:skip] ] ]
--m #m #IH * /2/
-[ #s #v whd in ⊢ (?%?→?); #Hv
- lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy
- >(?:abs s v = lam y s (v⌈y⌉))
- [| whd in ⊢ (???%); >nominal_eta // @(not_to_not … Hy) @in_list_to_in_list_append_r ]
- @Hlam
- [ @(not_to_not … Hy) @in_list_to_in_list_append_r
- | @(not_to_not … Hy) @in_list_to_in_list_append_l ]
- @IH [| @Hv | >eq_height_vopen % ]
-| #v1 #v2 whd in ⊢ (?%?→?); #Hv @Happ
- [ @IH [| @Hv | @le_max_1 ] | @IH [| @Hv | @le_max_2 ] ] ]
-qed.
-
-corollary pretm_ind_plus :
- ∀P:pretm → Type[0].
- (∀x:𝔸.P (par x)) →
- (∀n:ℕ.P (var n)) →
- (∀v1,v2. P v1 → P v2 → P (app v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) →
- ∀u.P u.
-#P #Hpar #Hvar #Happ #C #Hlam #u @pretm_ind_plus_aux /2/
-qed.
-
-(* maps a permutation to a list of terms *)
-definition Pi_map_list : (𝔸 → 𝔸) → list 𝔸 → list 𝔸 ≝ map 𝔸 𝔸 .
-
-(* interpretation "permutation of name list" 'middot p x = (Pi_map_list p x).*)
-
-(*
-inductive tm : pretm → Prop ≝
-| tm_par : ∀x:𝔸.tm (par x)
-| tm_app : ∀u,v.tm u → tm v → tm (app u v)
-| tm_lam : ∀x,s,u.tm u → tm (lam x s u).
-
-inductive ctx_aux : nat → pretm → Prop ≝
-| ctx_var : ∀n,k.n < k → ctx_aux k (var n)
-| ctx_par : ∀x,k.ctx_aux k (par x)
-| ctx_app : ∀u,v,k.ctx_aux k u → ctx_aux k v → ctx_aux k (app u v)
-(* è sostituibile da ctx_lam ? *)
-| ctx_abs : ∀s,u.ctx_aux (S k) u → ctx_aux k (abs s u).
-*)
-
-inductive tm_or_ctx (k:nat) : pretm → Type[0] ≝
-| toc_var : ∀n.n < k → tm_or_ctx k (var n)
-| toc_par : ∀x.tm_or_ctx k (par x)
-| toc_app : ∀u,v.tm_or_ctx k u → tm_or_ctx k v → tm_or_ctx k (app u v)
-| toc_lam : ∀x,s,u.tm_or_ctx k u → tm_or_ctx k (lam x s u).
-
-definition tm ≝ λt.tm_or_ctx O t.
-definition ctx ≝ λt.tm_or_ctx 1 t.
-
-definition check_tm ≝ λu,k.
- pretm_ind_plus ?
- (λ_.true)
- (λn.leb (S n) k)
- (λv1,v2,rv1,rv2.rv1 ∧ rv2)
- [] (λx,s,v,px,pC,rv.rv)
- u.
-
-axiom pretm_ind_plus_app : ∀P,u,v,C,H1,H2,H3,H4.
- pretm_ind_plus P H1 H2 H3 C H4 (app u v) =
- H3 u v (pretm_ind_plus P H1 H2 H3 C H4 u) (pretm_ind_plus P H1 H2 H3 C H4 v).
-
-axiom pretm_ind_plus_lam : ∀P,x,s,u,C,px,pC,H1,H2,H3,H4.
- pretm_ind_plus P H1 H2 H3 C H4 (lam x s (u⌈x⌉)) =
- H4 x s u px pC (pretm_ind_plus P H1 H2 H3 C H4 (u⌈x⌉)).
-
-record TM : Type[0] ≝ {
- pretm_of_TM :> pretm;
- tm_of_TM : check_tm pretm_of_TM O = true
-}.
-
-record CTX : Type[0] ≝ {
- pretm_of_CTX :> pretm;
- ctx_of_CTX : check_tm pretm_of_CTX 1 = true
-}.
-
-inductive tm2 : pretm → Type[0] ≝
-| tm_par : ∀x.tm2 (par x)
-| tm_app : ∀u,v.tm2 u → tm2 v → tm2 (app u v)
-| tm_lam : ∀x,s,u.x ∉ FV u → (∀y.y ∉ FV u → tm2 (u⌈y⌉)) → tm2 (lam x s (u⌈x⌉)).
-
-(*
-inductive tm' : pretm → Prop ≝
-| tm_par : ∀x.tm' (par x)
-| tm_app : ∀u,v.tm' u → tm' v → tm' (app u v)
-| tm_lam : ∀x,s,u,C.x ∉ FV u → x ∉ C → (∀y.y ∉ FV u → tm' (❴u❵y)) → tm' (lam x s (❴u❵x)).
-*)
-
-lemma pi_vclose_tm :
- ∀z1,z2,x,u.swap 𝔸 z1 z2·(νx.u) = (ν swap ? z1 z2 x.swap 𝔸 z1 z2 · u).
-#z1 #z2 #x #u
-change with (vclose_tm_aux ???) in match (vclose_tm ??);
-change with (vclose_tm_aux ???) in ⊢ (???%); lapply O elim u normalize //
-[ #n #k cases (leb k n) normalize %
-| #x0 #k cases (true_or_false (x0==z1)) #H1 >H1 normalize
- [ cases (true_or_false (x0==x)) #H2 >H2 normalize
- [ <(\P H2) >H1 normalize >(\b (refl ? z2)) %
- | >H1 normalize cases (true_or_false (x==z1)) #H3 >H3 normalize
- [ >(\P H3) in H2; >H1 #Hfalse destruct (Hfalse)
- | cases (true_or_false (x==z2)) #H4 >H4 normalize
- [ cases (true_or_false (z2==z1)) #H5 >H5 normalize //
- >(\P H5) in H4; >H3 #Hfalse destruct (Hfalse)
- | >(\bf ?) // @sym_not_eq @(\Pf H4) ]
- ]
- ]
- | cases (true_or_false (x0==x)) #H2 >H2 normalize
- [ <(\P H2) >H1 normalize >(\b (refl ??)) %
- | >H1 normalize cases (true_or_false (x==z1)) #H3 >H3 normalize
- [ cases (true_or_false (x0==z2)) #H4 >H4 normalize
- [ cases (true_or_false (z1==z2)) #H5 >H5 normalize //
- <(\P H5) in H4; <(\P H3) >H2 #Hfalse destruct (Hfalse)
- | >H4 % ]
- | cases (true_or_false (x0==z2)) #H4 >H4 normalize
- [ cases (true_or_false (x==z2)) #H5 >H5 normalize
- [ <(\P H5) in H4; >H2 #Hfalse destruct (Hfalse)
- | >(\bf ?) // @sym_not_eq @(\Pf H3) ]
- | cases (true_or_false (x==z2)) #H5 >H5 normalize
- [ >H1 %
- | >H2 % ]
- ]
- ]
- ]
- ]
-]
-qed.
-
-lemma pi_vopen_tm :
- ∀z1,z2,x,u.swap 𝔸 z1 z2·(u⌈x⌉) = (swap 𝔸 z1 z2 · u⌈swap 𝔸 z1 z2 x⌉).
-#z1 #z2 #x #u
-change with (vopen_tm_aux ???) in match (vopen_tm ??);
-change with (vopen_tm_aux ???) in ⊢ (???%); lapply O elim u normalize //
-#n #k cases (true_or_false (eqb n k)) #H1 >H1 normalize //
-cases (true_or_false (leb n k)) #H2 >H2 normalize //
-qed.
-
-lemma pi_lam :
- ∀z1,z2,x,s,u.swap 𝔸 z1 z2 · lam x s u = lam (swap 𝔸 z1 z2 x) s (swap 𝔸 z1 z2 · u).
-#z1 #z2 #x #s #u whd in ⊢ (???%); <(pi_vclose_tm …) %
-qed.
-
-lemma eqv_FV : ∀z1,z2,u.FV (swap 𝔸 z1 z2 · u) = Pi_map_list (swap 𝔸 z1 z2) (FV u).
-#z1 #z2 #u elim u //
-[ #s #v normalize //
-| #v1 #v2 normalize /2/ ]
-qed.
-
-lemma swap_inv_tm : ∀z1,z2,u.swap 𝔸 z1 z2 · (swap 𝔸 z1 z2 · u) = u.
-#z1 #z2 #u elim u [1,3,4:normalize //]
-#x whd in ⊢ (??%?); >swap_inv %
-qed.
-
-lemma eqv_in_list : ∀x,l,z1,z2.x ∈ l → swap 𝔸 z1 z2 x ∈ Pi_map_list (swap 𝔸 z1 z2) l.
-#x #l #z1 #z2 #Hin elim Hin
-[ #x0 #l0 %
-| #x1 #x2 #l0 #Hin #IH %2 @IH ]
-qed.
-
-lemma eqv_tm2 : ∀u.tm2 u → ∀z1,z2.tm2 ((swap ? z1 z2)·u).
-#u #Hu #z1 #z2 letin p ≝ (swap ? z1 z2) elim Hu /2/
-#x #s #v #Hx #Hv #IH >pi_lam >pi_vopen_tm %3
-[ @(not_to_not … Hx) -Hx #Hx
- <(swap_inv ? z1 z2 x) <(swap_inv_tm z1 z2 v) >eqv_FV @eqv_in_list //
-| #y #Hy <(swap_inv ? z1 z2 y)
- <pi_vopen_tm @IH @(not_to_not … Hy) -Hy #Hy <(swap_inv ? z1 z2 y)
- >eqv_FV @eqv_in_list //
-]
-qed.
-
-lemma vclose_vopen_aux : ∀x,u,k.vopen_tm_aux (vclose_tm_aux u x k) x k = u.
-#x #u elim u normalize //
-[ #n #k cases (true_or_false (leb k n)) #H >H whd in ⊢ (??%?);
- [ cases (true_or_false (eqb (S n) k)) #H1 >H1
- [ <(eqb_true_to_eq … H1) in H; #H lapply (leb_true_to_le … H) -H #H
- cases (le_to_not_lt … H) -H #H cases (H ?) %
- | whd in ⊢ (??%?); >lt_to_leb_false // @le_S_S /2/ ]
- | cases (true_or_false (eqb n k)) #H1 >H1 normalize
- [ >(eqb_true_to_eq … H1) in H; #H lapply (leb_false_to_not_le … H) -H
- * #H cases (H ?) %
- | >le_to_leb_true // @not_lt_to_le % #H2 >le_to_leb_true in H;
- [ #H destruct (H) | /2/ ]
- ]
- ]
-| #x0 #k cases (true_or_false (x0==x)) #H1 >H1 normalize // >(\P H1) >eqb_n_n % ]
-qed.
-
-lemma vclose_vopen : ∀x,u.((νx.u)⌈x⌉) = u. #x #u @vclose_vopen_aux
-qed.
-
-(*
-theorem tm_to_tm : ∀t.tm' t → tm t.
-#t #H elim H
-*)
-
-lemma in_list_singleton : ∀T.∀t1,t2:T.t1 ∈ [t2] → t1 = t2.
-#T #t1 #t2 #H @(in_list_inv_ind ??? H) /2/
-qed.
-
-lemma fresh_vclose_tm_aux : ∀u,x,k.x ∉ FV (vclose_tm_aux u x k).
-#u #x elim u //
-[ #n #k normalize cases (leb k n) normalize //
-| #x0 #k normalize cases (true_or_false (x0==x)) #H >H normalize //
- lapply (\Pf H) @not_to_not #Hin >(in_list_singleton ??? Hin) %
-| #v1 #v2 #IH1 #IH2 #k normalize % #Hin cases (in_list_append_to_or_in_list ???? Hin) /2/ ]
-qed.
-
-lemma fresh_vclose_tm : ∀u,x.x ∉ FV (νx.u). //
-qed.
-
-lemma check_tm_true_to_toc : ∀u,k.check_tm u k = true → tm_or_ctx k u.
-#u @(pretm_ind_plus ???? [ ] ? u)
-[ #x #k #_ %2
-| #n #k change with (leb (S n) k) in ⊢ (??%?→?); #H % @leb_true_to_le //
-| #v1 #v2 #rv1 #rv2 #k change with (pretm_ind_plus ???????) in ⊢ (??%?→?);
- >pretm_ind_plus_app #H cases (andb_true ?? H) -H #Hv1 #Hv2 %3
- [ @rv1 @Hv1 | @rv2 @Hv2 ]
-| #x #s #v #Hx #_ #rv #k change with (pretm_ind_plus ???????) in ⊢ (??%?→?);
- >pretm_ind_plus_lam // #Hv %4 @rv @Hv ]
-qed.
-
-lemma toc_to_check_tm_true : ∀u,k.tm_or_ctx k u → check_tm u k = true.
-#u #k #Hu elim Hu //
-[ #n #Hn change with (leb (S n) k) in ⊢ (??%?); @le_to_leb_true @Hn
-| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 change with (pretm_ind_plus ???????) in ⊢ (??%?);
- >pretm_ind_plus_app change with (check_tm v1 k ∧ check_tm v2 k) in ⊢ (??%?); /2/
-| #x #s #v #Hv #IH <(vclose_vopen x v) change with (pretm_ind_plus ???????) in ⊢ (??%?);
- >pretm_ind_plus_lam [| // | @fresh_vclose_tm ] >(vclose_vopen x v) @IH ]
-qed.
-
-lemma fresh_swap_tm : ∀z1,z2,u.z1 ∉ FV u → z2 ∉ FV u → swap 𝔸 z1 z2 · u = u.
-#z1 #z2 #u elim u
-[2: normalize in ⊢ (?→%→%→?); #x #Hz1 #Hz2 whd in ⊢ (??%?); >swap_other //
- [ @(not_to_not … Hz2) | @(not_to_not … Hz1) ] //
-|1: //
-| #s #v #IH normalize #Hz1 #Hz2 >IH // [@Hz2|@Hz1]
-| #v1 #v2 #IH1 #IH2 normalize #Hz1 #Hz2
- >IH1 [| @(not_to_not … Hz2) @in_list_to_in_list_append_l | @(not_to_not … Hz1) @in_list_to_in_list_append_l ]
- >IH2 // [@(not_to_not … Hz2) @in_list_to_in_list_append_r | @(not_to_not … Hz1) @in_list_to_in_list_append_r ]
-]
-qed.
-
-theorem tm_to_tm2 : ∀u.tm u → tm2 u.
-#t #Ht elim Ht
-[ #n #Hn cases (not_le_Sn_O n) #Hfalse cases (Hfalse Hn)
-| @tm_par
-| #u #v #Hu #Hv @tm_app
-| #x #s #u #Hu #IHu <(vclose_vopen x u) @tm_lam
- [ @fresh_vclose_tm
- | #y #Hy <(fresh_swap_tm x y (νx.u)) /2/ @fresh_vclose_tm ]
-]
-qed.
-
-theorem tm2_to_tm : ∀u.tm2 u → tm u.
-#u #pu elim pu /2/ #x #s #v #Hx #Hv #IH %4 @IH //
-qed.
-
-(* define PAR APP LAM *)
-definition PAR ≝ λx.mk_TM (par x) ?. // qed.
-definition APP ≝ λu,v:TM.mk_TM (app u v) ?.
-change with (pretm_ind_plus ???????) in match (check_tm ??); >pretm_ind_plus_app
-change with (check_tm ??) in match (pretm_ind_plus ???????); change with (check_tm ??) in match (pretm_ind_plus ???????) in ⊢ (??(??%)?);
-@andb_elim >(tm_of_TM u) >(tm_of_TM v) % qed.
-definition LAM ≝ λx,s.λu:TM.mk_TM (lam x s u) ?.
-change with (pretm_ind_plus ???????) in match (check_tm ??); <(vclose_vopen x u)
->pretm_ind_plus_lam [| // | @fresh_vclose_tm ]
-change with (check_tm ??) in match (pretm_ind_plus ???????); >vclose_vopen
-@(tm_of_TM u) qed.
-
-axiom vopen_tm_down : ∀u,x,k.tm_or_ctx (S k) u → tm_or_ctx k (u⌈x⌉).
-(* needs true_plus_false
-
-#u #x #k #Hu elim Hu
-[ #n #Hn normalize cases (true_or_false (eqb n O)) #H >H [%2]
- normalize >(?: leb n O = false) [|cases n in H; // >eqb_n_n #H destruct (H) ]
- normalize lapply Hn cases n in H; normalize [ #Hfalse destruct (Hfalse) ]
- #n0 #_ #Hn0 % @le_S_S_to_le //
-| #x0 %2
-| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 %3 //
-| #x0 #s #v #Hv #IH normalize @daemon
-]
-qed.
-*)
-
-definition vopen_TM ≝ λu:CTX.λx.mk_TM (u⌈x⌉) ?.
-@toc_to_check_tm_true @vopen_tm_down @check_tm_true_to_toc @ctx_of_CTX qed.
-
-axiom vclose_tm_up : ∀u,x,k.tm_or_ctx k u → tm_or_ctx (S k) (νx.u).
-
-definition vclose_TM ≝ λu:TM.λx.mk_CTX (νx.u) ?.
-@toc_to_check_tm_true @vclose_tm_up @check_tm_true_to_toc @tm_of_TM qed.
-
-interpretation "ln wf term variable open" 'open u x = (vopen_TM u x).
-interpretation "ln wf term variable close" 'nu x u = (vclose_TM u x).
-
-theorem tm_alpha : ∀x,y,s,u.x ∉ FV u → y ∉ FV u → lam x s (u⌈x⌉) = lam y s (u⌈y⌉).
-#x #y #s #u #Hx #Hy whd in ⊢ (??%%); @eq_f >nominal_eta // >nominal_eta //
-qed.
-
-lemma TM_to_tm2 : ∀u:TM.tm2 u.
-#u @tm_to_tm2 @check_tm_true_to_toc @tm_of_TM qed.
-
-theorem TM_ind_plus_weak :
- ∀P:pretm → Type[0].
- (∀x:𝔸.P (PAR x)) →
- (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C →
- (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) →
- ∀u:TM.P u.
-#P #Hpar #Happ #C #Hlam #u elim (TM_to_tm2 u) //
-[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …) IH1 IH2)
- @toc_to_check_tm_true @tm2_to_tm //
-| #x #s #v #Hx #pv #IH
- lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0
- >(?:lam x s (v⌈x⌉) = lam x0 s (v⌈x0⌉))
- [|@tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ]
- @(Hlam x0 s (mk_CTX v ?) ??)
- [ <(nominal_eta … Hx) @toc_to_check_tm_true @vclose_tm_up @tm2_to_tm @pv //
- | @(not_to_not … Hx0) @in_list_to_in_list_append_r
- | @(not_to_not … Hx0) @in_list_to_in_list_append_l
- | @IH ]
-]
-qed.
-
-lemma eq_mk_TM : ∀u,v.u = v → ∀pu,pv.mk_TM u pu = mk_TM v pv.
-#u #v #Heq >Heq #pu #pv %
-qed.
-
-lemma eq_P : ∀T:Type[0].∀t1,t2:T.t1 = t2 → ∀P:T → Type[0].P t1 → P t2. // qed.
-
-theorem TM_ind_plus :
- ∀P:TM → Type[0].
- (∀x:𝔸.P (PAR x)) →
- (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C →
- (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) →
- ∀u:TM.P u.
-#P #Hpar #Happ #C #Hlam * #u #pu
->(?:mk_TM u pu =
- mk_TM u (toc_to_check_tm_true … (tm2_to_tm … (tm_to_tm2 … (check_tm_true_to_toc … pu))))) [|%]
-elim (tm_to_tm2 u ?) //
-[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …) IH1 IH2)
-| #x #s #v #Hx #pv #IH
- lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0
- lapply (Hlam x0 s (mk_CTX v ?) ???)
- [2: @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0
- |4: @toc_to_check_tm_true <(nominal_eta x v) // @vclose_tm_up @tm2_to_tm @pv //
- | #y #Hy whd in match (vopen_TM ??);
- >(?:mk_TM (v⌈y⌉) ? = mk_TM (v⌈y⌉) (toc_to_check_tm_true (v⌈y⌉) O (tm2_to_tm (v⌈y⌉) (pv y Hy))))
- [@IH|%]
- | @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_r @Hx0
- | @eq_P @eq_mk_TM whd in match (vopen_TM ??); @tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ]
-]
-qed.
-
-notation
-"hvbox('nominal' u 'return' out 'with'
- [ 'xpar' ident x ⇒ f1
- | 'xapp' ident v1 ident v2 ident recv1 ident recv2 ⇒ f2
- | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident recw ⇒ f3 ])"
-with precedence 48
-for @{ TM_ind_plus $out (λ${ident x}:?.$f1)
- (λ${ident v1}:?.λ${ident v2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2)
- $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident recw}:?.$f3)
- $u }.
-
-(* include "basics/jmeq.ma".*)
-
-definition subst ≝ (λu:TM.λx,v.
- nominal u return (λ_.TM) with
- [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ PAR x0 ] (* u instead of PAR x0 does not work, u stays the same at every rec call! *)
- | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2
- | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]).
-
-lemma subst_def : ∀u,x,v.subst u x v =
- nominal u return (λ_.TM) with
- [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ PAR x0 ]
- | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2
- | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]. //
-qed.
-
-axiom TM_ind_plus_LAM :
- ∀x,s,u,out,f1,f2,C,f3,Hx1,Hx2.
- TM_ind_plus out f1 f2 C f3 (LAM x s (u⌈x⌉)) =
- f3 x s u Hx1 Hx2 (λy,Hy.TM_ind_plus ? f1 f2 C f3 ?).
-
-axiom TM_ind_plus_APP :
- ∀u1,u2,out,f1,f2,C,f3.
- TM_ind_plus out f1 f2 C f3 (APP u1 u2) =
- f2 u1 u2 (TM_ind_plus out f1 f2 C f3 ?) (TM_ind_plus out f1 f2 C f3 ?).
-
-lemma eq_mk_CTX : ∀u,v.u = v → ∀pu,pv.mk_CTX u pu = mk_CTX v pv.
-#u #v #Heq >Heq #pu #pv %
-qed.
-
-lemma vclose_vopen_TM : ∀x.∀u:TM.((νx.u)⌈x⌉) = u.
-#x * #u #pu @eq_mk_TM @vclose_vopen qed.
-
-lemma nominal_eta_CTX : ∀x.∀u:CTX.x ∉ FV u → (νx.(u⌈x⌉)) = u.
-#x * #u #pu #Hx @eq_mk_CTX @nominal_eta // qed.
-
-theorem TM_alpha : ∀x,y,s.∀u:CTX.x ∉ FV u → y ∉ FV u → LAM x s (u⌈x⌉) = LAM y s (u⌈y⌉).
-#x #y #s #u #Hx #Hy @eq_mk_TM @tm_alpha // qed.
-
-axiom in_vopen_CTX : ∀x,y.∀v:CTX.x ∈ FV (v⌈y⌉) → x = y ∨ x ∈ FV v.
-
-theorem subst_fresh : ∀u,v:TM.∀x.x ∉ FV u → subst u x v = u.
-#u #v #x @(TM_ind_plus … (x::FV v) … u)
-[ #x0 normalize in ⊢ (%→?); #Hx normalize in ⊢ (??%?);
- >(\bf ?) [| @(not_to_not … Hx) #Heq >Heq % ] %
-| #u1 #u2 #IH1 #IH2 normalize in ⊢ (%→?); #Hx
- >subst_def >TM_ind_plus_APP @eq_mk_TM @eq_f2 @eq_f
- [ <subst_def @IH1 @(not_to_not … Hx) @in_list_to_in_list_append_l
- | <subst_def @IH2 @(not_to_not … Hx) @in_list_to_in_list_append_r ]
-| #x0 #s #v0 #Hx0 #HC #IH #Hx >subst_def >TM_ind_plus_LAM [|@HC|@Hx0]
- @eq_f <subst_def @IH // @(not_to_not … Hx) -Hx #Hx
- change with (FV (νx0.(v0⌈x0⌉))) in ⊢ (???%); >nominal_eta_CTX //
- cases (in_vopen_CTX … Hx) // #Heq >Heq in HC; * #HC @False_ind @HC %
-]
-qed.
-
-example subst_LAM_same : ∀x,s,u,v. subst (LAM x s u) x v = LAM x s u.
-#x #s #u #v >subst_def <(vclose_vopen_TM x u)
-lapply (p_fresh … (FV (νx.u)@x::FV v)) letin x0 ≝ (N_fresh … (FV (νx.u)@x::FV v)) #Hx0
->(TM_alpha x x0)
-[| @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 | @fresh_vclose_tm ]
->TM_ind_plus_LAM [| @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_r @Hx0 | @(not_to_not … Hx0) -Hx0 #Hx0 @in_list_to_in_list_append_l @Hx0 ]
-@eq_f change with (subst ((νx.u)⌈x0⌉) x v) in ⊢ (??%?); @subst_fresh
-@(not_to_not … Hx0) #Hx0' cases (in_vopen_CTX … Hx0')
-[ #Heq >Heq @in_list_to_in_list_append_r %
-| #Hfalse @False_ind cases (fresh_vclose_tm u x) #H @H @Hfalse ]
-qed.
-
-(*
-notation > "Λ ident x. ident T [ident x] ↦ P"
- with precedence 48 for @{'foo (λ${ident x}.λ${ident T}.$P)}.
-
-notation < "Λ ident x. ident T [ident x] ↦ P"
- with precedence 48 for @{'foo (λ${ident x}:$Q.λ${ident T}:$R.$P)}.
-*)
-
-(*
-notation
-"hvbox('nominal' u 'with'
- [ 'xpar' ident x ⇒ f1
- | 'xapp' ident v1 ident v2 ⇒ f2
- | 'xlam' ident x # C s w ⇒ f3 ])"
-with precedence 48
-for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1)
- (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2)
- $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }.
-*)
-
-(*
-notation
-"hvbox('nominal' u 'with'
- [ 'xpar' ident x ^ f1
- | 'xapp' ident v1 ident v2 ^ f2 ])"
-(* | 'xlam' ident x # C s w ^ f3 ]) *)
-with precedence 48
-for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1)
- (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2)
- $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }.
-*)
-notation
-"hvbox('nominal' u 'with'
- [ 'xpar' ident x ^ f1
- | 'xapp' ident v1 ident v2 ^ f2 ])"
-with precedence 48
-for @{ tm2_ind_plus ? (λ${ident x}:?.$f1)
- (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2)
- $C (λ${ident x}:?.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }.
-
-axiom in_Env : 𝔸 × tp → Env → Prop.
-notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}.
-interpretation "Env membership" 'lefttriangle x l = (in_Env x l).
-
-
-
-inductive judg : list tp → tm → tp → Prop ≝
-| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t
-| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u
-| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u).
-
-definition Env := list (𝔸 × tp).
-
-axiom vclose_env : Env → list tp.
-axiom vclose_tm : Env → tm → tm.
-axiom Lam : 𝔸 → tp → tm → tm.
-definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T.
-definition dom ≝ λG:Env.map ?? (fst ??) G.
-
-definition sctx ≝ 𝔸 × tm.
-axiom swap_tm : 𝔸 → 𝔸 → tm → tm.
-definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M.
-
-axiom in_list : ∀A:Type[0].A → list A → Prop.
-interpretation "list membership" 'mem x l = (in_list ? x l).
-interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)).
-
-axiom in_Env : 𝔸 × tp → Env → Prop.
-notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}.
-interpretation "Env membership" 'lefttriangle x l = (in_Env x l).
-
-(* axiom Lookup : 𝔸 → Env → option tp. *)
-
-(* forma alto livello del judgment
- t_abs* : ∀G,T,X,M,U.
- (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) →
- Judg G (Lam X T (M[X])) (arr T U) *)
-
-(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *)
-
-axiom Judg_ind : ∀P:Env → tm → tp → Prop.
- (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) →
- (∀G,M,N,T,U.
- Judg G M (arr T U) → Judg G N T →
- P G M (arr T U) → P G N T → P G (app M N) U) →
- (∀G,T1,T2,X,M1.
- (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) →
- (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) →
- P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) →
- ∀G,M,T.Judg G M T → P G M T.
-
-axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T.
-axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U.
-axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U).
-
-definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2.
-interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2).
-
-axiom daemon : ∀P:Prop.P.
-
-theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T.
-#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?)
-@(Judg_ind … HJ)
-[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub //
-| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2
- [| @IH1 // | @IH2 // ]
-| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH
- [ (* trivial property of Lam *) @daemon
- | (* trivial property of subenv *) @daemon ]
-]
-qed.
-
-(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e
- ricorsione *)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||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 "basics/lists/list.ma".
-include "basics/deqsets.ma".
-include "binding/names.ma".
-include "binding/fp.ma".
-
-definition alpha : Nset ≝ X. check alpha
-notation "𝔸" non associative with precedence 90 for @{'alphabet}.
-interpretation "set of names" 'alphabet = alpha.
-
-inductive tp : Type[0] ≝
-| top : tp
-| arr : tp → tp → tp.
-inductive pretm : Type[0] ≝
-| var : nat → pretm
-| par : 𝔸 → pretm
-| abs : tp → pretm → pretm
-| app : pretm → pretm → pretm.
-
-let rec Nth T n (l:list T) on n ≝
- match l with
- [ nil ⇒ None ?
- | cons hd tl ⇒ match n with
- [ O ⇒ Some ? hd
- | S n0 ⇒ Nth T n0 tl ] ].
-
-let rec vclose_tm_aux u x k ≝ match u with
- [ var n ⇒ if (leb k n) then var (S n) else u
- | par x0 ⇒ if (x0 == x) then (var k) else u
- | app v1 v2 ⇒ app (vclose_tm_aux v1 x k) (vclose_tm_aux v2 x k)
- | abs s v ⇒ abs s (vclose_tm_aux v x (S k)) ].
-definition vclose_tm ≝ λu,x.vclose_tm_aux u x O.
-
-definition vopen_var ≝ λn,x,k.match eqb n k with
- [ true ⇒ par x
- | false ⇒ match leb n k with
- [ true ⇒ var n
- | false ⇒ var (pred n) ] ].
-
-let rec vopen_tm_aux u x k ≝ match u with
- [ var n ⇒ vopen_var n x k
- | par x0 ⇒ u
- | app v1 v2 ⇒ app (vopen_tm_aux v1 x k) (vopen_tm_aux v2 x k)
- | abs s v ⇒ abs s (vopen_tm_aux v x (S k)) ].
-definition vopen_tm ≝ λu,x.vopen_tm_aux u x O.
-
-let rec FV u ≝ match u with
- [ par x ⇒ [x]
- | app v1 v2 ⇒ FV v1@FV v2
- | abs s v ⇒ FV v
- | _ ⇒ [ ] ].
-
-definition lam ≝ λx,s,u.abs s (vclose_tm u x).
-
-let rec Pi_map_tm p u on u ≝ match u with
-[ par x ⇒ par (p x)
-| var _ ⇒ u
-| app v1 v2 ⇒ app (Pi_map_tm p v1) (Pi_map_tm p v2)
-| abs s v ⇒ abs s (Pi_map_tm p v) ].
-
-interpretation "permutation of tm" 'middot p x = (Pi_map_tm p x).
-
-notation "hvbox(u⌈x⌉)"
- with precedence 45
- for @{ 'open $u $x }.
-
-(*
-notation "hvbox(u⌈x⌉)"
- with precedence 45
- for @{ 'open $u $x }.
-notation "❴ u ❵ x" non associative with precedence 90 for @{ 'open $u $x }.
-*)
-interpretation "ln term variable open" 'open u x = (vopen_tm u x).
-notation < "hvbox(ν x break . u)"
- with precedence 20
-for @{'nu $x $u }.
-notation > "ν list1 x sep , . term 19 u" with precedence 20
- for ${ fold right @{$u} rec acc @{'nu $x $acc)} }.
-interpretation "ln term variable close" 'nu x u = (vclose_tm u x).
-
-let rec tm_height u ≝ match u with
-[ app v1 v2 ⇒ S (max (tm_height v1) (tm_height v2))
-| abs s v ⇒ S (tm_height v)
-| _ ⇒ O ].
-
-theorem le_n_O_rect_Type0 : ∀n:nat. n ≤ O → ∀P: nat →Type[0]. P O → P n.
-#n (cases n) // #a #abs cases (?:False) /2/ qed.
-
-theorem nat_rect_Type0_1 : ∀n:nat.∀P:nat → Type[0].
-(∀m.(∀p. p < m → P p) → P m) → P n.
-#n #P #H
-cut (∀q:nat. q ≤ n → P q) /2/
-(elim n)
- [#q #HleO (* applica male *)
- @(le_n_O_rect_Type0 ? HleO)
- @H #p #ltpO cases (?:False) /2/ (* 3 *)
- |#p #Hind #q #HleS
- @H #a #lta @Hind @le_S_S_to_le /2/
- ]
-qed.
-
-lemma leb_false_to_lt : ∀n,m. leb n m = false → m < n.
-#n elim n
-[ #m normalize #H destruct(H)
-| #n0 #IH * // #m normalize #H @le_S_S @IH // ]
-qed.
-
-lemma nominal_eta_aux : ∀x,u.x ∉ FV u → ∀k.vclose_tm_aux (vopen_tm_aux u x k) x k = u.
-#x #u elim u
-[ #n #_ #k normalize cases (decidable_eq_nat n k) #Hnk
- [ >Hnk >eqb_n_n whd in ⊢ (??%?); >(\b ?) //
- | >(not_eq_to_eqb_false … Hnk) normalize cases (true_or_false (leb n k)) #Hleb
- [ >Hleb normalize >(?:leb k n = false) //
- @lt_to_leb_false @not_eq_to_le_to_lt /2/
- | >Hleb normalize >(?:leb k (pred n) = true) normalize
- [ cases (leb_false_to_lt … Hleb) //
- | @le_to_leb_true cases (leb_false_to_lt … Hleb) normalize /2/ ] ] ]
-| #y normalize in ⊢ (%→?→?); #Hy whd in ⊢ (?→??%?); >(\bf ?) // @(not_to_not … Hy) //
-| #s #v #IH normalize #Hv #k >IH // @Hv
-| #v1 #v2 #IH1 #IH2 normalize #Hv1v2 #k
- >IH1 [ >IH2 // | @(not_to_not … Hv1v2) @in_list_to_in_list_append_l ]
- @(not_to_not … Hv1v2) @in_list_to_in_list_append_r ]
-qed.
-
-corollary nominal_eta : ∀x,u.x ∉ FV u → (νx.u⌈x⌉) = u.
-#x #u #Hu @nominal_eta_aux //
-qed.
-
-lemma eq_height_vopen_aux : ∀v,x,k.tm_height (vopen_tm_aux v x k) = tm_height v.
-#v #x elim v
-[ #n #k normalize cases (eqb n k) // cases (leb n k) //
-| #u #k %
-| #s #u #IH #k normalize >IH %
-| #u1 #u2 #IH1 #IH2 #k normalize >IH1 >IH2 % ]
-qed.
-
-corollary eq_height_vopen : ∀v,x.tm_height (v⌈x⌉) = tm_height v.
-#v #x @eq_height_vopen_aux
-qed.
-
-theorem pretm_ind_plus_aux :
- ∀P:pretm → Type[0].
- (∀x:𝔸.P (par x)) →
- (∀n:ℕ.P (var n)) →
- (∀v1,v2. P v1 → P v2 → P (app v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) →
- ∀n,u.tm_height u ≤ n → P u.
-#P #Hpar #Hvar #Happ #C #Hlam #n change with ((λn.?) n); @(nat_rect_Type0_1 n ??)
-#m cases m
-[ #_ * /2/
- [ normalize #s #v #Hfalse cases (?:False) cases (not_le_Sn_O (tm_height v)) /2/
- | #v1 #v2 whd in ⊢ (?%?→?); #Hfalse cases (?:False) cases (not_le_Sn_O (S (max ??))) /2/ ] ]
--m #m #IH * /2/
-[ #s #v whd in ⊢ (?%?→?); #Hv
- lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy
- >(?:abs s v = lam y s (v⌈y⌉))
- [| whd in ⊢ (???%); >nominal_eta // @(not_to_not … Hy) @in_list_to_in_list_append_r ]
- @Hlam
- [ @(not_to_not … Hy) @in_list_to_in_list_append_r
- | @(not_to_not … Hy) @in_list_to_in_list_append_l ]
- @IH [| @Hv | >eq_height_vopen % ]
-| #v1 #v2 whd in ⊢ (?%?→?); #Hv @Happ
- [ @IH [| @Hv | // ] | @IH [| @Hv | // ] ] ]
-qed.
-
-corollary pretm_ind_plus :
- ∀P:pretm → Type[0].
- (∀x:𝔸.P (par x)) →
- (∀n:ℕ.P (var n)) →
- (∀v1,v2. P v1 → P v2 → P (app v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s,v.x ∉ FV v → x ∉ C → P (v⌈x⌉) → P (lam x s (v⌈x⌉))) →
- ∀u.P u.
-#P #Hpar #Hvar #Happ #C #Hlam #u @pretm_ind_plus_aux /2/
-qed.
-
-(* maps a permutation to a list of terms *)
-definition Pi_map_list : (𝔸 → 𝔸) → list 𝔸 → list 𝔸 ≝ map 𝔸 𝔸 .
-
-(* interpretation "permutation of name list" 'middot p x = (Pi_map_list p x).*)
-
-(*
-inductive tm : pretm → Prop ≝
-| tm_par : ∀x:𝔸.tm (par x)
-| tm_app : ∀u,v.tm u → tm v → tm (app u v)
-| tm_lam : ∀x,s,u.tm u → tm (lam x s u).
-
-inductive ctx_aux : nat → pretm → Prop ≝
-| ctx_var : ∀n,k.n < k → ctx_aux k (var n)
-| ctx_par : ∀x,k.ctx_aux k (par x)
-| ctx_app : ∀u,v,k.ctx_aux k u → ctx_aux k v → ctx_aux k (app u v)
-(* è sostituibile da ctx_lam ? *)
-| ctx_abs : ∀s,u.ctx_aux (S k) u → ctx_aux k (abs s u).
-*)
-
-inductive tm_or_ctx (k:nat) : pretm → Type[0] ≝
-| toc_var : ∀n.n < k → tm_or_ctx k (var n)
-| toc_par : ∀x.tm_or_ctx k (par x)
-| toc_app : ∀u,v.tm_or_ctx k u → tm_or_ctx k v → tm_or_ctx k (app u v)
-| toc_lam : ∀x,s,u.tm_or_ctx k u → tm_or_ctx k (lam x s u).
-
-definition tm ≝ λt.tm_or_ctx O t.
-definition ctx ≝ λt.tm_or_ctx 1 t.
-
-record TM : Type[0] ≝ {
- pretm_of_TM :> pretm;
- tm_of_TM : tm pretm_of_TM
-}.
-
-record CTX : Type[0] ≝ {
- pretm_of_CTX :> pretm;
- ctx_of_CTX : ctx pretm_of_CTX
-}.
-
-inductive tm2 : pretm → Type[0] ≝
-| tm_par : ∀x.tm2 (par x)
-| tm_app : ∀u,v.tm2 u → tm2 v → tm2 (app u v)
-| tm_lam : ∀x,s,u.x ∉ FV u → (∀y.y ∉ FV u → tm2 (u⌈y⌉)) → tm2 (lam x s (u⌈x⌉)).
-
-(*
-inductive tm' : pretm → Prop ≝
-| tm_par : ∀x.tm' (par x)
-| tm_app : ∀u,v.tm' u → tm' v → tm' (app u v)
-| tm_lam : ∀x,s,u,C.x ∉ FV u → x ∉ C → (∀y.y ∉ FV u → tm' (❴u❵y)) → tm' (lam x s (❴u❵x)).
-*)
-
-axiom swap_inj : ∀N.∀z1,z2,x,y.swap N z1 z2 x = swap N z1 z2 y → x = y.
-
-lemma pi_vclose_tm :
- ∀z1,z2,x,u.swap 𝔸 z1 z2·(νx.u) = (ν swap ? z1 z2 x.swap 𝔸 z1 z2 · u).
-#z1 #z2 #x #u
-change with (vclose_tm_aux ???) in match (vclose_tm ??);
-change with (vclose_tm_aux ???) in ⊢ (???%); lapply O elim u
-[3:whd in ⊢ (?→?→(?→ ??%%)→?→??%%); //
-|4:whd in ⊢ (?→?→(?→??%%)→(?→??%%)→?→??%%); //
-| #n #k whd in ⊢ (??(??%)%); cases (leb k n) normalize %
-| #x0 #k cases (true_or_false (x0==z1)) #H1 >H1 whd in ⊢ (??%%);
- [ cases (true_or_false (x0==x)) #H2 >H2 whd in ⊢ (??(??%)%);
- [ <(\P H2) >H1 whd in ⊢ (??(??%)%); >(\b ?) // >(\b ?) //
- | >H2 whd in match (swap ????); >H1
- whd in match (if false then var k else ?);
- whd in match (if true then z2 else ?); >(\bf ?)
- [ >(\P H1) >swap_left %
- | <(swap_inv ? z1 z2 z2) in ⊢ (?(??%?)); % #H3
- lapply (swap_inj … H3) >swap_right #H4 <H4 in H2; >H1 #H destruct (H) ]
- ]
- | >(?:(swap ? z1 z2 x0 == swap ? z1 z2 x) = (x0 == x))
- [| cases (true_or_false (x0==x)) #H2 >H2
- [ >(\P H2) @(\b ?) %
- | @(\bf ?) % #H >(swap_inj … H) in H2; >(\b ?) // #H0 destruct (H0) ] ]
- cases (true_or_false (x0==x)) #H2 >H2 whd in ⊢ (??(??%)%);
- [ <(\P H2) >H1 >(\b (refl ??)) %
- | >H1 >H2 % ]
- ]
- ]
-qed.
-
-lemma pi_vopen_tm :
- ∀z1,z2,x,u.swap 𝔸 z1 z2·(u⌈x⌉) = (swap 𝔸 z1 z2 · u⌈swap 𝔸 z1 z2 x⌉).
-#z1 #z2 #x #u
-change with (vopen_tm_aux ???) in match (vopen_tm ??);
-change with (vopen_tm_aux ???) in ⊢ (???%); lapply O elim u //
-[2: #s #v whd in ⊢ ((?→??%%)→?→??%%); //
-|3: #v1 #v2 whd in ⊢ ((?→??%%)→(?→??%%)→?→??%%); /2/ ]
-#n #k whd in ⊢ (??(??%)%); cases (true_or_false (eqb n k)) #H1 >H1 //
-cases (true_or_false (leb n k)) #H2 >H2 normalize //
-qed.
-
-lemma pi_lam :
- ∀z1,z2,x,s,u.swap 𝔸 z1 z2 · lam x s u = lam (swap 𝔸 z1 z2 x) s (swap 𝔸 z1 z2 · u).
-#z1 #z2 #x #s #u whd in ⊢ (???%); <(pi_vclose_tm …) %
-qed.
-
-lemma eqv_FV : ∀z1,z2,u.FV (swap 𝔸 z1 z2 · u) = Pi_map_list (swap 𝔸 z1 z2) (FV u).
-#z1 #z2 #u elim u //
-[ #s #v #H @H
-| #v1 #v2 whd in ⊢ (??%%→??%%→??%%); #H1 #H2 >H1 >H2
- whd in ⊢ (???(????%)); /2/ ]
-qed.
-
-lemma swap_inv_tm : ∀z1,z2,u.swap 𝔸 z1 z2 · (swap 𝔸 z1 z2 · u) = u.
-#z1 #z2 #u elim u
-[1: #n %
-|3: #s #v whd in ⊢ (?→??%%); //
-|4: #v1 #v2 #Hv1 #Hv2 whd in ⊢ (??%%); // ]
-#x whd in ⊢ (??%?); >swap_inv %
-qed.
-
-lemma eqv_in_list : ∀x,l,z1,z2.x ∈ l → swap 𝔸 z1 z2 x ∈ Pi_map_list (swap 𝔸 z1 z2) l.
-#x #l #z1 #z2 #Hin elim Hin
-[ #x0 #l0 %
-| #x1 #x2 #l0 #Hin #IH %2 @IH ]
-qed.
-
-lemma eqv_tm2 : ∀u.tm2 u → ∀z1,z2.tm2 ((swap ? z1 z2)·u).
-#u #Hu #z1 #z2 letin p ≝ (swap ? z1 z2) elim Hu /2/
-#x #s #v #Hx #Hv #IH >pi_lam >pi_vopen_tm %3
-[ @(not_to_not … Hx) -Hx #Hx
- <(swap_inv ? z1 z2 x) <(swap_inv_tm z1 z2 v) >eqv_FV @eqv_in_list //
-| #y #Hy <(swap_inv ? z1 z2 y)
- <pi_vopen_tm @IH @(not_to_not … Hy) -Hy #Hy <(swap_inv ? z1 z2 y)
- >eqv_FV @eqv_in_list //
-]
-qed.
-
-lemma vclose_vopen_aux : ∀x,u,k.vopen_tm_aux (vclose_tm_aux u x k) x k = u.
-#x #u elim u [1,3,4:normalize //]
-[ #n #k cases (true_or_false (leb k n)) #H >H whd in ⊢ (??%?);
- [ cases (true_or_false (eqb (S n) k)) #H1 >H1
- [ <(eqb_true_to_eq … H1) in H; #H lapply (leb_true_to_le … H) -H #H
- cases (le_to_not_lt … H) -H #H cases (H ?) %
- | whd in ⊢ (??%?); >lt_to_leb_false // @le_S_S /2/ ]
- | cases (true_or_false (eqb n k)) #H1 >H1 normalize
- [ >(eqb_true_to_eq … H1) in H; #H lapply (leb_false_to_not_le … H) -H
- * #H cases (H ?) %
- | >le_to_leb_true // @not_lt_to_le % #H2 >le_to_leb_true in H;
- [ #H destruct (H) | /2/ ]
- ]
- ]
-| #x0 #k whd in ⊢ (??(?%??)?); cases (true_or_false (x0==x))
- #H1 >H1 normalize // >(\P H1) >eqb_n_n % ]
-qed.
-
-lemma vclose_vopen : ∀x,u.((νx.u)⌈x⌉) = u. #x #u @vclose_vopen_aux
-qed.
-
-(*
-theorem tm_to_tm : ∀t.tm' t → tm t.
-#t #H elim H
-*)
-
-lemma in_list_singleton : ∀T.∀t1,t2:T.t1 ∈ [t2] → t1 = t2.
-#T #t1 #t2 #H @(in_list_inv_ind ??? H) /2/
-qed.
-
-lemma fresh_vclose_tm_aux : ∀u,x,k.x ∉ FV (vclose_tm_aux u x k).
-#u #x elim u //
-[ #n #k normalize cases (leb k n) normalize //
-| #x0 #k whd in ⊢ (?(???(?%))); cases (true_or_false (x0==x)) #H >H normalize //
- lapply (\Pf H) @not_to_not #Hin >(in_list_singleton ??? Hin) %
-| #v1 #v2 #IH1 #IH2 #k normalize % #Hin cases (in_list_append_to_or_in_list ???? Hin) -Hin #Hin
- [ cases (IH1 k) -IH1 #IH1 @IH1 @Hin | cases (IH2 k) -IH2 #IH2 @IH2 @Hin ]
-qed.
-
-lemma fresh_vclose_tm : ∀u,x.x ∉ FV (νx.u). //
-qed.
-
-lemma fresh_swap_tm : ∀z1,z2,u.z1 ∉ FV u → z2 ∉ FV u → swap 𝔸 z1 z2 · u = u.
-#z1 #z2 #u elim u
-[2: normalize in ⊢ (?→%→%→?); #x #Hz1 #Hz2 whd in ⊢ (??%?); >swap_other //
- [ @(not_to_not … Hz2) | @(not_to_not … Hz1) ] //
-|1: //
-| #s #v #IH normalize #Hz1 #Hz2 >IH // [@Hz2|@Hz1]
-| #v1 #v2 #IH1 #IH2 normalize #Hz1 #Hz2
- >IH1 [| @(not_to_not … Hz2) @in_list_to_in_list_append_l | @(not_to_not … Hz1) @in_list_to_in_list_append_l ]
- >IH2 // [@(not_to_not … Hz2) @in_list_to_in_list_append_r | @(not_to_not … Hz1) @in_list_to_in_list_append_r ]
-]
-qed.
-
-theorem tm_to_tm2 : ∀u.tm u → tm2 u.
-#t #Ht elim Ht
-[ #n #Hn cases (not_le_Sn_O n) #Hfalse cases (Hfalse Hn)
-| @tm_par
-| #u #v #Hu #Hv @tm_app
-| #x #s #u #Hu #IHu <(vclose_vopen x u) @tm_lam
- [ @fresh_vclose_tm
- | #y #Hy <(fresh_swap_tm x y (νx.u)) /2/ @fresh_vclose_tm ]
-]
-qed.
-
-theorem tm2_to_tm : ∀u.tm2 u → tm u.
-#u #pu elim pu /2/ #x #s #v #Hx #Hv #IH %4 @IH //
-qed.
-
-definition PAR ≝ λx.mk_TM (par x) ?. // qed.
-definition APP ≝ λu,v:TM.mk_TM (app u v) ?./2/ qed.
-definition LAM ≝ λx,s.λu:TM.mk_TM (lam x s u) ?./2/ qed.
-
-axiom vopen_tm_down : ∀u,x,k.tm_or_ctx (S k) u → tm_or_ctx k (u⌈x⌉).
-(* needs true_plus_false
-
-#u #x #k #Hu elim Hu
-[ #n #Hn normalize cases (true_or_false (eqb n O)) #H >H [%2]
- normalize >(?: leb n O = false) [|cases n in H; // >eqb_n_n #H destruct (H) ]
- normalize lapply Hn cases n in H; normalize [ #Hfalse destruct (Hfalse) ]
- #n0 #_ #Hn0 % @le_S_S_to_le //
-| #x0 %2
-| #v1 #v2 #Hv1 #Hv2 #IH1 #IH2 %3 //
-| #x0 #s #v #Hv #IH normalize @daemon
-]
-qed.
-*)
-
-definition vopen_TM ≝ λu:CTX.λx.mk_TM (u⌈x⌉) (vopen_tm_down …). @ctx_of_CTX qed.
-
-axiom vclose_tm_up : ∀u,x,k.tm_or_ctx k u → tm_or_ctx (S k) (νx.u).
-
-definition vclose_TM ≝ λu:TM.λx.mk_CTX (νx.u) (vclose_tm_up …). @tm_of_TM qed.
-
-interpretation "ln wf term variable open" 'open u x = (vopen_TM u x).
-interpretation "ln wf term variable close" 'nu x u = (vclose_TM u x).
-
-theorem tm_alpha : ∀x,y,s,u.x ∉ FV u → y ∉ FV u → lam x s (u⌈x⌉) = lam y s (u⌈y⌉).
-#x #y #s #u #Hx #Hy whd in ⊢ (??%%); @eq_f >nominal_eta // >nominal_eta //
-qed.
-
-theorem TM_ind_plus :
-(* non si può dare il principio in modo dipendente (almeno utilizzando tm2)
- la "prova" purtroppo è in Type e non si può garantire che sia esattamente
- quella che ci aspetteremmo
- *)
- ∀P:pretm → Type[0].
- (∀x:𝔸.P (PAR x)) →
- (∀v1,v2:TM.P v1 → P v2 → P (APP v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s.∀v:CTX.x ∉ FV v → x ∉ C →
- (∀y.y ∉ FV v → P (v⌈y⌉)) → P (LAM x s (v⌈x⌉))) →
- ∀u:TM.P u.
-#P #Hpar #Happ #C #Hlam * #u #pu elim (tm_to_tm2 u pu) //
-[ #v1 #v2 #pv1 #pv2 #IH1 #IH2 @(Happ (mk_TM …) (mk_TM …)) /2/
-| #x #s #v #Hx #pv #IH
- lapply (p_fresh … (C@FV v)) letin x0 ≝ (N_fresh … (C@FV v)) #Hx0
- >(?:lam x s (v⌈x⌉) = lam x0 s (v⌈x0⌉))
- [|@tm_alpha // @(not_to_not … Hx0) @in_list_to_in_list_append_r ]
- @(Hlam x0 s (mk_CTX v ?) ??)
- [ <(nominal_eta … Hx) @vclose_tm_up @tm2_to_tm @pv //
- | @(not_to_not … Hx0) @in_list_to_in_list_append_r
- | @(not_to_not … Hx0) @in_list_to_in_list_append_l
- | @IH ]
-]
-qed.
-
-notation
-"hvbox('nominal' u 'return' out 'with'
- [ 'xpar' ident x ⇒ f1
- | 'xapp' ident v1 ident v2 ident recv1 ident recv2 ⇒ f2
- | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident recw ⇒ f3 ])"
-with precedence 48
-for @{ TM_ind_plus $out (λ${ident x}:?.$f1)
- (λ${ident v1}:?.λ${ident v2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2)
- $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident recw}:?.$f3)
- $u }.
-
-(* include "basics/jmeq.ma".*)
-
-definition subst ≝ (λu:TM.λx,v.
- nominal u return (λ_.TM) with
- [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ u ]
- | xapp v1 v2 recv1 recv2 ⇒ APP recv1 recv2
- | xlam ❨y # x::FV v❩ s w py1 py2 recw ⇒ LAM y s (recw y py1) ]).
-
-lemma fasfd : ∀s,v. pretm_of_TM (subst (LAM O s (PAR 1)) O v) = pretm_of_TM (LAM O s (PAR 1)).
-#s #v normalize in ⊢ (??%?);
-
-
-theorem tm2_ind_plus :
-(* non si può dare il principio in modo dipendente (almeno utilizzando tm2) *)
- ∀P:pretm → Type[0].
- (∀x:𝔸.P (par x)) →
- (∀v1,v2.tm2 v1 → tm2 v2 → P v1 → P v2 → P (app v1 v2)) →
- ∀C:list 𝔸.
- (∀x,s,v.x ∉ FV v → x ∉ C → (∀y.y ∉ FV v → tm2 (v⌈y⌉)) →
- (∀y.y ∉ FV v → P (v⌈y⌉)) → P (lam x s (v⌈x⌉))) →
- ∀u.tm2 u → P u.
-#P #Hpar #Happ #C #Hlam #u #pu elim pu /2/
-#x #s #v #px #pv #IH
-lapply (p_fresh … (C@FV v)) letin y ≝ (N_fresh … (C@FV v)) #Hy
->(?:lam x s (v⌈x⌉) = lam y s (v⌈y⌉)) [| @tm_alpha // @(not_to_not … Hy) @in_list_to_in_list_append_r ]
-@Hlam /2/ lapply Hy -Hy @not_to_not #Hy
-[ @in_list_to_in_list_append_r @Hy | @in_list_to_in_list_append_l @Hy ]
-qed.
-
-definition check_tm ≝
- λu.pretm_ind_plus ? (λ_.true) (λ_.false)
- (λv1,v2,r1,r2.r1 ∧ r2) [ ] (λx,s,v,pv1,pv2,rv.rv) u.
-
-(*
-lemma check_tm_complete : ∀u.tm u → check_tm u = true.
-#u #pu @(tm2_ind_plus … [ ] … (tm_to_tm2 ? pu)) //
-[ #v1 #v2 #pv1 #pv2 #IH1 #IH2
-| #x #s #v #Hx1 #Hx2 #Hv #IH
-*)
-
-notation
-"hvbox('nominal' u 'return' out 'with'
- [ 'xpar' ident x ⇒ f1
- | 'xapp' ident v1 ident v2 ident pv1 ident pv2 ident recv1 ident recv2 ⇒ f2
- | 'xlam' ❨ident y # C❩ ident s ident w ident py1 ident py2 ident pw ident recw ⇒ f3 ])"
-with precedence 48
-for @{ tm2_ind_plus $out (λ${ident x}:?.$f1)
- (λ${ident v1}:?.λ${ident v2}:?.λ${ident pv1}:?.λ${ident pv2}:?.λ${ident recv1}:?.λ${ident recv2}:?.$f2)
- $C (λ${ident y}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.λ${ident pw}:?.λ${ident recw}:?.$f3)
- ? (tm_to_tm2 ? $u) }.
-(* notation
-"hvbox('nominal' u 'with'
- [ 'xlam' ident x # C ident s ident w ⇒ f3 ])"
-with precedence 48
-for @{ tm2_ind_plus ???
- $C (λ${ident x}:?.λ${ident s}:?.λ${ident w}:?.λ${ident py1}:?.λ${ident py2}:?.
- λ${ident pw}:?.λ${ident recw}:?.$f3) $u (tm_to_tm2 ??) }.
-*)
-
-
-definition subst ≝ (λu.λpu:tm u.λx,v.
- nominal pu return (λ_.pretm) with
- [ xpar x0 ⇒ match x == x0 with [ true ⇒ v | false ⇒ u ]
- | xapp v1 v2 pv1 pv2 recv1 recv2 ⇒ app recv1 recv2
- | xlam ❨y # x::FV v❩ s w py1 py2 pw recw ⇒ lam y s (recw y py1) ]).
-
-lemma fasfd : ∀x,s,u,p1,v. subst (lam x s u) p1 x v = lam x s u.
-#x #s #u #p1 #v
-
-
-definition subst ≝ λu.λpu:tm u.λx,y.
- tm2_ind_plus ?
- (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ])
- (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2)
- (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1))
- u (tm_to_tm2 … pu).
-check subst
-definition subst ≝ λu.λpu:tm u.λx,v.
- nominal u with
- [ xlam y # (x::FV v) s w ^ ? ].
-
-(*
-notation > "Λ ident x. ident T [ident x] ↦ P"
- with precedence 48 for @{'foo (λ${ident x}.λ${ident T}.$P)}.
-
-notation < "Λ ident x. ident T [ident x] ↦ P"
- with precedence 48 for @{'foo (λ${ident x}:$Q.λ${ident T}:$R.$P)}.
-*)
-
-(*
-notation
-"hvbox('nominal' u 'with'
- [ 'xpar' ident x ⇒ f1
- | 'xapp' ident v1 ident v2 ⇒ f2
- | 'xlam' ident x # C s w ⇒ f3 ])"
-with precedence 48
-for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1)
- (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2)
- $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }.
-*)
-
-(*
-notation
-"hvbox('nominal' u 'with'
- [ 'xpar' ident x ^ f1
- | 'xapp' ident v1 ident v2 ^ f2 ])"
-(* | 'xlam' ident x # C s w ^ f3 ]) *)
-with precedence 48
-for @{ tm2_ind_plus ? (λ${ident x}:$Tx.$f1)
- (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2)
- $C (λ${ident x}:$Tx.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }.
-*)
-notation
-"hvbox('nominal' u 'with'
- [ 'xpar' ident x ^ f1
- | 'xapp' ident v1 ident v2 ^ f2 ])"
-with precedence 48
-for @{ tm2_ind_plus ? (λ${ident x}:?.$f1)
- (λ${ident v1}:$Tv1.λ${ident v2}:$Tv2.λ${ident pv1}:$Tpv1.λ${ident pv2}:$Tpv2.λ${ident recv1}:$Trv1.λ${ident recv2}:$Trv2.$f2)
- $C (λ${ident x}:?.λ${ident s}:$Ts.λ${ident w}:$Tw.λ${ident py1}:$Tpy1.λ${ident py2}:$Tpy2.λ${ident pw}:$Tpw.λ${ident recw}:$Trw.$f3) $u (tm_to_tm2 ??) }.
-
-
-definition subst ≝ λu.λpu:tm u.λx,v.
- nominal u with
- [ xpar x0 ^ match x == x0 with [ true ⇒ v | false ⇒ u ]
- | xapp v1 v2 ^ ? ].
- | xlam y # (x::FV v) s w ^ ? ].
-
-
- (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ])
- (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2)
- (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1))
- u (tm_to_tm2 … pu).
-
-
-*)
-definition subst ≝ λu.λpu:tm u.λx,v.
- tm2_ind_plus ?
- (* par x0 *) (λx0.match x == x0 with [ true ⇒ v | false ⇒ u ])
- (* app v1 v2 *) (λv1,v2,pv1,pv2,recv1,recv2.app recv1 recv2)
- (* lam y#(x::FV v) s w *) (x::FV v) (λy,s,w,py1,py2,pw,recw.lam y s (recw y py1))
- u (tm_to_tm2 … pu).
-
-check subst
-
-
-axiom in_Env : 𝔸 × tp → Env → Prop.
-notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}.
-interpretation "Env membership" 'lefttriangle x l = (in_Env x l).
-
-
-
-inductive judg : list tp → tm → tp → Prop ≝
-| t_var : ∀g,n,t.Nth ? n g = Some ? t → judg g (var n) t
-| t_app : ∀g,m,n,t,u.judg g m (arr t u) → judg g n t → judg g (app m n) u
-| t_abs : ∀g,t,m,u.judg (t::g) m u → judg g (abs t m) (arr t u).
-
-definition Env := list (𝔸 × tp).
-
-axiom vclose_env : Env → list tp.
-axiom vclose_tm : Env → tm → tm.
-axiom Lam : 𝔸 → tp → tm → tm.
-definition Judg ≝ λG,M,T.judg (vclose_env G) (vclose_tm G M) T.
-definition dom ≝ λG:Env.map ?? (fst ??) G.
-
-definition sctx ≝ 𝔸 × tm.
-axiom swap_tm : 𝔸 → 𝔸 → tm → tm.
-definition sctx_app : sctx → 𝔸 → tm ≝ λM0,Y.let 〈X,M〉 ≝ M0 in swap_tm X Y M.
-
-axiom in_list : ∀A:Type[0].A → list A → Prop.
-interpretation "list membership" 'mem x l = (in_list ? x l).
-interpretation "list non-membership" 'notmem x l = (Not (in_list ? x l)).
-
-axiom in_Env : 𝔸 × tp → Env → Prop.
-notation "X ◃ G" non associative with precedence 45 for @{'lefttriangle $X $G}.
-interpretation "Env membership" 'lefttriangle x l = (in_Env x l).
-
-let rec FV M ≝ match M with
- [ par X ⇒ [X]
- | app M1 M2 ⇒ FV M1@FV M2
- | abs T M0 ⇒ FV M0
- | _ ⇒ [ ] ].
-
-(* axiom Lookup : 𝔸 → Env → option tp. *)
-
-(* forma alto livello del judgment
- t_abs* : ∀G,T,X,M,U.
- (∀Y ∉ supp(M).Judg (〈Y,T〉::G) (M[Y]) U) →
- Judg G (Lam X T (M[X])) (arr T U) *)
-
-(* prima dimostrare, poi perfezionare gli assiomi, poi dimostrarli *)
-
-axiom Judg_ind : ∀P:Env → tm → tp → Prop.
- (∀X,G,T.〈X,T〉 ◃ G → P G (par X) T) →
- (∀G,M,N,T,U.
- Judg G M (arr T U) → Judg G N T →
- P G M (arr T U) → P G N T → P G (app M N) U) →
- (∀G,T1,T2,X,M1.
- (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → Judg (〈Y,T1〉::G) (sctx_app M1 Y) T2) →
- (∀Y.Y ∉ (FV (Lam X T1 (sctx_app M1 X))) → P (〈Y,T1〉::G) (sctx_app M1 Y) T2) →
- P G (Lam X T1 (sctx_app M1 X)) (arr T1 T2)) →
- ∀G,M,T.Judg G M T → P G M T.
-
-axiom t_par : ∀X,G,T.〈X,T〉 ◃ G → Judg G (par X) T.
-axiom t_app2 : ∀G,M,N,T,U.Judg G M (arr T U) → Judg G N T → Judg G (app M N) U.
-axiom t_Lam : ∀G,X,M,T,U.Judg (〈X,T〉::G) M U → Judg G (Lam X T M) (arr T U).
-
-definition subenv ≝ λG1,G2.∀x.x ◃ G1 → x ◃ G2.
-interpretation "subenv" 'subseteq G1 G2 = (subenv G1 G2).
-
-axiom daemon : ∀P:Prop.P.
-
-theorem weakening : ∀G1,G2,M,T.G1 ⊆ G2 → Judg G1 M T → Judg G2 M T.
-#G1 #G2 #M #T #Hsub #HJ lapply Hsub lapply G2 -G2 change with (∀G2.?)
-@(Judg_ind … HJ)
-[ #X #G #T0 #Hin #G2 #Hsub @t_par @Hsub //
-| #G #M0 #N #T0 #U #HM0 #HN #IH1 #IH2 #G2 #Hsub @t_app2
- [| @IH1 // | @IH2 // ]
-| #G #T1 #T2 #X #M1 #HM1 #IH #G2 #Hsub @t_Lam @IH
- [ (* trivial property of Lam *) @daemon
- | (* trivial property of subenv *) @daemon ]
-]
-qed.
-
-(* Serve un tipo Tm per i termini localmente chiusi e i suoi principi di induzione e
- ricorsione *)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||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 "basics/logic.ma".
-include "basics/lists/in.ma".
-include "basics/types.ma".
-
-(*interpretation "list membership" 'mem x l = (in_list ? x l).*)
-
-record Nset : Type[1] ≝
-{
- (* carrier is specified as a coercion: when an object X of type Nset is
- given, but something of type Type is expected, Matita will insert a
- hidden coercion: the user sees "X", but really means "carrier X" *)
- carrier :> DeqSet;
- N_fresh : list carrier → carrier;
- p_fresh : ∀l.N_fresh l ∉ l
-}.
-
-definition maxlist ≝
- λl.foldr ?? (λx,acc.max x acc) 0 l.
-
-definition natfresh ≝ λl.S (maxlist l).
-
-lemma le_max_1 : ∀x,y.x ≤ max x y. /2/
-qed.
-
-lemma le_max_2 : ∀x,y.y ≤ max x y. /2/
-qed.
-
-lemma le_maxlist : ∀l,x.x ∈ l → x ≤ maxlist l.
-#l elim l
-[#x #Hx @False_ind cases (not_in_list_nil ? x) #H1 /2/
-|#y #tl #IH #x #H1 change with (max ??) in ⊢ (??%);
- cases (in_list_cons_case ???? H1);#H2;
- [ >H2 @le_max_1
- | whd in ⊢ (??%); lapply (refl ? (leb y (maxlist tl)));
- cases (leb y (maxlist tl)) in ⊢ (???% → %);#H3
- [ @IH //
- | lapply (IH ? H2) #H4
- lapply (leb_false_to_not_le … H3) #H5
- lapply (not_le_to_lt … H5) #H6
- @(transitive_le … H4)
- @(transitive_le … H6) %2 %
- ]
- ]
-]
-qed.
-
-(* prove freshness for nat *)
-lemma lt_l_natfresh_l : ∀l,x.x ∈ l → x < natfresh l.
-#l #x #H1 @le_S_S /2/
-qed.
-
-(*naxiom p_Xfresh : ∀l.∀x:Xcarr.x ∈ l → x ≠ ntm (Xfresh l) ∧ x ≠ ntp (Xfresh l).*)
-lemma p_natfresh : ∀l.natfresh l ∉ l.
-#l % #H1 lapply (lt_l_natfresh_l … H1) #H2
-cases (lt_to_not_eq … H2) #H3 @H3 %
-qed.
-
-include "basics/finset.ma".
-
-definition X : Nset ≝ mk_Nset DeqNat ….
-[ @natfresh
-| @p_natfresh
-]
-qed.
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+(* Logic system *)
+
+include "basics/pts.ma".
+include "hints_declaration.ma".
+
+inductive Imply (A,B:Prop) : Prop ≝
+| Imply_intro: (A → B) → Imply A B.
+
+definition Imply_elim ≝ λA,B:Prop.λf:Imply A B. λa:A.
+ match f with [ Imply_intro g ⇒ g a].
+
+inductive And (A,B:Prop) : Prop ≝
+| And_intro: A → B → And A B.
+
+definition And_elim_l ≝ λA,B.λc:And A B.
+ match c with [ And_intro a b ⇒ a ].
+
+definition And_elim_r ≝ λA,B.λc:And A B.
+ match c with [ And_intro a b ⇒ b ].
+
+inductive Or (A,B:Prop) : Prop ≝
+| Or_intro_l: A → Or A B
+| Or_intro_r: B → Or A B.
+
+definition Or_elim ≝ λA,B,C:Prop.λc:Or A B.λfa: A → C.λfb: B → C.
+ match c with
+ [ Or_intro_l a ⇒ fa a
+ | Or_intro_r b ⇒ fb b].
+
+inductive Top : Prop :=
+| Top_intro : Top.
+
+inductive Bot : Prop := .
+
+definition Bot_elim ≝ λP:Prop.λx:Bot.
+ match x in Bot return λx.P with [].
+
+definition Not := λA:Prop.Imply A Bot.
+
+definition Not_intro : ∀A.(A → Bot) → Not A ≝ λA.
+ Imply_intro A Bot.
+
+definition Not_elim : ∀A.Not A → A → Bot ≝ λA.
+ Imply_elim ? Bot.
+
+definition Discharge := λA:Prop.λa:A.
+ a.
+
+axiom Raa : ∀A.(Not A → Bot) → A.
+
+axiom sort : Type[0].
+
+inductive Exists (A:Type[0]) (P:A→Prop) : Prop ≝
+ Exists_intro: ∀w:A. P w → Exists A P.
+
+definition Exists_elim ≝
+ λA:Type[0].λP:A→Prop.λC:Prop.λc:Exists A P.λH:(Πx.P x → C).
+ match c with [ Exists_intro w p ⇒ H w p ].
+
+inductive Forall (A:Type[0]) (P:A→Prop) : Prop ≝
+ Forall_intro: (∀n:A. P n) → Forall A P.
+
+definition Forall_elim ≝
+ λA:Type[0].λP:A→Prop.λn:A.λf:Forall A P.match f with [ Forall_intro g ⇒ g n ].
+
+(* Dummy proposition *)
+axiom unit : Prop.
+
+(* Notations *)
+notation "hbox(a break ⇒ b)" right associative with precedence 20
+for @{ 'Imply $a $b }.
+interpretation "Imply" 'Imply a b = (Imply a b).
+interpretation "constructive or" 'or x y = (Or x y).
+interpretation "constructive and" 'and x y = (And x y).
+notation "⊤" non associative with precedence 90 for @{'Top}.
+interpretation "Top" 'Top = Top.
+notation "⊥" non associative with precedence 90 for @{'Bot}.
+interpretation "Bot" 'Bot = Bot.
+interpretation "Not" 'not a = (Not a).
+notation "✶" non associative with precedence 90 for @{'unit}.
+interpretation "dummy prop" 'unit = unit.
+notation > "\exists list1 ident x sep , . term 19 Px" with precedence 20
+for ${ fold right @{$Px} rec acc @{'myexists (λ${ident x}.$acc)} }.
+notation < "hvbox(\exists ident i break . p)" with precedence 20
+for @{ 'myexists (\lambda ${ident i} : $ty. $p) }.
+interpretation "constructive ex" 'myexists \eta.x = (Exists sort x).
+notation > "\forall ident x.break term 19 Px" with precedence 20
+for @{ 'Forall (λ${ident x}.$Px) }.
+notation < "\forall ident x.break term 19 Px" with precedence 20
+for @{ 'Forall (λ${ident x}:$tx.$Px) }.
+interpretation "Forall" 'Forall \eta.Px = (Forall ? Px).
+
+(* Variables *)
+axiom A : Prop.
+axiom B : Prop.
+axiom C : Prop.
+axiom D : Prop.
+axiom E : Prop.
+axiom F : Prop.
+axiom G : Prop.
+axiom H : Prop.
+axiom I : Prop.
+axiom J : Prop.
+axiom K : Prop.
+axiom L : Prop.
+axiom M : Prop.
+axiom N : Prop.
+axiom O : Prop.
+axiom x: sort.
+axiom y: sort.
+axiom z: sort.
+axiom w: sort.
+
+(* Every formula user provided annotates its proof:
+ `A` becomes `(show A ?)` *)
+definition show : ΠA.A→A ≝ λA:Prop.λa:A.a.
+
+(* When something does not fit, this daemon is used *)
+axiom cast: ΠA,B:Prop.B → A.
+
+(* begin a proof: draws the root *)
+notation > "'prove' p" non associative with precedence 19
+for @{ 'prove $p }.
+interpretation "prove KO" 'prove p = (cast ? ? (show p ?)).
+interpretation "prove OK" 'prove p = (show p ?).
+
+(* Leaves *)
+notation < "\infrule (t\atop ⋮) a ?" with precedence 19
+for @{ 'leaf_ok $a $t }.
+interpretation "leaf OK" 'leaf_ok a t = (show a t).
+notation < "\infrule (t\atop ⋮) mstyle color #ff0000 (a) ?" with precedence 19
+for @{ 'leaf_ko $a $t }.
+interpretation "leaf KO" 'leaf_ko a t = (cast ? ? (show a t)).
+
+(* discharging *)
+notation < "[ a ] \sup mstyle color #ff0000 (H)" with precedence 19
+for @{ 'discharge_ko_1 $a $H }.
+interpretation "discharge_ko_1" 'discharge_ko_1 a H =
+ (show a (cast ? ? (Discharge ? H))).
+notation < "[ mstyle color #ff0000 (a) ] \sup mstyle color #ff0000 (H)" with precedence 19
+for @{ 'discharge_ko_2 $a $H }.
+interpretation "discharge_ko_2" 'discharge_ko_2 a H =
+ (cast ? ? (show a (cast ? ? (Discharge ? H)))).
+
+notation < "[ a ] \sup H" with precedence 19
+for @{ 'discharge_ok_1 $a $H }.
+interpretation "discharge_ok_1" 'discharge_ok_1 a H =
+ (show a (Discharge ? H)).
+notation < "[ mstyle color #ff0000 (a) ] \sup H" with precedence 19
+for @{ 'discharge_ok_2 $a $H }.
+interpretation "discharge_ok_2" 'discharge_ok_2 a H =
+ (cast ? ? (show a (Discharge ? H))).
+
+notation > "'discharge' [H]" with precedence 19
+for @{ 'discharge $H }.
+interpretation "discharge KO" 'discharge H = (cast ? ? (Discharge ? H)).
+interpretation "discharge OK" 'discharge H = (Discharge ? H).
+
+(* ⇒ introduction *)
+notation < "\infrule hbox(\emsp b \emsp) ab (mstyle color #ff0000 (⇒\sub\i \emsp) ident H) " with precedence 19
+for @{ 'Imply_intro_ko_1 $ab (λ${ident H}:$p.$b) }.
+interpretation "Imply_intro_ko_1" 'Imply_intro_ko_1 ab \eta.b =
+ (show ab (cast ? ? (Imply_intro ? ? b))).
+
+notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (mstyle color #ff0000 (⇒\sub\i \emsp) ident H) " with precedence 19
+for @{ 'Imply_intro_ko_2 $ab (λ${ident H}:$p.$b) }.
+interpretation "Imply_intro_ko_2" 'Imply_intro_ko_2 ab \eta.b =
+ (cast ? ? (show ab (cast ? ? (Imply_intro ? ? b)))).
+
+notation < "maction (\infrule hbox(\emsp b \emsp) ab (⇒\sub\i \emsp ident H) ) (\vdots)" with precedence 19
+for @{ 'Imply_intro_ok_1 $ab (λ${ident H}:$p.$b) }.
+interpretation "Imply_intro_ok_1" 'Imply_intro_ok_1 ab \eta.b =
+ (show ab (Imply_intro ? ? b)).
+
+notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (⇒\sub\i \emsp ident H) " with precedence 19
+for @{ 'Imply_intro_ok_2 $ab (λ${ident H}:$p.$b) }.
+interpretation "Imply_intro_ok_2" 'Imply_intro_ok_2 ab \eta.b =
+ (cast ? ? (show ab (Imply_intro ? ? b))).
+
+notation > "⇒#'i' [ident H] term 90 b" with precedence 19
+for @{ 'Imply_intro $b (λ${ident H}.show $b ?) }.
+
+interpretation "Imply_intro KO" 'Imply_intro b pb =
+ (cast ? (Imply unit b) (Imply_intro ? b pb)).
+interpretation "Imply_intro OK" 'Imply_intro b pb =
+ (Imply_intro ? b pb).
+
+(* ⇒ elimination *)
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b mstyle color #ff0000 (⇒\sub\e) " with precedence 19
+for @{ 'Imply_elim_ko_1 $ab $a $b }.
+interpretation "Imply_elim_ko_1" 'Imply_elim_ko_1 ab a b =
+ (show b (cast ? ? (Imply_elim ? ? (cast ? ? ab) (cast ? ? a)))).
+
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) mstyle color #ff0000 (⇒\sub\e) " with precedence 19
+for @{ 'Imply_elim_ko_2 $ab $a $b }.
+interpretation "Imply_elim_ko_2" 'Imply_elim_ko_2 ab a b =
+ (cast ? ? (show b (cast ? ? (Imply_elim ? ? (cast ? ? ab) (cast ? ? a))))).
+
+notation < "maction (\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (⇒\sub\e) ) (\vdots)" with precedence 19
+for @{ 'Imply_elim_ok_1 $ab $a $b }.
+interpretation "Imply_elim_ok_1" 'Imply_elim_ok_1 ab a b =
+ (show b (Imply_elim ? ? ab a)).
+
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) (⇒\sub\e) " with precedence 19
+for @{ 'Imply_elim_ok_2 $ab $a $b }.
+interpretation "Imply_elim_ok_2" 'Imply_elim_ok_2 ab a b =
+ (cast ? ? (show b (Imply_elim ? ? ab a))).
+
+notation > "⇒#'e' term 90 ab term 90 a" with precedence 19
+for @{ 'Imply_elim (show $ab ?) (show $a ?) }.
+interpretation "Imply_elim KO" 'Imply_elim ab a =
+ (cast ? ? (Imply_elim ? ? (cast (Imply unit unit) ? ab) (cast unit ? a))).
+interpretation "Imply_elim OK" 'Imply_elim ab a =
+ (Imply_elim ? ? ab a).
+
+(* ∧ introduction *)
+notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) ab mstyle color #ff0000 (∧\sub\i)" with precedence 19
+for @{ 'And_intro_ko_1 $a $b $ab }.
+interpretation "And_intro_ko_1" 'And_intro_ko_1 a b ab =
+ (show ab (cast ? ? (And_intro ? ? a b))).
+
+notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) mstyle color #ff0000 (ab) mstyle color #ff0000 (∧\sub\i)" with precedence 19
+for @{ 'And_intro_ko_2 $a $b $ab }.
+interpretation "And_intro_ko_2" 'And_intro_ko_2 a b ab =
+ (cast ? ? (show ab (cast ? ? (And_intro ? ? a b)))).
+
+notation < "maction (\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) ab (∧\sub\i)) (\vdots)" with precedence 19
+for @{ 'And_intro_ok_1 $a $b $ab }.
+interpretation "And_intro_ok_1" 'And_intro_ok_1 a b ab =
+ (show ab (And_intro ? ? a b)).
+
+notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) mstyle color #ff0000 (ab) (∧\sub\i)" with precedence 19
+for @{ 'And_intro_ok_2 $a $b $ab }.
+interpretation "And_intro_ok_2" 'And_intro_ok_2 a b ab =
+ (cast ? ? (show ab (And_intro ? ? a b))).
+
+notation > "∧#'i' term 90 a term 90 b" with precedence 19
+for @{ 'And_intro (show $a ?) (show $b ?) }.
+interpretation "And_intro KO" 'And_intro a b = (cast ? ? (And_intro ? ? a b)).
+interpretation "And_intro OK" 'And_intro a b = (And_intro ? ? a b).
+
+(* ∧ elimination *)
+notation < "\infrule hbox(\emsp ab \emsp) a mstyle color #ff0000 (∧\sub(\e_\l))" with precedence 19
+for @{ 'And_elim_l_ko_1 $ab $a }.
+interpretation "And_elim_l_ko_1" 'And_elim_l_ko_1 ab a =
+ (show a (cast ? ? (And_elim_l ? ? (cast ? ? ab)))).
+
+notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) mstyle color #ff0000 (∧\sub(\e_\l))" with precedence 19
+for @{ 'And_elim_l_ko_2 $ab $a }.
+interpretation "And_elim_l_ko_2" 'And_elim_l_ko_2 ab a =
+ (cast ? ? (show a (cast ? ? (And_elim_l ? ? (cast ? ? ab))))).
+
+notation < "maction (\infrule hbox(\emsp ab \emsp) a (∧\sub(\e_\l))) (\vdots)" with precedence 19
+for @{ 'And_elim_l_ok_1 $ab $a }.
+interpretation "And_elim_l_ok_1" 'And_elim_l_ok_1 ab a =
+ (show a (And_elim_l ? ? ab)).
+
+notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) (∧\sub(\e_\l))" with precedence 19
+for @{ 'And_elim_l_ok_2 $ab $a }.
+interpretation "And_elim_l_ok_2" 'And_elim_l_ok_2 ab a =
+ (cast ? ? (show a (And_elim_l ? ? ab))).
+
+notation > "∧#'e_l' term 90 ab" with precedence 19
+for @{ 'And_elim_l (show $ab ?) }.
+interpretation "And_elim_l KO" 'And_elim_l a = (cast ? ? (And_elim_l ? ? (cast (And unit unit) ? a))).
+interpretation "And_elim_l OK" 'And_elim_l a = (And_elim_l ? ? a).
+
+notation < "\infrule hbox(\emsp ab \emsp) a mstyle color #ff0000 (∧\sub(\e_\r))" with precedence 19
+for @{ 'And_elim_r_ko_1 $ab $a }.
+interpretation "And_elim_r_ko_1" 'And_elim_r_ko_1 ab a =
+ (show a (cast ? ? (And_elim_r ? ? (cast ? ? ab)))).
+
+notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) mstyle color #ff0000 (∧\sub(\e_\r))" with precedence 19
+for @{ 'And_elim_r_ko_2 $ab $a }.
+interpretation "And_elim_r_ko_2" 'And_elim_r_ko_2 ab a =
+ (cast ? ? (show a (cast ? ? (And_elim_r ? ? (cast ? ? ab))))).
+
+notation < "maction (\infrule hbox(\emsp ab \emsp) a (∧\sub(\e_\r))) (\vdots)" with precedence 19
+for @{ 'And_elim_r_ok_1 $ab $a }.
+interpretation "And_elim_r_ok_1" 'And_elim_r_ok_1 ab a =
+ (show a (And_elim_r ? ? ab)).
+
+notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) (∧\sub(\e_\r))" with precedence 19
+for @{ 'And_elim_r_ok_2 $ab $a }.
+interpretation "And_elim_r_ok_2" 'And_elim_r_ok_2 ab a =
+ (cast ? ? (show a (And_elim_r ? ? ab))).
+
+notation > "∧#'e_r' term 90 ab" with precedence 19
+for @{ 'And_elim_r (show $ab ?) }.
+interpretation "And_elim_r KO" 'And_elim_r a = (cast ? ? (And_elim_r ? ? (cast (And unit unit) ? a))).
+interpretation "And_elim_r OK" 'And_elim_r a = (And_elim_r ? ? a).
+
+(* ∨ introduction *)
+notation < "\infrule hbox(\emsp a \emsp) ab mstyle color #ff0000 (∨\sub(\i_\l))" with precedence 19
+for @{ 'Or_intro_l_ko_1 $a $ab }.
+interpretation "Or_intro_l_ko_1" 'Or_intro_l_ko_1 a ab =
+ (show ab (cast ? ? (Or_intro_l ? ? a))).
+
+notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) mstyle color #ff0000 (∨\sub(\i_\l))" with precedence 19
+for @{ 'Or_intro_l_ko_2 $a $ab }.
+interpretation "Or_intro_l_ko_2" 'Or_intro_l_ko_2 a ab =
+ (cast ? ? (show ab (cast ? ? (Or_intro_l ? ? a)))).
+
+notation < "maction (\infrule hbox(\emsp a \emsp) ab (∨\sub(\i_\l))) (\vdots)" with precedence 19
+for @{ 'Or_intro_l_ok_1 $a $ab }.
+interpretation "Or_intro_l_ok_1" 'Or_intro_l_ok_1 a ab =
+ (show ab (Or_intro_l ? ? a)).
+
+notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) (∨\sub(\i_\l))" with precedence 19
+for @{ 'Or_intro_l_ok_2 $a $ab }.
+interpretation "Or_intro_l_ok_2" 'Or_intro_l_ok_2 a ab =
+ (cast ? ? (show ab (Or_intro_l ? ? a))).
+
+notation > "∨#'i_l' term 90 a" with precedence 19
+for @{ 'Or_intro_l (show $a ?) }.
+interpretation "Or_intro_l KO" 'Or_intro_l a = (cast ? (Or ? unit) (Or_intro_l ? ? a)).
+interpretation "Or_intro_l OK" 'Or_intro_l a = (Or_intro_l ? ? a).
+
+notation < "\infrule hbox(\emsp a \emsp) ab mstyle color #ff0000 (∨\sub(\i_\r))" with precedence 19
+for @{ 'Or_intro_r_ko_1 $a $ab }.
+interpretation "Or_intro_r_ko_1" 'Or_intro_r_ko_1 a ab =
+ (show ab (cast ? ? (Or_intro_r ? ? a))).
+
+notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) mstyle color #ff0000 (∨\sub(\i_\r))" with precedence 19
+for @{ 'Or_intro_r_ko_2 $a $ab }.
+interpretation "Or_intro_r_ko_2" 'Or_intro_r_ko_2 a ab =
+ (cast ? ? (show ab (cast ? ? (Or_intro_r ? ? a)))).
+
+notation < "maction (\infrule hbox(\emsp a \emsp) ab (∨\sub(\i_\r))) (\vdots)" with precedence 19
+for @{ 'Or_intro_r_ok_1 $a $ab }.
+interpretation "Or_intro_r_ok_1" 'Or_intro_r_ok_1 a ab =
+ (show ab (Or_intro_r ? ? a)).
+
+notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) (∨\sub(\i_\r))" with precedence 19
+for @{ 'Or_intro_r_ok_2 $a $ab }.
+interpretation "Or_intro_r_ok_2" 'Or_intro_r_ok_2 a ab =
+ (cast ? ? (show ab (Or_intro_r ? ? a))).
+
+notation > "∨#'i_r' term 90 a" with precedence 19
+for @{ 'Or_intro_r (show $a ?) }.
+interpretation "Or_intro_r KO" 'Or_intro_r a = (cast ? (Or unit ?) (Or_intro_r ? ? a)).
+interpretation "Or_intro_r OK" 'Or_intro_r a = (Or_intro_r ? ? a).
+
+(* ∨ elimination *)
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) c (mstyle color #ff0000 (∨\sub\e \emsp) ident Ha \emsp ident Hb)" with precedence 19
+for @{ 'Or_elim_ko_1 $ab $c (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) }.
+interpretation "Or_elim_ko_1" 'Or_elim_ko_1 ab c \eta.ac \eta.bc =
+ (show c (cast ? ? (Or_elim ? ? ? (cast ? ? ab) (cast ? ? ac) (cast ? ? bc)))).
+
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) mstyle color #ff0000 (c) (mstyle color #ff0000 (∨\sub\e) \emsp ident Ha \emsp ident Hb)" with precedence 19
+for @{ 'Or_elim_ko_2 $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }.
+interpretation "Or_elim_ko_2" 'Or_elim_ko_2 ab \eta.ac \eta.bc c =
+ (cast ? ? (show c (cast ? ? (Or_elim ? ? ? (cast ? ? ab) (cast ? ? ac) (cast ? ? bc))))).
+
+notation < "maction (\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) c (∨\sub\e \emsp ident Ha \emsp ident Hb)) (\vdots)" with precedence 19
+for @{ 'Or_elim_ok_1 $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }.
+interpretation "Or_elim_ok_1" 'Or_elim_ok_1 ab \eta.ac \eta.bc c =
+ (show c (Or_elim ? ? ? ab ac bc)).
+
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) mstyle color #ff0000 (c) (∨\sub\e \emsp ident Ha \emsp ident Hb)" with precedence 19
+for @{ 'Or_elim_ok_2 $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }.
+interpretation "Or_elim_ok_2" 'Or_elim_ok_2 ab \eta.ac \eta.bc c =
+ (cast ? ? (show c (Or_elim ? ? ? ab ac bc))).
+
+definition unit_to ≝ λx:Prop.unit → x.
+
+notation > "∨#'e' term 90 ab [ident Ha] term 90 cl [ident Hb] term 90 cr" with precedence 19
+for @{ 'Or_elim (show $ab ?) (λ${ident Ha}.show $cl ?) (λ${ident Hb}.show $cr ?) }.
+interpretation "Or_elim KO" 'Or_elim ab ac bc =
+ (cast ? ? (Or_elim ? ? ?
+ (cast (Or unit unit) ? ab)
+ (cast (unit_to unit) (unit_to ?) ac)
+ (cast (unit_to unit) (unit_to ?) bc))).
+interpretation "Or_elim OK" 'Or_elim ab ac bc = (Or_elim ? ? ? ab ac bc).
+
+(* ⊤ introduction *)
+notation < "\infrule \nbsp ⊤ mstyle color #ff0000 (⊤\sub\i)" with precedence 19
+for @{'Top_intro_ko_1}.
+interpretation "Top_intro_ko_1" 'Top_intro_ko_1 =
+ (show ? (cast ? ? Top_intro)).
+
+notation < "\infrule \nbsp mstyle color #ff0000 (⊤) mstyle color #ff0000 (⊤\sub\i)" with precedence 19
+for @{'Top_intro_ko_2}.
+interpretation "Top_intro_ko_2" 'Top_intro_ko_2 =
+ (cast ? ? (show ? (cast ? ? Top_intro))).
+
+notation < "maction (\infrule \nbsp ⊤ (⊤\sub\i)) (\vdots)" with precedence 19
+for @{'Top_intro_ok_1}.
+interpretation "Top_intro_ok_1" 'Top_intro_ok_1 = (show ? Top_intro).
+
+notation < "maction (\infrule \nbsp ⊤ (⊤\sub\i)) (\vdots)" with precedence 19
+for @{'Top_intro_ok_2 }.
+interpretation "Top_intro_ok_2" 'Top_intro_ok_2 = (cast ? ? (show ? Top_intro)).
+
+notation > "⊤#'i'" with precedence 19 for @{ 'Top_intro }.
+interpretation "Top_intro KO" 'Top_intro = (cast ? ? Top_intro).
+interpretation "Top_intro OK" 'Top_intro = Top_intro.
+
+(* ⊥ introduction *)
+notation < "\infrule b a mstyle color #ff0000 (⊥\sub\e)" with precedence 19
+for @{'Bot_elim_ko_1 $a $b}.
+interpretation "Bot_elim_ko_1" 'Bot_elim_ko_1 a b =
+ (show a (Bot_elim ? (cast ? ? b))).
+
+notation < "\infrule b mstyle color #ff0000 (a) mstyle color #ff0000 (⊥\sub\e)" with precedence 19
+for @{'Bot_elim_ko_2 $a $b}.
+interpretation "Bot_elim_ko_2" 'Bot_elim_ko_2 a b =
+ (cast ? ? (show a (Bot_elim ? (cast ? ? b)))).
+
+notation < "maction (\infrule b a (⊥\sub\e)) (\vdots)" with precedence 19
+for @{'Bot_elim_ok_1 $a $b}.
+interpretation "Bot_elim_ok_1" 'Bot_elim_ok_1 a b =
+ (show a (Bot_elim ? b)).
+
+notation < "\infrule b mstyle color #ff0000 (a) (⊥\sub\e)" with precedence 19
+for @{'Bot_elim_ok_2 $a $b}.
+interpretation "Bot_elim_ok_2" 'Bot_elim_ok_2 a b =
+ (cast ? ? (show a (Bot_elim ? b))).
+
+notation > "⊥#'e' term 90 b" with precedence 19
+for @{ 'Bot_elim (show $b ?) }.
+interpretation "Bot_elim KO" 'Bot_elim a = (Bot_elim ? (cast ? ? a)).
+interpretation "Bot_elim OK" 'Bot_elim a = (Bot_elim ? a).
+
+(* ¬ introduction *)
+notation < "\infrule hbox(\emsp b \emsp) ab (mstyle color #ff0000 (\lnot\sub(\emsp\i)) \emsp ident H)" with precedence 19
+for @{ 'Not_intro_ko_1 $ab (λ${ident H}:$p.$b) }.
+interpretation "Not_intro_ko_1" 'Not_intro_ko_1 ab \eta.b =
+ (show ab (cast ? ? (Not_intro ? (cast ? ? b)))).
+
+notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (mstyle color #ff0000 (\lnot\sub(\emsp\i)) \emsp ident H)" with precedence 19
+for @{ 'Not_intro_ko_2 $ab (λ${ident H}:$p.$b) }.
+interpretation "Not_intro_ko_2" 'Not_intro_ko_2 ab \eta.b =
+ (cast ? ? (show ab (cast ? ? (Not_intro ? (cast ? ? b))))).
+
+notation < "maction (\infrule hbox(\emsp b \emsp) ab (\lnot\sub(\emsp\i) \emsp ident H) ) (\vdots)" with precedence 19
+for @{ 'Not_intro_ok_1 $ab (λ${ident H}:$p.$b) }.
+interpretation "Not_intro_ok_1" 'Not_intro_ok_1 ab \eta.b =
+ (show ab (Not_intro ? b)).
+
+notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (\lnot\sub(\emsp\i) \emsp ident H) " with precedence 19
+for @{ 'Not_intro_ok_2 $ab (λ${ident H}:$p.$b) }.
+interpretation "Not_intro_ok_2" 'Not_intro_ok_2 ab \eta.b =
+ (cast ? ? (show ab (Not_intro ? b))).
+
+notation > "¬#'i' [ident H] term 90 b" with precedence 19
+for @{ 'Not_intro (λ${ident H}.show $b ?) }.
+interpretation "Not_intro KO" 'Not_intro a = (cast ? ? (Not_intro ? (cast ? ? a))).
+interpretation "Not_intro OK" 'Not_intro a = (Not_intro ? a).
+
+(* ¬ elimination *)
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b mstyle color #ff0000 (\lnot\sub(\emsp\e)) " with precedence 19
+for @{ 'Not_elim_ko_1 $ab $a $b }.
+interpretation "Not_elim_ko_1" 'Not_elim_ko_1 ab a b =
+ (show b (cast ? ? (Not_elim ? (cast ? ? ab) (cast ? ? a)))).
+
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) mstyle color #ff0000 (\lnot\sub(\emsp\e)) " with precedence 19
+for @{ 'Not_elim_ko_2 $ab $a $b }.
+interpretation "Not_elim_ko_2" 'Not_elim_ko_2 ab a b =
+ (cast ? ? (show b (cast ? ? (Not_elim ? (cast ? ? ab) (cast ? ? a))))).
+
+notation < "maction (\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (\lnot\sub(\emsp\e)) ) (\vdots)" with precedence 19
+for @{ 'Not_elim_ok_1 $ab $a $b }.
+interpretation "Not_elim_ok_1" 'Not_elim_ok_1 ab a b =
+ (show b (Not_elim ? ab a)).
+
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) (\lnot\sub(\emsp\e)) " with precedence 19
+for @{ 'Not_elim_ok_2 $ab $a $b }.
+interpretation "Not_elim_ok_2" 'Not_elim_ok_2 ab a b =
+ (cast ? ? (show b (Not_elim ? ab a))).
+
+notation > "¬#'e' term 90 ab term 90 a" with precedence 19
+for @{ 'Not_elim (show $ab ?) (show $a ?) }.
+interpretation "Not_elim KO" 'Not_elim ab a =
+ (cast ? ? (Not_elim unit (cast ? ? ab) (cast ? ? a))).
+interpretation "Not_elim OK" 'Not_elim ab a =
+ (Not_elim ? ab a).
+
+(* RAA *)
+notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (\RAA) \emsp ident x)" with precedence 19
+for @{ 'RAA_ko_1 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "RAA_ko_1" 'RAA_ko_1 Px Pn =
+ (show Pn (cast ? ? (Raa ? (cast ? ? Px)))).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (mstyle color #ff0000 (\RAA) \emsp ident x)" with precedence 19
+for @{ 'RAA_ko_2 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "RAA_ko_2" 'RAA_ko_2 Px Pn =
+ (cast ? ? (show Pn (cast ? ? (Raa ? (cast ? ? Px))))).
+
+notation < "maction (\infrule hbox(\emsp Px \emsp) Pn (\RAA \emsp ident x)) (\vdots)" with precedence 19
+for @{ 'RAA_ok_1 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "RAA_ok_1" 'RAA_ok_1 Px Pn =
+ (show Pn (Raa ? Px)).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (\RAA \emsp ident x)" with precedence 19
+for @{ 'RAA_ok_2 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "RAA_ok_2" 'RAA_ok_2 Px Pn =
+ (cast ? ? (show Pn (Raa ? Px))).
+
+notation > "'RAA' [ident H] term 90 b" with precedence 19
+for @{ 'Raa (λ${ident H}.show $b ?) }.
+interpretation "RAA KO" 'Raa p = (cast ? unit (Raa ? (cast ? (unit_to ?) p))).
+interpretation "RAA OK" 'Raa p = (Raa ? p).
+
+(* ∃ introduction *)
+notation < "\infrule hbox(\emsp Pn \emsp) Px mstyle color #ff0000 (∃\sub\i)" with precedence 19
+for @{ 'Exists_intro_ko_1 $Pn $Px }.
+interpretation "Exists_intro_ko_1" 'Exists_intro_ko_1 Pn Px =
+ (show Px (cast ? ? (Exists_intro ? ? ? (cast ? ? Pn)))).
+
+notation < "\infrule hbox(\emsp Pn \emsp) mstyle color #ff0000 (Px) mstyle color #ff0000 (∃\sub\i)" with precedence 19
+for @{ 'Exists_intro_ko_2 $Pn $Px }.
+interpretation "Exists_intro_ko_2" 'Exists_intro_ko_2 Pn Px =
+ (cast ? ? (show Px (cast ? ? (Exists_intro ? ? ? (cast ? ? Pn))))).
+
+notation < "maction (\infrule hbox(\emsp Pn \emsp) Px (∃\sub\i)) (\vdots)" with precedence 19
+for @{ 'Exists_intro_ok_1 $Pn $Px }.
+interpretation "Exists_intro_ok_1" 'Exists_intro_ok_1 Pn Px =
+ (show Px (Exists_intro ? ? ? Pn)).
+
+notation < "\infrule hbox(\emsp Pn \emsp) mstyle color #ff0000 (Px) (∃\sub\i)" with precedence 19
+for @{ 'Exists_intro_ok_2 $Pn $Px }.
+interpretation "Exists_intro_ok_2" 'Exists_intro_ok_2 Pn Px =
+ (cast ? ? (show Px (Exists_intro ? ? ? Pn))).
+
+notation >"∃#'i' {term 90 t} term 90 Pt" non associative with precedence 19
+for @{'Exists_intro $t (λw.? w) (show $Pt ?)}.
+interpretation "Exists_intro KO" 'Exists_intro t P Pt =
+ (cast ? ? (Exists_intro sort P t (cast ? ? Pt))).
+interpretation "Exists_intro OK" 'Exists_intro t P Pt =
+ (Exists_intro sort P t Pt).
+
+(* ∃ elimination *)
+notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) c (mstyle color #ff0000 (∃\sub\e) \emsp ident n \emsp ident HPn)" with precedence 19
+for @{ 'Exists_elim_ko_1 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ko_1" 'Exists_elim_ko_1 ExPx Pc c =
+ (show c (cast ? ? (Exists_elim ? ? ? (cast ? ? ExPx) (cast ? ? Pc)))).
+
+notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) mstyle color #ff0000 (c) (mstyle color #ff0000 (∃\sub\e) \emsp ident n \emsp ident HPn)" with precedence 19
+for @{ 'Exists_elim_ko_2 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ko_2" 'Exists_elim_ko_2 ExPx Pc c =
+ (cast ? ? (show c (cast ? ? (Exists_elim ? ? ? (cast ? ? ExPx) (cast ? ? Pc))))).
+
+notation < "maction (\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) c (∃\sub\e \emsp ident n \emsp ident HPn)) (\vdots)" with precedence 19
+for @{ 'Exists_elim_ok_1 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ok_1" 'Exists_elim_ok_1 ExPx Pc c =
+ (show c (Exists_elim ? ? ? ExPx Pc)).
+
+notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) mstyle color #ff0000 (c) (∃\sub\e \emsp ident n \emsp ident HPn)" with precedence 19
+for @{ 'Exists_elim_ok_2 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ok_2" 'Exists_elim_ok_2 ExPx Pc c =
+ (cast ? ? (show c (Exists_elim ? ? ? ExPx Pc))).
+
+definition ex_concl := λx:sort → Prop.Πy:sort.unit → x y.
+definition ex_concl_dummy := Πy:sort.unit → unit.
+definition fake_pred := λx:sort.unit.
+
+notation >"∃#'e' term 90 ExPt {ident t} [ident H] term 90 c" non associative with precedence 19
+for @{'Exists_elim (λx.? x) (show $ExPt ?) (λ${ident t}:sort.λ${ident H}.show $c ?)}.
+interpretation "Exists_elim KO" 'Exists_elim P ExPt c =
+ (cast ? ? (Exists_elim sort P ?
+ (cast (Exists ? P) ? ExPt)
+ (cast ex_concl_dummy (ex_concl ?) c))).
+interpretation "Exists_elim OK" 'Exists_elim P ExPt c =
+ (Exists_elim sort P ? ExPt c).
+
+(* ∀ introduction *)
+
+notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (∀\sub\i) \emsp ident x)" with precedence 19
+for @{ 'Forall_intro_ko_1 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ko_1" 'Forall_intro_ko_1 Px Pn =
+ (show Pn (cast ? ? (Forall_intro ? ? (cast ? ? Px)))).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000(Pn) (mstyle color #ff0000 (∀\sub\i) \emsp ident x)" with precedence 19
+for @{ 'Forall_intro_ko_2 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ko_2" 'Forall_intro_ko_2 Px Pn =
+ (cast ? ? (show Pn (cast ? ? (Forall_intro ? ? (cast ? ? Px))))).
+
+notation < "maction (\infrule hbox(\emsp Px \emsp) Pn (∀\sub\i \emsp ident x)) (\vdots)" with precedence 19
+for @{ 'Forall_intro_ok_1 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ok_1" 'Forall_intro_ok_1 Px Pn =
+ (show Pn (Forall_intro ? ? Px)).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (∀\sub\i \emsp ident x)" with precedence 19
+for @{ 'Forall_intro_ok_2 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ok_2" 'Forall_intro_ok_2 Px Pn =
+ (cast ? ? (show Pn (Forall_intro ? ? Px))).
+
+notation > "∀#'i' {ident y} term 90 Px" non associative with precedence 19
+for @{ 'Forall_intro (λ_.?) (λ${ident y}.show $Px ?) }.
+interpretation "Forall_intro KO" 'Forall_intro P Px =
+ (cast ? ? (Forall_intro sort P (cast ? ? Px))).
+interpretation "Forall_intro OK" 'Forall_intro P Px =
+ (Forall_intro sort P Px).
+
+(* ∀ elimination *)
+notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (∀\sub\e))" with precedence 19
+for @{ 'Forall_elim_ko_1 $Px $Pn }.
+interpretation "Forall_elim_ko_1" 'Forall_elim_ko_1 Px Pn =
+ (show Pn (cast ? ? (Forall_elim ? ? ? (cast ? ? Px)))).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000(Pn) (mstyle color #ff0000 (∀\sub\e))" with precedence 19
+for @{ 'Forall_elim_ko_2 $Px $Pn }.
+interpretation "Forall_elim_ko_2" 'Forall_elim_ko_2 Px Pn =
+ (cast ? ? (show Pn (cast ? ? (Forall_elim ? ? ? (cast ? ? Px))))).
+
+notation < "maction (\infrule hbox(\emsp Px \emsp) Pn (∀\sub\e)) (\vdots)" with precedence 19
+for @{ 'Forall_elim_ok_1 $Px $Pn }.
+interpretation "Forall_elim_ok_1" 'Forall_elim_ok_1 Px Pn =
+ (show Pn (Forall_elim ? ? ? Px)).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (∀\sub\e)" with precedence 19
+for @{ 'Forall_elim_ok_2 $Px $Pn }.
+interpretation "Forall_elim_ok_2" 'Forall_elim_ok_2 Px Pn =
+ (cast ? ? (show Pn (Forall_elim ? ? ? Px))).
+
+notation > "∀#'e' {term 90 t} term 90 Pn" non associative with precedence 19
+for @{ 'Forall_elim (λ_.?) $t (show $Pn ?) }.
+interpretation "Forall_elim KO" 'Forall_elim P t Px =
+ (cast ? unit (Forall_elim sort P t (cast ? ? Px))).
+interpretation "Forall_elim OK" 'Forall_elim P t Px =
+ (Forall_elim sort P t Px).
+
+(* already proved lemma *)
+definition hide_args : ΠA:Type[0].A→A := λA:Type[0].λa:A.a.
+notation < "t" non associative with precedence 90 for @{'hide_args $t}.
+interpretation "hide 0 args" 'hide_args t = (hide_args ? t).
+interpretation "hide 1 args" 'hide_args t = (hide_args ? t ?).
+interpretation "hide 2 args" 'hide_args t = (hide_args ? t ? ?).
+interpretation "hide 3 args" 'hide_args t = (hide_args ? t ? ? ?).
+interpretation "hide 4 args" 'hide_args t = (hide_args ? t ? ? ? ?).
+interpretation "hide 5 args" 'hide_args t = (hide_args ? t ? ? ? ? ?).
+interpretation "hide 6 args" 'hide_args t = (hide_args ? t ? ? ? ? ? ?).
+interpretation "hide 7 args" 'hide_args t = (hide_args ? t ? ? ? ? ? ? ?).
+
+(* more args crashes the pattern matcher *)
+
+(* already proved lemma, 0 assumptions *)
+definition Lemma : ΠA.A→A ≝ λA:Prop.λa:A.a.
+
+notation < "\infrule
+ (\infrule
+ (\emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma_ko_1 $p ($H : $_) }.
+interpretation "lemma_ko_1" 'lemma_ko_1 p H =
+ (show p (cast ? ? (Lemma ? (cast ? ? H)))).
+
+notation < "\infrule
+ (\infrule
+ (\emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma_ko_2 $p ($H : $_) }.
+interpretation "lemma_ko_2" 'lemma_ko_2 p H =
+ (cast ? ? (show p (cast ? ?
+ (Lemma ? (cast ? ? H))))).
+
+
+notation < "\infrule
+ (\infrule
+ (\emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma_ok_1 $p ($H : $_) }.
+interpretation "lemma_ok_1" 'lemma_ok_1 p H =
+ (show p (Lemma ? H)).
+
+notation < "\infrule
+ (\infrule
+ (\emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma_ok_2 $p ($H : $_) }.
+interpretation "lemma_ok_2" 'lemma_ok_2 p H =
+ (cast ? ? (show p (Lemma ? H))).
+
+notation > "'lem' 0 term 90 l" non associative with precedence 19
+for @{ 'Lemma (hide_args ? $l : ?) }.
+interpretation "lemma KO" 'Lemma l =
+ (cast ? ? (Lemma unit (cast unit ? l))).
+interpretation "lemma OK" 'Lemma l = (Lemma ? l).
+
+
+(* already proved lemma, 1 assumption *)
+definition Lemma1 : ΠA,B. (A ⇒ B) → A → B ≝
+ λA,B:Prop.λf:A⇒B.λa:A.
+ Imply_elim A B f a.
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma1_ko_1 $a $p ($H : $_) }.
+interpretation "lemma1_ko_1" 'lemma1_ko_1 a p H =
+ (show p (cast ? ? (Lemma1 ? ? (cast ? ? H) (cast ? ? a)))).
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma1_ko_2 $a $p ($H : $_) }.
+interpretation "lemma1_ko_2" 'lemma1_ko_2 a p H =
+ (cast ? ? (show p (cast ? ?
+ (Lemma1 ? ? (cast ? ? H) (cast ? ? a))))).
+
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma1_ok_1 $a $p ($H : $_) }.
+interpretation "lemma1_ok_1" 'lemma1_ok_1 a p H =
+ (show p (Lemma1 ? ? H a)).
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma1_ok_2 $a $p ($H : $_) }.
+interpretation "lemma1_ok_2" 'lemma1_ok_2 a p H =
+ (cast ? ? (show p (Lemma1 ? ? H a))).
+
+
+notation > "'lem' 1 term 90 l term 90 p" non associative with precedence 19
+for @{ 'Lemma1 (hide_args ? $l : ?) (show $p ?) }.
+interpretation "lemma 1 KO" 'Lemma1 l p =
+ (cast ? ? (Lemma1 unit unit (cast (Imply unit unit) ? l) (cast unit ? p))).
+interpretation "lemma 1 OK" 'Lemma1 l p = (Lemma1 ? ? l p).
+
+(* already proved lemma, 2 assumptions *)
+definition Lemma2 : ΠA,B,C. (A ⇒ B ⇒ C) → A → B → C ≝
+ λA,B,C:Prop.λf:A⇒B⇒C.λa:A.λb:B.
+ Imply_elim B C (Imply_elim A (B⇒C) f a) b.
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma2_ko_1 $a $b $p ($H : $_) }.
+interpretation "lemma2_ko_1" 'lemma2_ko_1 a b p H =
+ (show p (cast ? ? (Lemma2 ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b)))).
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma2_ko_2 $a $b $p ($H : $_) }.
+interpretation "lemma2_ko_2" 'lemma2_ko_2 a b p H =
+ (cast ? ? (show p (cast ? ?
+ (Lemma2 ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b))))).
+
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma2_ok_1 $a $b $p ($H : $_) }.
+interpretation "lemma2_ok_1" 'lemma2_ok_1 a b p H =
+ (show p (Lemma2 ? ? ? H a b)).
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma2_ok_2 $a $b $p ($H : $_) }.
+interpretation "lemma2_ok_2" 'lemma2_ok_2 a b p H =
+ (cast ? ? (show p (Lemma2 ? ? ? H a b))).
+
+notation > "'lem' 2 term 90 l term 90 p term 90 q" non associative with precedence 19
+for @{ 'Lemma2 (hide_args ? $l : ?) (show $p ?) (show $q ?) }.
+interpretation "lemma 2 KO" 'Lemma2 l p q =
+ (cast ? ? (Lemma2 unit unit unit (cast (Imply unit (Imply unit unit)) ? l) (cast unit ? p) (cast unit ? q))).
+interpretation "lemma 2 OK" 'Lemma2 l p q = (Lemma2 ? ? ? l p q).
+
+(* already proved lemma, 3 assumptions *)
+definition Lemma3 : ΠA,B,C,D. (A ⇒ B ⇒ C ⇒ D) → A → B → C → D ≝
+ λA,B,C,D:Prop.λf:A⇒B⇒C⇒D.λa:A.λb:B.λc:C.
+ Imply_elim C D (Imply_elim B (C⇒D) (Imply_elim A (B⇒C⇒D) f a) b) c.
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp\emsp\emsp c \emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma3_ko_1 $a $b $c $p ($H : $_) }.
+interpretation "lemma3_ko_1" 'lemma3_ko_1 a b c p H =
+ (show p (cast ? ?
+ (Lemma3 ? ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b) (cast ? ? c)))).
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp\emsp\emsp c \emsp)
+ (╲ mstyle mathsize normal (mstyle color #ff0000 (H)) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma3_ko_2 $a $b $c $p ($H : $_) }.
+interpretation "lemma3_ko_2" 'lemma3_ko_2 a b c p H =
+ (cast ? ? (show p (cast ? ?
+ (Lemma3 ? ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b) (cast ? ? c))))).
+
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp\emsp\emsp c \emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ p \nbsp"
+non associative with precedence 19
+for @{ 'lemma3_ok_1 $a $b $c $p ($H : $_) }.
+interpretation "lemma3_ok_1" 'lemma3_ok_1 a b c p H =
+ (show p (Lemma3 ? ? ? ? H a b c)).
+
+notation < "\infrule
+ (\infrule
+ (\emsp a \emsp\emsp\emsp b \emsp\emsp\emsp c \emsp)
+ (╲ mstyle mathsize normal (H) ╱) \nbsp)
+ mstyle color #ff0000 (p) \nbsp"
+non associative with precedence 19
+for @{ 'lemma3_ok_2 $a $b $c $p ($H : $_) }.
+interpretation "lemma3_ok_2" 'lemma3_ok_2 a b c p H =
+ (cast ? ? (show p (Lemma3 ? ? ? ? H a b c))).
+
+notation > "'lem' 3 term 90 l term 90 p term 90 q term 90 r" non associative with precedence 19
+for @{ 'Lemma3 (hide_args ? $l : ?) (show $p ?) (show $q ?) (show $r ?) }.
+interpretation "lemma 3 KO" 'Lemma3 l p q r =
+ (cast ? ? (Lemma3 unit unit unit unit (cast (Imply unit (Imply unit (Imply unit unit))) ? l) (cast unit ? p) (cast unit ? q) (cast unit ? r))).
+interpretation "lemma 3 OK" 'Lemma3 l p q r = (Lemma3 ? ? ? ? l p q r).
--- /dev/null
+include "reverse_complexity/speedup.ma".
+
+definition aes : (nat → nat) → (nat → nat) → Prop ≝ λf,g.
+ ∃n.∀m. n ≤ m → f m = g m.
+
+lemma CF_almost: ∀f,g,s. CF s g → aes g f → CF s f.
+#f #g #s #CFg * #n lapply CFg -CFg lapply g -g
+elim n
+ [#g #CFg #H @(ext_CF … g) [#m @H // |//]
+ |#i #Hind #g #CFg #H
+ @(Hind (λx. if eqb i x then f i else g x))
+ [@CF_if
+ [@(O_to_CF … MSC) [@le_to_O @(MSC_in … CFg)] @CF_eqb //
+ |@(O_to_CF … MSC) [@le_to_O @(MSC_in … CFg)] @CF_const
+ |@CFg
+ ]
+ |#m #leim cases (le_to_or_lt_eq … leim)
+ [#ltim lapply (lt_to_not_eq … ltim) #noteqim
+ >(not_eq_to_eqb_false … noteqim) @H @ltim
+ |#eqim >eqim >eqb_n_n //
+ ]
+ ]
+ ]
+qed.
\ No newline at end of file
--- /dev/null
+include "basics/types.ma".
+include "arithmetics/nat.ma".
+
+(********************************** pairing ***********************************)
+axiom pair: nat → nat → nat.
+axiom fst : nat → nat.
+axiom snd : nat → nat.
+
+interpretation "abstract pair" 'pair f g = (pair f g).
+
+axiom fst_pair: ∀a,b. fst 〈a,b〉 = a.
+axiom snd_pair: ∀a,b. snd 〈a,b〉 = b.
+axiom surj_pair: ∀x. ∃a,b. x = 〈a,b〉.
+
+axiom le_fst : ∀p. fst p ≤ p.
+axiom le_snd : ∀p. snd p ≤ p.
+axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉.
+
+lemma expand_pair: ∀x. x = 〈fst x, snd x〉.
+#x lapply (surj_pair x) * #a * #b #Hx >Hx >fst_pair >snd_pair //
+qed.
+
+(************************************* U **************************************)
+axiom U: nat → nat →nat → option nat.
+
+axiom monotonic_U: ∀i,x,n,m,y.n ≤m →
+ U i x n = Some ? y → U i x m = Some ? y.
+
+lemma unique_U: ∀i,x,n,m,yn,ym.
+ U i x n = Some ? yn → U i x m = Some ? ym → yn = ym.
+#i #x #n #m #yn #ym #Hn #Hm cases (decidable_le n m)
+ [#lenm lapply (monotonic_U … lenm Hn) >Hm #HS destruct (HS) //
+ |#ltmn lapply (monotonic_U … n … Hm) [@lt_to_le @not_le_to_lt //]
+ >Hn #HS destruct (HS) //
+ ]
+qed.
+
+definition code_for ≝ λf,i.∀x.
+ ∃n.∀m. n ≤ m → U i x m = f x.
+
+definition terminate ≝ λi,x,r. ∃y. U i x r = Some ? y.
+
+notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}.
+
+lemma terminate_dec: ∀i,x,n. {i ⊙ x} ↓ n ∨ ¬ {i ⊙ x} ↓ n.
+#i #x #n normalize cases (U i x n)
+ [%2 % * #y #H destruct|#y %1 %{y} //]
+qed.
+
+lemma monotonic_terminate: ∀i,x,n,m.
+ n ≤ m → {i ⊙ x} ↓ n → {i ⊙ x} ↓ m.
+#i #x #n #m #lenm * #z #H %{z} @(monotonic_U … H) //
+qed.
+
+definition termb ≝ λi,x,t.
+ match U i x t with [None ⇒ false |Some y ⇒ true].
+
+lemma termb_true_to_term: ∀i,x,t. termb i x t = true → {i ⊙ x} ↓ t.
+#i #x #t normalize cases (U i x t) normalize [#H destruct | #y #_ %{y} //]
+qed.
+
+lemma term_to_termb_true: ∀i,x,t. {i ⊙ x} ↓ t → termb i x t = true.
+#i #x #t * #y #H normalize >H //
+qed.
+
+definition out ≝ λi,x,r.
+ match U i x r with [ None ⇒ 0 | Some z ⇒ z].
+
+definition bool_to_nat: bool → nat ≝
+ λb. match b with [true ⇒ 1 | false ⇒ 0].
+
+coercion bool_to_nat.
+
+definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉.
+
+lemma pU_vs_U_Some : ∀i,x,r,y. pU i x r = 〈1,y〉 ↔ U i x r = Some ? y.
+#i #x #r #y % normalize
+ [cases (U i x r) normalize
+ [#H cut (0=1) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H @H]
+ #H1 destruct
+ |#a #H cut (a=y) [lapply (eq_f … snd … H) >snd_pair >snd_pair #H1 @H1]
+ #H1 //
+ ]
+ |#H >H //]
+qed.
+
+lemma pU_vs_U_None : ∀i,x,r. pU i x r = 〈0,0〉 ↔ U i x r = None ?.
+#i #x #r % normalize
+ [cases (U i x r) normalize //
+ #a #H cut (1=0) [lapply (eq_f … fst … H) >fst_pair >fst_pair #H1 @H1]
+ #H1 destruct
+ |#H >H //]
+qed.
+
+lemma fst_pU: ∀i,x,r. fst (pU i x r) = termb i x r.
+#i #x #r normalize cases (U i x r) normalize >fst_pair //
+qed.
+
+lemma snd_pU: ∀i,x,r. snd (pU i x r) = out i x r.
+#i #x #r normalize cases (U i x r) normalize >snd_pair //
+qed.
+
+
+definition total ≝ λf.λx:nat. Some nat (f x).
+
+
--- /dev/null
+include "reverse_complexity/complexity.ma".
+include "reverse_complexity/sigma_diseq.ma".
+
+include alias "reverse_complexity/basics.ma".
+
+lemma bigop_prim_rec: ∀a,b,c,p,f,x.
+ bigop (b x-a x) (λi:ℕ.p 〈i+a x,x〉) ? (c 〈a x,x〉) plus (λi:ℕ.f 〈i+a x,x〉) =
+ prim_rec c
+ (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉
+ then plus (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (b x-a x) 〈a x ,x〉.
+#a #b #c #p #f #x normalize elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma bigop_prim_rec_dec: ∀a,b,c,p,f,x.
+ bigop (b x-a x) (λi:ℕ.p 〈b x -i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x-i,x〉) =
+ prim_rec c
+ (λi.if p 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉
+ then plus (f 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (b x-a x) 〈b x ,x〉.
+#a #b #c #p #f #x normalize elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma bigop_prim_rec_dec1: ∀a,b,c,p,f,x.
+ bigop (S(b x)-a x) (λi:ℕ.p 〈b x - i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x- i,x〉) =
+ prim_rec c
+ (λi.if p 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉
+ then plus (f 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (S(b x)-a x) 〈b x,x〉.
+#a #b #c #p #f #x elim (S(b x)-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma sum_prim_rec1: ∀a,b,p,f,x.
+ ∑_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉
+ then f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉 + fst (snd i)
+ else fst (snd i)) (b x-a x) 〈a x ,x〉.
+#a #b #p #f #x elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma bigop_plus_c: ∀k,p,f,c.
+ c k + bigop k (λi.p i) ? 0 plus (λi.f i) =
+ bigop k (λi.p i) ? (c k) plus (λi.f i).
+#k #p #f elim k [normalize //]
+#i #Hind #c cases (true_or_false (p i)) #Hcase
+[>bigop_Strue // >bigop_Strue // <associative_plus >(commutative_plus ? (f i))
+ >associative_plus @eq_f @Hind
+|>bigop_Sfalse // >bigop_Sfalse //
+]
+qed.
+
+
+(*********************************** maximum **********************************)
+
+lemma max_gen: ∀a,b,p,f,x. a ≤b →
+ max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) = max_{i < b | leb a i ∧ p 〈i,x〉 }(f 〈i,x〉).
+#a #b #p #f #x @(bigop_I_gen ????? MaxA)
+qed.
+
+lemma max_prim_rec_base: ∀a,b,p,f,x. a ≤b →
+ max_{i < b| p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x.
+#a #b #p #f #x #leab >max_gen // elim b
+ [normalize //
+ |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma max_prim_rec: ∀a,b,p,f,x. a ≤b →
+ max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if leb a (fst i) ∧ p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x.
+#a #b #p #f #x #leab >max_gen // elim b
+ [normalize //
+ |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair
+ cases (true_or_false (leb a i ∧ p 〈i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma max_prim_rec1: ∀a,b,p,f,x.
+ max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉
+ then max (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (b x-a x) 〈a x ,x〉.
+#a #b #p #f #x elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+(* the argument is 〈b-a,〈a,x〉〉 *)
+
+definition max_unary_pr ≝ λp,f.unary_pr (λx.0)
+ (λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let a ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ if p 〈k + a,x〉 then max (f 〈k+a,x〉) r else r).
+
+lemma max_unary_pr1: ∀a,b,p,f,x.
+ max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) =
+ ((max_unary_pr p f) ∘ (λx.〈b x - a x,〈a x,x〉〉)) x.
+#a #b #p #f #x normalize >fst_pair >snd_pair @max_prim_rec1
+qed.
+
+definition aux_compl ≝ λhp,hf.λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let a ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hp 〈k+a,x〉 + hf 〈k+a,x〉 + (* was MSC r*) MSC i .
+
+definition aux_compl1 ≝ λhp,hf.λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let a ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r.
+
+lemma aux_compl1_def: ∀k,r,m,hp,hf.
+ aux_compl1 hp hf 〈k,〈r,m〉〉 =
+ let a ≝ fst m in
+ let x ≝ snd m in
+ hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r.
+#k #r #m #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair //
+qed.
+
+lemma aux_compl1_def1: ∀k,r,a,x,hp,hf.
+ aux_compl1 hp hf 〈k,〈r,〈a,x〉〉〉 = hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r.
+#k #r #a #x #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair
+>fst_pair >snd_pair //
+qed.
+
+axiom Oaux_compl: ∀hp,hf. O (aux_compl1 hp hf) (aux_compl hp hf).
+
+lemma MSC_max: ∀f,h,x. CF h f → MSC (max_{i < x}(f i)) ≤ max_{i < x}(h i).
+#f #h #x #hf elim x // #i #Hind >bigop_Strue [|//] >bigop_Strue [|//]
+whd in match (max ??);
+cases (true_or_false (leb (f i) (bigop i (λi0:ℕ.true) ? 0 max(λi0:ℕ.f i0))))
+#Hcase >Hcase
+ [@(transitive_le … Hind) @le_maxr //
+ |@(transitive_le … (MSC_out … hf i)) @le_maxl //
+ ]
+qed.
+
+lemma CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s.
+ CF ha a → CF hb b → CF hp p → CF hf f →
+ O s (λx.ha x + hb x +
+ (∑_{i ∈[a x ,b x[ }
+ (hp 〈i,x〉 + hf 〈i,x〉 + max_{i ∈ [a x, b x [ }(hf 〈i,x〉)))) →
+ CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)).
+#a #b #p #f #ha #hb #hp #hf #s #CFa #CFb #CFp #CFf #HO
+@ext_CF1 [|#x @max_unary_pr1]
+@(CF_comp ??? (λx.ha x + hb x))
+ [|@CF_comp_pair
+ [@CF_minus [@(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //]
+ |@CF_comp_pair
+ [@(O_to_CF … CFa) @O_plus_l //
+ | @(O_to_CF … CF_id) @O_plus_r @le_to_O @(MSC_in … CFb)
+ ]
+ ]
+ |@(CF_prim_rec … MSC … (aux_compl1 hp hf))
+ [@CF_const
+ |@(O_to_CF … (Oaux_compl … ))
+ @CF_if
+ [@(CF_comp p … MSC … CFp)
+ [@CF_comp_pair
+ [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd]
+ |@CF_comp_snd @CF_comp_snd @CF_snd]
+ |@le_to_O #x normalize >commutative_plus >associative_plus @le_plus //
+ ]
+ |@CF_max1
+ [@(CF_comp f … MSC … CFf)
+ [@CF_comp_pair
+ [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd]
+ |@CF_comp_snd @CF_comp_snd @CF_snd]
+ |@le_to_O #x normalize >commutative_plus //
+ ]
+ |@CF_comp_fst @(monotonic_CF … CF_snd) normalize //
+ ]
+ |@CF_comp_fst @(monotonic_CF … CF_snd) normalize //
+ ]
+ |@O_refl
+ ]
+ |@(O_trans … HO) -HO
+ @(O_trans ? (λx:ℕ
+ .ha x+hb x
+ +bigop (b x-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus
+ (λi:ℕ
+ .(λi0:ℕ
+ .hp 〈i0,x〉+hf 〈i0,x〉
+ +bigop (b x-a x) (λi1:ℕ.true) ? 0 max
+ (λi1:ℕ.(λi2:ℕ.hf 〈i2,x〉) (i1+a x))) (i+a x))))
+ [
+ @le_to_O #n @le_plus // whd in match (unary_pr ???);
+ >fst_pair >snd_pair >bigop_prim_rec elim (b n - a n)
+ [normalize //
+ |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >aux_compl1_def1
+ >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus
+ [-Hind @le_plus // normalize >fst_pair >snd_pair
+ @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 )
+ (λi1:ℕ.hf 〈i1+a n,n〉)))
+ [elim x [normalize @MSC_le]
+ #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair
+ >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize
+ [cases (true_or_false (leb (f 〈x0+a n,n〉)
+ (prim_rec (λx00:ℕ.O)
+ (λi:ℕ
+ .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉)
+ (fst (snd i))
+ then fst (snd i)
+ else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize
+ [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ |@(transitive_le … (MSC_out … CFf …)) @(le_maxl … (le_n …))
+ ]
+ |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ ]
+ |@(le_maxr … (le_n …))
+ ]
+ |@(transitive_le … Hind) -Hind
+ generalize in match (bigop x (λi:ℕ.true) ? 0 max
+ (λi1:ℕ.(λi2:ℕ.hf 〈i2,n〉) (i1+a n))); #c
+ generalize in match (hf 〈x+a n,n〉); #c1
+ elim x [//] #x0 #Hind
+ >prim_rec_S >prim_rec_S normalize >fst_pair >fst_pair >snd_pair
+ >snd_pair >snd_pair >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair
+ >fst_pair @le_plus
+ [@le_plus // cases (true_or_false (leb c1 c)) #Hcase
+ >Hcase normalize // @lt_to_le @not_le_to_lt @(leb_false_to_not_le ?? Hcase)
+ |@Hind
+ ]
+ ]
+ ]
+ |@O_plus [@O_plus_l //] @le_to_O #x
+ <bigop_plus_c @le_plus // @(transitive_le … (MSC_pair …)) @le_plus
+ [@MSC_out @CFa | @MSC_out @(O_to_CF MSC … (CF_const x)) @le_to_O @(MSC_in … CFb)]
+ ]
+qed.
+
+axiom daemon : ∀P:Prop.P.
+axiom O_extl: ∀s1,s2,f. (∀x.s1 x = s2 x) → O s1 f → O s2 f.
+
+lemma CF_max2: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s.
+ CF ha a → CF hb b → CF hp p → CF hf f →
+ O s (λx.ha x + hb x +
+ (b x - a x)*max_{i ∈ [a x, b x [ }(hp 〈i,x〉 + hf 〈i,x〉)) →
+ CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)).
+#a #b #p #f #ha #hb #hp #hf #s #CFa #CFb #CFp #CFf #Os
+@(O_to_CF … Os (CF_max … CFa CFb CFp CFf ?)) @O_plus
+ [@O_plus_l @O_refl
+ |@O_plus_r @O_ext2 [|| #x @(bigop_op … plusAC)]
+ @O_plus
+ [@le_to_O normalize #x @sigma_to_Max
+ |@le_to_O #x @transitive_le [|@sigma_const] @le_times //
+ @Max_le #i #lti #_ @(transitive_le ???? (le_MaxI … ))
+ [@le_plus_n | // | @lt_minus_to_plus_r // | //]
+ ]
+ ]
+qed.
+
+(*
+lemma CF_max_monotonic: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s.
+ CF ha a → CF hb b → CF hp p → CF hf f →
+ O s (λx.ha x + hb x +
+ (b x - a x)*max_{i ∈ [a x, b x [ }(hp 〈i,x〉 + hf 〈i,x〉)) →
+ CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)).
+#a #b #p #f #ha #hb #hp #hf #s #CFa #CFb #CFp #CFf #Os
+@(O_to_CF … Os (CF_max … CFa CFb CFp CFf ?)) @O_plus
+ [@O_plus_l @O_refl
+ |@O_plus_r @O_ext2 [|| #x @(bigop_op … plusAC)]
+ @O_plus
+ [@le_to_O normalize #x @sigma_to_Max
+ |@le_to_O #x @transitive_le [|@sigma_const] @le_times //
+ @Max_le #i #lti #_ @(transitive_le ???? (le_MaxI … ))
+ [@le_plus_n | // | @lt_minus_to_plus_r // | //]
+ ]
+ ]
+qed.
+*)
+
+(* old
+axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s.
+ CF ha a → CF hb b → CF hp p → CF hf f →
+ O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) →
+ CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). *)
+
+(******************************** minimization ********************************)
+
+let rec my_minim a f x k on k ≝
+ match k with
+ [O ⇒ a
+ |S p ⇒ if eqb (my_minim a f x p) (a+p)
+ then if f 〈a+p,x〉 then a+p else S(a+p)
+ else (my_minim a f x p) ].
+
+lemma my_minim_S : ∀a,f,x,k.
+ my_minim a f x (S k) =
+ if eqb (my_minim a f x k) (a+k)
+ then if f 〈a+k,x〉 then a+k else S(a+k)
+ else (my_minim a f x k) .
+// qed.
+
+lemma my_minim_fa : ∀a,f,x,k. f〈a,x〉 = true → my_minim a f x k = a.
+#a #f #x #k #H elim k // #i #Hind normalize >Hind
+cases (true_or_false (eqb a (a+i))) #Hcase >Hcase normalize //
+<(eqb_true_to_eq … Hcase) >H //
+qed.
+
+lemma my_minim_nfa : ∀a,f,x,k. f〈a,x〉 = false →
+my_minim a f x (S k) = my_minim (S a) f x k.
+#a #f #x #k #H elim k
+ [normalize <plus_n_O >H >eq_to_eqb_true //
+ |#i #Hind >my_minim_S >Hind >my_minim_S <plus_n_Sm //
+ ]
+qed.
+
+lemma my_min_eq : ∀a,f,x,k.
+ min k a (λi.f 〈i,x〉) = my_minim a f x k.
+#a #f #x #k lapply a -a elim k // #i #Hind #a normalize in ⊢ (??%?);
+cases (true_or_false (f 〈a,x〉)) #Hcase >Hcase
+ [>(my_minim_fa … Hcase) // | >Hind @sym_eq @(my_minim_nfa … Hcase) ]
+qed.
+
+(* cambiare qui: togliere S *)
+
+
+definition minim_unary_pr ≝ λf.unary_pr (λx.S(fst x))
+ (λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let b ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ if f 〈b-k,x〉 then b-k else r).
+
+definition compl_minim_unary ≝ λhf:nat → nat.λi.
+ let k ≝ fst i in
+ let b ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉.
+
+definition compl_minim_unary_aux ≝ λhf,i.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let b ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hf 〈b-k,x〉 + MSC i.
+
+lemma compl_minim_unary_aux_def : ∀hf,k,r,b,x.
+ compl_minim_unary_aux hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈r,〈b,x〉〉〉.
+#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair //
+qed.
+
+lemma compl_minim_unary_def : ∀hf,k,r,b,x.
+ compl_minim_unary hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉.
+#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair //
+qed.
+
+lemma compl_minim_unary_aux_def2 : ∀hf,k,r,x.
+ compl_minim_unary_aux hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈r,x〉〉.
+#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair //
+qed.
+
+lemma compl_minim_unary_def2 : ∀hf,k,r,x.
+ compl_minim_unary hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈S(fst x),snd x〉〉.
+#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair //
+qed.
+
+lemma min_aux: ∀a,f,x,k. min (S k) (a x) (λi:ℕ.f 〈i,x〉) =
+ ((minim_unary_pr f) ∘ (λx.〈S k,〈a x + k,x〉〉)) x.
+#a #f #x #k whd in ⊢ (???%); >fst_pair >snd_pair
+lapply a -a (* @max_prim_rec1 *)
+elim k
+ [normalize #a >fst_pair >snd_pair >fst_pair >snd_pair >snd_pair >fst_pair
+ <plus_n_O <minus_n_O >fst_pair //
+ |#i #Hind #a normalize in ⊢ (??%?); >prim_rec_S >fst_pair >snd_pair
+ >fst_pair >snd_pair >snd_pair >fst_pair <plus_n_Sm <(Hind (λx.S (a x)))
+ whd in ⊢ (???%); >minus_S_S <(minus_plus_m_m (a x) i) //
+qed.
+
+lemma minim_unary_pr1: ∀a,b,f,x.
+ μ_{i ∈[a x,b x]}(f 〈i,x〉) =
+ if leb (a x) (b x) then ((minim_unary_pr f) ∘ (λx.〈S (b x) - a x,〈b x,x〉〉)) x
+ else (a x).
+#a #b #f #x cases (true_or_false (leb (a x) (b x))) #Hcase >Hcase
+ [cut (b x = a x + (b x - a x))
+ [>commutative_plus <plus_minus_m_m // @leb_true_to_le // ]
+ #Hcut whd in ⊢ (???%); >minus_Sn_m [|@leb_true_to_le //]
+ >min_aux whd in ⊢ (??%?); <Hcut //
+ |>eq_minus_O // @not_le_to_lt @leb_false_to_not_le //
+ ]
+qed.
+
+axiom sum_inv: ∀a,b,f.∑_{i ∈ [a,S b[ }(f i) = ∑_{i ∈ [a,S b[ }(f (a + b - i)).
+
+(*
+#a #b #f @(bigop_iso … plusAC) whd %{(λi.b - a - i)} %{(λi.b - a -i)} %
+ [%[#i #lti #_ normalize @eq_f >commutative_plus <plus_minus_associative
+ [2: @le_minus_to_plus_r //
+ [// @eq_f @@sym_eq @plus_to_minus
+ |#i #Hi #_ % [% [@le_S_S
+*)
+
+example sum_inv_check : ∑_{i ∈ [3,6[ }(i*i) = ∑_{i ∈ [3,6[ }((8-i)*(8-i)).
+normalize // qed.
+
+(* rovesciamo la somma *)
+
+lemma CF_mu: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s.
+ CF sa a → CF sb b → CF sf f →
+ O s (λx.sa x + sb x + ∑_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉 + MSC 〈b x - i,〈S(b x),x〉〉)) →
+ CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)).
+#a #b #f #ha #hb #hf #s #CFa #CFb #CFf #HO
+@ext_CF1 [|#x @minim_unary_pr1]
+@CF_if
+ [@CF_le @(O_to_CF … HO)
+ [@(O_to_CF … CFa) @O_plus_l @O_plus_l @O_refl
+ |@(O_to_CF … CFb) @O_plus_l @O_plus_r @O_refl
+ ]
+ |@(CF_comp ??? (λx.ha x + hb x))
+ [|@CF_comp_pair
+ [@CF_minus [@CF_compS @(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //]
+ |@CF_comp_pair
+ [@(O_to_CF … CFb) @O_plus_r //
+ |@(O_to_CF … CF_id) @O_plus_r @le_to_O @(MSC_in … CFb)
+ ]
+ ]
+ |@(CF_prim_rec_gen … MSC … (compl_minim_unary_aux hf))
+ [@((λx:ℕ.fst (snd x)
+ +compl_minim_unary hf
+ 〈fst x,
+ 〈unary_pr fst
+ (λi:ℕ
+ .(let (k:ℕ) ≝fst i in
+ let (r:ℕ) ≝fst (snd i) in
+ let (b:ℕ) ≝fst (snd (snd i)) in
+ let (x:ℕ) ≝snd (snd (snd i)) in if f 〈b-S k,x〉 then b-S k else r ))
+ 〈fst x,snd (snd x)〉,
+ snd (snd x)〉〉))
+ |@CF_compS @CF_fst
+ |@CF_if
+ [@(CF_comp f … MSC … CFf)
+ [@CF_comp_pair
+ [@CF_minus [@CF_comp_fst @CF_comp_snd @CF_snd|@CF_fst]
+ |@CF_comp_snd @CF_comp_snd @CF_snd]
+ |@le_to_O #x normalize >commutative_plus //
+ ]
+ |@(O_to_CF … MSC)
+ [@le_to_O #x normalize //
+ |@CF_minus
+ [@CF_comp_fst @CF_comp_snd @CF_snd |@CF_fst]
+ ]
+ |@CF_comp_fst @(monotonic_CF … CF_snd) normalize //
+ ]
+ |@O_plus
+ [@O_plus_l @O_refl
+ |@O_plus_r @O_ext2 [||#x >compl_minim_unary_aux_def2 @refl]
+ @O_trans [||@le_to_O #x >compl_minim_unary_def2 @le_n]
+ @O_plus [@O_plus_l //]
+ @O_plus_r
+ @O_trans [|@le_to_O #x @MSC_pair] @O_plus
+ [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_fst …))
+ >fst_pair @le_n]
+ @O_trans [|@le_to_O #x @MSC_pair] @O_plus
+ [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …))
+ >snd_pair @(transitive_le ???? (le_fst …)) >fst_pair
+ normalize >snd_pair >fst_pair lapply (surj_pair x)
+ * #x1 * #x2 #Hx >Hx >fst_pair >snd_pair elim x1 //
+ #i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair
+ cases (f ?) [@le_S // | //]]
+ @le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …)) >snd_pair
+ >(expand_pair (snd (snd x))) in ⊢ (?%?); @le_pair //
+ ]
+ ]
+ |cut (O s (λx.ha x + hb x +
+ ∑_{i ∈[a x ,S(b x)[ }(hf 〈a x+b x-i,x〉 + MSC 〈b x -(a x+b x-i),〈S(b x),x〉〉)))
+ [@(O_ext2 … HO) #x @eq_f @sum_inv] -HO #HO
+ @(O_trans … HO) -HO
+ @(O_trans ? (λx:ℕ.ha x+hb x
+ +bigop (S(b x)-a x) (λi:ℕ.true) ?
+ (MSC 〈b x,x〉) plus (λi:ℕ.(λj.hf j + MSC 〈b x - fst j,〈S(b (snd j)),snd j〉〉) 〈b x- i,x〉)))
+ [@le_to_O #n @le_plus // whd in match (unary_pr ???);
+ >fst_pair >snd_pair >(bigop_prim_rec_dec1 a b ? (λi.true))
+ (* it is crucial to recall that the index is bound by S(b x) *)
+ cut (S (b n) - a n ≤ S (b n)) [//]
+ elim (S(b n) - a n)
+ [normalize //
+ |#x #Hind #lex >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair
+ >compl_minim_unary_def >prim_rec_S >fst_pair >snd_pair >fst_pair
+ >snd_pair >fst_pair >snd_pair >fst_pair whd in ⊢ (??%); >commutative_plus
+ @le_plus [2:@Hind @le_S @le_S_S_to_le @lex] -Hind >snd_pair
+ >minus_minus_associative // @le_S_S_to_le //
+ ]
+ |@O_plus [@O_plus_l //] @O_ext2 [||#x @bigop_plus_c]
+ @O_plus
+ [@O_plus_l @O_trans [|@le_to_O #x @MSC_pair]
+ @O_plus
+ [@O_plus_r @le_to_O @(MSC_out … CFb)
+ |@O_plus_r @le_to_O @(MSC_in … CFb)
+ ]
+ |@O_plus_r @(O_ext2 … (O_refl …)) #x @same_bigop
+ [//|#i #H #_ @eq_f2 [@eq_f @eq_f2 //|>fst_pair @eq_f @eq_f2 //]
+ ]
+ ]
+ ]
+ ]
+ |@(O_to_CF … CFa) @(O_trans … HO) @O_plus_l @O_plus_l @O_refl
+ ]
+
+qed.
+
+lemma CF_mu2: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s.
+ CF sa a → CF sb b → CF sf f →
+ O s (λx.sa x + sb x + ∑_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉 + MSC〈S(b x),x〉)) →
+ CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)).
+#a #b #f #sa #sb #sf #s #CFa #CFb #CFf #HO
+@(O_to_CF … HO (CF_mu … CFa CFb CFf ?)) @O_plus [@O_plus_l @O_refl]
+@O_plus_r @O_ext2 [|| #x @(bigop_op … plusAC)]
+@O_plus [@le_to_O #x @le_sigma //]
+@(O_trans ? (λx.∑_{i ∈[a x ,S(b x)[ }(MSC(b x -i)+MSC 〈S(b x),x〉)))
+ [@le_to_O #x @le_sigma //]
+@O_ext2 [|| #x @(bigop_op … plusAC)] @O_plus
+ [@le_to_O #x @le_sigma // #i #lti #_ @(transitive_le … (MSC 〈S (b x),x〉)) //
+ @monotonic_MSC @(transitive_le … (S(b x))) // @le_S //
+ |@le_to_O #x @le_sigma //
+ ]
+qed.
+
+(* uses MSC_S *)
+
+lemma CF_mu3: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s. (∀x.sf x > 0) →
+ CF sa a → CF sb b → CF sf f →
+ O s (λx.sa x + sb x + ∑_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉 + MSC〈b x,x〉)) →
+ CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)).
+#a #b #f #sa #sb #sf #s #sfpos #CFa #CFb #CFf #HO
+@(O_to_CF … HO (CF_mu2 … CFa CFb CFf ?)) @O_plus [@O_plus_l @O_refl]
+@O_plus_r @O_ext2 [|| #x @(bigop_op … plusAC)]
+@O_plus [@le_to_O #x @le_sigma //]
+@le_to_O #x @le_sigma // #x #lti #_ >MSC_pair_eq >MSC_pair_eq <associative_plus
+@le_plus // @(transitive_le … (MSC_sublinear … )) /2 by monotonic_lt_plus_l/
+qed.
+
+lemma CF_mu4: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s. (∀x.sf x > 0) →
+ CF sa a → CF sb b → CF sf f →
+ O s (λx.sa x + sb x + (S(b x) - a x)*Max_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉)) →
+ CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)).
+#a #b #f #sa #sb #sf #s #sfpos #CFa #CFb #CFf #HO
+@(O_to_CF … HO (CF_mu3 … sfpos CFa CFb CFf ?)) @O_plus [@O_plus_l @O_refl]
+@O_ext2 [|| #x @(bigop_op … plusAC)] @O_plus_r @O_plus
+ [@le_to_O #x @sigma_to_Max
+ |lapply (MSC_in … CFf) #Hle
+ %{1} %{0} #n #_ @(transitive_le … (sigma_const …))
+ >(commutative_times 1) <times_n_1
+ cases (decidable_le (S (b n)) (a n)) #H
+ [>(eq_minus_O … H) //
+ |lapply (le_S_S_to_le … (not_le_to_lt … H)) -H #H
+ @le_times // @(transitive_le … (Hle … ))
+ cut (b n = b n - a n + a n) [<plus_minus_m_m // ]
+ #Hcut >Hcut in ⊢ (?%?); @(le_Max (λi.sf 〈i+a n,n〉)) /2/
+ ]
+ ]
+qed.
+
+(*
+axiom CF_mu: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s.
+ CF sa a → CF sb b → CF sf f →
+ O s (λx.sa x + sb x + ∑_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉)) →
+ CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)). *)
+
+
+
\ No newline at end of file
--- /dev/null
+include "reverse_complexity/basics.ma".
+include "reverse_complexity/big_O.ma".
+
+(********************************* complexity *********************************)
+
+(* We assume operations have a minimal structural complexity MSC.
+For instance, for time complexity, MSC is equal to the size of input.
+For space complexity, MSC is typically 0, since we only measure the space
+required in addition to dimension of the input. *)
+
+axiom MSC : nat → nat.
+axiom MSC_sublinear : ∀n. MSC (S n) ≤ S (MSC n).
+(* axiom MSC_S: O MSC (λx.MSC (S x)) . *)
+axiom MSC_le: ∀n. MSC n ≤ n.
+
+axiom monotonic_MSC: monotonic ? le MSC.
+axiom MSC_pair_eq: ∀a,b. MSC 〈a,b〉 = MSC a + MSC b.
+
+
+lemma MSC_pair: ∀a,b. MSC 〈a,b〉 ≤ MSC a + MSC b. // qed.
+
+(* C s i means i is running in O(s) *)
+
+definition C ≝ λs,i.∃c.∃a.∀x.a ≤ x → ∃y.
+ U i x (c*(s x)) = Some ? y.
+
+(* C f s means f ∈ O(s) where MSC ∈O(s) *)
+definition CF ≝ λs,f.∃i.code_for (total f) i ∧ C s i.
+
+axiom MSC_in: ∀f,h. CF h f → ∀x. MSC x ≤ h x.
+axiom MSC_out: ∀f,h. CF h f → ∀x. MSC (f x) ≤ h x.
+
+
+lemma ext_CF : ∀f,g,s. (∀n. f n = g n) → CF s f → CF s g.
+#f #g #s #Hext * #i * #Hcode #HC %{i} %
+ [#x cases (Hcode x) #a #H %{a} whd in match (total ??); <Hext @H | //]
+qed.
+
+lemma ext_CF1 : ∀f,g,s. (∀n. f n = g n) → CF s g → CF s f.
+#f #g #s #Hext * #i * #Hcode #HC %{i} %
+ [#x cases (Hcode x) #a #H %{a} whd in match (total ??); >Hext @H | //]
+qed.
+
+lemma monotonic_CF: ∀s1,s2,f.(∀x. s1 x ≤ s2 x) → CF s1 f → CF s2 f.
+#s1 #s2 #f #H * #i * #Hcode #Hs1
+%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1 %{c} %{a} #n #lean
+cases(Hs1 n lean) #y #Hy %{y} @(monotonic_U …Hy) @le_times //
+qed.
+
+lemma O_to_CF: ∀s1,s2,f.O s2 s1 → CF s1 f → CF s2 f.
+#s1 #s2 #f #H * #i * #Hcode #Hs1
+%{i} % [//] cases Hs1 #c * #a -Hs1 #Hs1
+cases H #c1 * #a1 #Ha1 %{(c*c1)} %{(a+a1)} #n #lean
+cases(Hs1 n ?) [2:@(transitive_le … lean) //] #y #Hy %{y} @(monotonic_U …Hy)
+>associative_times @le_times // @Ha1 @(transitive_le … lean) //
+qed.
+
+lemma timesc_CF: ∀s,f,c.CF (λx.c*s x) f → CF s f.
+#s #f #c @O_to_CF @O_times_c
+qed.
+
+(********************************* composition ********************************)
+axiom CF_comp: ∀f,g,sf,sg,sh. CF sg g → CF sf f →
+ O sh (λx. sg x + sf (g x)) → CF sh (f ∘ g).
+
+lemma CF_comp_ext: ∀f,g,h,sh,sf,sg. CF sg g → CF sf f →
+ (∀x.f(g x) = h x) → O sh (λx. sg x + sf (g x)) → CF sh h.
+#f #g #h #sh #sf #sg #Hg #Hf #Heq #H @(ext_CF (f ∘ g))
+ [#n normalize @Heq | @(CF_comp … H) //]
+qed.
+
+(************************************* smn ************************************)
+axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉).
+
+(****************************** constructibility ******************************)
+
+definition constructible ≝ λs. CF s s.
+
+lemma constr_comp : ∀s1,s2. constructible s1 → constructible s2 →
+ (∀x. x ≤ s2 x) → constructible (s2 ∘ s1).
+#s1 #s2 #Hs1 #Hs2 #Hle @(CF_comp … Hs1 Hs2) @O_plus @le_to_O #x [@Hle | //]
+qed.
+
+lemma ext_constr: ∀s1,s2. (∀x.s1 x = s2 x) →
+ constructible s1 → constructible s2.
+#s1 #s2 #Hext #Hs1 @(ext_CF … Hext) @(monotonic_CF … Hs1) #x >Hext //
+qed.
+
+
+(***************************** primitive recursion ****************************)
+
+(* no arguments *)
+
+let rec prim_rec0 (k:nat) (h:nat →nat) n on n ≝
+ match n with
+ [ O ⇒ k
+ | S a ⇒ h 〈a, prim_rec0 k h a〉].
+
+lemma prim_rec0_S: ∀k,h,n.
+ prim_rec0 k h (S n) = h 〈n, prim_rec0 k h n〉.
+// qed.
+
+let rec prim_rec0_compl (k,sk:nat) (h,sh:nat →nat) n on n ≝
+ match n with
+ [ O ⇒ sk
+ | S a ⇒ prim_rec0_compl k sk h sh a + sh (prim_rec0 k h a)].
+
+axiom CF_prim_rec0_gen: ∀k,h,sk,sh,sh1,sf. CF sh h →
+ O sh1 (λx.snd x + sh 〈fst x,prim_rec0 k h (fst x)〉) →
+ O sf (prim_rec0 sk sh1) →
+ CF sf (prim_rec0 k h).
+
+lemma CF_prim_rec0: ∀k,h,sk,sh,sf. CF sh h →
+ O sf (prim_rec0 sk (λx. snd x + sh 〈fst x,prim_rec0 k h (fst x)〉))
+ → CF sf (prim_rec0 k h).
+#k #h #sk #sh #sf #Hh #HO @(CF_prim_rec0_gen … Hh … HO) @O_refl
+qed.
+
+(* with arguments. m is a vector of arguments *)
+let rec prim_rec (k,h:nat →nat) n m on n ≝
+ match n with
+ [ O ⇒ k m
+ | S a ⇒ h 〈a,〈prim_rec k h a m, m〉〉].
+
+lemma prim_rec_S: ∀k,h,n,m.
+ prim_rec k h (S n) m = h 〈n,〈prim_rec k h n m, m〉〉.
+// qed.
+
+lemma prim_rec_le: ∀k1,k2,h1,h2. (∀x.k1 x ≤ k2 x) →
+(∀x,y.x ≤y → h1 x ≤ h2 y) →
+ ∀x,m. prim_rec k1 h1 x m ≤ prim_rec k2 h2 x m.
+#k1 #k2 #h1 #h2 #lek #leh #x #m elim x //
+#n #Hind normalize @leh @le_pair // @le_pair //
+qed.
+
+definition unary_pr ≝ λk,h,x. prim_rec k h (fst x) (snd x).
+
+lemma prim_rec_unary: ∀k,h,a,b.
+ prim_rec k h a b = unary_pr k h 〈a,b〉.
+#k #h #a #b normalize >fst_pair >snd_pair //
+qed.
+
+let rec prim_rec_compl (k,h,sk,sh:nat →nat) n m on n ≝
+ match n with
+ [ O ⇒ sk m
+ | S a ⇒ prim_rec_compl k h sk sh a m + sh 〈a,〈prim_rec k h a m,m〉〉].
+
+axiom CF_prim_rec_gen: ∀k,h,sk,sh,sh1. CF sk k → CF sh h →
+ O sh1 (λx. fst (snd x) +
+ sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉) →
+ CF (unary_pr sk sh1) (unary_pr k h).
+
+lemma CF_prim_rec: ∀k,h,sk,sh,sf. CF sk k → CF sh h →
+ O sf (unary_pr sk
+ (λx. fst (snd x) +
+ sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉))
+ → CF sf (unary_pr k h).
+#k #h #sk #sh #sf #Hk #Hh #Osf @(O_to_CF … Osf) @(CF_prim_rec_gen … Hk Hh)
+@O_refl
+qed.
+
+lemma constr_prim_rec: ∀s1,s2. constructible s1 → constructible s2 →
+ (∀n,r,m. 2 * r ≤ s2 〈n,〈r,m〉〉) → constructible (unary_pr s1 s2).
+#s1 #s2 #Hs1 #Hs2 #Hincr @(CF_prim_rec … Hs1 Hs2) whd %{2} %{0}
+#x #_ lapply (surj_pair x) * #a * #b #eqx >eqx whd in match (unary_pr ???);
+>fst_pair >snd_pair
+whd in match (unary_pr ???); >fst_pair >snd_pair elim a
+ [normalize //
+ |#n #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair
+ >prim_rec_S @transitive_le [| @(monotonic_le_plus_l … Hind)]
+ @transitive_le [| @(monotonic_le_plus_l … (Hincr n ? b))]
+ whd in match (unary_pr ???); >fst_pair >snd_pair //
+ ]
+qed.
+
+(**************************** primitive operations*****************************)
+
+definition id ≝ λx:nat.x.
+
+axiom CF_id: CF MSC id.
+axiom CF_const: ∀k.CF MSC (λx.k).
+axiom CF_comp_fst: ∀h,f. CF h f → CF h (fst ∘ f).
+axiom CF_comp_snd: ∀h,f. CF h f → CF h (snd ∘ f).
+axiom CF_comp_pair: ∀h,f,g. CF h f → CF h g → CF h (λx. 〈f x,g x〉).
+
+lemma CF_fst: CF MSC fst.
+@(ext_CF (fst ∘ id)) [#n //] @(CF_comp_fst … CF_id)
+qed.
+
+lemma CF_snd: CF MSC snd.
+@(ext_CF (snd ∘ id)) [#n //] @(CF_comp_snd … CF_id)
+qed.
+
+(**************************** arithmetic operations ***************************)
+
+axiom CF_compS: ∀h,f. CF h f → CF h (S ∘ f).
+axiom CF_le: ∀h,f,g. CF h f → CF h g → CF h (λx. leb (f x) (g x)).
+axiom CF_eqb: ∀h,f,g. CF h f → CF h g → CF h (λx.eqb (f x) (g x)).
+axiom CF_max1: ∀h,f,g. CF h f → CF h g → CF h (λx. max (f x) (g x)).
+axiom CF_plus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x + g x).
+axiom CF_minus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x - g x).
+
+(******************************** if then else ********************************)
+
+lemma if_prim_rec: ∀b:nat → bool. ∀f,g:nat → nat.∀x:nat.
+ if b x then f x else g x = prim_rec g (f ∘ snd ∘ snd) (b x) x.
+#b #f #g #x cases (b x) normalize //
+qed.
+
+lemma CF_if: ∀b:nat → bool. ∀f,g,s. CF s b → CF s f → CF s g →
+ CF s (λx. if b x then f x else g x).
+#b #f #g #s #CFb #CFf #CFg @(ext_CF (λx.unary_pr g (f ∘ snd ∘ snd) 〈b x,x〉))
+ [#n @sym_eq normalize >fst_pair >snd_pair @if_prim_rec
+ |@(CF_comp ??? s)
+ [|@CF_comp_pair // @(O_to_CF … CF_id) @le_to_O @MSC_in //
+ |@(CF_prim_rec_gen ??? (λx.MSC x + s(snd (snd x))) … CFg) [3:@O_refl|]
+ @(CF_comp … (λx.MSC x + s(snd x)) … CF_snd)
+ [@(CF_comp … s … CF_snd CFf) @O_refl
+ |@O_plus
+ [@O_plus_l @O_refl
+ |@O_plus
+ [@O_plus_l @le_to_O #x @monotonic_MSC //
+ |@O_plus_r @O_refl
+ ]
+ ]
+ ]
+ |%{6} %{0} #n #_ normalize in ⊢ (??%); <plus_n_O cases (b n) normalize
+ >snd_pair >fst_pair normalize
+ [@le_plus [//] >fst_pair >fst_pair >snd_pair >fst_pair
+ @le_plus [//] >snd_pair >snd_pair >snd_pair >snd_pair
+ <associative_plus <associative_plus @le_plus [|//]
+ @(transitive_le … (MSC_pair …)) >associative_plus @le_plus
+ [@(transitive_le ???? (MSC_in … CFf n)) @monotonic_MSC //
+ |@(transitive_le … (MSC_pair …)) @le_plus [|@(MSC_in … CFf)]
+ normalize @MSC_out @CFg
+ ]
+ |@le_plus //
+ ]
+ ]
+ ]
+qed.
+
+(********************************* simulation *********************************)
+
+definition pU_unary ≝ λp. pU (fst p) (fst (snd p)) (snd (snd p)).
+
+axiom sU : nat → nat.
+axiom CF_U : CF sU pU_unary.
+
+axiom monotonic_sU: ∀i1,i2,x1,x2,s1,s2. i1 ≤ i2 → x1 ≤ x2 → s1 ≤ s2 →
+ sU 〈i1,〈x1,s1〉〉 ≤ sU 〈i2,〈x2,s2〉〉.
+
+lemma monotonic_sU_aux : ∀x1,x2. fst x1 ≤ fst x2 → fst (snd x1) ≤ fst (snd x2) →
+snd (snd x1) ≤ snd (snd x2) → sU x1 ≤ sU x2.
+#x1 #x2 cases (surj_pair x1) #a1 * #y #eqx1 >eqx1 -eqx1 cases (surj_pair y)
+#b1 * #c1 #eqy >eqy -eqy
+cases (surj_pair x2) #a2 * #y2 #eqx2 >eqx2 -eqx2 cases (surj_pair y2)
+#b2 * #c2 #eqy2 >eqy2 -eqy2 >fst_pair >snd_pair >fst_pair >snd_pair
+>fst_pair >snd_pair >fst_pair >snd_pair @monotonic_sU
+qed.
+
+axiom sU_pos: ∀x. 0 < sU x.
+
+axiom sU_le: ∀i,x,s. s ≤ sU 〈i,〈x,s〉〉.
+
+lemma sU_le_i: ∀i,x,s. MSC i ≤ sU 〈i,〈x,s〉〉.
+#i #x #s @(transitive_le ???? (MSC_in … CF_U 〈i,〈x,s〉〉)) @monotonic_MSC //
+qed.
+
+lemma sU_le_x: ∀i,x,s. MSC x ≤ sU 〈i,〈x,s〉〉.
+#i #x #s @(transitive_le ???? (MSC_in … CF_U 〈i,〈x,s〉〉)) @monotonic_MSC
+@(transitive_le … 〈x,s〉) //
+qed.
+
+
+
+
+definition termb_unary ≝ λx:ℕ.termb (fst x) (fst (snd x)) (snd (snd x)).
+definition out_unary ≝ λx:ℕ.out (fst x) (fst (snd x)) (snd (snd x)).
+
+lemma CF_termb: CF sU termb_unary.
+@(ext_CF (fst ∘ pU_unary)) [2: @CF_comp_fst @CF_U]
+#n whd in ⊢ (??%?); whd in ⊢ (??(?%)?); >fst_pair %
+qed.
+
+lemma CF_out: CF sU out_unary.
+@(ext_CF (snd ∘ pU_unary)) [2: @CF_comp_snd @CF_U]
+#n whd in ⊢ (??%?); whd in ⊢ (??(?%)?); >snd_pair %
+qed.
+
include "arithmetics/nat.ma".
-include "arithmetics/log.ma".
-(* include "arithmetics/ord.ma". *)
+include "arithmetics/log.ma".
include "arithmetics/bigops.ma".
include "arithmetics/bounded_quantifiers.ma".
-include "arithmetics/pidgeon_hole.ma".
+include "arithmetics/pidgeon_hole.ma".
include "basics/sets.ma".
include "basics/types.ma".
#n >commutative_max //
qed.
+alias id "max" = "cic:/matita/arithmetics/nat/max#def:2".
+alias id "mk_Aop" = "cic:/matita/arithmetics/bigops/Aop#con:0:1:2".
definition MaxA ≝
mk_Aop nat 0 max Max0 Max0r (λa,b,c.sym_eq … (Max_assoc a b c)).
#f #g #s #Osf #Ofg @(O_plus … Osf) @(O_trans … Osf) //
qed.
-(*
-lemma O_ff: ∀f,s. O s f → O s (f+f).
-#f #s #Osf /2/
-qed. *)
-
lemma O_ext2: ∀f,g,s. O s f → (∀x.f x = g x) → O s g.
#f #g #s * #c * #a #Osf #eqfg %{c} %{a} #n #lean <eqfg @Osf //
qed.
-
definition not_O ≝ λf,g.∀c,n0.∃n. n0 ≤ n ∧ c* (f n) < g n .
-(* this is the only classical result *)
-axiom not_O_def: ∀f,g. ¬ O f g → not_O f g.
-
(******************************* small O notation *****************************)
(* o f g means g ∈ o(f) *)
(************************ basic complexity notions ****************************)
-(* u is the deterministic configuration relation of the universal machine (one
- step)
-
-axiom u: nat → option nat.
-
-let rec U c n on n ≝
- match n with
- [ O ⇒ None ?
- | S m ⇒ match u c with
- [ None ⇒ Some ? c (* halting case *)
- | Some c1 ⇒ U c1 m
- ]
- ].
-
-lemma halt_U: ∀i,n,y. u i = None ? → U i n = Some ? y → y = i.
-#i #n #y #H cases n
- [normalize #H1 destruct |#m normalize >H normalize #H1 destruct //]
-qed.
-
-lemma Some_to_halt: ∀n,i,y. U i n = Some ? y → u y = None ? .
-#n elim n
- [#i #y normalize #H destruct (H)
- |#m #Hind #i #y normalize
- cut (u i = None ? ∨ ∃c. u i = Some ? c)
- [cases (u i) [/2/ | #c %2 /2/ ]]
- *[#H >H normalize #H1 destruct (H1) // |* #c #H >H normalize @Hind ]
- ]
-qed. *)
-
axiom U: nat → nat → nat → option nat.
-(*
-lemma monotonici_U: ∀y,n,m,i.
- U i m = Some ? y → U i (n+m) = Some ? y.
-#y #n #m elim m
- [#i normalize #H destruct
- |#p #Hind #i <plus_n_Sm normalize
- cut (u i = None ? ∨ ∃c. u i = Some ? c)
- [cases (u i) [/2/ | #c %2 /2/ ]]
- *[#H1 >H1 normalize // |* #c #H >H normalize #H1 @Hind //]
- ]
-qed. *)
axiom monotonic_U: ∀i,x,n,m,y.n ≤m →
U i x n = Some ? y → U i x m = Some ? y.
-(* #i #n #m #y #lenm #H >(plus_minus_m_m m n) // @monotonici_U //
-qed. *)
-
-(* axiom U: nat → nat → option nat. *)
-(* axiom monotonic_U: ∀i,n,m,y.n ≤m →
- U i n = Some ? y → U i m = Some ? y. *)
lemma unique_U: ∀i,x,n,m,yn,ym.
U i x n = Some ? yn → U i x m = Some ? ym → yn = ym.
∃n.∀m. n ≤ m → U i x m = f x.
definition terminate ≝ λi,x,r. ∃y. U i x r = Some ? y.
-notation "[i,x] ↓ r" with precedence 60 for @{terminate $i $x $r}.
+notation "{i ⊙ x} ↓ r" with precedence 60 for @{terminate $i $x $r}.
definition lang ≝ λi,x.∃r,y. U i x r = Some ? y ∧ 0 < y.
Max_{i < S (of_size n) | eqb (|i|) n}|(f i)|.
// qed.
-(*
-definition Max_const : ∀f,p,n,a. a < n → p a →
- ∀n. f n = g n →
- Max_{i < n | p n}(f n) = *)
-
lemma size_f_size : ∀f,n. size_f (f ∘ size) n = |(f n)|.
#f #n @le_to_le_to_eq
[@Max_le #a #lta #Ha normalize >(eqb_true_to_eq … Ha) //
#n @Max_le #a #lta #Ha <(eqb_true_to_eq … Ha) //
qed.
-(* definition def ≝ λf:nat → option nat.λx.∃y. f x = Some ? y.*)
-
(* C s i means that the complexity of i is O(s) *)
-definition C ≝ λs,i.∃c.∃a.∀x.a ≤ |x| → ∃y.
- U i x (c*(s(|x|))) = Some ? y.
+definition C ≝ λs,i.∃c.∃a.∀x.a ≤ |x| → {i ⊙ x} ↓ (c*(s(|x|))).
definition CF ≝ λs,f.∃i.code_for f i ∧ C s i.
%{y} @(monotonic_U …Hy) >associative_times @le_times // @H @(le_maxl … Hmax)
qed.
+(*********************** The hierachy theorem (left) **************************)
+
+theorem hierarchy_theorem_left: ∀s1,s2:nat→nat.
+ O(s1) ⊆ O(s2) → CF s1 ⊆ CF s2.
+#s1 #s2 #HO #f * #i * #Hcode * #c * #a #Hs1_i %{i} % //
+cases (sub_O_to_O … HO) -HO #c1 * #b #Hs1s2
+%{(c*c1)} %{(max a b)} #x #lemax
+cases (Hs1_i x ?) [2: @(le_maxl …lemax)]
+#y #Hy %{y} @(monotonic_U … Hy) >associative_times
+@le_times // @Hs1s2 @(le_maxr … lemax)
+qed.
+
(************************** The diagonal language *****************************)
(* the diagonal language used for the hierarchy theorem *)
U (fst i) i (s (|i|)) = Some ? 0.
lemma equiv_diag: ∀s,i.
- diag s i ↔ [fst i,i] ↓ s (|i|) ∧ ¬lang (fst i) i.
+ diag s i ↔ {fst i ⊙ i} ↓ s(|i|) ∧ ¬lang (fst i) i.
#s #i %
[whd in ⊢ (%→?); #H % [%{0} //] % * #x * #y *
#H1 #Hy cut (0 = y) [@(unique_U … H H1)] #eqy /2/
#H3 @(absurd ?? H3) @H2 @H3
qed.
-(* axiom weak_pad : ∀a,∃a0.∀n. a0 < n → ∃b. |〈a,b〉| = n. *)
+let rec f_img (f:nat →nat) n on n ≝
+ match n with
+ [O ⇒ [ ]
+ |S m ⇒ f m::f_img f m
+ ].
+
+(* a few lemma to prove injective_to_exists. This is a general result; a nice
+example of the pidgeon hole pricniple *)
+
+lemma f_img_to_exists:
+ ∀f.∀n,a. a ∈ f_img f n → ∃b. b < n ∧ f b = a.
+#f #n #a elim n normalize [@False_ind]
+#m #Hind *
+ [#Ha %{m} /2/ |#H cases(Hind H) #b * #Hb #Ha %{b} % // @le_S //]
+qed.
+
+lemma length_f_img: ∀f,n. |f_img f n| = n.
+#f #n elim n // normalize //
+qed.
+
+lemma unique_f_img: ∀f,n. injective … f → unique ? (f_img f n).
+#f #n #Hinj elim n normalize //
+#m #Hind % // % #H lapply(f_img_to_exists …H) * #b * #ltbm
+#eqbm @(absurd … ltbm) @le_to_not_lt >(Hinj… eqbm) //
+qed.
+
+lemma injective_to_exists: ∀f. injective nat nat f →
+ ∀n.(∀i.i < n → f i < n) → ∀a. a < n → ∃b. b<n ∧ f b = a.
+#f #finj #n #H1 #a #ltan @(f_img_to_exists f n a)
+@(eq_length_to_mem_all … (length_f_img …) (unique_f_img …finj …) …ltan)
+#x #Hx cases(f_img_to_exists … Hx) #b * #ltbn #eqx <eqx @H1 //
+qed.
+
lemma weak_pad1 :∀n,a.∃b. n ≤ 〈a,b〉.
#n #a
cut (∀i.decidable (〈a,i〉 < n))
|#H lapply(not_forall_to_exists … Hdec H)
* #b * #H1 #H2 %{b} @not_lt_to_le @H2
]
-qed.
+qed.
lemma pad : ∀n,a. ∃b. n ≤ |〈a,b〉|.
#n #a cases (weak_pad1 (of_size n) a) #b #Hb
qed.
lemma o_to_ex: ∀s1,s2. o s1 s2 → ∀i. C s2 i →
- ∃b.[i, 〈i,b〉] ↓ s1 (|〈i,b〉|).
+ ∃b.{i ⊙ 〈i,b〉} ↓ s1 (|〈i,b〉|).
#s1 #s2 #H #i * #c * #x0 #H1
cases (H c) #n0 #H2 cases (pad (max x0 n0) i) #b #Hmax
%{b} cases (H1 〈i,b〉 ?)
[ None ⇒ None ?
| Some y ⇒ f y ].
-(* axiom CFU: ∀h,g,s. CF s (to_Some h) → CF s (to_Some (of_size ∘ g)) →
- CF (Slow s) (λx.U (h x) (g x)). *)
-
axiom sU2: nat → nat → nat.
axiom sU: nat → nat → nat → nat.
-(* axiom CFU: CF sU (λx.U (fst x) (snd x)). *)
-
axiom CFU_new: ∀h,g,f,s.
CF s (to_Some h) → CF s (to_Some g) → CF s (to_Some f) →
O s (λx. sU (size_f h x) (size_f g x) (size_f f x)) →
axiom superlinear_sU: ∀i,x,r. r ≤ sU i x r.
+(* not used *)
definition sU_space ≝ λi,x,r.i+x+r.
-definition sU_time ≝ λi,x,r.i+x+(i^2)*r*(log 2 r).
-
-(*
-axiom CF_comp: ∀f,g,s1, s2. CF s1 f → CF s2 g →
- CF (λx.s2 x + s1 (size (deopt (g x)))) (opt_comp f g).
-
-(* axiom CF_comp: ∀f,g,s1, s2. CF s1 f → CF s2 g →
- CF (s1 ∘ (λx. size (deopt (g x)))) (opt_comp f g). *)
-
-axiom CF_comp_strong: ∀f,g,s1,s2. CF s1 f → CF s2 g →
- CF (s1 ∘ s2) (opt_comp f g). *)
+definition sU_time ≝ λi,x,r.i+x+(i^2)*r*(log 2 r).
definition IF ≝ λb,f,g:nat →option nat. λx.
match b x with
axiom CF_fst: ∀f,s. CF s (λx.Some ? (f x)) → CF s (λx.Some ? (fst (f x))).
-definition minimal ≝ λs. CF s (λn. Some ? n) ∧ ∀c. CF s (λn. Some ? c).
-
-
-(*
-axiom le_snd: ∀n. |snd n| ≤ |n|.
-axiom daemon: ∀P:Prop.P. *)
+definition sufficiently_large ≝ λs. CF s (λn. Some ? n) ∧ ∀c. CF s (λn. Some ? c).
definition constructible ≝ λs. CF s (λx.Some ? (of_size (s (|x|)))).
-lemma diag_s: ∀s. minimal s → constructible s →
+lemma diag_s: ∀s. sufficiently_large s → constructible s →
CF (λx.sU x x (s x)) (diag_cf s).
#s * #Hs_id #Hs_c #Hs_constr
cut (O (λx:ℕ.sU x x (s x)) s) [%{1} %{0} #n //]
|%{1} %{0} #n #_ >commutative_times <times_n_1
@monotonic_sU // >size_f_size >size_of_size //
]
-qed.
-
-(*
-lemma diag_s: ∀s. minimal s → constructible s →
- CF (λx.s x + sU x x (s x)) (diag_cf s).
-#s * #Hs_id #Hs_c #Hs_constr
-@ext_CF [2: #n @sym_eq @diag_cf_def | skip]
-@IF_CF_new [2,3:@(monotonic_CF … (Hs_c ?)) @O_plus_l //]
-@CFU_new
- [@CF_fst @(monotonic_CF … Hs_id) @O_plus_l //
- |@(monotonic_CF … Hs_id) @O_plus_l //
- |@(monotonic_CF … Hs_constr) @O_plus_l //
- |@O_plus_r %{1} %{0} #n #_ >commutative_times <times_n_1
- @monotonic_sU // >size_f_size >size_of_size //
- ]
-qed. *)
-
-(* proof with old axioms
-lemma diag_s: ∀s. minimal s → constructible s →
- CF (λx.s x + sU x x (s x)) (diag_cf s).
-#s * #Hs_id #Hs_c #Hs_constr
-@ext_CF [2: #n @sym_eq @diag_cf_def | skip]
-@(monotonic_CF ???? (IF_CF (λi:ℕ.U (pair (fst i) i) (|of_size (s (|i|))|))
- … (λi.s i + s i + s i + (sU (size_f fst i) (size_f (λi.i) i) (size_f (λi.of_size (s (|i|))) i))) … (Hs_c 1) (Hs_c 0) … ))
- [2: @CFU [@CF_fst // | // |@Hs_constr]
- |@(O_ext2 (λn:ℕ.s n+sU (size_f fst n) n (s n) + (s n+s n+s n+s n)))
- [2: #i >size_f_size >size_of_size >size_f_id //]
- @O_absorbr
- [%{1} %{0} #n #_ >commutative_times <times_n_1 @le_plus //
- @monotonic_sU //
- |@O_plus_l @(O_plus … (O_refl s)) @(O_plus … (O_refl s))
- @(O_plus … (O_refl s)) //
- ]
-qed.
-*)
-
-(*************************** The hierachy theorem *****************************)
-
-(*
-theorem hierarchy_theorem_right: ∀s1,s2:nat→nat.
- O s1 idN → constructible s1 →
- not_O s2 s1 → ¬ CF s1 ⊆ CF s2.
-#s1 #s2 #Hs1 #monos1 #H % #H1
-@(absurd … (CF s2 (diag_cf s1)))
- [@H1 @diag_s // |@(diag1_not_s1 … H)]
-qed.
-*)
-
-theorem hierarchy_theorem_left: ∀s1,s2:nat→nat.
- O(s1) ⊆ O(s2) → CF s1 ⊆ CF s2.
-#s1 #s2 #HO #f * #i * #Hcode * #c * #a #Hs1_i %{i} % //
-cases (sub_O_to_O … HO) -HO #c1 * #b #Hs1s2
-%{(c*c1)} %{(max a b)} #x #lemax
-cases (Hs1_i x ?) [2: @(le_maxl …lemax)]
-#y #Hy %{y} @(monotonic_U … Hy) >associative_times
-@le_times // @Hs1s2 @(le_maxr … lemax)
-qed.
-
+qed.
\ No newline at end of file
--- /dev/null
+include "arithmetics/sigma_pi.ma".
+
+(************************* notation for minimization **************************)
+notation "μ_{ ident i < n } p"
+ with precedence 80 for @{min $n 0 (λ${ident i}.$p)}.
+
+notation "μ_{ ident i ≤ n } p"
+ with precedence 80 for @{min (S $n) 0 (λ${ident i}.$p)}.
+
+notation "μ_{ ident i ∈ [a,b[ } p"
+ with precedence 80 for @{min ($b-$a) $a (λ${ident i}.$p)}.
+
+notation "μ_{ ident i ∈ [a,b] } p"
+ with precedence 80 for @{min (S $b-$a) $a (λ${ident i}.$p)}.
+
+(************************************ MAX *************************************)
+notation "Max_{ ident i < n | p } f"
+ with precedence 80
+for @{'bigop $n max 0 (λ${ident i}. $p) (λ${ident i}. $f)}.
+
+notation "Max_{ ident i < n } f"
+ with precedence 80
+for @{'bigop $n max 0 (λ${ident i}.true) (λ${ident i}. $f)}.
+
+notation "Max_{ ident j ∈ [a,b[ } f"
+ with precedence 80
+for @{'bigop ($b-$a) max 0 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
+ (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
+
+notation "Max_{ ident j ∈ [a,b[ | p } f"
+ with precedence 80
+for @{'bigop ($b-$a) max 0 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a)))
+ (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
+
+lemma Max_assoc: ∀a,b,c. max (max a b) c = max a (max b c).
+#a #b #c normalize cases (true_or_false (leb a b)) #leab >leab normalize
+ [cases (true_or_false (leb b c )) #lebc >lebc normalize
+ [>(le_to_leb_true a c) // @(transitive_le ? b) @leb_true_to_le //
+ |>leab //
+ ]
+ |cases (true_or_false (leb b c )) #lebc >lebc normalize //
+ >leab normalize >(not_le_to_leb_false a c) // @lt_to_not_le
+ @(transitive_lt ? b) @not_le_to_lt @leb_false_to_not_le //
+ ]
+qed.
+
+lemma Max0 : ∀n. max 0 n = n.
+// qed.
+
+lemma Max0r : ∀n. max n 0 = n.
+#n >commutative_max //
+qed.
+
+definition MaxA ≝
+ mk_Aop nat 0 max Max0 Max0r (λa,b,c.sym_eq … (Max_assoc a b c)).
+
+definition MaxAC ≝ mk_ACop nat 0 MaxA commutative_max.
+
+lemma le_Max: ∀f,p,n,a. a < n → p a = true →
+ f a ≤ Max_{i < n | p i}(f i).
+#f #p #n #a #ltan #pa
+>(bigop_diff p ? 0 MaxAC f a n) // @(le_maxl … (le_n ?))
+qed.
+
+lemma le_MaxI: ∀f,p,n,m,a. m ≤ a → a < n → p a = true →
+ f a ≤ Max_{i ∈ [m,n[ | p i}(f i).
+#f #p #n #m #a #lema #ltan #pa
+>(bigop_diff ? ? 0 MaxAC (λi.f (i+m)) (a-m) (n-m))
+ [<plus_minus_m_m // @(le_maxl … (le_n ?))
+ |<plus_minus_m_m //
+ |/2 by monotonic_lt_minus_l/
+ ]
+qed.
+
+lemma Max_le: ∀f,p,n,b.
+ (∀a.a < n → p a = true → f a ≤ b) → Max_{i < n | p i}(f i) ≤ b.
+#f #p #n elim n #b #H //
+#b0 #H1 cases (true_or_false (p b)) #Hb
+ [>bigop_Strue [2:@Hb] @to_max [@H1 // | @H #a #ltab #pa @H1 // @le_S //]
+ |>bigop_Sfalse [2:@Hb] @H #a #ltab #pa @H1 // @le_S //
+ ]
+qed.
+
+
+(************************* a couple of technical lemmas ***********************)
+lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0.
+#a elim a // #n #Hind *
+ [#H @False_ind /2 by absurd/ | #b normalize #H @Hind @le_S_S_to_le /2/]
+qed.
+
+lemma sigma_const: ∀c,a,b.
+ ∑_{i ∈ [a,b[ }c ≤ (b-a)*c.
+#c #a #b elim (b-a) // #n #Hind normalize @le_plus //
+qed.
+
+lemma sigma_to_Max: ∀h,a,b.
+ ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*Max_{i ∈ [a,b[ }(h i).
+#h #a #b elim (b-a)
+ [//
+ |#n #Hind >bigop_Strue [2://] whd in ⊢ (??%);
+ @le_plus
+ [>bigop_Strue // @(le_maxl … (le_n …))
+ |@(transitive_le … Hind) @le_times // >bigop_Strue //
+ @(le_maxr … (le_n …))
+ ]
+ ]
+qed.
+
+lemma sigma_bound1: ∀h,a,b. monotonic nat le h →
+ ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h b.
+#h #a #b #H
+@(transitive_le … (sigma_to_Max …)) @le_times //
+@Max_le #i #lti #_ @H @lt_to_le @lt_minus_to_plus_r //
+qed.
+
+lemma sigma_bound_decr1: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) →
+ ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a.
+#h #a #b #H
+@(transitive_le … (sigma_to_Max …)) @le_times //
+@Max_le #i #lti #_ @H // @lt_minus_to_plus_r //
+qed.
+
+lemma sigma_bound: ∀h,a,b. monotonic nat le h →
+ ∑_{i ∈ [a,S b[ }(h i) ≤ (S b-a)*h b.
+#h #a #b #H cases (decidable_le a b)
+ [#leab cut (b = pred (S b - a + a))
+ [<plus_minus_m_m // @le_S //] #Hb >Hb in match (h b);
+ generalize in match (S b -a);
+ #n elim n
+ [//
+ |#m #Hind >bigop_Strue [2://] @le_plus
+ [@H @le_n |@(transitive_le … Hind) @le_times [//] @H //]
+ ]
+ |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba
+ cut (S b -a = 0) [@minus_to_0 //] #Hcut >Hcut //
+ ]
+qed.
+
+lemma sigma_bound_decr: ∀h,a,b. (∀a1,a2. a1 ≤ a2 → a2 < b → h a2 ≤ h a1) →
+ ∑_{i ∈ [a,b[ }(h i) ≤ (b-a)*h a.
+#h #a #b #H cases (decidable_le a b)
+ [#leab cut ((b -a) +a ≤ b) [/2 by le_minus_to_plus_r/] generalize in match (b -a);
+ #n elim n
+ [//
+ |#m #Hind >bigop_Strue [2://] #Hm
+ cut (m+a ≤ b) [@(transitive_le … Hm) //] #Hm1
+ @le_plus [@H // |@(transitive_le … (Hind Hm1)) //]
+ ]
+ |#ltba lapply (not_le_to_lt … ltba) -ltba #ltba
+ cut (b -a = 0) [@minus_to_0 @lt_to_le @ltba] #Hcut >Hcut //
+ ]
+qed.
+
\ No newline at end of file
--- /dev/null
+include "reverse_complexity/bigops_compl.ma".
+
+(********************************* the speedup ********************************)
+
+definition min_input ≝ λh,i,x. μ_{y ∈ [S i,x] } (termb i y (h (S i) y)).
+
+lemma min_input_def : ∀h,i,x.
+ min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)).
+// qed.
+
+lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i.
+#h #i #x #lexi >min_input_def
+cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut //
+qed.
+
+lemma min_input_to_terminate: ∀h,i,x.
+ min_input h i x = x → {i ⊙ x} ↓ (h (S i) x).
+#h #i #x #Hminx
+cases (decidable_le (S i) x) #Hix
+ [cases (true_or_false (termb i x (h (S i) x))) #Hcase
+ [@termb_true_to_term //
+ |<Hminx in Hcase; #H lapply (fmin_false (λx.termb i x (h (S i) x)) (x-i) (S i) H)
+ >min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?);
+ <plus_n_Sm <plus_minus_m_m [2: @lt_to_le //]
+ #Habs @False_ind /2/
+ ]
+ |@False_ind >min_input_i in Hminx;
+ [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //]
+ ]
+qed.
+
+lemma min_input_to_lt: ∀h,i,x.
+ min_input h i x = x → i < x.
+#h #i #x #Hminx cases (decidable_le (S i) x) //
+#ltxi @False_ind >min_input_i in Hminx;
+ [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //]
+qed.
+
+lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 →
+ min_input h i x = x → min_input h i x1 = x.
+#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex))
+ [@(fmin_true … (sym_eq … Hminx)) //
+ |@(min_input_to_lt … Hminx)
+ |#j #H1 <Hminx @lt_min_to_false //
+ |@plus_minus_m_m @le_S_S @(transitive_le … lex) @lt_to_le
+ @(min_input_to_lt … Hminx)
+ ]
+qed.
+
+definition g ≝ λh,u,x.
+ S (max_{i ∈[u,x[ | eqb (min_input h i x) x} (out i x (h (S i) x))).
+
+lemma g_def : ∀h,u,x. g h u x =
+ S (max_{i ∈[u,x[ | eqb (min_input h i x) x} (out i x (h (S i) x))).
+// qed.
+
+lemma le_u_to_g_1 : ∀h,u,x. x ≤ u → g h u x = 1.
+#h #u #x #lexu >g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/]
+#eq0 >eq0 normalize // qed.
+
+lemma g_lt : ∀h,i,x. min_input h i x = x →
+ out i x (h (S i) x) < g h 0 x.
+#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/
+qed.
+
+lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0.
+#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase
+ [#H %2 @H | #H %1 @H]
+qed.
+
+definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x.
+interpretation "almost equal" 'napart f g = (almost_equal f g).
+
+lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧
+ max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0.
+#h #u elim u
+ [normalize % #H cases (H u) #x * #_ * #H1 @H1 //
+ |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx
+ cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase
+ [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax
+ [2: #H %{x} % // <minus_n_O @H]
+ #Hneq0 (* if x is not enough we retry with nu=x *)
+ cases (Hind x) #x1 * #ltx1
+ >bigop_Sfalse
+ [#H %{x1} % [@transitive_lt //| <minus_n_O @H]
+ |@not_eq_to_eqb_false >(le_to_min_input … (eqb_true_to_eq … Hcase))
+ [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1]
+ ]
+ |>bigop_Sfalse [2:@Hcase] #H %{x} % // <minus_n_O @H
+ ]
+ ]
+qed.
+
+lemma condition_1: ∀h,u.g h 0 ≈ g h u.
+#h #u @(not_to_not … (eventually_cancelled h u))
+#H #nu cases (H (max u nu)) #x * #ltx #Hdiff
+%{x} % [@(le_to_lt_to_lt … ltx) @(le_maxr … (le_n …))] @(not_to_not … Hdiff)
+#H @(eq_f ?? S) >(bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA)
+ [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//]
+qed.
+
+(******************************** Condition 2 *********************************)
+
+lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y.
+#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found //
+ [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //]
+ |#y #leiy #lty @(lt_min_to_false ????? lty) //
+ ]
+qed.
+
+lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. i<x ∧ {i ⊙ x} ↓ h (S i) x.
+#h #i whd in ⊢(%→?); #H % #H1 cases (exists_to_exists_min … H1) #y #Hminy
+lapply (g_lt … Hminy)
+lapply (min_input_to_terminate … Hminy) * #r #termy
+cases (H y) -H #ny #Hy
+cut (r = g h 0 y) [@(unique_U … ny … termy) @Hy //] #Hr
+whd in match (out ???); >termy >Hr
+#H @(absurd ? H) @le_to_not_lt @le_n
+qed.
+
+(**************************** complexity of g *********************************)
+
+definition unary_g ≝ λh.λux. g h (fst ux) (snd ux).
+definition auxg ≝
+ λh,ux. max_{i ∈[fst ux,snd ux[ | eqb (min_input h i (snd ux)) (snd ux)}
+ (out i (snd ux) (h (S i) (snd ux))).
+
+lemma compl_g1 : ∀h,s. CF s (auxg h) → CF s (unary_g h).
+#h #s #H1 @(CF_compS ? (auxg h) H1)
+qed.
+
+definition aux1g ≝
+ λh,ux. max_{i ∈[fst ux,snd ux[ | (λp. eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) 〈i,ux〉}
+ ((λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) 〈i,ux〉).
+
+lemma eq_aux : ∀h,x.aux1g h x = auxg h x.
+#h #x @same_bigop
+ [#n #_ >fst_pair >snd_pair // |#n #_ #_ >fst_pair >snd_pair //]
+qed.
+
+lemma compl_g2 : ∀h,s1,s2,s.
+ CF s1
+ (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))) →
+ CF s2
+ (λp:ℕ.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p)))) →
+ O s (λx.MSC x + (snd x - fst x)*Max_{i ∈[fst x ,snd x[ }(s1 〈i,x〉+s2 〈i,x〉)) →
+ CF s (auxg h).
+#h #s1 #s2 #s #Hs1 #Hs2 #HO @(ext_CF (aux1g h))
+ [#n whd in ⊢ (??%%); @eq_aux]
+@(CF_max2 … CF_fst CF_snd Hs1 Hs2 …) @(O_trans … HO)
+@O_plus [@O_plus @O_plus_l // | @O_plus_r //]
+qed.
+
+lemma compl_g3 : ∀h,s.
+ CF s (λp:ℕ.min_input h (fst p) (snd (snd p))) →
+ CF s (λp:ℕ.eqb (min_input h (fst p) (snd (snd p))) (snd (snd p))).
+#h #s #H @(CF_eqb … H) @(CF_comp … CF_snd CF_snd) @(O_trans … (le_to_O … (MSC_in … H)))
+@O_plus // %{1} %{0} #n #_ >commutative_times <times_n_1 @monotonic_MSC //
+qed.
+
+definition min_input_aux ≝ λh,p.
+ μ_{y ∈ [S (fst p),snd (snd p)] }
+ ((λx.termb (fst (snd x)) (fst x) (h (S (fst (snd x))) (fst x))) 〈y,p〉).
+
+lemma min_input_eq : ∀h,p.
+ min_input_aux h p =
+ min_input h (fst p) (snd (snd p)).
+#h #p >min_input_def whd in ⊢ (??%?); >minus_S_S @min_f_g #i #_ #_
+whd in ⊢ (??%%); >fst_pair >snd_pair //
+qed.
+
+definition termb_aux ≝ λh.
+ termb_unary ∘ λp.〈fst (snd p),〈fst p,h (S (fst (snd p))) (fst p)〉〉.
+
+lemma compl_g4 : ∀h,s1,s. (∀x. 0 < s1 x) →
+ (CF s1
+ (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))) →
+ (O s (λx.MSC x + ((snd(snd x) - (fst x)))*Max_{i ∈[S(fst x) ,S(snd (snd x))[ }(s1 〈i,x〉))) →
+ CF s (λp:ℕ.min_input h (fst p) (snd (snd p))).
+#h #s1 #s #pos_s1 #Hs1 #HO @(ext_CF (min_input_aux h))
+ [#n whd in ⊢ (??%%); @min_input_eq]
+@(CF_mu4 … MSC MSC … pos_s1 … Hs1)
+ [@CF_compS @CF_fst
+ |@CF_comp_snd @CF_snd
+ |@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //]
+qed.
+
+(* ??? *)
+
+lemma coroll: ∀s1:nat→nat. (∀n. monotonic ℕ le (λa:ℕ.s1 〈a,n〉)) →
+O (λx.(snd (snd x)-fst x)*(s1 〈snd (snd x),x〉))
+ (λx.∑_{i ∈[S(fst x) ,S(snd (snd x))[ }(s1 〈i,x〉)).
+#s1 #Hs1 %{1} %{0} #n #_ >commutative_times <times_n_1
+@(transitive_le … (sigma_bound …)) [@Hs1|>minus_S_S //]
+qed.
+
+lemma coroll2: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) →
+O (λx.(snd x - fst x)*s1 〈fst x,x〉) (λx.∑_{i ∈[fst x,snd x[ }(s1 〈i,x〉)).
+#s1 #Hs1 %{1} %{0} #n #_ >commutative_times <times_n_1
+@(transitive_le … (sigma_bound_decr …)) [2://] @Hs1
+qed.
+
+lemma coroll3: ∀s1:nat→nat. (∀n,a,b. a ≤ b → b < snd n → s1 〈b,n〉 ≤ s1 〈a,n〉) →
+O (λx.(snd x - fst x)*s1 〈fst x,x〉)
+ (λx.(snd x - fst x)*Max_{i ∈[fst x,snd x[ }(s1 〈i,x〉)).
+#s1 #Hs1 @le_to_O #i @le_times // @Max_le #a #lta #_ @Hs1 //
+@lt_minus_to_plus_r //
+qed.
+
+(**************************** end of technical lemmas *************************)
+
+lemma compl_g5 : ∀h,s1.
+ (∀n. 0 < s1 n) →
+ (∀n. monotonic ℕ le (λa:ℕ.s1 〈a,n〉)) →
+ (CF s1
+ (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))) →
+ CF (λx.MSC x + (snd (snd x)-fst x)*s1 〈snd (snd x),x〉)
+ (λp:ℕ.min_input h (fst p) (snd (snd p))).
+#h #s1 #Hpos #Hmono #Hs1 @(compl_g4 … Hpos Hs1) @O_plus
+[@O_plus_l // |@O_plus_r @le_to_O #n @le_times //
+@Max_le #i #lti #_ @Hmono @le_S_S_to_le @lt_minus_to_plus_r @lti
+qed.
+
+lemma compl_g6: ∀h.
+ constructible (λx. h (fst x) (snd x)) →
+ (CF (λx. sU 〈max (fst (snd x)) (snd (snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉)
+ (λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))).
+#h #hconstr @(ext_CF (termb_aux h))
+ [#n normalize >fst_pair >snd_pair >fst_pair >snd_pair // ]
+@(CF_comp … (λx.MSC x + h (S (fst (snd x))) (fst x)) … CF_termb)
+ [@CF_comp_pair
+ [@CF_comp_fst @(monotonic_CF … CF_snd) #x //
+ |@CF_comp_pair
+ [@(monotonic_CF … CF_fst) #x //
+ |@(ext_CF ((λx.h (fst x) (snd x))∘(λx.〈S (fst (snd x)),fst x〉)))
+ [#n normalize >fst_pair >snd_pair %]
+ @(CF_comp … MSC …hconstr)
+ [@CF_comp_pair [@CF_compS @CF_comp_fst // |//]
+ |@O_plus @le_to_O [//|#n >fst_pair >snd_pair //]
+ ]
+ ]
+ ]
+ |@O_plus
+ [@O_plus
+ [@(O_trans … (λx.MSC (fst x) + MSC (max (fst (snd x)) (snd (snd x)))))
+ [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx
+ >fst_pair >snd_pair @(transitive_le … (MSC_pair …))
+ >distributive_times_plus @le_plus [//]
+ cases (surj_pair b) #c * #d #eqb >eqb
+ >fst_pair >snd_pair @(transitive_le … (MSC_pair …))
+ whd in ⊢ (??%); @le_plus
+ [@monotonic_MSC @(le_maxl … (le_n …))
+ |>commutative_times <times_n_1 @monotonic_MSC @(le_maxr … (le_n …))
+ ]
+ |@O_plus [@le_to_O #x @sU_le_x |@le_to_O #x @sU_le_i]
+ ]
+ |@le_to_O #n @sU_le
+ ]
+ |@le_to_O #x @monotonic_sU // @(le_maxl … (le_n …)) ]
+ ]
+qed.
+
+definition big : nat →nat ≝ λx.
+ let m ≝ max (fst x) (snd x) in 〈m,m〉.
+
+lemma big_def : ∀a,b. big 〈a,b〉 = 〈max a b,max a b〉.
+#a #b normalize >fst_pair >snd_pair // qed.
+
+lemma le_big : ∀x. x ≤ big x.
+#x cases (surj_pair x) #a * #b #eqx >eqx @le_pair >fst_pair >snd_pair
+[@(le_maxl … (le_n …)) | @(le_maxr … (le_n …))]
+qed.
+
+definition faux2 ≝ λh.
+ (λx.MSC x + (snd (snd x)-fst x)*
+ (λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉).
+
+lemma compl_g7: ∀h.
+ constructible (λx. h (fst x) (snd x)) →
+ (∀n. monotonic ? le (h n)) →
+ CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉)
+ (λp:ℕ.min_input h (fst p) (snd (snd p))).
+#h #hcostr #hmono @(monotonic_CF … (faux2 h))
+ [#n normalize >fst_pair >snd_pair //]
+@compl_g5 [#n @sU_pos | 3:@(compl_g6 h hcostr)] #n #x #y #lexy >fst_pair >snd_pair
+>fst_pair >snd_pair @monotonic_sU // @hmono @lexy
+qed.
+
+lemma compl_g71: ∀h.
+ constructible (λx. h (fst x) (snd x)) →
+ (∀n. monotonic ? le (h n)) →
+ CF (λx.MSC (big x) + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉)
+ (λp:ℕ.min_input h (fst p) (snd (snd p))).
+#h #hcostr #hmono @(monotonic_CF … (compl_g7 h hcostr hmono)) #x
+@le_plus [@monotonic_MSC //]
+cases (decidable_le (fst x) (snd(snd x)))
+ [#Hle @le_times // @monotonic_sU
+ |#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt]
+ ]
+qed.
+
+definition out_aux ≝ λh.
+ out_unary ∘ λp.〈fst p,〈snd(snd p),h (S (fst p)) (snd (snd p))〉〉.
+
+lemma compl_g8: ∀h.
+ constructible (λx. h (fst x) (snd x)) →
+ (CF (λx. sU 〈max (fst x) (snd x),〈snd(snd x),h (S (fst x)) (snd(snd x))〉〉)
+ (λp.out (fst p) (snd (snd p)) (h (S (fst p)) (snd (snd p))))).
+#h #hconstr @(ext_CF (out_aux h))
+ [#n normalize >fst_pair >snd_pair >fst_pair >snd_pair // ]
+@(CF_comp … (λx.h (S (fst x)) (snd(snd x)) + MSC x) … CF_out)
+ [@CF_comp_pair
+ [@(monotonic_CF … CF_fst) #x //
+ |@CF_comp_pair
+ [@CF_comp_snd @(monotonic_CF … CF_snd) #x //
+ |@(ext_CF ((λx.h (fst x) (snd x))∘(λx.〈S (fst x),snd(snd x)〉)))
+ [#n normalize >fst_pair >snd_pair %]
+ @(CF_comp … MSC …hconstr)
+ [@CF_comp_pair [@CF_compS // | @CF_comp_snd // ]
+ |@O_plus @le_to_O [//|#n >fst_pair >snd_pair //]
+ ]
+ ]
+ ]
+ |@O_plus
+ [@O_plus
+ [@le_to_O #n @sU_le
+ |@(O_trans … (λx.MSC (max (fst x) (snd x))))
+ [%{2} %{0} #x #_ cases (surj_pair x) #a * #b #eqx >eqx
+ >fst_pair >snd_pair @(transitive_le … (MSC_pair …))
+ whd in ⊢ (??%); @le_plus
+ [@monotonic_MSC @(le_maxl … (le_n …))
+ |>commutative_times <times_n_1 @monotonic_MSC @(le_maxr … (le_n …))
+ ]
+ |@le_to_O #x @(transitive_le ???? (sU_le_i … )) //
+ ]
+ ]
+ |@le_to_O #x @monotonic_sU [@(le_maxl … (le_n …))|//|//]
+ ]
+qed.
+
+lemma compl_g9 : ∀h.
+ constructible (λx. h (fst x) (snd x)) →
+ (∀n. monotonic ? le (h n)) →
+ (∀n,a,b. a ≤ b → b ≤ n → h b n ≤ h a n) →
+ CF (λx. (S (snd x-fst x))*MSC 〈x,x〉 +
+ (snd x-fst x)*(S(snd x-fst x))*sU 〈x,〈snd x,h (S (fst x)) (snd x)〉〉)
+ (auxg h).
+#h #hconstr #hmono #hantimono
+@(compl_g2 h ??? (compl_g3 … (compl_g71 h hconstr hmono)) (compl_g8 h hconstr))
+@O_plus
+ [@O_plus_l @le_to_O #x >(times_n_1 (MSC x)) >commutative_times @le_times
+ [// | @monotonic_MSC // ]]
+@(O_trans … (coroll3 ??))
+ [#n #a #b #leab #ltb >fst_pair >fst_pair >snd_pair >snd_pair
+ cut (b ≤ n) [@(transitive_le … (le_snd …)) @lt_to_le //] #lebn
+ cut (max a n = n)
+ [normalize >le_to_leb_true [//|@(transitive_le … leab lebn)]] #maxa
+ cut (max b n = n) [normalize >le_to_leb_true //] #maxb
+ @le_plus
+ [@le_plus [>big_def >big_def >maxa >maxb //]
+ @le_times
+ [/2 by monotonic_le_minus_r/
+ |@monotonic_sU // @hantimono [@le_S_S // |@ltb]
+ ]
+ |@monotonic_sU // @hantimono [@le_S_S // |@ltb]
+ ]
+ |@le_to_O #n >fst_pair >snd_pair
+ cut (max (fst n) n = n) [normalize >le_to_leb_true //] #Hmax >Hmax
+ >associative_plus >distributive_times_plus
+ @le_plus [@le_times [@le_S // |>big_def >Hmax //] |//]
+ ]
+qed.
+
+definition c ≝ λx.(S (snd x-fst x))*MSC 〈x,x〉.
+
+definition sg ≝ λh,x.
+ let a ≝ fst x in
+ let b ≝ snd x in
+ c x + (b-a)*(S(b-a))*sU 〈x,〈snd x,h (S a) b〉〉.
+
+lemma sg_def1 : ∀h,a,b.
+ sg h 〈a,b〉 = c 〈a,b〉 +
+ (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉.
+#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair //
+qed.
+
+lemma sg_def : ∀h,a,b.
+ sg h 〈a,b〉 =
+ S (b-a)*MSC 〈〈a,b〉,〈a,b〉〉 + (b-a)*(S(b-a))*sU 〈〈a,b〉,〈b,h (S a) b〉〉.
+#h #a #b normalize >fst_pair >snd_pair //
+qed.
+
+lemma compl_g11 : ∀h.
+ constructible (λx. h (fst x) (snd x)) →
+ (∀n. monotonic ? le (h n)) →
+ (∀n,a,b. a ≤ b → b ≤ n → h b n ≤ h a n) →
+ CF (sg h) (unary_g h).
+#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham)
+qed.
+
+(**************************** closing the argument ****************************)
+
+let rec h_of_aux (r:nat →nat) (c,d,b:nat) on d : nat ≝
+ match d with
+ [ O ⇒ c
+ | S d1 ⇒ (S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) +
+ d*(S d)*sU 〈〈b-d,b〉,〈b,r (h_of_aux r c d1 b)〉〉].
+
+lemma h_of_aux_O: ∀r,c,b.
+ h_of_aux r c O b = c.
+// qed.
+
+lemma h_of_aux_S : ∀r,c,d,b.
+ h_of_aux r c (S d) b =
+ (S (S d))*(MSC 〈〈b-(S d),b〉,〈b-(S d),b〉〉) +
+ (S d)*(S (S d))*sU 〈〈b-(S d),b〉,〈b,r(h_of_aux r c d b)〉〉.
+// qed.
+
+definition h_of ≝ λr,p.
+ let m ≝ max (fst p) (snd p) in
+ h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (snd p - fst p) (snd p).
+
+lemma h_of_O: ∀r,a,b. b ≤ a →
+ h_of r 〈a,b〉 = let m ≝ max a b in MSC 〈〈m,m〉,〈m,m〉〉.
+#r #a #b #Hle normalize >fst_pair >snd_pair >(minus_to_0 … Hle) //
+qed.
+
+lemma h_of_def: ∀r,a,b.h_of r 〈a,b〉 =
+ let m ≝ max a b in
+ h_of_aux r (MSC 〈〈m,m〉,〈m,m〉〉) (b - a) b.
+#r #a #b normalize >fst_pair >snd_pair //
+qed.
+
+lemma mono_h_of_aux: ∀r.(∀x. x ≤ r x) → monotonic ? le r →
+ ∀d,d1,c,c1,b,b1.c ≤ c1 → d ≤ d1 → b ≤ b1 →
+ h_of_aux r c d b ≤ h_of_aux r c1 d1 b1.
+#r #Hr #monor #d #d1 lapply d -d elim d1
+ [#d #c #c1 #b #b1 #Hc #Hd @(le_n_O_elim ? Hd) #leb
+ >h_of_aux_O >h_of_aux_O //
+ |#m #Hind #d #c #c1 #b #b1 #lec #led #leb cases (le_to_or_lt_eq … led)
+ [#ltd @(transitive_le … (Hind … lec ? leb)) [@le_S_S_to_le @ltd]
+ >h_of_aux_S @(transitive_le ???? (le_plus_n …))
+ >(times_n_1 (h_of_aux r c1 m b1)) in ⊢ (?%?);
+ >commutative_times @le_times [//|@(transitive_le … (Hr ?)) @sU_le]
+ |#Hd >Hd >h_of_aux_S >h_of_aux_S
+ cut (b-S m ≤ b1 - S m) [/2 by monotonic_le_minus_l/] #Hb1
+ @le_plus [@le_times //]
+ [@monotonic_MSC @le_pair @le_pair //
+ |@le_times [//] @monotonic_sU
+ [@le_pair // |// |@monor @Hind //]
+ ]
+ ]
+ ]
+qed.
+
+lemma mono_h_of2: ∀r.(∀x. x ≤ r x) → monotonic ? le r →
+ ∀i,b,b1. b ≤ b1 → h_of r 〈i,b〉 ≤ h_of r 〈i,b1〉.
+#r #Hr #Hmono #i #a #b #leab >h_of_def >h_of_def
+cut (max i a ≤ max i b)
+ [@to_max
+ [@(le_maxl … (le_n …))|@(transitive_le … leab) @(le_maxr … (le_n …))]]
+#Hmax @(mono_h_of_aux r Hr Hmono)
+ [@monotonic_MSC @le_pair @le_pair @Hmax |/2 by monotonic_le_minus_l/ |@leab]
+qed.
+
+axiom h_of_constr : ∀r:nat →nat.
+ (∀x. x ≤ r x) → monotonic ? le r → constructible r →
+ constructible (h_of r).
+
+lemma speed_compl: ∀r:nat →nat.
+ (∀x. x ≤ r x) → monotonic ? le r → constructible r →
+ CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))).
+#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …))
+ [#x cases (surj_pair x) #a * #b #eqx >eqx
+ >sg_def cases (decidable_le b a)
+ [#leba >(minus_to_0 … leba) normalize in ⊢ (?%?);
+ <plus_n_O <plus_n_O >h_of_def
+ cut (max a b = a)
+ [normalize cases (le_to_or_lt_eq … leba)
+ [#ltba >(lt_to_leb_false … ltba) %
+ |#eqba <eqba >(le_to_leb_true … (le_n ?)) % ]]
+ #Hmax >Hmax normalize >(minus_to_0 … leba) normalize
+ @monotonic_MSC @le_pair @le_pair //
+ |#ltab >h_of_def >h_of_def
+ cut (max a b = b)
+ [normalize >(le_to_leb_true … ) [%] @lt_to_le @not_le_to_lt @ltab]
+ #Hmax >Hmax
+ cut (max (S a) b = b)
+ [whd in ⊢ (??%?); >(le_to_leb_true … ) [%] @not_le_to_lt @ltab]
+ #Hmax1 >Hmax1
+ cut (∃d.b - a = S d)
+ [%{(pred(b-a))} >S_pred [//] @lt_plus_to_minus_r @not_le_to_lt @ltab]
+ * #d #eqd >eqd
+ cut (b-S a = d) [//] #eqd1 >eqd1 >h_of_aux_S >eqd1
+ cut (b - S d = a)
+ [@plus_to_minus >commutative_plus @minus_to_plus
+ [@lt_to_le @not_le_to_lt // | //]] #eqd2 >eqd2
+ normalize //
+ ]
+ |#n #a #b #leab #lebn >h_of_def >h_of_def
+ cut (max a n = n)
+ [normalize >le_to_leb_true [%|@(transitive_le … leab lebn)]] #Hmaxa
+ cut (max b n = n)
+ [normalize >(le_to_leb_true … lebn) %] #Hmaxb
+ >Hmaxa >Hmaxb @Hmono @(mono_h_of_aux r … Hr Hmono) // /2 by monotonic_le_minus_r/
+ |#n #a #b #leab @Hmono @(mono_h_of2 … Hr Hmono … leab)
+ |@(constr_comp … Hconstr Hr) @(ext_constr (h_of r))
+ [#x cases (surj_pair x) #a * #b #eqx >eqx >fst_pair >snd_pair //]
+ @(h_of_constr r Hr Hmono Hconstr)
+ ]
+qed.
+
+lemma speed_compl_i: ∀r:nat →nat.
+ (∀x. x ≤ r x) → monotonic ? le r → constructible r →
+ ∀i. CF (λx.h_of r 〈i,x〉) (λx.g (λi,x. r(h_of r 〈i,x〉)) i x).
+#r #Hr #Hmono #Hconstr #i
+@(ext_CF (λx.unary_g (λi,x. r(h_of r 〈i,x〉)) 〈i,x〉))
+ [#n whd in ⊢ (??%%); @eq_f @sym_eq >fst_pair >snd_pair %]
+@smn @(ext_CF … (speed_compl r Hr Hmono Hconstr)) #n //
+qed.
+
+(**************************** the speedup theorem *****************************)
+theorem pseudo_speedup:
+ ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r →
+ ∃f.∀sf. CF sf f → ∃g,sg. f ≈ g ∧ CF sg g ∧ O sf (r ∘ sg).
+(* ∃m,a.∀n. a≤n → r(sg a) < m * sf n. *)
+#r #Hr #Hmono #Hconstr
+(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *)
+%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #i *
+#Hcodei #HCi
+(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *)
+%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))}
+(* sg is (λx.h_of r 〈i,x〉) *)
+%{(λx. h_of r 〈S i,x〉)}
+lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg
+%[%[@condition_1 |@Hg]
+ |cases Hg #H1 * #j * #Hcodej #HCj
+ lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *)
+ cases HCi #m * #a #Ha %{m} %{(max (S i) a)} #n #ltin @lt_to_le @not_le_to_lt
+ @(not_to_not … Hcond2) -Hcond2 #Hlesf %{n} %
+ [@(transitive_le … ltin) @(le_maxl … (le_n …))]
+ cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))]
+ #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf) //
+ ]
+qed.
+
+theorem pseudo_speedup':
+ ∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r →
+ ∃f.∀sf. CF sf f → ∃g,sg. f ≈ g ∧ CF sg g ∧
+ (* ¬ O (r ∘ sg) sf. *)
+ ∃m,a.∀n. a≤n → r(sg a) < m * sf n.
+#r #Hr #Hmono #Hconstr
+(* f is (g (λi,x. r(h_of r 〈i,x〉)) 0) *)
+%{(g (λi,x. r(h_of r 〈i,x〉)) 0)} #sf * #i *
+#Hcodei #HCi
+(* g is (g (λi,x. r(h_of r 〈i,x〉)) (S i)) *)
+%{(g (λi,x. r(h_of r 〈i,x〉)) (S i))}
+(* sg is (λx.h_of r 〈S i,x〉) *)
+%{(λx. h_of r 〈S i,x〉)}
+lapply (speed_compl_i … Hr Hmono Hconstr (S i)) #Hg
+%[%[@condition_1 |@Hg]
+ |cases Hg #H1 * #j * #Hcodej #HCj
+ lapply (condition_2 … Hcodei) #Hcond2 (* @not_to_not *)
+ cases HCi #m * #a #Ha
+ %{m} %{(max (S i) a)} #n #ltin @not_le_to_lt @(not_to_not … Hcond2) -Hcond2
+ #Hlesf %{n} % [@(transitive_le … ltin) @(le_maxl … (le_n …))]
+ cases (Ha n ?) [2: @(transitive_le … ltin) @(le_maxr … (le_n …))]
+ #y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf)
+ @Hmono @(mono_h_of2 … Hr Hmono … ltin)
+ ]
+qed.
+
\ No newline at end of file
include "arithmetics/sigma_pi.ma".
include "arithmetics/bounded_quantifiers.ma".
include "reverse_complexity/big_O.ma".
-include "basics/core_notation/napart_2.ma".
-(************************* notation for minimization *****************************)
+(************************* notation for minimization **************************)
notation "μ_{ ident i < n } p"
with precedence 80 for @{min $n 0 (λ${ident i}.$p)}.
axiom le_snd : ∀p. snd p ≤ p.
axiom le_pair: ∀a,a1,b,b1. a ≤ a1 → b ≤ b1 → 〈a,b〉 ≤ 〈a1,b1〉.
+lemma expand_pair: ∀x. x = 〈fst x, snd x〉.
+#x lapply (surj_pair x) * #a * #b #Hx >Hx >fst_pair >snd_pair //
+qed.
+
(************************************* U **************************************)
axiom U: nat → nat →nat → option nat.
definition bool_to_nat: bool → nat ≝
λb. match b with [true ⇒ 1 | false ⇒ 0].
-
+
coercion bool_to_nat.
definition pU : nat → nat → nat → nat ≝ λi,x,r.〈termb i x r,out i x r〉.
#i #x #r normalize cases (U i x r) normalize >snd_pair //
qed.
-(********************************* the speedup ********************************)
-
-definition min_input ≝ λh,i,x. μ_{y ∈ [S i,x] } (termb i y (h (S i) y)).
-
-lemma min_input_def : ∀h,i,x.
- min_input h i x = min (x -i) (S i) (λy.termb i y (h (S i) y)).
-// qed.
-
-lemma min_input_i: ∀h,i,x. x ≤ i → min_input h i x = S i.
-#h #i #x #lexi >min_input_def
-cut (x - i = 0) [@sym_eq /2 by eq_minus_O/] #Hcut //
-qed.
-lemma min_input_to_terminate: ∀h,i,x.
- min_input h i x = x → {i ⊙ x} ↓ (h (S i) x).
-#h #i #x #Hminx
-cases (decidable_le (S i) x) #Hix
- [cases (true_or_false (termb i x (h (S i) x))) #Hcase
- [@termb_true_to_term //
- |<Hminx in Hcase; #H lapply (fmin_false (λx.termb i x (h (S i) x)) (x-i) (S i) H)
- >min_input_def in Hminx; #Hminx >Hminx in ⊢ (%→?);
- <plus_n_Sm <plus_minus_m_m [2: @lt_to_le //]
- #Habs @False_ind /2/
- ]
- |@False_ind >min_input_i in Hminx;
- [#eqix >eqix in Hix; * /2/ | @le_S_S_to_le @not_le_to_lt //]
- ]
-qed.
-
-lemma min_input_to_lt: ∀h,i,x.
- min_input h i x = x → i < x.
-#h #i #x #Hminx cases (decidable_le (S i) x) //
-#ltxi @False_ind >min_input_i in Hminx;
- [#eqix >eqix in ltxi; * /2/ | @le_S_S_to_le @not_le_to_lt //]
-qed.
-
-lemma le_to_min_input: ∀h,i,x,x1. x ≤ x1 →
- min_input h i x = x → min_input h i x1 = x.
-#h #i #x #x1 #lex #Hminx @(min_exists … (le_S_S … lex))
- [@(fmin_true … (sym_eq … Hminx)) //
- |@(min_input_to_lt … Hminx)
- |#j #H1 <Hminx @lt_min_to_false //
- |@plus_minus_m_m @le_S_S @(transitive_le … lex) @lt_to_le
- @(min_input_to_lt … Hminx)
- ]
-qed.
-
-definition g ≝ λh,u,x.
- S (max_{i ∈[u,x[ | eqb (min_input h i x) x} (out i x (h (S i) x))).
-
-lemma g_def : ∀h,u,x. g h u x =
- S (max_{i ∈[u,x[ | eqb (min_input h i x) x} (out i x (h (S i) x))).
-// qed.
-
-lemma le_u_to_g_1 : ∀h,u,x. x ≤ u → g h u x = 1.
-#h #u #x #lexu >g_def cut (x-u = 0) [/2 by minus_le_minus_minus_comm/]
-#eq0 >eq0 normalize // qed.
-
-lemma g_lt : ∀h,i,x. min_input h i x = x →
- out i x (h (S i) x) < g h 0 x.
-#h #i #x #H @le_S_S @(le_MaxI … i) /2 by min_input_to_lt/
-qed.
-
-lemma max_neq0 : ∀a,b. max a b ≠ 0 → a ≠ 0 ∨ b ≠ 0.
-#a #b whd in match (max a b); cases (true_or_false (leb a b)) #Hcase >Hcase
- [#H %2 @H | #H %1 @H]
-qed.
-
-definition almost_equal ≝ λf,g:nat → nat. ¬ ∀nu.∃x. nu < x ∧ f x ≠ g x.
-interpretation "almost equal" 'napart f g = (almost_equal f g).
-
-lemma eventually_cancelled: ∀h,u.¬∀nu.∃x. nu < x ∧
- max_{i ∈ [0,u[ | eqb (min_input h i x) x} (out i x (h (S i) x)) ≠ 0.
-#h #u elim u
- [normalize % #H cases (H u) #x * #_ * #H1 @H1 //
- |#u0 @not_to_not #Hind #nu cases (Hind nu) #x * #ltx
- cases (true_or_false (eqb (min_input h (u0+O) x) x)) #Hcase
- [>bigop_Strue [2:@Hcase] #Hmax cases (max_neq0 … Hmax) -Hmax
- [2: #H %{x} % // <minus_n_O @H]
- #Hneq0 (* if x is not enough we retry with nu=x *)
- cases (Hind x) #x1 * #ltx1
- >bigop_Sfalse
- [#H %{x1} % [@transitive_lt //| <minus_n_O @H]
- |@not_eq_to_eqb_false >(le_to_min_input … (eqb_true_to_eq … Hcase))
- [@lt_to_not_eq @ltx1 | @lt_to_le @ltx1]
- ]
- |>bigop_Sfalse [2:@Hcase] #H %{x} % // <minus_n_O @H
- ]
- ]
-qed.
-
-lemma condition_1: ∀h,u.g h 0 ≈ g h u.
-#h #u @(not_to_not … (eventually_cancelled h u))
-#H #nu cases (H (max u nu)) #x * #ltx #Hdiff
-%{x} % [@(le_to_lt_to_lt … ltx) @(le_maxr … (le_n …))] @(not_to_not … Hdiff)
-#H @(eq_f ?? S) >(bigop_sumI 0 u x (λi:ℕ.eqb (min_input h i x) x) nat 0 MaxA)
- [>H // |@lt_to_le @(le_to_lt_to_lt …ltx) /2 by le_maxr/ |//]
-qed.
-
-(******************************** Condition 2 *********************************)
definition total ≝ λf.λx:nat. Some nat (f x).
-
-lemma exists_to_exists_min: ∀h,i. (∃x. i < x ∧ {i ⊙ x} ↓ h (S i) x) → ∃y. min_input h i y = y.
-#h #i * #x * #ltix #Hx %{(min_input h i x)} @min_spec_to_min @found //
- [@(f_min_true (λy:ℕ.termb i y (h (S i) y))) %{x} % [% // | @term_to_termb_true //]
- |#y #leiy #lty @(lt_min_to_false ????? lty) //
- ]
-qed.
-
-lemma condition_2: ∀h,i. code_for (total (g h 0)) i → ¬∃x. i<x ∧ {i ⊙ x} ↓ h (S i) x.
-#h #i whd in ⊢(%→?); #H % #H1 cases (exists_to_exists_min … H1) #y #Hminy
-lapply (g_lt … Hminy)
-lapply (min_input_to_terminate … Hminy) * #r #termy
-cases (H y) -H #ny #Hy
-cut (r = g h 0 y) [@(unique_U … ny … termy) @Hy //] #Hr
-whd in match (out ???); >termy >Hr
-#H @(absurd ? H) @le_to_not_lt @le_n
-qed.
(********************************* complexity *********************************)
[#x cases (Hcode x) #a #H %{a} whd in match (total ??); <Hext @H | //]
qed.
+lemma ext_CF1 : ∀f,g,s. (∀n. f n = g n) → CF s g → CF s f.
+#f #g #s #Hext * #HO * #i * #Hcode #HC % // %{i} %
+ [#x cases (Hcode x) #a #H %{a} whd in match (total ??); >Hext @H | //]
+qed.
+
lemma monotonic_CF: ∀s1,s2,f.(∀x. s1 x ≤ s2 x) → CF s1 f → CF s2 f.
#s1 #s2 #f #H * #HO * #i * #Hcode #Hs1 %
[cases HO #c * #a -HO #HO %{c} %{a} #n #lean @(transitive_le … (HO n lean))
(* primitve recursion *)
+(* no arguments *)
+
+let rec prim_rec0 (k:nat) (h:nat →nat) n on n ≝
+ match n with
+ [ O ⇒ k
+ | S a ⇒ h 〈a, prim_rec0 k h a〉].
+
+lemma prim_rec0_S: ∀k,h,n.
+ prim_rec0 k h (S n) = h 〈n, prim_rec0 k h n〉.
+// qed.
+
+let rec prim_rec0_compl (k,sk:nat) (h,sh:nat →nat) n on n ≝
+ match n with
+ [ O ⇒ sk
+ | S a ⇒ prim_rec0_compl k sk h sh a + sh (prim_rec0 k h a)].
+
+axiom CF_prim_rec0_gen: ∀k,h,sk,sh,sh1,sf. CF sh h →
+ O sh1 (λx.snd x + sh 〈fst x,prim_rec0 k h (fst x)〉) →
+ O sf (prim_rec0 sk sh1) →
+ CF sf (prim_rec0 k h).
+
+lemma CF_prim_rec0: ∀k,h,sk,sh,sf. CF sh h →
+ O sf (prim_rec0 sk (λx. snd x + sh 〈fst x,prim_rec0 k h (fst x)〉))
+ → CF sf (prim_rec0 k h).
+#k #h #sk #sh #sf #Hh #HO @(CF_prim_rec0_gen … Hh … HO) @O_refl
+qed.
+
+(* with argument(s) m *)
let rec prim_rec (k,h:nat →nat) n m on n ≝
match n with
[ O ⇒ k m
prim_rec k h (S n) m = h 〈n,〈prim_rec k h n m, m〉〉.
// qed.
+lemma prim_rec_le: ∀k1,k2,h1,h2. (∀x.k1 x ≤ k2 x) →
+(∀x,y.x ≤y → h1 x ≤ h2 y) →
+ ∀x,m. prim_rec k1 h1 x m ≤ prim_rec k2 h2 x m.
+#k1 #k2 #h1 #h2 #lek #leh #x #m elim x //
+#n #Hind normalize @leh @le_pair // @le_pair //
+qed.
+
definition unary_pr ≝ λk,h,x. prim_rec k h (fst x) (snd x).
+lemma prim_rec_unary: ∀k,h,a,b.
+ prim_rec k h a b = unary_pr k h 〈a,b〉.
+#k #h #a #b normalize >fst_pair >snd_pair //
+qed.
+
+
let rec prim_rec_compl (k,h,sk,sh:nat →nat) n m on n ≝
match n with
[ O ⇒ sk m
| S a ⇒ prim_rec_compl k h sk sh a m + sh (prim_rec k h a m)].
-
-axiom CF_prim_rec: ∀k,h,sk,sh,sf. CF sk k → CF sh h →
- O sf (unary_pr sk (λx. fst (snd x) + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉))
- → CF sf (unary_pr k h).
-(* falso ????
-lemma prim_rec_O: ∀k1,h1,k2,h2. O k1 k2 → O h1 h2 →
- O (unary_pr k1 h1) (unary_pr k2 h2).
-#k1 #h1 #k2 #h2 #HO1 #HO2 whd *)
+axiom CF_prim_rec_gen: ∀k,h,sk,sh,sh1. CF sk k → CF sh h →
+ O sh1 (λx. fst (snd x) + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉) →
+ CF (unary_pr sk sh1) (unary_pr k h).
+lemma CF_prim_rec: ∀k,h,sk,sh,sf. CF sk k → CF sh h →
+ O sf (unary_pr sk (λx. fst (snd x) + sh 〈fst x,〈unary_pr k h 〈fst x,snd (snd x)〉,snd (snd x)〉〉))
+ → CF sf (unary_pr k h).
+#k #h #sk #sh #sf #Hk #Hh #Osf @(O_to_CF … Osf) @(CF_prim_rec_gen … Hk Hh) @O_refl
+qed.
(**************************** primitive operations*****************************)
definition id ≝ λx:nat.x.
axiom CF_id: CF MSC id.
+axiom CF_const: ∀k.CF MSC (λx.k).
axiom CF_compS: ∀h,f. CF h f → CF h (S ∘ f).
axiom CF_comp_fst: ∀h,f. CF h f → CF h (fst ∘ f).
axiom CF_comp_snd: ∀h,f. CF h f → CF h (snd ∘ f).
lemma CF_snd: CF MSC snd.
@(ext_CF (snd ∘ id)) [#n //] @(CF_comp_snd … CF_id)
-qed.
+qed.
(************************************** eqb ***********************************)
(*********************************** maximum **********************************)
+alias symbol "pair" (instance 13) = "abstract pair".
+alias symbol "pair" (instance 12) = "abstract pair".
+alias symbol "and" (instance 11) = "boolean and".
+alias symbol "plus" (instance 8) = "natural plus".
+alias symbol "pair" (instance 7) = "abstract pair".
+alias symbol "plus" (instance 6) = "natural plus".
+alias symbol "pair" (instance 5) = "abstract pair".
+alias id "max" = "cic:/matita/arithmetics/nat/max#def:2".
+alias symbol "minus" (instance 3) = "natural minus".
+lemma max_gen: ∀a,b,p,f,x. a ≤b →
+ max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) = max_{i < b | leb a i ∧ p 〈i,x〉 }(f 〈i,x〉).
+#a #b #p #f #x @(bigop_I_gen ????? MaxA)
+qed.
+
+lemma max_prim_rec_base: ∀a,b,p,f,x. a ≤b →
+ max_{i < b| p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x.
+#a #b #p #f #x #leab >max_gen // elim b
+ [normalize //
+ |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma max_prim_rec: ∀a,b,p,f,x. a ≤b →
+ max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if leb a (fst i) ∧ p 〈fst i,x〉 then max (f 〈fst i,snd (snd i)〉) (fst (snd i)) else fst (snd i)) b x.
+#a #b #p #f #x #leab >max_gen // elim b
+ [normalize //
+ |#i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair
+ cases (true_or_false (leb a i ∧ p 〈i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+(* FG: aliases added by matita compiled with OCaml 4.0.5, bah ??? *)
+alias symbol "pair" (instance 15) = "abstract pair".
+alias symbol "minus" (instance 14) = "natural minus".
+alias symbol "plus" (instance 11) = "natural plus".
+alias symbol "pair" (instance 10) = "abstract pair".
+alias symbol "plus" (instance 13) = "natural plus".
+alias symbol "pair" (instance 12) = "abstract pair".
+alias symbol "plus" (instance 8) = "natural plus".
+alias symbol "pair" (instance 7) = "abstract pair".
+alias symbol "plus" (instance 6) = "natural plus".
+alias symbol "pair" (instance 5) = "abstract pair".
+alias id "max" = "cic:/matita/arithmetics/nat/max#def:2".
+alias symbol "minus" (instance 3) = "natural minus".
+lemma max_prim_rec1: ∀a,b,p,f,x.
+ max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉
+ then max (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (b x-a x) 〈a x ,x〉.
+#a #b #p #f #x elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma sum_prim_rec1: ∀a,b,p,f,x.
+ ∑_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉
+ then plus (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (b x-a x) 〈a x ,x〉.
+#a #b #p #f #x elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+(* FG: aliases added by matita compiled with OCaml 4.0.5, bah ??? *)
+alias symbol "pair" (instance 15) = "abstract pair".
+alias symbol "minus" (instance 14) = "natural minus".
+alias symbol "plus" (instance 11) = "natural plus".
+alias symbol "pair" (instance 10) = "abstract pair".
+alias symbol "plus" (instance 13) = "natural plus".
+alias symbol "pair" (instance 12) = "abstract pair".
+alias symbol "plus" (instance 8) = "natural plus".
+alias symbol "pair" (instance 7) = "abstract pair".
+alias symbol "pair" (instance 6) = "abstract pair".
+alias symbol "plus" (instance 4) = "natural plus".
+alias symbol "pair" (instance 3) = "abstract pair".
+alias symbol "minus" (instance 2) = "natural minus".
+lemma bigop_prim_rec: ∀a,b,c,p,f,x.
+ bigop (b x-a x) (λi:ℕ.p 〈i+a x,x〉) ? (c 〈a x,x〉) plus (λi:ℕ.f 〈i+a x,x〉) =
+ prim_rec c
+ (λi.if p 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉
+ then plus (f 〈fst i +fst (snd (snd i)),snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (b x-a x) 〈a x ,x〉.
+#a #b #c #p #f #x normalize elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈i+a x,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+(* FG: aliases added by matita compiled with OCaml 4.0.5, bah ??? *)
+alias symbol "pair" (instance 15) = "abstract pair".
+alias symbol "minus" (instance 14) = "natural minus".
+alias symbol "minus" (instance 11) = "natural minus".
+alias symbol "pair" (instance 10) = "abstract pair".
+alias symbol "minus" (instance 13) = "natural minus".
+alias symbol "pair" (instance 12) = "abstract pair".
+alias symbol "minus" (instance 8) = "natural minus".
+alias symbol "pair" (instance 7) = "abstract pair".
+alias symbol "pair" (instance 6) = "abstract pair".
+alias symbol "minus" (instance 4) = "natural minus".
+alias symbol "pair" (instance 3) = "abstract pair".
+alias symbol "minus" (instance 2) = "natural minus".
+lemma bigop_prim_rec_dec: ∀a,b,c,p,f,x.
+ bigop (b x-a x) (λi:ℕ.p 〈b x -i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x-i,x〉) =
+ prim_rec c
+ (λi.if p 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉
+ then plus (f 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (b x-a x) 〈b x ,x〉.
+#a #b #c #p #f #x normalize elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+lemma bigop_prim_rec_dec1: ∀a,b,c,p,f,x.
+ bigop (S(b x)-a x) (λi:ℕ.p 〈b x - i,x〉) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈b x- i,x〉) =
+ prim_rec c
+ (λi.if p 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉
+ then plus (f 〈fst (snd (snd i)) - (fst i),snd (snd (snd i))〉) (fst (snd i))
+ else fst (snd i)) (S(b x)-a x) 〈b x,x〉.
+#a #b #c #p #f #x elim (S(b x)-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ cases (true_or_false (p 〈b x - i,x〉)) #Hcase >Hcase
+ [<Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed.
+
+(*
+lemma bigop_aux_1: ∀k,c,f.
+ bigop (S k) (λi:ℕ.true) ? c plus (λi:ℕ.f i) =
+ f 0 + bigop k (λi:ℕ.true) ? c plus (λi:ℕ.f (S i)).
+#k #c #f elim k [normalize //] #i #Hind >bigop_Strue //
+
+lemma bigop_prim_rec_dec: ∀a,b,c,f,x.
+ bigop (b x-a x) (λi:ℕ.true) ? (c 〈b x,x〉) plus (λi:ℕ.f 〈i+a x,x〉) =
+ prim_rec c
+ (λi. plus (f 〈fst (snd (snd i)) - fst i,snd (snd (snd i))〉) (fst (snd i)))
+ (b x-a x) 〈b x ,x〉.
+#a #b #c #f #x normalize elim (b x-a x)
+ [normalize //
+ |#i #Hind >prim_rec_S
+ >fst_pair >snd_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ <Hind >bigop_Strue // |<Hind >bigop_Sfalse // ]
+ ]
+qed. *)
+
+lemma bigop_plus_c: ∀k,p,f,c.
+ c k + bigop k (λi.p i) ? 0 plus (λi.f i) =
+ bigop k (λi.p i) ? (c k) plus (λi.f i).
+#k #p #f elim k [normalize //]
+#i #Hind #c cases (true_or_false (p i)) #Hcase
+[>bigop_Strue // >bigop_Strue // <associative_plus >(commutative_plus ? (f i))
+ >associative_plus @eq_f @Hind
+|>bigop_Sfalse // >bigop_Sfalse //
+]
+qed.
+
+(* the argument is 〈b-a,〈a,x〉〉 *)
+
+(* FG: aliases added by matita compiled with OCaml 4.0.5, bah ??? *)
+alias symbol "plus" (instance 3) = "natural plus".
+alias symbol "pair" (instance 2) = "abstract pair".
+alias id "max" = "cic:/matita/arithmetics/nat/max#def:2".
+alias symbol "plus" (instance 5) = "natural plus".
+alias symbol "pair" (instance 4) = "abstract pair".
+definition max_unary_pr ≝ λp,f.unary_pr (λx.0)
+ (λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let a ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ if p 〈k + a,x〉 then max (f 〈k+a,x〉) r else r).
+
+lemma max_unary_pr1: ∀a,b,p,f,x.
+ max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉) =
+ ((max_unary_pr p f) ∘ (λx.〈b x - a x,〈a x,x〉〉)) x.
+#a #b #p #f #x normalize >fst_pair >snd_pair @max_prim_rec1
+qed.
+
+(*
+lemma max_unary_pr: ∀a,b,p,f,x.
+ max_{i ∈[a,b[ | p 〈i,x〉 }(f 〈i,x〉) =
+ prim_rec (λi.0)
+ (λi.if p 〈fst i +a,x〉 then max (f 〈fst i +a ,snd (snd i)〉) (fst (snd i)) else fst (snd i)) (b-a) x.
+*)
+
+(*
+definition unary_compl ≝ λp,f,hf.
+ unary_pr MSC
+ (λx:ℕ
+ .fst (snd x)
+ +hf
+ 〈fst x,
+ 〈unary_pr (λx0:ℕ.O)
+ (λi:ℕ
+ .(let (k:ℕ) ≝fst i in
+ let (r:ℕ) ≝fst (snd i) in
+ let (a:ℕ) ≝fst (snd (snd i)) in
+ let (x:ℕ) ≝snd (snd (snd i)) in
+ if p 〈k+a,x〉 then max (f 〈k+a,x〉) r else r )) 〈fst x,snd (snd x)〉,
+ snd (snd x)〉〉). *)
+
+(* FG: aliases added by matita compiled with OCaml 4.0.5, bah ??? *)
+alias symbol "plus" (instance 6) = "natural plus".
+alias symbol "pair" (instance 5) = "abstract pair".
+alias symbol "plus" (instance 4) = "natural plus".
+alias symbol "pair" (instance 3) = "abstract pair".
+alias symbol "plus" (instance 2) = "natural plus".
+alias symbol "plus" (instance 1) = "natural plus".
+definition aux_compl ≝ λhp,hf.λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let a ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hp 〈k+a,x〉 + hf 〈k+a,x〉 + (* was MSC r*) MSC i .
+
+definition aux_compl1 ≝ λhp,hf.λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let a ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r.
+
+lemma aux_compl1_def: ∀k,r,m,hp,hf.
+ aux_compl1 hp hf 〈k,〈r,m〉〉 =
+ let a ≝ fst m in
+ let x ≝ snd m in
+ hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r.
+#k #r #m #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair //
+qed.
+
+lemma aux_compl1_def1: ∀k,r,a,x,hp,hf.
+ aux_compl1 hp hf 〈k,〈r,〈a,x〉〉〉 = hp 〈k+a,x〉 + hf 〈k+a,x〉 + MSC r.
+#k #r #a #x #hp #hf normalize >fst_pair >snd_pair >snd_pair >fst_pair
+>fst_pair >snd_pair //
+qed.
+
+
+axiom Oaux_compl: ∀hp,hf. O (aux_compl1 hp hf) (aux_compl hp hf).
+
+(*
+definition IF ≝ λb,f,g:nat →option nat. λx.
+ match b x with
+ [None ⇒ None ?
+ |Some n ⇒ if (eqb n 0) then f x else g x].
+*)
+
+axiom CF_if: ∀b:nat → bool. ∀f,g,s. CF s b → CF s f → CF s g →
+ CF s (λx. if b x then f x else g x).
+
+(*
+lemma IF_CF: ∀b,f,g,sb,sf,sg. CF sb b → CF sf f → CF sg g →
+ CF (λn. sb n + sf n + sg n) (IF b f g).
+#b #f #g #sb #sf #sg #Hb #Hf #Hg @IF_CF_new
+ [@(monotonic_CF … Hb) @O_plus_l @O_plus_l //
+ |@(monotonic_CF … Hf) @O_plus_l @O_plus_r //
+ |@(monotonic_CF … Hg) @O_plus_r //
+ ]
+qed.
+*)
+
+axiom CF_le: ∀h,f,g. CF h f → CF h g → CF h (λx. leb (f x) (g x)).
+axiom CF_max1: ∀h,f,g. CF h f → CF h g → CF h (λx. max (f x) (g x)).
+axiom CF_plus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x + g x).
+axiom CF_minus: ∀h,f,g. CF h f → CF h g → CF h (λx. f x - g x).
+
+axiom MSC_prop: ∀f,h. CF h f → ∀x. MSC (f x) ≤ h x.
+
+lemma MSC_max: ∀f,h,x. CF h f → MSC (max_{i < x}(f i)) ≤ max_{i < x}(h i).
+#f #h #x #hf elim x // #i #Hind >bigop_Strue [|//] >bigop_Strue [|//]
+whd in match (max ??);
+cases (true_or_false (leb (f i) (bigop i (λi0:ℕ.true) ? 0 max(λi0:ℕ.f i0))))
+#Hcase >Hcase
+ [@(transitive_le … Hind) @le_maxr //
+ |@(transitive_le … (MSC_prop … hf i)) @le_maxl //
+ ]
+qed.
+
+lemma CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s.
+ CF ha a → CF hb b → CF hp p → CF hf f →
+ O s (λx.ha x + hb x +
+ (∑_{i ∈[a x ,b x[ }
+ (hp 〈i,x〉 + hf 〈i,x〉 + max_{i ∈ [a x, b x [ }(hf 〈i,x〉)))) →
+ CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)).
+#a #b #p #f #ha #hb #hp #hf #s #CFa #CFb #CFp #CFf #HO
+@ext_CF1 [|#x @max_unary_pr1]
+@(CF_comp ??? (λx.ha x + hb x))
+ [|@CF_comp_pair
+ [@CF_minus [@(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //]
+ |@CF_comp_pair
+ [@(O_to_CF … CFa) @O_plus_l //
+ | @(O_to_CF … CF_id) @O_plus_r @(proj1 … CFb)
+ ]
+ ]
+ |@(CF_prim_rec … MSC … (aux_compl1 hp hf))
+ [@CF_const
+ |@(O_to_CF … (Oaux_compl … ))
+ @CF_if
+ [@(CF_comp p … MSC … CFp)
+ [@CF_comp_pair
+ [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd]
+ |@CF_comp_snd @CF_comp_snd @CF_snd]
+ |@le_to_O #x normalize >commutative_plus >associative_plus @le_plus //
+ ]
+ |@CF_max1
+ [@(CF_comp f … MSC … CFf)
+ [@CF_comp_pair
+ [@CF_plus [@CF_fst| @CF_comp_fst @CF_comp_snd @CF_snd]
+ |@CF_comp_snd @CF_comp_snd @CF_snd]
+ |@le_to_O #x normalize >commutative_plus //
+ ]
+ |@CF_comp_fst @(monotonic_CF … CF_snd) normalize //
+ ]
+ |@CF_comp_fst @(monotonic_CF … CF_snd) normalize //
+ ]
+ |@O_refl
+ ]
+ |@(O_trans … HO) -HO
+ @(O_trans ? (λx:ℕ
+ .ha x+hb x
+ +bigop (b x-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus
+ (λi:ℕ
+ .(λi0:ℕ
+ .hp 〈i0,x〉+hf 〈i0,x〉
+ +bigop (b x-a x) (λi1:ℕ.true) ? 0 max
+ (λi1:ℕ.(λi2:ℕ.hf 〈i2,x〉) (i1+a x))) (i+a x))))
+ [
+ @le_to_O #n @le_plus // whd in match (unary_pr ???);
+ >fst_pair >snd_pair >bigop_prim_rec elim (b n - a n)
+ [normalize //
+ |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >aux_compl1_def1
+ >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus
+ [-Hind @le_plus // normalize >fst_pair >snd_pair
+ @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 )
+ (λi1:ℕ.hf 〈i1+a n,n〉)))
+ [elim x [normalize @MSC_le]
+ #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair
+ >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize
+ [cases (true_or_false (leb (f 〈x0+a n,n〉)
+ (prim_rec (λx00:ℕ.O)
+ (λi:ℕ
+ .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉)
+ (fst (snd i))
+ then fst (snd i)
+ else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize
+ [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ |@(transitive_le … (MSC_prop … CFf …)) @(le_maxl … (le_n …))
+ ]
+ |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ ]
+ |@(le_maxr … (le_n …))
+ ]
+ |@(transitive_le … Hind) -Hind
+ generalize in match (bigop x (λi:ℕ.true) ? 0 max
+ (λi1:ℕ.(λi2:ℕ.hf 〈i2,n〉) (i1+a n))); #c
+ generalize in match (hf 〈x+a n,n〉); #c1
+ elim x [//] #x0 #Hind
+ >prim_rec_S >prim_rec_S normalize >fst_pair >fst_pair >snd_pair
+ >snd_pair >snd_pair >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair
+ >fst_pair @le_plus
+ [@le_plus // cases (true_or_false (leb c1 c)) #Hcase
+ >Hcase normalize // @lt_to_le @not_le_to_lt @(leb_false_to_not_le ?? Hcase)
+ |@Hind
+ ]
+ ]
+ ]
+ |@O_plus [@O_plus_l //] @le_to_O #x
+ <bigop_plus_c @le_plus // @(transitive_le … (MSC_pair …)) @le_plus
+ [@MSC_prop @CFa | @MSC_prop @(O_to_CF MSC … (CF_const x)) @(proj1 … CFb)]
+ ]
+qed.
+
+(* old
axiom CF_max: ∀a,b.∀p:nat →bool.∀f,ha,hb,hp,hf,s.
CF ha a → CF hb b → CF hp p → CF hf f →
O s (λx.ha x + hb x + ∑_{i ∈[a x ,b x[ }(hp 〈i,x〉 + hf 〈i,x〉)) →
- CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)).
-
+ CF s (λx.max_{i ∈[a x,b x[ | p 〈i,x〉 }(f 〈i,x〉)). *)
+
(******************************** minimization ********************************)
-axiom CF_mu: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s.
+alias symbol "plus" (instance 2) = "natural plus".
+alias symbol "plus" (instance 5) = "natural plus".
+alias symbol "plus" (instance 1) = "natural plus".
+alias symbol "plus" (instance 4) = "natural plus".
+alias symbol "pair" (instance 3) = "abstract pair".
+alias id "O" = "cic:/matita/arithmetics/nat/nat#con:0:1:0".
+let rec my_minim a f x k on k ≝
+ match k with
+ [O ⇒ a
+ |S p ⇒ if eqb (my_minim a f x p) (a+p)
+ then if f 〈a+p,x〉 then a+p else S(a+p)
+ else (my_minim a f x p) ].
+
+lemma my_minim_S : ∀a,f,x,k.
+ my_minim a f x (S k) =
+ if eqb (my_minim a f x k) (a+k)
+ then if f 〈a+k,x〉 then a+k else S(a+k)
+ else (my_minim a f x k) .
+// qed.
+
+lemma my_minim_fa : ∀a,f,x,k. f〈a,x〉 = true → my_minim a f x k = a.
+#a #f #x #k #H elim k // #i #Hind normalize >Hind
+cases (true_or_false (eqb a (a+i))) #Hcase >Hcase normalize //
+<(eqb_true_to_eq … Hcase) >H //
+qed.
+
+lemma my_minim_nfa : ∀a,f,x,k. f〈a,x〉 = false →
+my_minim a f x (S k) = my_minim (S a) f x k.
+#a #f #x #k #H elim k
+ [normalize <plus_n_O >H >eq_to_eqb_true //
+ |#i #Hind >my_minim_S >Hind >my_minim_S <plus_n_Sm //
+ ]
+qed.
+
+lemma my_min_eq : ∀a,f,x,k.
+ min k a (λi.f 〈i,x〉) = my_minim a f x k.
+#a #f #x #k lapply a -a elim k // #i #Hind #a normalize in ⊢ (??%?);
+cases (true_or_false (f 〈a,x〉)) #Hcase >Hcase
+ [>(my_minim_fa … Hcase) // | >Hind @sym_eq @(my_minim_nfa … Hcase) ]
+qed.
+
+(* cambiare qui: togliere S *)
+
+(* FG: aliases added by matita compiled with OCaml 4.0.5, bah ??? *)
+alias symbol "minus" (instance 1) = "natural minus".
+alias symbol "minus" (instance 3) = "natural minus".
+alias symbol "pair" (instance 2) = "abstract pair".
+definition minim_unary_pr ≝ λf.unary_pr (λx.S(fst x))
+ (λi.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let b ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ if f 〈b-k,x〉 then b-k else r).
+
+definition compl_minim_unary ≝ λhf:nat → nat.λi.
+ let k ≝ fst i in
+ let b ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉.
+
+definition compl_minim_unary_aux ≝ λhf,i.
+ let k ≝ fst i in
+ let r ≝ fst (snd i) in
+ let b ≝ fst (snd (snd i)) in
+ let x ≝ snd (snd (snd i)) in
+ hf 〈b-k,x〉 + MSC i.
+
+lemma compl_minim_unary_aux_def : ∀hf,k,r,b,x.
+ compl_minim_unary_aux hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈r,〈b,x〉〉〉.
+#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair //
+qed.
+
+lemma compl_minim_unary_def : ∀hf,k,r,b,x.
+ compl_minim_unary hf 〈k,〈r,〈b,x〉〉〉 = hf 〈b-k,x〉 + MSC 〈k,〈S b,x〉〉.
+#hf #k #r #b #x normalize >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair //
+qed.
+
+lemma compl_minim_unary_aux_def2 : ∀hf,k,r,x.
+ compl_minim_unary_aux hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈r,x〉〉.
+#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair //
+qed.
+
+lemma compl_minim_unary_def2 : ∀hf,k,r,x.
+ compl_minim_unary hf 〈k,〈r,x〉〉 = hf 〈fst x-k,snd x〉 + MSC 〈k,〈S(fst x),snd x〉〉.
+#hf #k #r #x normalize >snd_pair >snd_pair >fst_pair //
+qed.
+
+lemma min_aux: ∀a,f,x,k. min (S k) (a x) (λi:ℕ.f 〈i,x〉) =
+ ((minim_unary_pr f) ∘ (λx.〈S k,〈a x + k,x〉〉)) x.
+#a #f #x #k whd in ⊢ (???%); >fst_pair >snd_pair
+lapply a -a (* @max_prim_rec1 *)
+elim k
+ [normalize #a >fst_pair >snd_pair >fst_pair >snd_pair >snd_pair >fst_pair
+ <plus_n_O <minus_n_O >fst_pair //
+ |#i #Hind #a normalize in ⊢ (??%?); >prim_rec_S >fst_pair >snd_pair
+ >fst_pair >snd_pair >snd_pair >fst_pair <plus_n_Sm <(Hind (λx.S (a x)))
+ whd in ⊢ (???%); >minus_S_S <(minus_plus_m_m (a x) i) //
+qed.
+
+lemma minim_unary_pr1: ∀a,b,f,x.
+ μ_{i ∈[a x,b x]}(f 〈i,x〉) =
+ if leb (a x) (b x) then ((minim_unary_pr f) ∘ (λx.〈S (b x) - a x,〈b x,x〉〉)) x
+ else (a x).
+#a #b #f #x cases (true_or_false (leb (a x) (b x))) #Hcase >Hcase
+ [cut (b x = a x + (b x - a x))
+ [>commutative_plus <plus_minus_m_m // @leb_true_to_le // ]
+ #Hcut whd in ⊢ (???%); >minus_Sn_m [|@leb_true_to_le //]
+ >min_aux whd in ⊢ (??%?); <Hcut //
+ |>eq_minus_O // @not_le_to_lt @leb_false_to_not_le //
+ ]
+qed.
+
+axiom sum_inv: ∀a,b,f.∑_{i ∈ [a,S b[ }(f i) = ∑_{i ∈ [a,S b[ }(f (a + b - i)).
+
+(*
+#a #b #f @(bigop_iso … plusAC) whd %{(λi.b - a - i)} %{(λi.b - a -i)} %
+ [%[#i #lti #_ normalize @eq_f >commutative_plus <plus_minus_associative
+ [2: @le_minus_to_plus_r //
+ [// @eq_f @@sym_eq @plus_to_minus
+ |#i #Hi #_ % [% [@le_S_S
+*)
+
+example sum_inv_check : ∑_{i ∈ [3,6[ }(i*i) = ∑_{i ∈ [3,6[ }((8-i)*(8-i)).
+normalize // qed.
+
+(* provo rovesciando la somma *)
+
+axiom daemon: ∀P:Prop.P.
+
+lemma CF_mu: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s.
+ CF sa a → CF sb b → CF sf f →
+ O s (λx.sa x + sb x + ∑_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉 + MSC 〈b x - i,〈S(b x),x〉〉)) →
+ CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)).
+#a #b #f #ha #hb #hf #s #CFa #CFb #CFf #HO
+@ext_CF1 [|#x @minim_unary_pr1]
+@CF_if
+ [@CF_le @(O_to_CF … HO)
+ [@(O_to_CF … CFa) @O_plus_l @O_plus_l @O_refl
+ |@(O_to_CF … CFb) @O_plus_l @O_plus_r @O_refl
+ ]
+ |@(CF_comp ??? (λx.ha x + hb x))
+ [|@CF_comp_pair
+ [@CF_minus [@CF_compS @(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //]
+ |@CF_comp_pair
+ [@(O_to_CF … CFb) @O_plus_r //
+ |@(O_to_CF … CF_id) @O_plus_r @(proj1 … CFb)
+ ]
+ ]
+ |@(CF_prim_rec_gen … MSC … (compl_minim_unary_aux hf))
+ [@((λx:ℕ.fst (snd x)
+ +compl_minim_unary hf
+ 〈fst x,
+ 〈unary_pr fst
+ (λi:ℕ
+ .(let (k:ℕ) ≝fst i in
+ let (r:ℕ) ≝fst (snd i) in
+ let (b:ℕ) ≝fst (snd (snd i)) in
+ let (x:ℕ) ≝snd (snd (snd i)) in if f 〈b-S k,x〉 then b-S k else r ))
+ 〈fst x,snd (snd x)〉,
+ snd (snd x)〉〉))
+ |@CF_compS @CF_fst
+ |@CF_if
+ [@(CF_comp f … MSC … CFf)
+ [@CF_comp_pair
+ [@CF_minus [@CF_comp_fst @CF_comp_snd @CF_snd|@CF_fst]
+ |@CF_comp_snd @CF_comp_snd @CF_snd]
+ |@le_to_O #x normalize >commutative_plus //
+ ]
+ |@(O_to_CF … MSC)
+ [@le_to_O #x normalize //
+ |@CF_minus
+ [@CF_comp_fst @CF_comp_snd @CF_snd |@CF_fst]
+ ]
+ |@CF_comp_fst @(monotonic_CF … CF_snd) normalize //
+ ]
+ |@O_plus
+ [@O_plus_l @O_refl
+ |@O_plus_r @O_ext2 [||#x >compl_minim_unary_aux_def2 @refl]
+ @O_trans [||@le_to_O #x >compl_minim_unary_def2 @le_n]
+ @O_plus [@O_plus_l //]
+ @O_plus_r
+ @O_trans [|@le_to_O #x @MSC_pair] @O_plus
+ [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_fst …))
+ >fst_pair @le_n]
+ @O_trans [|@le_to_O #x @MSC_pair] @O_plus
+ [@le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …))
+ >snd_pair @(transitive_le ???? (le_fst …)) >fst_pair
+ normalize >snd_pair >fst_pair lapply (surj_pair x)
+ * #x1 * #x2 #Hx >Hx >fst_pair >snd_pair elim x1 //
+ #i #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >fst_pair
+ cases (f ?) [@le_S // | //]]
+ @le_to_O #x @monotonic_MSC @(transitive_le ???? (le_snd …)) >snd_pair
+ >(expand_pair (snd (snd x))) in ⊢ (?%?); @le_pair //
+ ]
+ ]
+ |cut (O s (λx.ha x + hb x +
+ ∑_{i ∈[a x ,S(b x)[ }(hf 〈a x+b x-i,x〉 + MSC 〈b x -(a x+b x-i),〈S(b x),x〉〉)))
+ [@(O_ext2 … HO) #x @eq_f @sum_inv] -HO #HO
+ @(O_trans … HO) -HO
+ @(O_trans ? (λx:ℕ.ha x+hb x
+ +bigop (S(b x)-a x) (λi:ℕ.true) ?
+ (MSC 〈b x,x〉) plus (λi:ℕ.(λj.hf j + MSC 〈b x - fst j,〈S(b (snd j)),snd j〉〉) 〈b x- i,x〉)))
+ [@le_to_O #n @le_plus // whd in match (unary_pr ???);
+ >fst_pair >snd_pair >(bigop_prim_rec_dec1 a b ? (λi.true))
+ (* it is crucial to recall that the index is bound by S(b x) *)
+ cut (S (b n) - a n ≤ S (b n)) [//]
+ elim (S(b n) - a n)
+ [normalize //
+ |#x #Hind #lex >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair
+ >compl_minim_unary_def >prim_rec_S >fst_pair >snd_pair >fst_pair
+ >snd_pair >fst_pair >snd_pair >fst_pair whd in ⊢ (??%); >commutative_plus
+ @le_plus [2:@Hind @le_S @le_S_S_to_le @lex] -Hind >snd_pair
+ >minus_minus_associative // @le_S_S_to_le //
+ ]
+ |@O_plus [@O_plus_l //] @O_ext2 [||#x @bigop_plus_c]
+ @O_plus
+ [@O_plus_l @O_trans [|@le_to_O #x @MSC_pair]
+ @O_plus [@O_plus_r @le_to_O @(MSC_prop … CFb)|@O_plus_r @(proj1 … CFb)]
+ |@O_plus_r @(O_ext2 … (O_refl …)) #x @same_bigop
+ [//|#i #H #_ @eq_f2 [@eq_f @eq_f2 //|>fst_pair @eq_f @eq_f2 //]
+ ]
+ ]
+ ]
+ ]
+ |@(O_to_CF … CFa) @(O_trans … HO) @O_plus_l @O_plus_l @O_refl
+ ]
+qed.
+
+(*
+lemma CF_mu: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s.
CF sa a → CF sb b → CF sf f →
O s (λx.sa x + sb x + ∑_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉)) →
CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)).
+#a #b #f #ha #hb #hf #s #CFa #CFb #CFf #HO
+@ext_CF1 [|#x @minim_unary_pr1]
+@(CF_comp ??? (λx.ha x + hb x))
+ [|@CF_comp_pair
+ [@CF_minus [@CF_compS @(O_to_CF … CFb) @O_plus_r // |@(O_to_CF … CFa) @O_plus_l //]
+ |@CF_comp_pair
+ [@CF_max1 [@(O_to_CF … CFa) @O_plus_l // | @CF_compS @(O_to_CF … CFb) @O_plus_r //]
+ | @(O_to_CF … CF_id) @O_plus_r @(proj1 … CFb)
+ ]
+ ]
+ |@(CF_prim_rec … MSC … (compl_minim_unary_aux hf))
+ [@CF_fst
+ |@CF_if
+ [@(CF_comp f … MSC … CFf)
+ [@CF_comp_pair
+ [@CF_minus [@CF_comp_fst @CF_comp_snd @CF_snd|@CF_compS @CF_fst]
+ |@CF_comp_snd @CF_comp_snd @CF_snd]
+ |@le_to_O #x normalize >commutative_plus //
+ ]
+ |@(O_to_CF … MSC)
+ [@le_to_O #x normalize //
+ |@CF_minus
+ [@CF_comp_fst @CF_comp_snd @CF_snd |@CF_compS @CF_fst]
+ ]
+ |@CF_comp_fst @(monotonic_CF … CF_snd) normalize //
+ ]
+ |@O_refl
+ ]
+ |@(O_trans … HO) -HO
+ @(O_trans ? (λx:ℕ
+ .ha x+hb x
+ +bigop (S(b x)-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus (λi:ℕ.hf 〈i+a x,x〉)))
+ [@le_to_O #n @le_plus // whd in match (unary_pr ???);
+ >fst_pair >snd_pair >(bigop_prim_rec ? (λn.S(b n)) ? (λi.true)) elim (S(b n) - a n)
+ [normalize @monotonic_MSC @le_pair
+ |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >compl_minim_unary_def
+ >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus
+ [-Hind @le_plus // normalize >fst_pair >snd_pair
+ @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 )
+ (λi1:ℕ.hf 〈i1+a n,n〉)))
+ [elim x [normalize @MSC_le]
+ #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair
+ >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize
+ [cases (true_or_false (leb (f 〈x0+a n,n〉)
+ (prim_rec (λx00:ℕ.O)
+ (λi:ℕ
+ .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉)
+ (fst (snd i))
+ then fst (snd i)
+ else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize
+ [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ |@(transitive_le … (MSC_prop … CFf …)) @(le_maxl … (le_n …))
+ ]
+ |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ ]
+ |@(le_maxr … (le_n …))
+ ]
+
+
+ @(O_trans ? (λx:ℕ
+ .ha x+hb x
+ +bigop (b x-a x) (λi:ℕ.true) ? (MSC 〈a x,x〉) plus
+ (λi:ℕ
+ .(λi0:ℕ
+ .hp 〈i0,x〉+hf 〈i0,x〉
+ +bigop (b x-a x) (λi1:ℕ.true) ? 0 max
+ (λi1:ℕ.(λi2:ℕ.hf 〈i2,x〉) (i1+a x))) (i+a x))))
+ [
+ @le_to_O #n @le_plus // whd in match (unary_pr ???);
+ >fst_pair >snd_pair >bigop_prim_rec elim (b n - a n)
+ [normalize //
+ |#x #Hind >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >aux_compl1_def1
+ >prim_rec_S >fst_pair >snd_pair >fst_pair >snd_pair >fst_pair
+ >snd_pair normalize in ⊢ (??%); >commutative_plus @le_plus
+ [-Hind @le_plus // normalize >fst_pair >snd_pair
+ @(transitive_le ? (bigop x (λi1:ℕ.true) ? 0 (λn0:ℕ.λm:ℕ.if leb n0 m then m else n0 )
+ (λi1:ℕ.hf 〈i1+a n,n〉)))
+ [elim x [normalize @MSC_le]
+ #x0 #Hind >prim_rec_S >fst_pair >snd_pair >snd_pair >snd_pair
+ >fst_pair >fst_pair cases (p 〈x0+a n,n〉) normalize
+ [cases (true_or_false (leb (f 〈x0+a n,n〉)
+ (prim_rec (λx00:ℕ.O)
+ (λi:ℕ
+ .if p 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ then if leb (f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉)
+ (fst (snd i))
+ then fst (snd i)
+ else f 〈fst i+fst (snd (snd i)),snd (snd (snd i))〉
+ else fst (snd i) ) x0 〈a n,n〉))) #Hcase >Hcase normalize
+ [@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ |@(transitive_le … (MSC_prop … CFf …)) @(le_maxl … (le_n …))
+ ]
+ |@(transitive_le … Hind) -Hind @(le_maxr … (le_n …))
+ ]
+ |@(le_maxr … (le_n …))
+ ]
+ |@(transitive_le … Hind) -Hind
+ generalize in match (bigop x (λi:ℕ.true) ? 0 max
+ (λi1:ℕ.(λi2:ℕ.hf 〈i2,n〉) (i1+a n))); #c
+ generalize in match (hf 〈x+a n,n〉); #c1
+ elim x [//] #x0 #Hind
+ >prim_rec_S >prim_rec_S normalize >fst_pair >fst_pair >snd_pair
+ >snd_pair >snd_pair >snd_pair >snd_pair >snd_pair >fst_pair >fst_pair
+ >fst_pair @le_plus
+ [@le_plus // cases (true_or_false (leb c1 c)) #Hcase
+ >Hcase normalize // @lt_to_le @not_le_to_lt @(leb_false_to_not_le ?? Hcase)
+ |@Hind
+ ]
+ ]
+ ]
+ |@O_plus [@O_plus_l //] @le_to_O #x
+ <bigop_plus_c @le_plus // @(transitive_le … (MSC_pair …)) @le_plus
+ [@MSC_prop @CFa | @MSC_prop @(O_to_CF MSC … (CF_const x)) @(proj1 … CFb)]
+ ]
+qed.
+
+axiom CF_mu: ∀a,b.∀f:nat →bool.∀sa,sb,sf,s.
+ CF sa a → CF sb b → CF sf f →
+ O s (λx.sa x + sb x + ∑_{i ∈[a x ,S(b x)[ }(sf 〈i,x〉)) →
+ CF s (λx.μ_{i ∈[a x,b x] }(f 〈i,x〉)).*)
(************************************* smn ************************************)
axiom smn: ∀f,s. CF s f → ∀x. CF (λy.s 〈x,y〉) (λy.f 〈x,y〉).
(******************** complexity of g ********************)
-definition unary_g ≝ λh.λux. g h (fst ux) (snd ux).
+(*definition unary_g ≝ λh.λux. g h (fst ux) (snd ux).
definition auxg ≝
λh,ux. max_{i ∈[fst ux,snd ux[ | eqb (min_input h i (snd ux)) (snd ux)}
(out i (snd ux) (h (S i) (snd ux))).
[@CF_compS @CF_fst
|@CF_comp_snd @CF_snd
|@(O_trans … HO) @O_plus [@O_plus @O_plus_l // | @O_plus_r //]
-qed.
+qed.*)
(************************* a couple of technical lemmas ***********************)
lemma minus_to_0: ∀a,b. a ≤ b → minus a b = 0.
(**************************** end of technical lemmas *************************)
-lemma compl_g5 : ∀h,s1.(∀n. monotonic ℕ le (λa:ℕ.s1 〈a,n〉)) →
+(*lemma compl_g5 : ∀h,s1.(∀n. monotonic ℕ le (λa:ℕ.s1 〈a,n〉)) →
(CF s1
(λp.termb (fst (snd p)) (fst p) (h (S (fst (snd p))) (fst p)))) →
CF (λx.MSC x + (snd (snd x)-fst x)*s1 〈snd (snd x),x〉)
]
|@le_to_O #x @monotonic_sU // @(le_maxl … (le_n …)) ]
]
-qed.
+qed. *)
definition big : nat →nat ≝ λx.
let m ≝ max (fst x) (snd x) in 〈m,m〉.
(λx.MSC x + (snd (snd x)-fst x)*
(λx.sU 〈max (fst(snd x)) (snd(snd x)),〈fst x,h (S (fst (snd x))) (fst x)〉〉) 〈snd (snd x),x〉).
-lemma compl_g7: ∀h.
+(*lemma compl_g7: ∀h.
constructible (λx. h (fst x) (snd x)) →
(∀n. monotonic ? le (h n)) →
CF (λx.MSC x + (snd (snd x)-fst x)*sU 〈max (fst x) (snd x),〈snd (snd x),h (S (fst x)) (snd (snd x))〉〉)
[#Hle @le_times // @monotonic_sU
|#Hlt >(minus_to_0 … (lt_to_le … )) [// | @not_le_to_lt @Hlt]
]
-qed.
+qed.*)
definition out_aux ≝ λh.
out_unary ∘ λp.〈fst p,〈snd(snd p),h (S (fst p)) (snd (snd p))〉〉.
]
qed.
-lemma compl_g9 : ∀h.
+(*lemma compl_g9 : ∀h.
constructible (λx. h (fst x) (snd x)) →
(∀n. monotonic ? le (h n)) →
(∀n,a,b. a ≤ b → b ≤ n → h b n ≤ h a n) →
>associative_plus >distributive_times_plus
@le_plus [@le_times [@le_S // |>big_def >Hmax //] |//]
]
-qed.
+qed.*)
definition sg ≝ λh,x.
(S (snd x-fst x))*MSC 〈x,x〉 + (snd x-fst x)*(S(snd x-fst x))*sU 〈x,〈snd x,h (S (fst x)) (snd x)〉〉.
#h #a #b whd in ⊢ (??%?); >fst_pair >snd_pair //
qed.
-lemma compl_g11 : ∀h.
+(*lemma compl_g11 : ∀h.
constructible (λx. h (fst x) (snd x)) →
(∀n. monotonic ? le (h n)) →
(∀n,a,b. a ≤ b → b ≤ n → h b n ≤ h a n) →
CF (sg h) (unary_g h).
#h #hconstr #Hm #Ham @compl_g1 @(compl_g9 h hconstr Hm Ham)
-qed.
+qed.*)
(**************************** closing the argument ****************************)
]
qed.
+(*
lemma h_of_aux_constr :
∀r,c. constructible (λx.h_of_aux r c (fst x) (snd x)).
#r #c
(S d)*(MSC 〈〈b-d,b〉,〈b-d,b〉〉) +
d*(S d)*sU 〈〈b-d,b〉,〈b,r (fst (snd x))〉〉)))
[#n @sym_eq whd in match (unary_pr ???); @h_of_aux_prim_rec
- |@constr_prim_rec
+ |@constr_prim_rec*)
definition h_of ≝ λr,p.
let m ≝ max (fst p) (snd p) in
(∀x. x ≤ r x) → monotonic ? le r → constructible r →
constructible (h_of r).
-lemma speed_compl: ∀r:nat →nat.
+(*lemma speed_compl: ∀r:nat →nat.
(∀x. x ≤ r x) → monotonic ? le r → constructible r →
CF (h_of r) (unary_g (λi,x. r(h_of r 〈i,x〉))).
#r #Hr #Hmono #Hconstr @(monotonic_CF … (compl_g11 …))
@(ext_CF (λx.unary_g (λi,x. r(h_of r 〈i,x〉)) 〈i,x〉))
[#n whd in ⊢ (??%%); @eq_f @sym_eq >fst_pair >snd_pair %]
@smn @(ext_CF … (speed_compl r Hr Hmono Hconstr)) #n //
-qed.
+qed.*)
(**************************** the speedup theorem *****************************)
-theorem pseudo_speedup:
+(*theorem pseudo_speedup:
∀r:nat →nat. (∀x. x ≤ r x) → monotonic ? le r → constructible r →
∃f.∀sf. CF sf f → ∃g,sg. f ≈ g ∧ CF sg g ∧ O sf (r ∘ sg).
(* ∃m,a.∀n. a≤n → r(sg a) < m * sf n. *)
#y #Uy %{y} @(monotonic_U … Uy) @(transitive_le … Hlesf)
@Hmono @(mono_h_of2 … Hr Hmono … ltin)
]
-qed.
-
+qed.*)
+
\ No newline at end of file
+++ /dev/null
-include "turing/auxiliary_machines1.ma".
-include "turing/multi_to_mono/shift_trace_machines.ma".
-
-(******************************************************************************)
-(* mtiL: complete move L for tape i. We reaching the left border of trace i, *)
-(* add a blank if there is no more tape, then move the i-trace and finally *)
-(* come back to the head position. *)
-(******************************************************************************)
-
-(* we say that a tape is regular if for any trace after the first blank we
- only have other blanks *)
-
-definition all_blanks_in ≝ λsig,l.
- ∀x. mem ? x l → x = blank sig.
-
-definition regular_i ≝ λsig,n.λl:list (multi_sig sig n).λi.
- all_blanks_in ? (after_blank ? (trace sig n i l)).
-
-definition regular_trace ≝ λsig,n,a.λls,rs:list (multi_sig sig n).λi.
- Or (And (regular_i sig n (a::ls) i) (regular_i sig n rs i))
- (And (regular_i sig n ls i) (regular_i sig n (a::rs) i)).
-
-axiom regular_tail: ∀sig,n,l,i.
- regular_i sig n l i → regular_i sig n (tail ? l) i.
-
-axiom regular_extend: ∀sig,n,l,i.
- regular_i sig n l i → regular_i sig n (l@[all_blank sig n]) i.
-
-axiom all_blank_after_blank: ∀sig,n,l1,b,l2,i.
- nth i ? (vec … b) (blank ?) = blank ? →
- regular_i sig n (l1@b::l2) i → all_blanks_in ? (trace sig n i l2).
-
-lemma regular_trace_extl: ∀sig,n,a,ls,rs,i.
- regular_trace sig n a ls rs i →
- regular_trace sig n a (ls@[all_blank sig n]) rs i.
-#sig #n #a #ls #rs #i *
- [* #H1 #H2 % % // @(regular_extend … H1)
- |* #H1 #H2 %2 % // @(regular_extend … H1)
- ]
-qed.
-
-lemma regular_cons_hd_rs: ∀sig,n.∀a:multi_sig sig n.∀ls,rs1,rs2,i.
- regular_trace sig n a ls (rs1@rs2) i →
- regular_trace sig n a ls (rs1@((hd ? rs2 (all_blank …))::(tail ? rs2))) i.
-#sig #n #a #ls #rs1 #rs2 #i cases rs2 [2: #b #tl #H @H]
-*[* #H1 >append_nil #H2 %1 %
- [@H1 | whd in match (hd ???); @(regular_extend … rs1) //]
- |* #H1 >append_nil #H2 %2 %
- [@H1 | whd in match (hd ???); @(regular_extend … (a::rs1)) //]
- ]
-qed.
-
-lemma eq_trace_to_regular : ∀sig,n.∀a1,a2:multi_sig sig n.∀ls1,ls2,rs1,rs2,i.
- nth i ? (vec … a1) (blank ?) = nth i ? (vec … a2) (blank ?) →
- trace sig n i ls1 = trace sig n i ls2 →
- trace sig n i rs1 = trace sig n i rs2 →
- regular_trace sig n a1 ls1 rs1 i →
- regular_trace sig n a2 ls2 rs2 i.
-#sig #n #a1 #a2 #ls1 #ls2 #rs1 #rs2 #i #H1 #H2 #H3 #H4
-whd in match (regular_trace ??????); whd in match (regular_i ????);
-whd in match (regular_i ?? rs2 ?); whd in match (regular_i ?? ls2 ?);
-whd in match (regular_i ?? (a2::rs2) ?); whd in match (trace ????);
-<trace_def whd in match (trace ??? (a2::rs2)); <trace_def
-<H1 <H2 <H3 @H4
-qed.
-
-(******************************* move_to_blank_L ******************************)
-(* we compose machines together to reduce the number of output cases, and
- improve semantics *)
-
-definition move_to_blank_L ≝ λsig,n,i.
- (move_until ? L (no_blank sig n i)) · extend ? (all_blank sig n).
-
-(*
-definition R_move_to_blank_L ≝ λsig,n,i,t1,t2.
-(current ? t1 = None ? →
- t2 = midtape (multi_sig sig n) (left ? t1) (all_blank …) (right ? t1)) ∧
-∀ls,a,rs.t1 = midtape ? ls a rs →
- ((no_blank sig n i a = false) ∧ t2 = t1) ∨
- (∃b,ls1,ls2.
- (no_blank sig n i b = false) ∧
- (∀j.j ≤n → to_blank_i ?? j (ls1@b::ls2) = to_blank_i ?? j ls) ∧
- t2 = midtape ? ls2 b ((reverse ? (a::ls1))@rs)).
-*)
-
-definition R_move_to_blank_L ≝ λsig,n,i,t1,t2.
-(current ? t1 = None ? →
- t2 = midtape (multi_sig sig n) (left ? t1) (all_blank …) (right ? t1)) ∧
-∀ls,a,rs.
- t1 = midtape (multi_sig sig n) ls a rs →
- regular_i sig n (a::ls) i →
- (∀j. j ≠ i → regular_trace … a ls rs j) →
- (∃b,ls1,ls2.
- (regular_i sig n (ls1@b::ls2) i) ∧
- (∀j. j ≠ i → regular_trace …
- (hd ? (ls1@b::ls2) (all_blank …)) (tail ? (ls1@b::ls2)) rs j) ∧
- (no_blank sig n i b = false) ∧
- (hd (multi_sig sig n) (ls1@[b]) (all_blank …) = a) ∧ (* not implied by the next fact *)
- (∀j.j ≤n → to_blank_i ?? j (ls1@b::ls2) = to_blank_i ?? j (a::ls)) ∧
- t2 = midtape ? ls2 b ((reverse ? ls1)@rs)).
-
-theorem sem_move_to_blank_L: ∀sig,n,i.
- move_to_blank_L sig n i ⊨ R_move_to_blank_L sig n i.
-#sig #n #i
-@(sem_seq_app ??????
- (ssem_move_until_L ? (no_blank sig n i)) (sem_extend ? (all_blank sig n)))
-#tin #tout * #t1 * * #Ht1a #Ht1b * #Ht2a #Ht2b %
- [#Hcur >(Ht1a Hcur) in Ht2a; /2 by /
- |#ls #a #rs #Htin #Hreg #Hreg2 -Ht1a cases (Ht1b … Htin)
- [* #Hnb #Ht1 -Ht1b -Ht2a >Ht1 in Ht2b; >Htin #H
- %{a} %{[ ]} %{ls}
- %[%[%[%[%[@Hreg|@Hreg2]|@Hnb]|//]|//]|@H normalize % #H1 destruct (H1)]
- |*
- [(* we find the blank *)
- * #ls1 * #b * #ls2 * * * #H1 #H2 #H3 #Ht1
- >Ht1 in Ht2b; #Hout -Ht1b
- %{b} %{(a::ls1)} %{ls2}
- %[%[%[%[%[>H1 in Hreg; #H @H
- |#j #jneqi whd in match (hd ???); whd in match (tail ??);
- <H1 @(Hreg2 j jneqi)]|@H2] |//]|>H1 //]
- |@Hout normalize % normalize #H destruct (H)
- ]
- |* #b * #lss * * #H1 #H2 #Ht1 -Ht1b >Ht1 in Ht2a;
- whd in match (left ??); whd in match (right ??); #Hout
- %{(all_blank …)} %{(lss@[b])} %{[]}
- %[%[%[%[%[<H2 @regular_extend //
- |<H2 #j #jneqi whd in match (hd ???); whd in match (tail ??);
- @regular_trace_extl @Hreg2 //]
- |whd in match (no_blank ????); >blank_all_blank //]
- |<H2 //]
- |#j #lejn <H2 @sym_eq @to_blank_i_ext]
- |>reverse_append >reverse_single @Hout normalize //
- ]
- ]
- ]
-qed.
-
-(******************************************************************************)
-
-definition shift_i_L ≝ λsig,n,i.
- ncombf_r (multi_sig …) (shift_i sig n i) (all_blank sig n) ·
- mti sig n i ·
- extend ? (all_blank sig n).
-
-definition R_shift_i_L ≝ λsig,n,i,t1,t2.
- (∀a,ls,rs.
- t1 = midtape ? ls a rs →
- ((∃rs1,b,rs2,a1,rss.
- rs = rs1@b::rs2 ∧
- nth i ? (vec … b) (blank ?) = (blank ?) ∧
- (∀x. mem ? x rs1 → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧
- shift_l sig n i (a::rs1) (a1::rss) ∧
- t2 = midtape (multi_sig sig n) ((reverse ? (a1::rss))@ls) b rs2) ∨
- (∃b,rss.
- (∀x. mem ? x rs → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧
- shift_l sig n i (a::rs) (rss@[b]) ∧
- t2 = midtape (multi_sig sig n)
- ((reverse ? (rss@[b]))@ls) (all_blank sig n) [ ]))).
-
-definition R_shift_i_L_new ≝ λsig,n,i,t1,t2.
- (∀a,ls,rs.
- t1 = midtape ? ls a rs →
- ∃rs1,b,rs2,rss.
- b = hd ? rs2 (all_blank sig n) ∧
- nth i ? (vec … b) (blank ?) = (blank ?) ∧
- rs = rs1@rs2 ∧
- (∀x. mem ? x rs1 → nth i ? (vec … x) (blank ?) ≠ (blank ?)) ∧
- shift_l sig n i (a::rs1) rss ∧
- t2 = midtape (multi_sig sig n) ((reverse ? rss)@ls) b (tail ? rs2)).
-
-theorem sem_shift_i_L: ∀sig,n,i. shift_i_L sig n i ⊨ R_shift_i_L sig n i.
-#sig #n #i
-@(sem_seq_app ??????
- (sem_ncombf_r (multi_sig sig n) (shift_i sig n i)(all_blank sig n))
- (sem_seq ????? (ssem_mti sig n i)
- (sem_extend ? (all_blank sig n))))
-#tin #tout * #t1 * * #Ht1a #Ht1b * #t2 * * #Ht2a #Ht2b * #Htout1 #Htout2
-#a #ls #rs cases rs
- [#Htin %2 %{(shift_i sig n i a (all_blank sig n))} %{[ ]}
- %[%[#x @False_ind | @daemon]
- |lapply (Ht1a … Htin) -Ht1a -Ht1b #Ht1
- lapply (Ht2a … Ht1) -Ht2a -Ht2b #Ht2 >Ht2 in Htout1;
- >Ht1 whd in match (left ??); whd in match (right ??); #Htout @Htout //
- ]
- |#a1 #rs1 #Htin
- lapply (Ht1b … Htin) -Ht1a -Ht1b #Ht1
- lapply (Ht2b … Ht1) -Ht2a -Ht2b *
- [(* a1 is blank *) * #H1 #H2 %1
- %{[ ]} %{a1} %{rs1} %{(shift_i sig n i a a1)} %{[ ]}
- %[%[%[%[// |//] |#x @False_ind] | @daemon]
- |>Htout2 [>H2 >reverse_single @Ht1 |>H2 >Ht1 normalize % #H destruct (H)]
- ]
- |*
- [* #rs10 * #b * #rs2 * #rss * * * * #H1 #H2 #H3 #H4
- #Ht2 %1
- %{(a1::rs10)} %{b} %{rs2} %{(shift_i sig n i a a1)} %{rss}
- %[%[%[%[>H1 //|//] |@H3] |@daemon ]
- |>reverse_cons >associative_append
- >H2 in Htout2; #Htout >Htout [@Ht2| >Ht2 normalize % #H destruct (H)]
- ]
- |* #b * #rss * * #H1 #H2
- #Ht2 %2
- %{(shift_i sig n i b (all_blank sig n))} %{(shift_i sig n i a a1::rss)}
- %[%[@H1 |@daemon ]
- |>Ht2 in Htout1; #Htout >Htout //
- whd in match (left ??); whd in match (right ??);
- >reverse_append >reverse_single >associative_append >reverse_cons
- >associative_append //
- ]
- ]
- ]
- ]
-qed.
-
-theorem sem_shift_i_L_new: ∀sig,n,i.
- shift_i_L sig n i ⊨ R_shift_i_L_new sig n i.
-#sig #n #i
-@(Realize_to_Realize … (sem_shift_i_L sig n i))
-#t1 #t2 #H #a #ls #rs #Ht1 lapply (H a ls rs Ht1) *
- [* #rs1 * #b * #rs2 * #a1 * #rss * * * * #H1 #H2 #H3 #H4 #Ht2
- %{rs1} %{b} %{(b::rs2)} %{(a1::rss)}
- %[%[%[%[%[//|@H2]|@H1]|@H3]|@H4] | whd in match (tail ??); @Ht2]
- |* #b * #rss * * #H1 #H2 #Ht2
- %{rs} %{(all_blank sig n)} %{[]} %{(rss@[b])}
- %[%[%[%[%[//|@blank_all_blank]|//]|@H1]|@H2] | whd in match (tail ??); @Ht2]
- ]
-qed.
-
-
-(*******************************************************************************
-The following machine implements a full move of for a trace: we reach the left
-border, shift the i-th trace and come back to the head position. *)
-
-(* this exclude the possibility that traces do not overlap: the head must
-remain inside all traces *)
-
-definition mtiL ≝ λsig,n,i.
- move_to_blank_L sig n i ·
- shift_i_L sig n i ·
- move_until ? L (no_head sig n).
-
-definition Rmtil ≝ λsig,n,i,t1,t2.
- ∀ls,a,rs.
- t1 = midtape (multi_sig sig n) ls a rs →
- nth n ? (vec … a) (blank ?) = head ? →
- (∀i.regular_trace sig n a ls rs i) →
- (* next: we cannot be on rightof on trace i *)
- (nth i ? (vec … a) (blank ?) = (blank ?)
- → nth i ? (vec … (hd ? rs (all_blank …))) (blank ?) ≠ (blank ?)) →
- no_head_in … ls →
- no_head_in … rs →
- (∃ls1,a1,rs1.
- t2 = midtape (multi_sig …) ls1 a1 rs1 ∧
- (∀i.regular_trace … a1 ls1 rs1 i) ∧
- (∀j. j ≤ n → j ≠ i → to_blank_i ? n j (a1::ls1) = to_blank_i ? n j (a::ls)) ∧
- (∀j. j ≤ n → j ≠ i → to_blank_i ? n j rs1 = to_blank_i ? n j rs) ∧
- (to_blank_i ? n i ls1 = to_blank_i ? n i (a::ls)) ∧
- (to_blank_i ? n i (a1::rs1)) = to_blank_i ? n i rs).
-
-theorem sem_Rmtil: ∀sig,n,i. i < n → mtiL sig n i ⊨ Rmtil sig n i.
-#sig #n #i #lt_in
-@(sem_seq_app ??????
- (sem_move_to_blank_L … )
- (sem_seq ????? (sem_shift_i_L_new …)
- (ssem_move_until_L ? (no_head sig n))))
-#tin #tout * #t1 * * #_ #Ht1 * #t2 * #Ht2 * #_ #Htout
-(* we start looking into Rmitl *)
-#ls #a #rs #Htin (* tin is a midtape *)
-#Hhead #Hreg #no_rightof #Hnohead_ls #Hnohead_rs
-cut (regular_i sig n (a::ls) i)
- [cases (Hreg i) * //
- cases (true_or_false (nth i ? (vec … a) (blank ?) == (blank ?))) #Htest
- [#_ @daemon (* absurd, since hd rs non e' blank *)
- |#H #_ @daemon]] #Hreg1
-lapply (Ht1 … Htin Hreg1 ?) [#j #_ @Hreg] -Ht1 -Htin
-* #b * #ls1 * #ls2 * * * * * #reg_ls1_i #reg_ls1_j #Hno_blankb #Hhead #Hls1 #Ht1
-lapply (Ht2 … Ht1) -Ht2 -Ht1
-* #rs1 * #b0 * #rs2 * #rss * * * * * #Hb0 #Hb0blank #Hrs1 #Hrs1b #Hrss #Ht2
-(* we need to recover the position of the head of the emulated machine
- that is the head of ls1. This is somewhere inside rs1 *)
-cut (∃rs11. rs1 = (reverse ? ls1)@rs11)
- [cut (ls1 = [ ] ∨ ∃aa,tlls1. ls1 = aa::tlls1)
- [cases ls1 [%1 // | #aa #tlls1 %2 %{aa} %{tlls1} //]] *
- [#H1ls1 %{rs1} >H1ls1 //
- |* #aa * #tlls1 #H1ls1 >H1ls1 in Hrs1;
- cut (aa = a) [>H1ls1 in Hls1; #H @(to_blank_hd … H)] #eqaa >eqaa
- #Hrs1_aux cases (compare_append … (sym_eq … Hrs1_aux)) #l *
- [* #H1 #H2 %{l} @H1
- |(* this is absurd : if l is empty, the case is as before.
- if l is not empty then it must start with a blank, since it is the
- first character in rs2. But in this case we would have a blank
- inside ls1=a::tls1 that is absurd *)
- @daemon
- ]]]
- * #rs11 #H1
-cut (rs = rs11@rs2)
- [@(injective_append_l … (reverse … ls1)) >Hrs1 <associative_append <H1 //] #H2
-lapply (Htout … Ht2) -Htout -Ht2 *
- [(* the current character on trace i holds the head-mark.
- The case is absurd, since b0 is the head of rs2, that is a sublist of rs,
- and the head-mark is not in rs *)
- * #H3 @False_ind @(absurd (nth n ? (vec … b0) (blank sig) = head ?))
- [@(\P ?) @injective_notb @H3 ]
- @Hnohead_rs >H2 >trace_append @mem_append_l2
- lapply Hb0 cases rs2
- [whd in match (hd ???); #H >H in H3; whd in match (no_head ???);
- >all_blank_n normalize -H #H destruct (H); @False_ind
- |#c #r #H4 %1 >H4 //
- ]
- |*
- [(* we reach the head position *)
- (* cut (trace sig n j (a1::ls20)=trace sig n j (ls1@b::ls2)) *)
- * #ls10 * #a1 * #ls20 * * * #Hls20 #Ha1 #Hnh #Htout
- cut (∀j.j ≠ i →
- trace sig n j (reverse (multi_sig sig n) rs1@b::ls2) =
- trace sig n j (ls10@a1::ls20))
- [#j #ineqj >append_cons <reverse_cons >trace_def <map_append <reverse_map
- lapply (trace_shift_neq …lt_in ? (sym_not_eq … ineqj) … Hrss) [//] #Htr
- <(trace_def … (b::rs1)) <Htr >reverse_map >map_append @eq_f @Hls20 ]
- #Htracej
- cut (trace sig n i (reverse (multi_sig sig n) (rs1@[b0])@ls2) =
- trace sig n i (ls10@a1::ls20))
- [>trace_def <map_append <reverse_map <map_append <(trace_def … [b0])
- cut (trace sig n i [b0] = [blank ?]) [@daemon] #Hcut >Hcut
- lapply (trace_shift … lt_in … Hrss) [//] whd in match (tail ??); #Htr <Htr
- >reverse_map >map_append <trace_def <Hls20 %
- ]
- #Htracei
- cut (∀j. j ≠ i →
- (trace sig n j (reverse (multi_sig sig n) rs11) = trace sig n j ls10) ∧
- (trace sig n j (ls1@b::ls2) = trace sig n j (a1::ls20)))
- [@daemon (* si fa
- #j #ineqj @(first_P_to_eq ? (λx. x ≠ head ?))
- [lapply (Htracej … ineqj) >trace_def in ⊢ (%→?); <map_append
- >trace_def in ⊢ (%→?); <map_append #H @H
- | *) ] #H2
- cut ((trace sig n i (b0::reverse ? rs11) = trace sig n i (ls10@[a1])) ∧
- (trace sig n i (ls1@ls2) = trace sig n i ls20))
- [>H1 in Htracei; >reverse_append >reverse_single >reverse_append
- >reverse_reverse >associative_append >associative_append
- @daemon
- ] #H3
- cut (∀j. j ≠ i →
- trace sig n j (reverse (multi_sig sig n) ls10@rs2) = trace sig n j rs)
- [#j #jneqi @(injective_append_l … (trace sig n j (reverse ? ls1)))
- >map_append >map_append >Hrs1 >H1 >associative_append
- <map_append <map_append in ⊢ (???%); @eq_f
- <map_append <map_append @eq_f2 // @sym_eq
- <(reverse_reverse … rs11) <reverse_map <reverse_map in ⊢ (???%);
- @eq_f @(proj1 … (H2 j jneqi))] #Hrs_j
- %{ls20} %{a1} %{(reverse ? (b0::ls10)@tail (multi_sig sig n) rs2)}
- %[%[%[%[%[@Htout
- |#j cases (decidable_eq_nat j i)
- [#eqji >eqji (* by cases wether a1 is blank *)
- @daemon
- |#jneqi lapply (reg_ls1_j … jneqi) #H4
- >reverse_cons >associative_append >Hb0 @regular_cons_hd_rs
- @(eq_trace_to_regular … H4)
- [<hd_trace >(proj2 … (H2 … jneqi)) >hd_trace %
- |<tail_trace >(proj2 … (H2 … jneqi)) >tail_trace %
- |@sym_eq @Hrs_j //
- ]
- ]]
- |#j #lejn #jneqi <(Hls1 … lejn)
- >to_blank_i_def >to_blank_i_def @eq_f @sym_eq @(proj2 … (H2 j jneqi))]
- |#j #lejn #jneqi >reverse_cons >associative_append >Hb0
- <to_blank_hd_cons >to_blank_i_def >to_blank_i_def @eq_f @Hrs_j //]
- |<(Hls1 i) [2:@lt_to_le //]
- lapply (all_blank_after_blank … reg_ls1_i)
- [@(\P ?) @daemon] #allb_ls2
- whd in match (to_blank_i ????); <(proj2 … H3)
- @daemon ]
- |>reverse_cons >associative_append
- cut (to_blank_i sig n i rs = to_blank_i sig n i (rs11@[b0])) [@daemon]
- #Hcut >Hcut >(to_blank_i_chop … b0 (a1::reverse …ls10)) [2: @Hb0blank]
- >to_blank_i_def >to_blank_i_def @eq_f
- >trace_def >trace_def @injective_reverse >reverse_map >reverse_cons
- >reverse_reverse >reverse_map >reverse_append >reverse_single @sym_eq
- @(proj1 … H3)
- ]
- |(*we do not find the head: this is absurd *)
- * #b1 * #lss * * #H2 @False_ind
- cut (∀x0. mem ? x0 (trace sig n n (b0::reverse ? rss@ls2)) → x0 ≠ head ?)
- [@daemon] -H2 #H2
- lapply (trace_shift_neq sig n i n … lt_in … Hrss)
- [@lt_to_not_eq @lt_in | // ]
- #H3 @(absurd
- (nth n ? (vec … (hd ? (ls1@[b]) (all_blank sig n))) (blank ?) = head ?))
- [>Hhead //
- |@H2 >trace_def %2 <map_append @mem_append_l1 <reverse_map <trace_def
- >H3 >H1 >trace_def >reverse_map >reverse_cons >reverse_append
- >reverse_reverse >associative_append <map_append @mem_append_l2
- cases ls1 [%1 % |#x #ll %1 %]
- ]
- ]
- ]
-qed.