Theory Position
section ‹Positions (of terms, contexts, etc.)›
text ‹Positions are just list of natural numbers, and here we define
standard notions such as the prefix-relation, parallel positions, left-of, etc.
Note that we also instantiate lists in certain ways, such that
we can write $p^n$ for the n-fold concatenation of the position $p$.›
theory Position
imports
"HOL-Library.Infinite_Set"
"HOL-Library.Sublist"
Show.Shows_Literal
begin
type_synonym pos = "nat list"
definition less_eq_pos :: "pos ⇒ pos ⇒ bool" (infix "≤⇩p" 50) where
"p ≤⇩p q ⟷ (∃r. p @ r = q)"
definition less_pos :: "pos ⇒ pos ⇒ bool" (infix "<⇩p" 50) where
"p <⇩p q ⟷ p ≤⇩p q ∧ p ≠ q"
lemma less_eq_pos_eq_prefix:
"less_eq_pos = Sublist.prefix"
unfolding less_eq_pos_def Sublist.prefix_def by metis
lemma less_pos_eq_strict_prefix:
"less_pos = Sublist.strict_prefix"
unfolding less_pos_def less_eq_pos_def Sublist.strict_prefix_def Sublist.prefix_def by metis