StringSet.iter iter fnames
let rec read ich =
- let _ = Scanf.sscanf (input_line ich) "%s@:include \"%s@\"." init in
+ let line = input_line ich in
+ begin try Scanf.sscanf line "%s@:include \"%s@\"." init
+ with Scanf.Scan_failure _ ->
+ begin try Scanf.sscanf line "./%s@:include \"%s@\"." init
+ with Scanf.Scan_failure _ ->
+ begin try Scanf.sscanf line "%s@:(*%s@*)" (fun _ _ -> ())
+ with Scanf.Scan_failure _ ->
+ Printf.eprintf "unknown line: %s.\n" line
+ end
+ end
+ end;
read ich
let _ =
--- /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/notation/functions/weight_2.ma".
+include "basic_2/grammar/lenv_weight.ma".
+
+(* WEIGHT OF A RESTRICTED CLOSURE *******************************************)
+
+definition rfw: lenv → term → ? ≝ λL,T. ♯{L} + ♯{T}.
+
+interpretation "weight (restricted closure)" 'Weight L T = (rfw L T).
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was: flt_shift *)
+lemma rfw_shift: ∀a,I,K,V,T. ♯{K.ⓑ{I}V, T} < ♯{K, ⓑ{a,I}V.T}.
+normalize //
+qed.
+
+lemma rfw_tpair_sn: ∀I,L,V,T. ♯{L, V} < ♯{L, ②{I}V.T}.
+normalize in ⊢ (?→?→?→?→?%%); //
+qed.
+
+lemma rfw_tpair_dx: ∀I,L,V,T. ♯{L, T} < ♯{L, ②{I}V.T}.
+normalize in ⊢ (?→?→?→?→?%%); //
+qed.
+
+lemma rfw_lpair_sn: ∀I,L,V,T. ♯{L, V} < ♯{L.ⓑ{I}V, T}.
+normalize /3 width=1 by monotonic_lt_plus_l, monotonic_le_plus_r/
+qed.
+
+(* Basic_1: removed theorems 7:
+ flt_thead_sx flt_thead_dx flt_trans
+ flt_arith0 flt_arith1 flt_arith2 flt_wf_ind
+*)
+(* Basic_1: removed local theorems 1: q_ind *)
(* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********)
inductive lpx_sn (R:lenv→relation term): relation lenv ≝
-| lpx_sn_stom: lpx_sn R (⋆) (⋆)
+| lpx_sn_atom: lpx_sn R (⋆) (⋆)
| lpx_sn_pair: ∀I,K1,K2,V1,V2.
lpx_sn R K1 K2 → R K1 V1 V2 →
lpx_sn R (K1. ⓑ{I} V1) (K2. ⓑ{I} V2)
#R #HR #L1 #L2 #H @(TC_star_ind … L2 H) -L2
[ /2 width=1 by lpx_sn_refl/
| /3 width=1 by TC_reflexive, lpx_sn_refl/
-| /3 width=5/
+| /3 width=5 by lpx_sn_pair, step/
]
qed-.
lemma lpx_sn_LTC_TC_lpx_sn: ∀R. (∀L. reflexive … (R L)) →
∀L1,L2. lpx_sn (LTC … R) L1 L2 →
TC … (lpx_sn R) L1 L2.
-#R #HR #L1 #L2 #H elim H -L1 -L2 /2 width=1/
-/2 width=1 by TC_lpx_sn_pair/
+#R #HR #L1 #L2 #H elim H -L1 -L2
+/2 width=1 by TC_lpx_sn_pair, lpx_sn_atom, inj/
qed-.
(* Inversion lemmas on transitive closure ***********************************)
lemma TC_lpx_sn_inv_atom2: ∀R,L1. TC … (lpx_sn R) L1 (⋆) → L1 = ⋆.
#R #L1 #H @(TC_ind_dx … L1 H) -L1
-[ #L1 #H lapply (lpx_sn_inv_atom2 … H) -H //
-| #L1 #L #HL1 #_ #IHL2 destruct
- lapply (lpx_sn_inv_atom2 … HL1) -HL1 //
+[ /2 width=2 by lpx_sn_inv_atom2/
+| #L1 #L #HL1 #_ #IHL2 destruct /2 width=2 by lpx_sn_inv_atom2/
]
qed-.
∀I,L1,K2,V2. TC … (lpx_sn R) L1 (K2.ⓑ{I}V2) →
∃∃K1,V1. TC … (lpx_sn R) K1 K2 & LTC … R K1 V1 V2 & L1 = K1. ⓑ{I} V1.
#R #HR #I #L1 #K2 #V2 #H @(TC_ind_dx … L1 H) -L1
-[ #L1 #H elim (lpx_sn_inv_pair2 … H) -H /3 width=5/
+[ #L1 #H elim (lpx_sn_inv_pair2 … H) -H /3 width=5 by inj, ex3_2_intro/
| #L1 #L #HL1 #_ * #K #V #HK2 #HV2 #H destruct
elim (lpx_sn_inv_pair2 … HL1) -HL1 #K1 #V1 #HK1 #HV1 #H destruct
- lapply (HR … HV2 … HK1) -HR -HV2 #HV2 /3 width=5/
+ lapply (HR … HV2 … HK1) -HR -HV2 /3 width=5 by TC_strap, ex3_2_intro/
]
qed-.
[ #X #H >(TC_lpx_sn_inv_atom2 … H) -X //
| #L2 #I #V2 #IHL2 #X #H
elim (TC_lpx_sn_inv_pair2 … H) // -H -HR
- #L1 #V1 #HL12 #HV12 #H destruct /3 width=1/
+ #L1 #V1 #HL12 #HV12 #H destruct /3 width=1 by/
]
qed-.
lemma TC_lpx_sn_inv_atom1: ∀R,L2. TC … (lpx_sn R) (⋆) L2 → L2 = ⋆.
#R #L2 #H elim H -L2
-[ #L2 #H lapply (lpx_sn_inv_atom1 … H) -H //
-| #L #L2 #_ #HL2 #IHL1 destruct
- lapply (lpx_sn_inv_atom1 … HL2) -HL2 //
+[ /2 width=2 by lpx_sn_inv_atom1/
+| #L #L2 #_ #HL2 #IHL1 destruct /2 width=2 by lpx_sn_inv_atom1/
]
qed-.
∃∃K2,V2. TC … (lpx_sn R) K1 K2 & LTC … R K1 V1 V2 & L2 = K2. ⓑ{I} V2.
#R #HR #L1 #L2 #H @(TC_lpx_sn_ind … H) // -HR -L1 -L2
[ #J #K #W #H destruct
-| #I #L1 #L2 #V1 #V2 #HL12 #HV12 #_ #J #K #W #H destruct /2 width=5/
+| #I #L1 #L2 #V1 #V2 #HL12 #HV12 #_ #J #K #W #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
lemma TC_lpx_sn_inv_lpx_sn_LTC: ∀R. s_rs_trans … R (lpx_sn R) →
∀L1,L2. TC … (lpx_sn R) L1 L2 →
lpx_sn (LTC … R) L1 L2.
-#R #HR #L1 #L2 #H @(TC_lpx_sn_ind … H) // -HR -L1 -L2 /2 width=1/
-qed-.
+/3 width=4 by TC_lpx_sn_ind, lpx_sn_pair/ qed-.
(* Forward lemmas on transitive closure *************************************)
--- /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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ♯ { term 46 L , break term 46 T } )"
+ non associative with precedence 90
+ for @{ 'Weight $L $T }.
(* Basic properties *********************************************************)
lemma fqu_drop_lt: ∀G,L,K,T,U,e. 0 < e →
- ⇩[0, e] L ≡ K → ⇧[0, e] T ≡ U → fqu G L U G K T.
+ ⇩[0, e] L ≡ K → ⇧[0, e] T ≡ U → ⦃G, L, U⦄ ⊃ ⦃G, K, T⦄.
#G #L #K #T #U #e #He >(plus_minus_m_m e 1) /2 width=3 by fqu_drop/
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 "basic_2/relocation/lleq_lleq.ma".
+include "basic_2/relocation/fqu.ma".
+
+(* SUPCLOSURE ***************************************************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lleq_fqu_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃ ⦃G2, K2, U⦄ →
+ ∀L1. L1 ⋕[T] L2 →
+ ∃∃K1. ⦃G1, L1, T⦄ ⊃ ⦃G2, K1, U⦄ & K1 ⋕[U] K2.
+#G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U
+[ #I #G #L2 #V #L1 #H elim (lleq_inv_lref_dx … H I L2 V) -H //
+ #K1 #H1 #H2 lapply (ldrop_inv_O2 … H1) -H1
+ #H destruct /2 width=3 by fqu_lref_O, ex2_intro/
+| * [ #a ] #I #G #L2 #V #T #L1 #H
+ [ elim (lleq_inv_bind … H)
+ | elim (lleq_inv_flat … H)
+ ] -H
+ /2 width=3 by fqu_pair_sn, ex2_intro/
+| #a #I #G #L2 #V #T #L1 #H elim (lleq_inv_bind … H) -H
+ /2 width=3 by fqu_bind_dx, ex2_intro/
+| #I #G #L2 #V #T #L1 #H elim (lleq_inv_flat … H) -H
+ /2 width=3 by fqu_flat_dx, ex2_intro/
+| #G #L2 #K2 #T #U #e #HLK2 #HTU #L1 #HL12
+ elim (ldrop_O1_le (e+1) L1)
+ [ /3 width=11 by fqu_drop, lleq_inv_lift, ex2_intro/
+ | lapply (ldrop_fwd_length_le2 … HLK2) -K2
+ lapply (lleq_fwd_length … HL12) -T -U //
+ ]
+]
+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 "basic_2/relocation/fqu_fqu.ma".
+include "basic_2/relocation/fquq_alt.ma".
+
+(* OPTIONAL SUPCLOSURE ******************************************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lleq_fquq_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃⸮ ⦃G2, K2, U⦄ →
+ ∀L1. L1 ⋕[T] L2 →
+ ∃∃K1. ⦃G1, L1, T⦄ ⊃⸮ ⦃G2, K1, U⦄ & K1 ⋕[U] K2.
+#G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fquq_inv_gen … H) -H
+[ #H elim (lleq_fqu_trans … H … HL12) -L2 /3 width=3 by fqu_fquq, ex2_intro/
+| * #HG #HL #HT destruct /2 width=3 by ex2_intro/
+]
+qed-.
include "ground_2/lstar.ma".
include "basic_2/notation/relations/rdrop_4.ma".
include "basic_2/grammar/lenv_length.ma".
-include "basic_2/grammar/lenv_weight.ma".
+include "basic_2/grammar/cl_restricted_weight.ma".
include "basic_2/relocation/lift.ma".
(* LOCAL ENVIRONMENT SLICING ************************************************)
#L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H /2 width=1/
qed-.
+lemma ldrop_fwd_length_eq: ∀L1,L2,K1,K2,d,e. ⇩[d, e] L1 ≡ K1 → ⇩[d, e] L2 ≡ K2 →
+ |L1| = |L2| → |K1| = |K2|.
+#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
+lapply (ldrop_fwd_length … HLK1) -HLK1
+lapply (ldrop_fwd_length … HLK2) -HLK2
+/2 width=2 by injective_plus_r/ (**) (* full auto fails *)
+qed-.
+
lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize
[ /2 width=3/
]
qed-.
+lemma ldrop_fwd_rfw: ∀I,L,K,V,i. ⇩[O, i] L ≡ K.ⓑ{I}V → ♯{K, V} < ♯{L, #i}.
+#I #L #K #V #i #HLK lapply (ldrop_fwd_lw … HLK) -HLK
+normalize in ⊢ (%→?%%); /2 width=1 by le_S_S/
+qed-.
+
(* Advanced inversion lemmas ************************************************)
fact ldrop_inv_O2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → e = 0 → L1 = L2.
--- /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/notation/relations/lazyeq_3.ma".
+include "basic_2/relocation/ldrop.ma".
+
+(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
+
+inductive lleq: term → relation lenv ≝
+| lleq_sort: ∀L1,L2,k. |L1| = |L2| → lleq (⋆k) L1 L2
+| lleq_lref: ∀I,L1,L2,K1,K2,V,i.
+ ⇩[0, i] L1 ≡ K1.ⓑ{I}V → ⇩[0, i] L2 ≡ K2.ⓑ{I}V →
+ lleq V K1 K2 → lleq (#i) L1 L2
+| lleq_free: ∀L1,L2,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| → lleq (#i) L1 L2
+| lleq_gref: ∀L1,L2,p. |L1| = |L2| → lleq (§p) L1 L2
+| lleq_bind: ∀a,I,L1,L2,V,T.
+ lleq V L1 L2 → lleq T (L1.ⓑ{I}V) (L2.ⓑ{I}V) →
+ lleq (ⓑ{a,I}V.T) L1 L2
+| lleq_flat: ∀I,L1,L2,V,T.
+ lleq V L1 L2 → lleq T L1 L2 → lleq (ⓕ{I}V.T) L1 L2
+.
+
+interpretation
+ "lazy equivalence (local environment)"
+ 'LazyEq T L1 L2 = (lleq T L1 L2).
+
+(* Basic_properties *********************************************************)
+
+lemma lleq_sym: ∀T. symmetric … (lleq T).
+#T #L1 #L2 #H elim H -T -L1 -L2
+/2 width=7 by lleq_sort, lleq_lref, lleq_free, lleq_gref, lleq_bind, lleq_flat/
+qed-.
+
+lemma lleq_refl: ∀T. reflexive … (lleq T).
+#T #L @(f2_ind … rfw … L T)
+#n #IH #L * * /3 width=1 by lleq_sort, lleq_gref, lleq_bind, lleq_flat/
+#i #H elim (lt_or_ge i (|L|)) /2 width=1 by lleq_free/
+#HiL elim (ldrop_O1_lt … HiL) -HiL destruct /4 width=7 by lleq_lref, ldrop_fwd_rfw/
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact lleq_inv_lref_aux: ∀X,L1,L2. L1 ⋕[X] L2 → ∀i. X = #i →
+ (|L1| ≤ i ∧ |L2| ≤ i) ∨
+ ∃∃I,K1,K2,V. ⇩[0, i] L1 ≡ K1.ⓑ{I}V &
+ ⇩[0, i] L2 ≡ K2.ⓑ{I}V &
+ K1 ⋕[V] K2.
+#X #L1 #L2 * -X -L1 -L2
+[ #L1 #L2 #k #_ #j #H destruct
+| #I #L1 #L2 #K1 #K2 #V #i #HLK1 #HLK2 #HK12 #j #H destruct /3 width=7 by ex3_4_intro, or_intror/
+| #L1 #L2 #i #HL1 #HL2 #_ #j #H destruct /3 width=1 by or_introl, conj/
+| #L1 #L2 #p #_ #j #H destruct
+| #a #I #L1 #L2 #V #T #_ #_ #j #H destruct
+| #I #L1 #L2 #V #T #_ #_ #j #H destruct
+]
+qed-.
+
+lemma lleq_inv_lref: ∀L1,L2,i. L1 ⋕[#i] L2 →
+ (|L1| ≤ i ∧ |L2| ≤ i) ∨
+ ∃∃I,K1,K2,V. ⇩[0, i] L1 ≡ K1.ⓑ{I}V &
+ ⇩[0, i] L2 ≡ K2.ⓑ{I}V &
+ K1 ⋕[V] K2.
+/2 width=3 by lleq_inv_lref_aux/ qed-.
+
+fact lleq_inv_bind_aux: ∀X,L1,L2. L1 ⋕[X] L2 → ∀a,I,V,T. X = ⓑ{a,I}V.T →
+ L1 ⋕[V] L2 ∧ L1.ⓑ{I}V ⋕[T] L2.ⓑ{I}V.
+#X #L1 #L2 * -X -L1 -L2
+[ #L1 #L2 #k #_ #b #J #W #U #H destruct
+| #I #L1 #L2 #K1 #K2 #V #i #_ #_ #_ #b #J #W #U #H destruct
+| #L1 #L2 #i #_ #_ #_ #b #J #W #U #H destruct
+| #L1 #L2 #p #_ #b #J #W #U #H destruct
+| #a #I #L1 #L2 #V #T #HV #HT #b #J #W #U #H destruct /2 width=1 by conj/
+| #I #L1 #L2 #V #T #_ #_ #b #J #W #U #H destruct
+]
+qed-.
+
+lemma lleq_inv_bind: ∀a,I,L1,L2,V,T. L1 ⋕[ ⓑ{a,I}V.T] L2 →
+ L1 ⋕[V] L2 ∧ L1.ⓑ{I}V ⋕[T] L2.ⓑ{I}V.
+/2 width=4 by lleq_inv_bind_aux/ qed-.
+
+fact lleq_inv_flat_aux: ∀X,L1,L2. L1 ⋕[X] L2 → ∀I,V,T. X = ⓕ{I}V.T →
+ L1 ⋕[V] L2 ∧ L1 ⋕[T] L2.
+#X #L1 #L2 * -X -L1 -L2
+[ #L1 #L2 #k #_ #J #W #U #H destruct
+| #I #L1 #L2 #K1 #K2 #V #i #_ #_ #_ #J #W #U #H destruct
+| #L1 #L2 #i #_ #_ #_ #J #W #U #H destruct
+| #L1 #L2 #p #_ #J #W #U #H destruct
+| #a #I #L1 #L2 #V #T #_ #_ #J #W #U #H destruct
+| #I #L1 #L2 #V #T #HV #HT #J #W #U #H destruct /2 width=1 by conj/
+]
+qed-.
+
+lemma lleq_inv_flat: ∀I,L1,L2,V,T. L1 ⋕[ ⓕ{I}V.T] L2 →
+ L1 ⋕[V] L2 ∧ L1 ⋕[T] L2.
+/2 width=4 by lleq_inv_flat_aux/ qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lleq_fwd_length: ∀L1,L2,T. L1 ⋕[T] L2 → |L1| = |L2|.
+#L1 #L2 #T #H elim H -L1 -L2 -T //
+#I #L1 #L2 #K1 #K2 #V #i #HLK1 #HLK2 #_ #IHK12
+lapply (ldrop_fwd_length … HLK1) -HLK1
+lapply (ldrop_fwd_length … HLK2) -HLK2
+normalize //
+qed-.
+
+lemma lleq_fwd_ldrop_sn: ∀L1,L2,T. L1 ⋕[T] L2 → ∀K1,i. ⇩[0, i] L1 ≡ K1 →
+ ∃K2. ⇩[0, i] L2 ≡ K2.
+#L1 #L2 #T #H #K1 #i #HLK1 lapply (lleq_fwd_length … H) -H
+#HL12 lapply (ldrop_fwd_length_le2 … HLK1) -HLK1 /2 width=1 by ldrop_O1_le/ (**) (* full auto fails *)
+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 "basic_2/relocation/ldrop_ldrop.ma".
+include "basic_2/relocation/lleq.ma".
+
+(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lleq_inv_lref_dx: ∀L1,L2,i. L1 ⋕[#i] L2 →
+ ∀I,K2,V. ⇩[0, i] L2 ≡ K2.ⓑ{I}V →
+ ∃∃K1. ⇩[0, i] L1 ≡ K1.ⓑ{I}V & K1 ⋕[V] K2.
+#L1 #L2 #i #H #I #K2 #V #HLK2 elim (lleq_inv_lref … H) -H *
+[ #_ #H elim (lt_refl_false i)
+ /3 width=5 by ldrop_fwd_length_lt2, lt_to_le_to_lt/
+| #I0 #K1 #K0 #V0 #HLK1 #HLK0 #HK10 lapply (ldrop_mono … HLK0 … HLK2) -L2
+ #H destruct /2 width=3 by ex2_intro/
+]
+qed-.
+
+lemma lleq_inv_lift: ∀L1,L2,U. L1 ⋕[U] L2 →
+ ∀K1,K2,d,e. ⇩[d, e] L1 ≡ K1 → ⇩[d, e] L2 ≡ K2 →
+ ∀T. ⇧[d, e] T ≡ U → K1 ⋕[T] K2.
+#L1 #L2 #U #H elim H -L1 -L2 -U
+[ #L1 #L2 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H >(lift_inv_sort2 … H) -X
+ lapply (ldrop_fwd_length_eq … HLK1 HLK2 HL12) -L1 -L2 -d -e
+ /2 width=1 by lleq_sort/ (**) (* full auto fails *)
+| #I #L1 #L2 #K #K0 #W #i #HLK #HLK0 #HK0 #IHK0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_lref2 … H) -H
+ * #Hdei #H destruct [ -HK0 | -IHK0 ]
+ [ elim (ldrop_conf_lt … HLK1 … HLK) // -L1 #L1 #V #HKL1 #KL1 #HV0
+ elim (ldrop_conf_lt … HLK2 … HLK0) // -Hdei -L2 #L2 #V2 #HKL2 #K0L2 #HV02
+ lapply (lift_inj … HV02 … HV0) -HV02 #H destruct /3 width=11 by lleq_lref/
+ | lapply (ldrop_conf_ge … HLK1 … HLK ?) // -L1
+ lapply (ldrop_conf_ge … HLK2 … HLK0 ?) // -Hdei -L2
+ /2 width=7 by lleq_lref/
+ ]
+| #L1 #L2 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_lref2 … H) -H
+ * #_ #H destruct
+ lapply (ldrop_fwd_length_eq … HLK1 HLK2 HL12)
+ [ /4 width=3 by lleq_free, ldrop_fwd_length_le4, transitive_le/
+ | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
+ lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
+ /3 width=1 by lleq_free, le_plus_to_minus_r/
+ ]
+| #L1 #L2 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H >(lift_inv_gref2 … H) -X
+ lapply (ldrop_fwd_length_eq … HLK1 HLK2 HL12) -L1 -L2 -d -e
+ /2 width=1 by lleq_gref/ (**) (* full auto fails *)
+| #a #I #L1 #L2 #W #U #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_bind2 … H) -H
+ #V #T #HVW #HTU #H destruct /4 width=5 by lleq_bind, ldrop_skip/
+| #I #L1 #L2 #W #U #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_flat2 … H) -H
+ #V #T #HVW #HTU #H destruct /3 width=5 by lleq_flat/
+]
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma lleq_dec: ∀T,L1,L2. Decidable (L1 ⋕[T] L2).
+#T #L1 @(f2_ind … rfw … L1 T) -L1 -T
+#n #IH #L1 * *
+[ #k #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_sort/
+| #i #H1 #L2 elim (eq_nat_dec (|L1|) (|L2|))
+ [ #HL12 elim (lt_or_ge i (|L1|))
+ #HiL1 elim (lt_or_ge i (|L2|)) /3 width=1 by or_introl, lleq_free/
+ #HiL2 elim (ldrop_O1_lt … HiL2)
+ #I2 #K2 #V2 #HLK2 elim (ldrop_O1_lt … HiL1)
+ #I1 #K1 #V1 #HLK1 elim (eq_bind2_dec I2 I1)
+ [ #H2 elim (eq_term_dec V2 V1)
+ [ #H3 elim (IH K1 V1 … K2) destruct
+ /3 width=7 by lleq_lref, ldrop_fwd_rfw, or_introl/
+ ]
+ ]
+ -IH #H3 @or_intror
+ #H elim (lleq_inv_lref … H) -H *
+ [1,3,5: /3 width=4 by lt_to_le_to_lt, lt_refl_false/ ]
+ #I0 #X1 #X2 #V0 #HLX1 #HLX2 #HX12
+ lapply (ldrop_mono … HLX1 … HLK1) -HLX1 -HLK1
+ lapply (ldrop_mono … HLX2 … HLK2) -HLX2 -HLK2
+ #H #H0 destruct /2 width=1 by/
+ ]
+| #p #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_gref/
+| #a #I #V #T #Hn #L2 destruct
+ elim (IH L1 V … L2) /2 width=1 by/
+ elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V)) -IH /3 width=1 by or_introl, lleq_bind/
+ #H1 #H2 @or_intror
+ #H elim (lleq_inv_bind … H) -H /2 width=1 by/
+| #I #V #T #Hn #L2 destruct
+ elim (IH L1 V … L2) /2 width=1 by/
+ elim (IH L1 T … L2) -IH /3 width=1 by or_introl, lleq_flat/
+ #H1 #H2 @or_intror
+ #H elim (lleq_inv_flat … H) -H /2 width=1 by/
+]
+-n /4 width=2 by lleq_fwd_length, or_intror/
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem lleq_trans: ∀T. Transitive … (lleq T).
+#T #L1 #L #H elim H -T -L1 -L
+/4 width=4 by lleq_fwd_length, lleq_gref, lleq_sort, trans_eq/
+[ #I #L1 #L #K1 #K #V #i #HLK1 #HLK #_ #IHK1 #L2 #H elim (lleq_inv_lref … H) -H *
+ [ -HLK1 -IHK1 #HLi #_ elim (lt_refl_false i)
+ /3 width=5 by ldrop_fwd_length_lt2, lt_to_le_to_lt/
+ | #I0 #K0 #K2 #V0 #HLK0 #HLK2 #HK12 lapply (ldrop_mono … HLK0 … HLK) -L
+ #H destruct /3 width=7 by lleq_lref/
+ ]
+| #L1 #L #i #HL1i #HLi #HL #L2 #H lapply (lleq_fwd_length … H)
+ #HL2 elim (lleq_inv_lref … H) -H * /2 width=1 by lleq_free/
+| #a #I #L1 #L #V #T #_ #_ #IHV #IHT #L2 #H elim (lleq_inv_bind … H) -H
+ /3 width=1 by lleq_bind/
+| #I #L1 #L #V #T #_ #_ #IHV #IHT #L2 #H elim (lleq_inv_flat … H) -H
+ /3 width=1 by lleq_flat/
+]
+qed-.
+
+theorem lleq_canc_sn: ∀L,L1,L2,T. L ⋕[T] L1→ L ⋕[T] L2 → L1 ⋕[T] L2.
+/3 width=3 by lleq_trans, lleq_sym/ qed-.
+
+theorem lleq_canc_dx: ∀L1,L2,L,T. L1 ⋕[T] L → L2 ⋕[T] L → L1 ⋕[T] L2.
+/3 width=3 by lleq_trans, lleq_sym/ qed-.
class "orange"
[ { "relocation" * } {
[ { "structural successor for closures" * } {
- [ "fqu ( ⦃?,?,?⦄ ⊃ ⦃?,?,?⦄ )" "fquq ( ⦃?,?,?⦄ ⊃⸮ ⦃?,?,?⦄ )" "fquq_alt ( ⦃?,?,?⦄ ⊃⊃⸮ ⦃?,?,?⦄ )" * ]
+ [ "fquq ( ⦃?,?,?⦄ ⊃⸮ ⦃?,?,?⦄ )" "fquq_alt ( ⦃?,?,?⦄ ⊃⊃⸮ ⦃?,?,?⦄ )" "fquq_fquq" * ]
+ [ "fqu ( ⦃?,?,?⦄ ⊃ ⦃?,?,?⦄ )" "fqu_fqu" * ]
}
]
+(*
[ { "lazy equivalence for closures" * } {
[ "fleq ( ⦃?,?,?⦄ ⋕ ⦃?,?,?⦄ )" "fleq_fleq" * ]
}
]
+*)
[ { "lazy equivalence for local environments" * } {
[ "lleq ( ? ⋕[?] ? )" "lleq_fleq" * ]
}
(* *)
(**************************************************************************)
-include "basics/star.ma".
+include "basics/star1.ma".
include "ground_2/xoa_props.ma".
include "ground_2/notation.ma".