我有以下语言记录:
datatype "term" = Rcd "string ⇀ term"
fun id_term :: "term ⇒ term"
where "id_term (Rcd vals) = Rcd (map_option id_term ∘ vals)"
这不会通过终止检查程序,因为类型的size
函数始终为0.我没有看到如何在不将地图约束到有限域的情况下提供可计算的度量。
那么:我如何证明上述定义的终止?我怀疑我必须证明某些归纳谓词超过条款的基础,但我不确定如何做到这一点。
答案 0 :(得分:2)
您可以使用primrec
:
primrec id_term :: "term ⇒ term"
where "id_term (Rcd vals) = Rcd (map_option id_term ∘ vals)"
primrec
允许您使用所涉及类型构造函数的map
函数。在您的情况下,那是map_option
和op ∘
。
答案 1 :(得分:2)
对于非原始递归函数,您也可以使用function
命令,但事情变得有点复杂,因为您没有size
度量。从本质上讲,您必须定义一个子项关系并显示它是有充分根据的,然后您可以使用它来表明您的函数终止:
datatype "term" = Rcd "string ⇀ term"
inductive subterm :: "term ⇒ term ⇒ bool" where
"t ∈ ran f ⟹ subterm t (Rcd f)"
lemma accp_subterm: "Wellfounded.accp subterm t"
proof (induction t)
case (Rcd f)
have IH: "Wellfounded.accp subterm t" if "t ∈ ran f" for t
using Rcd[of "Some t" t] and that by (auto simp: eq_commute ran_def)
show ?case by (rule accpI) (auto intro: IH elim!: subterm.cases)
qed
definition subterm_rel where "subterm_rel = {(t, Rcd f) |f t. t ∈ ran f}"
lemma subterm_rel_altdef: "subterm_rel = {(s, t) |s t. subterm s t}"
by (auto simp: subterm_rel_def subterm.simps)
lemma subterm_relI [intro]: "t ∈ ran f ⟹ (t, Rcd f) ∈ subterm_rel"
by (simp add: subterm_rel_def)
lemma subterm_relI' [intro]: "Some t = f x ⟹ (t, Rcd f) ∈ subterm_rel"
by (auto simp: subterm_rel_def ran_def)
lemma wf_subterm_rel [simp, intro]: "wf subterm_rel"
using accp_subterm unfolding subterm_rel_altdef accp_eq_acc wf_acc_iff by simp
function id_term :: "term ⇒ term"
where "id_term (Rcd vals) = Rcd (map_option id_term ∘ vals)"
by pat_completeness simp_all
termination by (relation subterm_rel) auto