--- /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/list.ma".
+include "lambda-delta/notation.ma".
+
--- /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_______________________________________________________________ *)
+
+(* THE FORMAL SYSTEM λδ - MATITA SOURCE SCRIPTS
+ * Specification started: 2011 April 17
+ * - Patience on me so that I gain peace and perfection! -
+ * [ suggested invocation to start formal specifications with ]
+ *)
+
+include "lambda-delta/ground.ma".
+
+(* BINARY ITEMS *************************************************************)
+
+(* binary items *)
+inductive item2: Type[0] ≝
+ | Abbr: item2 (* abbreviation *)
+ | Abst: item2 (* abstraction *)
+ | Appl: item2 (* application *)
+ | Cast: item2 (* explicit type annotation *)
+.
+
+(* reduction-related categorization *****************************************)
+
+(* binary items entitled for theta reduction *)
+definition thable2 ≝ λI. I = Abbr.
+
+(* binary items entitled for zeta reduction *)
+definition zable2 ≝ λI. I = Abbr ∨ I = Cast.
--- /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 "lambda-delta/language/term.ma".
+
+(* LOCAL ENVIRONMENTS *******************************************************)
+
+(* local environments *)
+inductive lenv: Type[0] ≝
+ | LSort: lenv (* empty *)
+ | LCon2: lenv → item2 → term → lenv (* binary item construction *)
+.
+
+interpretation "sort (local environment)" 'Star = LSort.
+
+interpretation "environment construction (binary)" 'DCon L I T = (LCon2 L I T).
--- /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 "lambda-delta/ground.ma".
+
+(* SORT HIERARCHY ***********************************************************)
+
+(* sort hierarchy specifications *)
+record sh: Type[0] ≝ {
+ next: nat → nat; (* next sort in the hierarchy *)
+ next_lt: ∀k. k < next k (* strict monotonicity condition *)
+}.
--- /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 "lambda-delta/language/item2.ma".
+
+(* TERMS ********************************************************************)
+
+(* terms *)
+inductive term: Type[0] ≝
+ | TSort: nat → term (* sort: starting at 0 *)
+ | TLRef: nat → term (* reference by index: starting at 0 *)
+ | TCon2: item2 → term → term → term (* binary construction *)
+.
+
+interpretation "sort (term)" 'Star k = (TSort k).
+
+interpretation "local reference (term)" 'Weight i = (TLRef i).
+
+interpretation "term construction (binary)" 'SCon I T1 T2 = (TCon2 I T1 T2).
--- /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 "lambda-delta/language/lenv.ma".
+
+(* WEIGHTS ******************************************************************)
+
+(* the weight of a term *)
+let rec tw T ≝ match T with
+ [ TSort _ ⇒ 1
+ | TLRef _ ⇒ 1
+ | TCon2 _ V T ⇒ tw V + tw T + 1
+ ].
+
+interpretation "weight (term)" 'Weight T = (tw T).
+
+(* the weight of a local environment *)
+let rec lw L ≝ match L with
+ [ LSort ⇒ 0
+ | LCon2 L _ V ⇒ lw L + #V
+ ].
+
+interpretation "weight (local environment)" 'Weight L = (lw L).
+
+(* the weight of a closure *)
+definition cw ≝ λL,T. #L + #T.
+
+interpretation "weight (closure)" 'Weight L T = (cw L T).
--- /dev/null
+NAMING CONVENTIONS FOR METAVARIABLES
+
+I,J : item
+K,L : local environment
+T,U,V,W: term
+
+d : relocation depth
+e : relocation height
+h : sort hierarchy parameter
+i,j : local reference position index (de Bruijn's)
+k : sort index
--- /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_______________________________________________________________ *)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+(* language *****************************************************************)
+
+notation "hvbox( ⋆ )"
+ non associative with precedence 90
+ for @{ 'Star }.
+
+notation "hvbox( ⋆ k )"
+ non associative with precedence 90
+ for @{ 'Star $k }.
+
+notation "hvbox( ♭ (term 90 I) break T1 . break T )"
+ non associative with precedence 90
+ for @{ 'SCon $I $T1 $T }.
+
+notation "hvbox( T . break ♭ (term 90 I) break T1 )"
+ non associative with precedence 90
+ for @{ 'DCon $T $I $T1 }.
+
+notation "hvbox( # term 90 x )"
+ non associative with precedence 90
+ for @{ 'Weight $x }.
+
+notation "hvbox( # [ x , break y ] )"
+ non associative with precedence 90
+ for @{ 'Weight $x $y }.
+
+(* substitution *************************************************************)
+
+notation "hvbox( [ d , break e ] ↑ break T1 ≡ break T2 )"
+ non associative with precedence 45
+ for @{ 'RLift $T1 $d $e $T2 }.
+
+notation "hvbox( [ d , break e ] ← break (term 90 L) / break T1 ≡ break T2 )"
+ non associative with precedence 45
+ for @{ 'RSubst $L $T1 $d $e $T2 }.
--- /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 "lambda-delta/language/term.ma".
+
+(* RELOCATION ***************************************************************)
+
+inductive lift: term → nat → nat → term → Prop ≝
+ | lift_sort : ∀k,d,e. lift (⋆k) d e (⋆k)
+ | lift_lref_lt: ∀i,d,e. i < d → lift (#i) d e (#i)
+ | lift_lref_ge: ∀i,d,e. d ≤ i → lift (#i) d e (#(i + e))
+ | lift_con2 : ∀I,V1,V2,T1,T2,d,e.
+ lift V1 d e V2 → lift T1 (d + 1) e T2 →
+ lift (♭I V1. T1) d e (♭I V2. T2)
+.
+
+interpretation "relocation" 'RLift T1 d e T2 = (lift T1 d e T2).
--- /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 "lambda-delta/language/lenv.ma".
+include "lambda-delta/substitution/lift.ma".
+
+(* SUBSTITUTION *************************************************************)
+
+inductive subst: lenv → term → nat → nat → term → Prop ≝
+ | subst_sort : ∀L,k,d,e. subst L (⋆k) d e (⋆k)
+ | subst_lref_lt: ∀L,i,d,e. i < d → subst L (#i) d e (#i)
+ | subst_lref_O : ∀L,V,e. 0 < e → subst (L. ♭Abbr V) #0 0 e V
+ | subst_lref_S : ∀L,I,V,i,T1,T2,d,e.
+ d ≤ i → i < d + e → subst L #i d e T1 → [1,d]↑ T1 ≡ T2 →
+ subst (L. ♭I V) #(i + 1) (d + 1) e T2
+ | subst_lref_ge: ∀L,i,d,e. d + e ≤ i → subst L (#i) d e (#(i - e))
+ | subst_con2 : ∀L,I,V1,V2,T1,T2,d,e.
+ subst L V1 d e V2 → subst (L. ♭I V1) T1 (d + 1) e T2 →
+ subst L (♭I V1. T1) d e (♭I V2. T2)
+.
+
+interpretation "telescopic substritution" 'RSubst L T1 d e T2 = (subst L T1 d e T2).
+
+lemma subst_lift_inv: ∀d,e,T1,T2. [d,e]↑ T1 ≡ T2 → ∀L. [d,e]← L / T2 ≡ T1.
+#d #e #T1 #T2 #H elim H -H d e T1 T2 /2/
+#i #d #e #Hdi #L >(minus_plus_m_m i e) in ⊢ (? ? ? ? ? %) /3/
+qed.
+