--- /dev/null
+include "delayed_updating/notation/functions/nodelabel_d_2.ma".
+
+| label_d2: pnat → nat → label
+
+interpretation
+ "variable reference by depth with offset (label)"
+ 'NodeLabelD k d = (label_d2 k d).
+
+ | label_d2 k d ⇒ depth q
+
+lemma depth_d2_dx (p) (k) (d):
+ ♭p = ♭(p◖𝗱❨k,d❩).
+// qed.
+
+lemma depth_d2_sn (p) (k) (d):
+ ♭p = ♭(𝗱❨k,d❩◗p).
+// qed.
+
+ | label_d2 k d ⇒ structure q
+
+lemma structure_d2_dx (p) (k) (d):
+ ⊗p = ⊗(p◖𝗱❨k,d❩).
+// qed.
+
+lemma structure_d2_sn (p) (k) (d):
+ ⊗p = ⊗(𝗱❨k,d❩◗p).
+#p #k #d <structure_append //
+qed.
+
+lemma eq_inv_d2_dx_structure (d) (h) (q) (p):
+ q◖𝗱❨h,d❩ = ⊗p → ⊥.
+#d #h #q #p elim p -p [| * [ #k | #k #d ] #p #IH ]
+[ <structure_empty #H0 destruct
+| <structure_d_dx #H0 /2 width=1 by/
+| <structure_d2_dx #H0 /2 width=1 by/
+| <structure_m_dx #H0 /2 width=1 by/
+| <structure_L_dx #H0 destruct
+| <structure_A_dx #H0 destruct
+| <structure_S_dx #H0 destruct
+]
+qed-.
+
+lemma eq_inv_d2_sn_structure (d) (h) (q) (p):
+ (𝗱❨h,d❩◗q) = ⊗p → ⊥.
+#d #h #q #p >list_cons_comm #H0
+elim (eq_inv_append_structure … H0) -H0 #r1 #r2
+<list_cons_comm #H0 #H1 #H2 destruct
+elim (eq_inv_d2_dx_structure … H0)
+qed-.
+
+include "delayed_updating/notation/functions/tau_3.ma".
+
+interpretation
+ "inner variable reference by depth with offset (prototerm)"
+ 'Tau k d t = (prototerm_node_1_2 (label_d2 k d) label_m t).
+
+lemma in_comp_iref2 (t) (q) (k) (d):
+ q ϵ t → 𝗱❨k,d❩◗𝗺◗q ϵ 𝛕❨k,d❩.t.
+/2 width=3 by ex2_intro/ qed.
+
+lemma in_comp_inv_iref2 (t) (p) (k) (d):
+ p ϵ 𝛕❨k,d❩.t →
+ ∃∃q. 𝗱❨k,d❩◗𝗺◗q = p & q ϵ t.
+#t #p #k #d * #q #Hq #Hp
+/2 width=3 by ex2_intro/
+qed-.
+
+lemma ppc_iref2 (t) (k) (d):
+ (𝛕❨k,d❩.t) ϵ 𝐏.
+#t #k #d #p * #q #Hq #H0 destruct //
+qed.
+
+lemma pic_inv_d2_dx (p) (k) (d):
+ p◖𝗱❨k,d❩ ϵ 𝐈 → ⊥.
+#p #k #d @(insert_eq_1 … (p◖𝗱❨k,d❩))
+#q * -q [|*: #q ] #H0 destruct
+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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR DELAYED UPDATING ********************************************)
+
+notation "hvbox( 𝗱 ❨ break term 46 k, break term 46 d ❩ )"
+ non associative with precedence 70
+ for @{ 'NodeLabelD $k $d }.
--- /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 DELAYED UPDATING ********************************************)
+
+notation "hvbox( 𝛕 ❨ break term 46 k, break term 46 d ❩. break term 70 t )"
+ non associative with precedence 70
+ for @{ 'Tau $k $d $t }.
+++ /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 DELAYED UPDATING ********************************************)
-
-notation "hvbox( 𝗱 ❨ break term 46 k, break term 46 d ❩ )"
- non associative with precedence 70
- for @{ 'NodeLabelD $k $d }.
+++ /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 DELAYED UPDATING ********************************************)
-
-notation "hvbox( 𝛕 ❨ break term 46 k, break term 46 d ❩. break term 70 t )"
- non associative with precedence 70
- for @{ 'Tau $k $d $t }.
* Initial invocation: - Patience on me to gain peace and perfection! -
*)
-include "ground/arith/nat.ma".
+include "ground/arith/pnat.ma".
include "delayed_updating/notation/functions/nodelabel_d_1.ma".
-include "delayed_updating/notation/functions/nodelabel_d_2.ma".
include "delayed_updating/notation/functions/nodelabel_m_0.ma".
include "delayed_updating/notation/functions/edgelabel_l_0.ma".
include "delayed_updating/notation/functions/edgelabel_a_0.ma".
(* LABEL ********************************************************************)
inductive label: Type[0] ≝
-| label_d : pnat → label
-| label_d2: pnat → nat → label
-| label_m : label
-| label_L : label
-| label_A : label
-| label_S : label
+| label_d: pnat → label
+| label_m: label
+| label_L: label
+| label_A: label
+| label_S: label
.
interpretation
"variable reference by depth (label)"
'NodeLabelD k = (label_d k).
-interpretation
- "variable reference by depth with offset (label)"
- 'NodeLabelD k d = (label_d2 k d).
-
interpretation
"mark (label)"
'NodeLabelM = (label_m).
[ list_empty ⇒ 𝟎
| list_lcons l q ⇒
match l with
- [ label_d k ⇒ depth q
- | label_d2 k d ⇒ depth q
- | label_m ⇒ depth q
- | label_L ⇒ ↑(depth q)
- | label_A ⇒ depth q
- | label_S ⇒ depth q
+ [ label_d k ⇒ depth q
+ | label_m ⇒ depth q
+ | label_L ⇒ ↑(depth q)
+ | label_A ⇒ depth q
+ | label_S ⇒ depth q
]
].
♭p = ♭(p◖𝗱k).
// qed.
-lemma depth_d2_dx (p) (k) (d):
- ♭p = ♭(p◖𝗱❨k,d❩).
-// qed.
-
lemma depth_m_dx (p):
♭p = ♭(p◖𝗺).
// qed.
theorem depth_append (p) (q):
(♭p)+(♭q) = ♭(p●q).
#p #q elim q -q //
-* [ #k | #k #d ] #q #IH <list_append_lcons_sn
+* [ #k ] #q #IH <list_append_lcons_sn
[ <depth_d_dx <depth_d_dx //
-| <depth_d2_dx <depth_d2_dx //
| <depth_m_dx <depth_m_dx //
| <depth_L_dx <depth_L_dx //
| <depth_A_dx <depth_A_dx //
♭p = ♭(𝗱k◗p).
// qed.
-lemma depth_d2_sn (p) (k) (d):
- ♭p = ♭(𝗱❨k,d❩◗p).
-// qed.
-
lemma depth_m_sn (p):
♭p = ♭(𝗺◗p).
// qed.
#q * -q [|*: #q ] #H0 destruct
qed-.
-lemma pic_inv_d2_dx (p) (k) (d):
- p◖𝗱❨k,d❩ ϵ 𝐈 → ⊥.
-#p #k #d @(insert_eq_1 … (p◖𝗱❨k,d❩))
-#q * -q [|*: #q ] #H0 destruct
-qed-.
-
(* Constructions with path_lcons ********************************************)
lemma pic_m_sn (p):
[ list_empty ⇒ 𝐞
| list_lcons l q ⇒
match l with
- [ label_d k ⇒ structure q
- | label_d2 k d ⇒ structure q
- | label_m ⇒ structure q
- | label_L ⇒ (structure q)◖𝗟
- | label_A ⇒ (structure q)◖𝗔
- | label_S ⇒ (structure q)◖𝗦
+ [ label_d k ⇒ structure q
+ | label_m ⇒ structure q
+ | label_L ⇒ (structure q)◖𝗟
+ | label_A ⇒ (structure q)◖𝗔
+ | label_S ⇒ (structure q)◖𝗦
]
].
⊗p = ⊗(p◖𝗱k).
// qed.
-lemma structure_d2_dx (p) (k) (d):
- ⊗p = ⊗(p◖𝗱❨k,d❩).
-// qed.
-
lemma structure_m_dx (p):
⊗p = ⊗(p◖𝗺).
// qed.
theorem structure_idem (p):
⊗p = ⊗⊗p.
#p elim p -p //
-* [ #k | #k #d ] #p #IH //
+* [ #k ] #p #IH //
qed.
theorem structure_append (p) (q):
⊗p●⊗q = ⊗(p●q).
#p #q elim q -q //
-* [ #k | #k #d ] #q #IH //
+* [ #k ] #q #IH //
<list_append_lcons_sn //
qed.
#p #k <structure_append //
qed.
-lemma structure_d2_sn (p) (k) (d):
- ⊗p = ⊗(𝗱❨k,d❩◗p).
-#p #k #d <structure_append //
-qed.
-
lemma structure_m_sn (p):
⊗p = ⊗(𝗺◗p).
#p <structure_append //
lemma eq_inv_d_dx_structure (h) (q) (p):
q◖𝗱h = ⊗p → ⊥.
-#h #q #p elim p -p [| * [ #k | #k #d ] #p #IH ]
+#h #q #p elim p -p [| * [ #k ] #p #IH ]
[ <structure_empty #H0 destruct
| <structure_d_dx #H0 /2 width=1 by/
-| <structure_d2_dx #H0 /2 width=1 by/
-| <structure_m_dx #H0 /2 width=1 by/
-| <structure_L_dx #H0 destruct
-| <structure_A_dx #H0 destruct
-| <structure_S_dx #H0 destruct
-]
-qed-.
-
-lemma eq_inv_d2_dx_structure (d) (h) (q) (p):
- q◖𝗱❨h,d❩ = ⊗p → ⊥.
-#d #h #q #p elim p -p [| * [ #k | #k #d ] #p #IH ]
-[ <structure_empty #H0 destruct
-| <structure_d_dx #H0 /2 width=1 by/
-| <structure_d2_dx #H0 /2 width=1 by/
| <structure_m_dx #H0 /2 width=1 by/
| <structure_L_dx #H0 destruct
| <structure_A_dx #H0 destruct
lemma eq_inv_m_dx_structure (q) (p):
q◖𝗺 = ⊗p → ⊥.
-#q #p elim p -p [| * [ #k | #k #d ] #p #IH ]
+#q #p elim p -p [| * [ #k ] #p #IH ]
[ <structure_empty #H0 destruct
| <structure_d_dx #H0 /2 width=1 by/
-| <structure_d2_dx #H0 /2 width=1 by/
| <structure_m_dx #H0 /2 width=1 by/
| <structure_L_dx #H0 destruct
| <structure_A_dx #H0 destruct
lemma eq_inv_L_dx_structure (q) (p):
q◖𝗟 = ⊗p →
∃∃r1,r2. q = ⊗r1 & 𝐞 = ⊗r2 & r1●𝗟◗r2 = p.
-#q #p elim p -p [| * [ #k | #k #d ] #p #IH ]
+#q #p elim p -p [| * [ #k ] #p #IH ]
[ <structure_empty #H0 destruct
| <structure_d_dx #H0
elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
/2 width=5 by ex3_2_intro/
-| <structure_d2_dx #H0
- elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
- /2 width=5 by ex3_2_intro/
| <structure_m_dx #H0
elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
/2 width=5 by ex3_2_intro/
lemma eq_inv_A_dx_structure (q) (p):
q◖𝗔 = ⊗p →
∃∃r1,r2. q = ⊗r1 & 𝐞 = ⊗r2 & r1●𝗔◗r2 = p.
-#q #p elim p -p [| * [ #k | #k #d ] #p #IH ]
+#q #p elim p -p [| * [ #k ] #p #IH ]
[ <structure_empty #H0 destruct
| <structure_d_dx #H0
elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
/2 width=5 by ex3_2_intro/
-| <structure_d2_dx #H0
- elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
- /2 width=5 by ex3_2_intro/
| <structure_m_dx #H0
elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
/2 width=5 by ex3_2_intro/
lemma eq_inv_S_dx_structure (q) (p):
q◖𝗦 = ⊗p →
∃∃r1,r2. q = ⊗r1 & 𝐞 = ⊗r2 & r1●𝗦◗r2 = p.
-#q #p elim p -p [| * [ #k | #k #d ] #p #IH ]
+#q #p elim p -p [| * [ #k ] #p #IH ]
[ <structure_empty #H0 destruct
| <structure_d_dx #H0
elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
/2 width=5 by ex3_2_intro/
-| <structure_d2_dx #H0
- elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
- /2 width=5 by ex3_2_intro/
| <structure_m_dx #H0
elim IH -IH // -H0 #r1 #r2 #H1 #H0 #H2 destruct
/2 width=5 by ex3_2_intro/
theorem eq_inv_append_structure (p) (q) (r):
p●q = ⊗r →
∃∃r1,r2.p = ⊗r1 & q = ⊗r2 & r1●r2 = r.
-#p #q elim q -q [| * [ #k | #k #d ] #q #IH ] #r
+#p #q elim q -q [| * [ #k ] #q #IH ] #r
[ <list_append_empty_sn #H0 destruct
/2 width=5 by ex3_2_intro/
| #H0 elim (eq_inv_d_dx_structure … H0)
-| #H0 elim (eq_inv_d2_dx_structure … H0)
| #H0 elim (eq_inv_m_dx_structure … H0)
| #H0 elim (eq_inv_L_dx_structure … H0) -H0 #r1 #r2 #Hr1 #Hr2 #H0 destruct
elim (IH … Hr1) -IH -Hr1 #s1 #s2 #H1 #H2 #H3 destruct
elim (eq_inv_d_dx_structure … H0)
qed-.
-lemma eq_inv_d2_sn_structure (d) (h) (q) (p):
- (𝗱❨h,d❩◗q) = ⊗p → ⊥.
-#d #h #q #p >list_cons_comm #H0
-elim (eq_inv_append_structure … H0) -H0 #r1 #r2
-<list_cons_comm #H0 #H1 #H2 destruct
-elim (eq_inv_d2_dx_structure … H0)
-qed-.
-
lemma eq_inv_m_sn_structure (q) (p):
(𝗺 ◗q) = ⊗p → ⊥.
#q #p >list_cons_comm #H0
⊗p ϵ 𝐈.
#p elim p -p
[ <structure_empty //
-| * [ #k | #k #d ] #p #IH
+| * [ #k ] #p #IH
[ <structure_d_dx //
- | <structure_d2_dx //
| <structure_m_dx //
| <structure_L_dx //
| <structure_A_dx //
include "delayed_updating/notation/functions/m_hook_1.ma".
include "delayed_updating/notation/functions/hash_1.ma".
include "delayed_updating/notation/functions/tau_2.ma".
-include "delayed_updating/notation/functions/tau_3.ma".
include "delayed_updating/notation/functions/lamda_1.ma".
include "delayed_updating/notation/functions/at_2.ma".
"inner variable reference by depth (prototerm)"
'Tau k t = (prototerm_node_1_2 (label_d k) label_m t).
-interpretation
- "inner variable reference by depth with offset (prototerm)"
- 'Tau k d t = (prototerm_node_1_2 (label_d2 k d) label_m t).
-
interpretation
"name-free functional abstraction (prototerm)"
'Lamda t = (prototerm_node_1 label_L t).
q ϵ t → 𝗱k◗𝗺◗q ϵ 𝛕k.t.
/2 width=3 by ex2_intro/ qed.
-lemma in_comp_iref2 (t) (q) (k) (d):
- q ϵ t → 𝗱❨k,d❩◗𝗺◗q ϵ 𝛕❨k,d❩.t.
-/2 width=3 by ex2_intro/ qed.
-
(* Basic inversions *********************************************************)
lemma in_comp_inv_iref (t) (p) (k):
/2 width=3 by ex2_intro/
qed-.
-lemma in_comp_inv_iref2 (t) (p) (k) (d):
- p ϵ 𝛕❨k,d❩.t →
- ∃∃q. 𝗱❨k,d❩◗𝗺◗q = p & q ϵ t.
-#t #p #k #d * #q #Hq #Hp
-/2 width=3 by ex2_intro/
-qed-.
-
(* COMMENT
lemma prototerm_in_root_inv_lcons_oref:
∀p,l,n. l◗p ϵ ▵#n →
(𝛕k.t) ϵ 𝐏.
#t #k #p * #q #Hq #H0 destruct //
qed.
-
-lemma ppc_iref2 (t) (k) (d):
- (𝛕❨k,d❩.t) ϵ 𝐏.
-#t #k #d #p * #q #Hq #H0 destruct //
-qed.