1

I'm trying to prove that a transitive relation on elements of lists is equivalent to a transitive relation on lists (under some conditions).

Here is a first lemma:

lemma list_all2_rtrancl1:
  "(list_all2 P)⇧*⇧* xs ys ⟹
   list_all2 P⇧*⇧* xs ys"
  apply (induct rule: rtranclp_induct)
  apply (simp add: list.rel_refl)
  by (smt list_all2_trans rtranclp.rtrancl_into_rtrancl)

And here is a symmetric lemma:

lemma list_all2_rtrancl2:
  "(⋀x. P x x) ⟹
   list_all2 P⇧*⇧* xs ys ⟹
   (list_all2 P)⇧*⇧* xs ys"
  apply (erule list_all2_induct)
  apply simp

I guess that a relation should be reflexive. But maybe I should use another assumptions. The lemma could be proven given the assumption that P is transitive, however P is not transitive. I'm stuck. Could you suggest what assumptions to choose and how to prove this lemma?

It seems that nitpick gives me a wrong counterexample for the specific case of the last lemma (xs = [0] and ys = [2]):

lemma list_all2_rtrancl2_example:
  "list_all2 (λx y. x = y ∨ Suc x = y)⇧*⇧* xs ys ⟹
   (list_all2 (λx y. x = y ∨ Suc x = y))⇧*⇧* xs ys"
  nitpick

I can prove that the lemma holds for this example:

lemma list_all2_rtrancl2_example_0_2:
  "list_all2 (λx y. x = y ∨ Suc x = y)⇧*⇧* [0] [2] ⟹
   (list_all2 (λx y. x = y ∨ Suc x = y))⇧*⇧* [0] [2]"
  apply (rule_tac ?b="[1]" in converse_rtranclp_into_rtranclp; simp)
  apply (rule_tac ?b="[2]" in converse_rtranclp_into_rtranclp; simp)
  done
Denis
  • 1,167
  • 1
  • 10
  • 30

1 Answers1

1

It may be feasible to use listrel instead of list_all2. Indeed, as shown below, they are equivalent (see set_listrel_eq_list_all2). However, there are several theorems in the standard library about listrel that do not have their equivalents for list_all2.


lemma set_listrel_eq_list_all2: 
  "listrel {(x, y). r x y} = {(xs, ys). list_all2 r xs ys}"
  using list_all2_conv_all_nth listrel_iff_nth by fastforce

lemma listrel_tclosure_1: "(listrel r)⇧* ⊆ listrel (r⇧*)"
  by 
    (
      simp add: 
        listrel_rtrancl_eq_rtrancl_listrel1 
        listrel_subset_rtrancl_listrel1 
        rtrancl_subset_rtrancl
    )

lemma listrel_tclosure_2: "refl r ⟹ listrel (r⇧*) ⊆ (listrel r)⇧*"
  by 
    (
      simp add: 
        listrel1_subset_listrel 
        listrel_rtrancl_eq_rtrancl_listrel1 
        rtrancl_mono
    )

context 
  includes lifting_syntax
begin

lemma listrel_list_all2_transfer[transfer_rule]:
  "((=) ===> (=) ===> (=) ===> (=)) 
  (λr xs ys. (xs, ys) ∈ listrel {(x, y). r x y}) list_all2"
  unfolding rel_fun_def using set_listrel_eq_list_all2 listrel_iff_nth by blast

end

lemma list_all2_rtrancl_1:
  "(list_all2 r)⇧*⇧* xs ys ⟹ list_all2 r⇧*⇧* xs ys"
proof transfer
  fix r :: "'a ⇒ 'a ⇒ bool" and xs :: "'a list" and ys:: "'a list"
  assume "(λxs ys. (xs, ys) ∈ listrel {(x, y). r x y})⇧*⇧* xs ys"
  then have "(xs, ys) ∈ (listrel {(x, y). r x y})⇧*"
    unfolding rtranclp_def rtrancl_def by auto  
  then have "(xs, ys) ∈ listrel ({(x, y). r x y}⇧*)" 
    using listrel_tclosure_1 by auto
  then show "(xs, ys) ∈ listrel {(x, y). r⇧*⇧* x y}"
    unfolding rtranclp_def rtrancl_def by auto  
qed

lemma list_all2_rtrancl_2:
  "reflp r ⟹ list_all2 r⇧*⇧* xs ys ⟹ (list_all2 r)⇧*⇧* xs ys"
proof transfer
  fix r :: "'a ⇒ 'a ⇒ bool" and xs :: "'a list" and ys :: "'a list"
  assume as_reflp: "reflp r" and p_in_lr: "(xs, ys) ∈ listrel {(x, y). r⇧*⇧* x y}"
  from as_reflp have refl: "refl {(x, y). r x y}" 
    using reflp_refl_eq by fastforce
  from p_in_lr have "(xs, ys) ∈ listrel ({(x, y). r x y}⇧*)"
    unfolding rtranclp_def rtrancl_def by auto
  with refl have "(xs, ys) ∈ (listrel {(x, y). r x y})⇧*"
    using listrel_tclosure_2 by auto
  then show "(λxs ys. (xs, ys) ∈ listrel {(x, y). r x y})⇧*⇧* xs ys" 
    unfolding rtranclp_def rtrancl_def by auto
qed

A direct proof for list_all2 is also provided (legacy):

  1. list_all2_induct is applied to the lists; the base case is trivial. Thence, it remains to show that (L P)* x#xs y#ys if (L (P*)) xs ys, (L P)* xs ys and P* x y.
  2. The idea is that it is possible to find zs (e.g. xs) such that (L P) xs zs and (L P)+ zs ys.
  3. Then, given that P* x y and P x x, by induction based on the transitive properties of P*, (L P) x#xs y#zs. Therefore, also, (L P)* x#xs y#zs.
  4. Also, given that (L P)+ zs ys and P y y, by induction, (L P)+ y#zs y#ys. Thus, also, (L P)* y#zs y#ys.
  5. From 3 and 4 conclude (L P)* x#xs y#ys.

lemma list_all2_rtrancl2:
  assumes as_r: "(⋀x. P x x)" 
  shows "(list_all2 P⇧*⇧*) xs ys ⟹ (list_all2 P)⇧*⇧* xs ys"
proof(induction rule: list_all2_induct)
  case Nil then show ?case by simp
next
  case (Cons x xs y ys) show ?case
  proof -
    from as_r have lp_xs_xs: "list_all2 P xs xs" by (rule list_all2_refl)
    from Cons.hyps(1) have x_xs_y_zs: "(list_all2 P)⇧*⇧* (x#xs) (y#xs)"
    proof(induction rule: rtranclp_induct)
      case base then show ?case by simp
    next
      case (step y z) then show ?case 
      proof -
        have rt_step_2: "(list_all2 P)⇧*⇧* (y#xs) (z#xs)" 
          by (rule r_into_rtranclp, rule list_all2_Cons[THEN iffD2]) 
            (simp add: step.hyps(2) lp_xs_xs)
        from step.IH rt_step_2 show ?thesis by (rule rtranclp_trans) 
      qed      
    qed
    from Cons.IH have "(list_all2 P)⇧*⇧* (y#xs) (y#ys)"
    proof(induction rule: rtranclp_induct)
      case base then show ?case by simp
    next
      case (step ya za) show ?case
      proof -
        have rt_step_2: "(list_all2 P)⇧*⇧* (y#ya) (y#za)" 
          by (rule r_into_rtranclp, rule list_all2_Cons[THEN iffD2])     
            (simp add: step.hyps(2) as_r)
        from step.IH rt_step_2 show ?thesis by (rule rtranclp_trans)
      qed
    qed
    with x_xs_y_zs show ?thesis by simp
  qed
qed

As a side note, in my view (I know very little about nitpick), nitpick should not provide invalid counterexamples without any warning. I believe, usually, when nitpick 'suspects' that a counterexample may be invalid it notifies the user that the example is 'potentially spurious'. It may be useful to submit a bug report if this issue has not been recorded elsewhere.


Isabelle version: Isabelle2020

  • Thanks a lot for help! I've sent a bug report on nitpick to the mailing list. – Denis Oct 25 '18 at 06:16
  • Thanks! There are a lot of new things for me in the proof. I don't understand inductive sets and the transfer package. It's a good reason to study them. – Denis Dec 24 '18 at 06:24
  • @Denis I am surprised that you mentioned that you do not understand inductive sets. From your other questions, it seems that you understand inductively defined predicates. There is an apparent correspondence between predicates and sets in HOL: `"P x ⟷ x ∈ {x. P x}"`. Of course, there are many books/lecture notes that provide theoretical background for the inductive sets, e.g. see this [link](https://www.cl.cam.ac.uk/teaching/1213/DiscMathII/Ch5.pdf). – user9716869 - supports Ukraine Dec 24 '18 at 11:32
  • @Denis A good example that demonstrates the correspondence between the inductively defined sets and inductively defined predicates in HOL can be found in section 7.1 in the book “A Proof Assistant for Higher-Order Logic” (PA) and section 4.5.1 in the book “Concrete Semantics with Isabelle/HOL” (CS). PA presents an example of the construction of the set of even numbers using “inductive_set” and CS presents an example of the construction of a predicate that determines whether or not a natural number is even using “inductive”. – user9716869 - supports Ukraine Dec 24 '18 at 11:32