Theory Lambda_Free_RPOs.Lambda_Free_Term
section ‹Lambda-Free Higher-Order Terms›
theory Lambda_Free_Term
imports Lambda_Free_Util
abbrevs ">s" = ">⇩s"
and ">h" = ">⇩h⇩d"
and "≤≥h" = "≤≥⇩h⇩d"
begin
text ‹
This theory defines ‹λ›-free higher-order terms and related locales.
›
subsection ‹Precedence on Symbols›
locale gt_sym =
fixes
gt_sym :: "'s ⇒ 's ⇒ bool" (infix ">⇩s" 50)
assumes
gt_sym_irrefl: "¬ f >⇩s f" and
gt_sym_trans: "h >⇩s g ⟹ g >⇩s f ⟹ h >⇩s f" and
gt_sym_total: "f >⇩s g ∨ g >⇩s f ∨ g = f" and
gt_sym_wf: "wfP (λf g. g >⇩s f)"
begin
lemma gt_sym_antisym: "f >⇩s g ⟹ ¬ g >⇩s f"
by (metis gt_sym_irrefl gt_sym_trans)
end
subsection ‹Heads›