使用集合理解的函数的终止证明



考虑以下愚蠢的Isabelle对树和子树的定义:

datatype tree = Leaf int
              | Node tree tree
fun children :: "tree ⇒ tree set" where
"children (Leaf _) = {}" |
"children (Node a b) = {a, b}"
lemma children_decreasing_size:
  assumes "c ∈ children t"
  shows   "size c < size t"
using assms
by (induction t, auto)
function subtrees :: "tree ⇒ tree set" where
"subtrees t = { s | c s. c ∈ children t ∧ s ∈ subtrees c }"
by auto
termination
apply (relation "measure size", simp)

subtrees的终止证明在这一点上被卡住了,尽管递归调用只在子级上进行,这些子级严格地小于建立良好的size关系(如平凡引理所示)。

证明状态如下:

goal (1 subgoal):
 1. ⋀t x xa xb. (xa, t) ∈ measure size

当然,这是不可能证明的,因为xat的子代的信息已经丢失。我做错什么了吗?我能做些什么来保存证据吗?我注意到,我可以使用列表而不是集合来制定相同的定义:

fun children_list :: "tree ⇒ tree list" where
"children_list (Leaf _) = []" |
"children_list (Node a b) = [a, b]"
function subtrees_list :: "tree ⇒ tree list" where
"subtrees_list t = concat (map subtrees_list (children_list t))"
by auto
termination
apply (relation "measure size", simp)

并获得一个信息更丰富、可证明的终止目标:

goal (1 subgoal):
 1. ⋀t x.
       x ∈ set (children_list t) ⟹
       (x, t) ∈ measure size

这是Isabelle的一些限制吗?我应该通过不使用集合来解决这个问题?

subtrees的集合理解中对c : children t的限制不会出现在终止证明义务中,因为函数包对&一无所知。同余规则可以用来实现这一点。在这种情况下,您可以将conj_cong本地声明为[fundef_cong],以实质上模拟从左到右的求值顺序(尽管在HOL中没有求值这样的东西)。例如,

context notes conj_cong[fundef_cong] begin
fun subtrees :: ...
termination ...
end

上下文块确保声明conj_cong[fundef_cong]仅对此函数定义有效。

带列表的版本之所以有效,是因为它使用了函数map,默认情况下有一个同余规则。如果您在集合上使用了一元绑定操作(而不是集合理解),那么集合也应该如此。

最新更新