2

I'm trying to prove that a transitive relation on elements of finite maps is equivalent to a transitive relation on finite maps itself.

Here is a helper lemma, which shows that relations on finite maps are transitive if relations on their elements are transitive:

lemma fmrel_trans:
  "(⋀x y z. x ∈ fmran' xm ⟹ P x y ⟹ Q y z ⟹ R x z) ⟹
   fmrel P xm ym ⟹ fmrel Q ym zm ⟹ fmrel R xm zm"
  unfolding fmrel_iff
  by (metis fmdomE fmdom_notD fmran'I option.rel_inject(2) option.rel_sel)

Here is a first lemma, which I successfully proved:

lemma trancl_to_fmrel:
  "(fmrel f)⇧+⇧+ xm ym ⟹ fmrel f⇧+⇧+ xm ym"
  apply (induct rule: tranclp_induct)
  apply (simp add: fmap.rel_mono_strong)
  apply (rule_tac ?P="f⇧+⇧+" and ?Q="f" and ?ym="y" in fmrel_trans; auto)
  done

And here is a symmetric lemma, which I can't prove:

lemma fmrel_to_trancl:
  "fmrel r⇧+⇧+ xm ym ⟹
   (⋀x. r x x) ⟹
   (fmrel r)⇧+⇧+ xm ym"

Equivalently this lemma can be stated as

lemma fmrel_tranclp_induct:
  "fmrel r⇧+⇧+ a b ⟹
   (⋀x. r x x) ⟹
   (⋀y. fmrel r a y ⟹ P y) ⟹
   (⋀y z. fmrel r⇧+⇧+ a y ⟹ fmrel r y z ⟹ P y ⟹ P z) ⟹ P b"

or

lemma fmrel_tranclp_trans_induct:
  "fmrel r⇧+⇧+ a b ⟹
   (⋀x. r x x) ⟹
   (⋀x y. fmrel r x y ⟹ P x y) ⟹
   (⋀x y z. fmrel r⇧+⇧+ x y ⟹ P x y ⟹ fmrel r⇧+⇧+ y z ⟹ P y z ⟹ P x z) ⟹ P a b"

Proving any of these 3 lemmas I can prove the rest.

The question is very similar to How to lift a transitive relation from elements to lists? But the proof in that question is based on the induction rule list_all2_induct. I can't find a similar rule for fmrel. I tried to prove something like this, but with no success:

lemma fmrel_induct
  [consumes 1, case_names Nil Cons, induct set: fmrel]:
  assumes P: "fmrel P xs ys"
  assumes Nil: "R fmempty fmempty"
  assumes Cons: "⋀k x xs y ys.
    ⟦P x y; fmrel P xs ys; fmlookup xs k = None; fmlookup ys k = None; R xs ys⟧ ⟹
    R (fmupd k x xs) (fmupd k y ys)"
  shows "R xs ys"

I also tried to replace fmrel by list_all2 in the lemmas, but with no success:

lemma fmrel_to_list_all2:
  "fmrel f xm ym ⟹
   xs = map snd (sorted_list_of_fmap xm) ⟹
   ys = map snd (sorted_list_of_fmap ym) ⟹
   list_all2 f xs ys"

The idea is that keys (domains) of xm and ym are equal. And fmrel is equivalent to list_all2 on sorted values (ranges) of the maps.

Could you help me to prove fmrel_to_trancl?

Denis
  • 1,167
  • 1
  • 10
  • 30
  • Dear Denis, I would like to add these theorems to the Finite_Map theory. Please let me know how you'd like to be attributed. – larsrh Dec 11 '18 at 16:53

1 Answers1

2

There exist several methodologies that can help you to achieve your goal:

  1. Perform lifting gradually but in a semi-automated manner using the functionality of the transfer package of Isabelle.
  2. Prove the theorems directly as they are stated for finite maps.
  3. As noted by larsrh, the theorems that you are looking to prove hold for general BNFs, not only for lists/finite maps. Therefore, in general, it may be sensible to think about whether it would be worth to augment the BNF infrastructure with similar theorems.

Below I present a skeleton of a solution using the first methodology and a complete solution using the second methodology. However, it goes without saying that the code presented below is not 'production ready'.


The code listing below demonstrates the first methodology. The theorems of interest are first transferred from list to alist and only then from alist to fmap. I do not provide a complete solution, but it should be relatively easy to infer it using the code listing below as a skeleton of a solution (if necessary, I can provide the missing details upon request).

First, the relevant theorems are proven for alist:

subsection ‹Further results about lists›

abbreviation "dmf ≡ distinct ∘ map fst"

lemma double_length_induct:
  "(⋀xs ys. ∀xs' ys'. 
  length xs' < length xs ⟶ length ys' < length ys ⟶ P xs' ys' ⟹ 
  P xs ys) ⟹ 
  P xs ys"
  sorry

lemma list_all2_sym: "list_all2 P xs ys ⟷ list_all2 (λy x. P x y) ys xs"
  sorry


subsection ‹Extension of the theory @{text Multiset}›

lemma list_all2_reorder_right_invariance:
  assumes rel: "list_all2 R xs ys" and ms_y: "mset ys' = mset ys"
  shows "∃xs'. list_all2 R xs' ys' ∧ mset xs' = mset xs"
  sorry


subsection ‹Further results about permutations›

lemma perm_map_of:
  assumes "dmf xs" and "dmf ys"
  shows "xs <~~> ys ⟹ map_of xs k = map_of ys k"
  sorry

lemma perm_map_of_none: 
  "xs <~~> ys ⟹ map_of xs k = None ⟷ map_of ys k = None"
  sorry

lemma map_of_eq_mset:
  assumes "dmf xs" and "dmf ys"
  shows "map_of xs = map_of ys ⟷ mset xs = mset ys"
  sorry

lemma map_of_eq_perm:
  assumes "dmf xs" and "dmf ys"
  shows "map_of xs = map_of ys ⟷ xs <~~> ys"
  sorry

definition perm_list_all2 :: 
  "('a ⇒ 'b ⇒ bool) ⇒ 'a list ⇒ 'b list ⇒ bool" where
  "perm_list_all2 P xs ys = (∃ys'. ys <~~> ys' ∧ list_all2 P xs ys')"

lemma perm_list_all2_def_alt: 
  "perm_list_all2 P xs ys = (∃xs'. xs <~~> xs' ∧ list_all2 P xs' ys)"  
  unfolding perm_list_all2_def
  sorry

lemma perm_list_all2_sym: 
  "perm_list_all2 P xs ys = perm_list_all2 (λy x. P x y) ys xs"
  sorry

lemma list_all2_to_perm_list_all2_2: 
  "list_all2 P xs ys ⟹ 
  xs <~~> xs' ⟹ 
  ys <~~> ys'⟹ 
  perm_list_all2 P xs' ys'" 
  sorry

lemma perm_list_all2_to_perm_list_all2: 
  "perm_list_all2 P xs ys ⟹ 
  xs <~~> xs' ⟹ 
  ys <~~> ys'⟹ 
  perm_list_all2 P xs' ys'" 
  sorry

lemma perm_list_all2_lengthD:
  "perm_list_all2 P xs ys ⟹ length xs = length ys"
  sorry

lemma perm_list_all2_Nil[iff, code]: "perm_list_all2 P [] ys = (ys = [])"
  sorry

lemma perm_list_all2_Cons:
  "P x y ⟹ perm_list_all2 P xs ys ⟹ perm_list_all2 P (x # xs) (y # ys) "
  sorry


subsection ‹Permutation of AList›

lemma update_new_imp_append:
  "map_of xs k = None ⟹ AList.update k v xs = xs @ [(k, v)]"
  sorry

lemma map_of_distinct_to_none: 
  "dmf xs ⟹ xs = ys @ [(ky, vy)] @ ys' ⟹ map_of ys ky = None"
  sorry

abbreviation "pred_snd P ≡ (λx y. fst x = fst y ∧ P (snd x) (snd y))" 

definition listvalrel :: 
  "('aval ⇒ 'bval ⇒ bool) ⇒ 
  ('key × 'aval) list ⇒ 
  ('key × 'bval) list ⇒ 
  bool" 
  where
  "listvalrel P xs ys = list_all2 (pred_snd P) xs ys"

lemma listvalrel_map_fst: "listvalrel P xs ys ⟹ map fst xs = map fst ys"
  unfolding listvalrel_def
  sorry


subsection ‹Permutation of AList›

lemma dmf_perm_imp_dmf: "dmf xs ⟹ xs <~~> xs' ⟹ dmf xs'"
  sorry

lemma perm_update:
  assumes "dmf xs" and "dmf ys"
  shows "xs <~~> ys ⟹ AList.update k v xs <~~> AList.update k v ys"
  sorry

definition perm_listvalrel :: 
"('aval ⇒ 'bval ⇒ bool) ⇒ 
  ('key × 'aval) list ⇒ 
  ('key × 'bval) list ⇒ 
  bool"
  where  
  "perm_listvalrel P xs ys = perm_list_all2 (pred_snd P) xs ys"

lemma perm_listvalrel_def_alt: 
  "perm_listvalrel P xs ys = (∃ys'. ys <~~> ys' ∧ listvalrel P xs ys')"  
  unfolding perm_listvalrel_def listvalrel_def by (simp add: perm_list_all2_def)

lemma perm_listvalrel_to_perm_listvalrel: 
  "perm_listvalrel P xs ys ⟹ 
  xs <~~> xs' ⟹ 
  ys <~~> ys'⟹ 
  perm_listvalrel P xs' ys'" 
  sorry

lemma perm_listvalrel_lengthD[intro?]:
  "perm_listvalrel P xs ys ⟹ length xs = length ys"
  sorry

lemma perm_listvalrel_Nil[iff, code]: 
  "perm_listvalrel P [] ys = (ys = [])"
  sorry

lemma perm_listvalrel_Cons:
  "pred_snd P x y ⟹ 
  perm_listvalrel P xs ys ⟹ 
  perm_listvalrel P (x # xs) (y # ys) "
  sorry

lemma map_of_remove1:
  "dmf xs ⟹ map_of xs k = Some v ⟹ map_of (remove1 (k, v) xs) k = None"
  sorry

lemma map_of_Cons:
  "k ≠ k' ⟹ map_of (xss @ xse) k = map_of (xss @ (k', v') # xse) k"
  sorry

lemma ro_imp_ro_rm1:
  assumes dmf_xs: "dmf (xss @ (k', v') # xse)" 
    and dmf_ys: "dmf (yss @ (k', w') # yse)" 
    and ro: 
      "⋀k. rel_option P 
      (map_of (xss @ (k', v') # xse) k) (map_of (yss @ (k', w') # yse) k)"
  shows "⋀k. rel_option P (map_of (xss @ xse) k) (map_of (yss @ yse) k)"
proof -
  fix k
  show "rel_option P (map_of (xss @ xse) k) (map_of (yss @ yse) k)"
  proof(cases "k ≠ k'")
    case True show ?thesis sorry
  next
    case False show ?thesis
    proof -
      from dmf_xs dmf_ys have 
        "distinct (xss @ (k', v') # xse)" and "distinct (yss @ (k', w') # yse)"
        by (metis comp_apply distinct_map)+
      then have 
        xss_xse: "xss @ xse = remove1 (k', v') (xss @ (k', v') # xse)" and
        yss_yse: "yss @ yse = remove1 (k', w') (yss @ (k', w') # yse)"
        by (simp add: remove1_append)+
      have 
        k'v'_in_set: "(k', v') ∈ List.set (xss @ (k', v') # xse)" and 
        k'w'_in_set: "(k', w') ∈ List.set (yss @ (k', w') # yse)" 
        by auto
      have 
        mo_v': "map_of (xss @ (k', v') # xse) k' = Some v'" and  
        mo_w': "map_of (yss @ (k', w') # yse) k' = Some w'"
        subgoal
          apply(rule map_of_eq_Some_iff[THEN iffD2]) using dmf_xs by auto
          apply(rule map_of_eq_Some_iff[THEN iffD2]) using dmf_ys by auto          
      have 
        xss_xse_rm: "map_of (remove1 (k', v') (xss @ (k', v') # xse)) k' = None" 
        and 
        yss_yse_rm: "map_of (remove1 (k', w') (yss @ (k', w') # yse)) k' = None"
        subgoal
          apply(rule map_of_remove1) using dmf_xs mo_v' by auto
          apply(rule map_of_remove1) using dmf_ys mo_w' by auto
      have "map_of (xss@xse) k = None" and "map_of (yss@yse) k = None"
        subgoal
          using False xss_xse_rm xss_xse by simp
          using False yss_yse_rm yss_yse by simp
      thus "rel_option P (map_of (xss@xse) k) (map_of (yss@yse) k)"
        by (metis rel_option_None1)  
    qed
  qed
qed

lemma perm_listvalrel_eq_rel_option:
  assumes dmf_xs: "dmf xs" and dmf_ys: "dmf ys"  
  shows 
    "perm_listvalrel P xs ys ⟷ (∀k. rel_option P (map_of xs k) (map_of ys k))"
proof
  assume "perm_listvalrel P xs ys"  
  then obtain ys' where 
    ys_ys': "ys <~~> ys'" and la2_psP: "list_all2 (pred_snd P) xs ys'"
    unfolding perm_listvalrel_def_alt listvalrel_def by clarsimp
  from dmf_ys ys_ys' have dys': "dmf ys'" by (rule dmf_perm_imp_dmf)
  from dmf_ys dys' ys_ys' have "map_of ys k = map_of ys' k" for k 
    by (rule perm_map_of)
  moreover have "rel_option P (map_of xs k) (map_of ys' k)" for k
  proof(cases "map_of xs k = None")
    case True show ?thesis sorry
  next
    case False show ?thesis
    proof -
      from False obtain v where v: "map_of xs k = Some v" by clarsimp
      then obtain n where n: "n < length xs ∧ xs!n = (k, v)" 
        by (meson in_set_conv_nth map_of_SomeD)
      with la2_psP have n_l_ys': "n < length ys'" 
        using list_all2_lengthD by force
      with la2_psP have psP: "pred_snd P (xs!n) (ys'!n)" 
        by (simp add: list_all2_conv_all_nth)
      with n have "fst (ys'!n) = k" by simp
      with dys' ys_ys' n_l_ys' have "map_of ys' k = Some (snd (ys'!n))" by auto
      moreover with psP n have "P v (snd (ys'!n))" by simp
      ultimately show "rel_option P (map_of xs k) (map_of ys' k)" 
        using v by simp
    qed
  qed
  ultimately show "∀k. rel_option P (map_of xs k) (map_of ys k)" by simp
next
  assume "(∀k. rel_option P (map_of xs k) (map_of ys k))"
  with dmf_xs dmf_ys show "perm_listvalrel P xs ys"
  proof(induction xs ys rule: double_length_induct)
    case (1 xs ys) show ?case
    proof(cases "xs = []")
      case True show ?thesis sorry
    next
      case False show ?thesis
      proof -
        from False obtain xse x where x_xse: "xs = x # xse" 
          by (meson remdups_adj.cases)
        then obtain k v where k_v: "x = (k, v)" by fastforce
        with x_xse have xs_split: "xs = [] @ (k, v) # xse" by simp 
        from k_v dmf_xs x_xse have v: "map_of xs k = Some v" by simp
        then have "map_of ys k ≠ None" 
          by (metis option.rel_distinct(2) "1.prems"(3))
        then obtain w where w: "map_of ys k = Some w" by clarsimp
        then have "(k, w) ∈ List.set ys" by (auto dest: map_of_SomeD)
        then obtain yss yse where ys_split: "ys = yss @ (k, w) # yse"
          using split_list by fastforce
        from xs_split ys_split "1.prems"(3) have ro_split:
          "∀k'. 
          rel_option P 
          (map_of ([] @ (k, v)#xse) k') (map_of (yss @ (k, w) # yse) k')"
          by simp
        have 
          d_xs_split: "dmf ([] @ (k, v) # xse)" and
          d_ys_split: "dmf (yss @ (k, w) # yse)"
          subgoal          
            using xs_split "1.prems"(1) by (rule subst)
            using ys_split "1.prems"(2) by (rule subst)
        then have d_rm_xs: "dmf ([] @ xse)" and d_rm_ys: "dmf (yss @ yse)" 
          by simp+
        from d_xs_split d_ys_split ro_split[rule_format] have ro_rm:
          "rel_option P (map_of ([] @ xse) k) (map_of (yss @ yse) k)" for k
          by (rule ro_imp_ro_rm1)
        with x_xse have l_rm_xs: "length ([] @ xse) < length xs" by simp
        with ys_split have l_rm_ys: "length (yss @ yse) < length ys" by simp
        with ro_split v w have "pred_snd P (k, v) (k, w)"
          by (metis "1.prems"(3) fst_conv option.rel_inject(2) snd_conv)
        moreover from l_rm_xs l_rm_ys d_rm_xs d_rm_ys ro_rm have 
          "perm_listvalrel P ([] @ xse) (yss @ yse)"
          by (rule "1.IH"[rule_format])
        ultimately have 
          "perm_listvalrel P ((k, v) # [] @ xse) ((k, w) # yss @ yse)"
          by (rule perm_listvalrel_Cons)
        with k_v x_xse have "perm_listvalrel P xs ((k, w) # yss @ yse)" by simp
        moreover have "((k, w) # yss @ yse) <~~> ys"
          unfolding ys_split by (rule perm_append_Cons)
        ultimately show "perm_listvalrel P xs ys"
          using perm_listvalrel_to_perm_listvalrel by blast
      qed
    qed
  qed
qed


subsection ‹Further results of DAList›

lift_definition alist_all2 :: 
  "(('akey × 'a) ⇒ ('bkey × 'b) ⇒ bool) ⇒ 
  ('akey, 'a) alist ⇒ 
  ('bkey, 'b) alist ⇒ 
  bool" 
  is List.list_all2 .

abbreviation alength :: "('akey, 'aval) alist ⇒ nat" where
  "alength ≡ size"

lemma alength_transfer[transfer_rule]: 
  includes lifting_syntax 
  shows "((pcr_alist (=) (=)) ===> (=)) List.length alength"
  sorry

lemma neq_empty_conv:
  "xs ≠ DAList.empty = 
  (∃ky vy ys. 
  xs = DAList.update ky vy ys ∧ 
  alength ys < alength xs ∧ 
  DAList.lookup ys ky = None)"
  sorry

lemma alength_induct:
  "(⋀xs::('key, 'val) alist. 
  ∀ys::('key, 'val) alist. alength ys < alength xs ⟶ P ys ⟹ P xs) ⟹ 
  P xs"
  sorry

lemma aupdate_induct_3[case_names Nil update, induct type: alist]:
  "P DAList.empty ⟹ 
  (⋀kx vx xs. 
  P xs ⟹ DAList.lookup xs kx = None ⟹ P (DAList.update kx vx xs)
  ) ⟹ 
  P xs"
  sorry

lemma alist_all2_update[iff]:
  assumes "DAList.lookup xs kx = None" and "DAList.lookup ys ky = None"
  shows 
    "alist_all2 P (DAList.update kx vx xs) (DAList.update ky vy ys) = 
    (P (kx, vx) (ky, vy) ∧ alist_all2 P xs ys)"
  sorry

lemma alist_all2_update1:
  assumes "DAList.lookup xs kx = None" 
  shows 
    "alist_all2 P (DAList.update kx vx xs) ys = 
    (∃kz vz zs.
    DAList.lookup zs kz = None ∧
    ys = (DAList.update kz vz zs) ∧ 
    P (kx, vx) (kz, vz) ∧ 
    alist_all2 P xs zs)"
proof(insert assms, transfer)
  fix xs :: "('b × 'a) list"
    and ys :: "('c × 'd) list"
    and kx P vx
  assume mo_xs_none: "map_of xs kx = None" and dmf_ys: "dmf ys" 
  from mo_xs_none have xs_append: "AList.update kx vx xs = xs @ [(kx, vx)]"
    by (simp add: update_new_imp_append)
  show 
    "list_all2 P (AList.update kx vx xs) ys = 
    (∃kz vz. ∃zs∈Collect dmf. 
    map_of zs kz = None ∧ 
    ys = AList.update kz vz zs ∧ 
    P (kx, vx) (kz, vz) ∧ 
    list_all2 P xs zs)"
  sorry
qed


subsection ‹Permutation of DAList›

lift_definition mset ::  "('key, 'val) alist ⇒ ('key × 'val) multiset" 
  is Multiset.mset .

lift_definition aperm :: 
  "('key, 'val) alist ⇒ ('key, 'val) alist ⇒ bool"  ("_ <~~>a _"  [50, 50] 50)
  is Permutation.perm .

lemma aperm_trans[intro]: "xs <~~>a ys ⟹ ys <~~>a zs ⟹ xs <~~>a zs"
  by transfer auto

lemma aperm_refl[iff]: "l <~~>a l"
  by transfer simp

lemma aperm_sym: "xs <~~>a ys ⟹ ys <~~>a xs"
  by transfer (rule perm_sym) 

lemma aperm_aperm_snd: "x <~~>a y ⟹ aperm x = aperm y"
  sorry

lemma aperm_update[intro!]: 
  "xs <~~>a ys ⟹ DAList.update k v xs <~~>a DAList.update k v ys"
  sorry

lemma aperm_imp_lookup_none:
  "xs <~~>a ys ⟹ DAList.lookup xs k = None ⟷ DAList.lookup ys k = None"
  sorry

lift_definition perm_alist_all2 :: 
  "(('akey × 'a) ⇒ ('bkey × 'b) ⇒ bool) ⇒ 
  ('akey, 'a) alist ⇒ 
  ('bkey, 'b) alist ⇒ 
  bool" 
  is perm_list_all2 .

lemma perm_alist_all2_def_alt:
  "perm_alist_all2 P xs ys = (∃ys'. ys <~~>a ys' ∧ alist_all2 P xs ys')"
  sorry

lemma perm_alist_all2_sym: 
  "perm_alist_all2 P xs ys = perm_alist_all2 (λy x. P x y) ys xs"
  sorry

lemma alist_all2_to_perm_perm_alist_all2: 
  "alist_all2 P xs ys ⟹ 
  xs <~~>a xs' ⟹ 
  ys <~~>a ys'⟹ 
  perm_alist_all2 P xs' ys'"
  sorry

lemma perm_alist_all2_Nil: 
  "perm_alist_all2 P DAList.empty ys = (ys = DAList.empty)"
  sorry

lemma perm_alist_all2_update:
  assumes "DAList.lookup xs kx = None"  and "DAList.lookup ys ky = None"
  shows 
    "P (kx, vx) (ky, vy) ⟹ 
    perm_alist_all2 P xs ys ⟹ 
    perm_alist_all2 P (DAList.update kx vx xs) (DAList.update ky vy ys)"
  sorry

lemma perm_alist_all2_update1:
  assumes "DAList.lookup xs kx = None" 
  shows 
    "perm_alist_all2 P (DAList.update kx vx xs) ys = 
    (∃kz vz zs. 
    DAList.lookup zs kz = None ∧ 
    ys <~~>a (DAList.update kz vz zs) ∧ 
    P (kx, vx) (kz, vz) ∧ 
    perm_alist_all2 P xs zs)"
  sorry

lemma perm_alist_all2_update2:
  assumes "DAList.lookup ys ky = None" 
  shows 
    "perm_alist_all2 P xs (DAList.update ky vy ys) = 
    (∃kz vz zs. 
    DAList.lookup zs kz = None ∧ 
    xs <~~>a (DAList.update kz vz zs) ∧ 
    P (kz, vz) (ky, vy) ∧ 
    perm_alist_all2 P zs ys)"
  sorry

lemma perm_alist_all2_induct[consumes 1, case_names Nil update]:
  assumes P: "perm_alist_all2 P xs ys"
    and Nil: "R DAList.empty DAList.empty"
  assumes update:
    "⋀kx vx xs ky vy ys ys'.
    ⟦
    DAList.lookup xs kx = None;
    DAList.lookup ys ky = None;
    P (kx, vx) (ky, vy);
    perm_alist_all2 P xs ys;
    R xs ys;
    (DAList.update ky vy ys) <~~>a ys'
    ⟧ ⟹ 
    R (DAList.update kx vx xs) ys'"
  shows "R xs ys"
  using P
  sorry

lift_definition perm_alistvalrel :: 
  "('aval ⇒ 'bval ⇒ bool) ⇒ 
  ('key, 'aval) alist ⇒ 
  ('key, 'bval) alist ⇒ 
  bool" 
  is perm_listvalrel .

lemma perm_alistvalrel_def_alt: 
  "perm_alistvalrel P xs ys = perm_alist_all2 (pred_snd P) xs ys"
  sorry

lemma perm_alistvalrel_update:
  assumes "DAList.lookup xs kx = None" and "DAList.lookup ys ky = None"
  shows 
    "pred_snd P (kx, vx) (ky, vy) ⟹ 
    perm_alistvalrel P xs ys ⟹ 
    perm_alistvalrel P (DAList.update kx vx xs) (DAList.update ky vy ys)"
  sorry

lemma perm_alistvalrel_update1:
  assumes "DAList.lookup xs k = None" 
  shows 
    "perm_alistvalrel P (DAList.update k vx xs) ys = 
    (∃vz zs. 
    DAList.lookup zs k = None ∧ 
    ys <~~>a (DAList.update k vz zs) ∧ 
    P vx vz ∧ 
    perm_alistvalrel P xs zs)"
  sorry

lemma perm_alistvalrel_update2:
  assumes "DAList.lookup ys k = None" 
  shows 
    "perm_alistvalrel P xs (DAList.update k vy ys) = 
    (∃vz zs. 
    DAList.lookup zs k = None ∧ 
    xs <~~>a (DAList.update k vz zs) ∧ 
    P vz vy ∧ 
    perm_alistvalrel P zs ys)"
  sorry

lemma perm_alistvalrel_induct[consumes 1, case_names Nil update]:
  assumes P: "perm_alistvalrel P xs ys"
    and Nil: "R DAList.empty DAList.empty"
  assumes update:
    "⋀k vx xs vy ys ys'.
    ⟦ 
    P vx vy; 
    perm_alistvalrel P xs ys;
    R xs ys;
    (DAList.update k vy ys) <~~>a ys'
    ⟧ ⟹ 
    R (DAList.update k vx xs) ys'"
  shows "R xs ys"
  using P
proof -
  from update have update':
    "⋀kx vx xs ky vy ys ys'.
    ⟦
    DAList.lookup xs kx = None;
    DAList.lookup ys ky = None;
    pred_snd P (kx, vx) (ky, vy);
    perm_alistvalrel P xs ys;
    R xs ys;
    (DAList.update ky vy ys) <~~>a ys'
    ⟧ ⟹ 
    R (DAList.update kx vx xs) ys'"
    by auto
  then show "R xs ys"
    apply(insert assms update')
    unfolding perm_alistvalrel_def_alt by (rule perm_alist_all2_induct)
qed

Then the theorems can be transferred to fmap:

lemma perm_eq_fmap_of_list:
  assumes "dmf xs" and "dmf ys" 
  shows "xs <~~> ys ⟷ fmap_of_list xs = fmap_of_list ys"
  sorry

lemma exists_distinct_fst_fmap_of_list: 
  "∃xa. dmf xa ∧ fmap_of_list xa = m"
  sorry

lift_definition fmap_of_alist_impl :: "('a, 'b) alist ⇒ ('a, 'b) fmap"
  is fmap_of_list .

lemma perm_eq_fmap_of_alist: 
  "xs <~~>a ys ⟷ fmap_of_alist_impl xs = fmap_of_alist_impl ys"
  sorry

lemma exists_distinct_fst_fmap_of_alist: "∃xs. fmap_of_alist_impl xs = m"
  sorry

definition fmap_of_alists :: "('a, 'b) alist set ⇒ ('a, 'b) fmap" where
  "fmap_of_alists X = fmap_of_alist_impl (SOME x. x ∈ X)"

definition alists_of_fmap :: "('a, 'b) fmap ⇒ ('a, 'b) alist set" where
  "alists_of_fmap y = {x. fmap_of_alist_impl x = y}"

interpretation fmap : quot_type aperm fmap_of_alists alists_of_fmap
  sorry

abbreviation "fmap_of_alist ≡ fmap.abs"
abbreviation "alist_of_fmap ≡ fmap.rep"

definition cr_fmaplist where "cr_fmaplist = (λx. (=) (fmap_of_alist x))"

lemma fmap_of_alist_def_impl: "fmap_of_alist y = fmap_of_alist_impl y"
  sorry

lemma alist_of_fmap_def_alt: "alist_of_fmap y = (SOME x. fmap_of_alist x = y)"
  sorry

lemma fmap_alist_rep_abs: "fmap_of_alist (alist_of_fmap a) = a"
  sorry

lemma Quotient_fmaplist: 
  "Quotient aperm fmap_of_alist alist_of_fmap cr_fmaplist"
  sorry

locale fmap_fmaplist 
begin

lemma reflp_aperm: "reflp aperm" by (simp add: reflpI)

setup_lifting Quotient_fmaplist reflp_aperm

lemma fmap_of_alist_imp_eq_lookup:
  "m = fmap_of_alist ml ⟹ fmlookup m k = DAList.lookup ml k"
  sorry

lemma cr_fmaplist_imp_eq_lookup: 
  "cr_fmaplist ml m ⟹ fmlookup m k = DAList.lookup ml k"
  sorry

lemma eq_lookup_imp_cr_fmaplist: 
  "∀k. fmlookup m k = DAList.lookup ml k ⟹ cr_fmaplist ml m"
  sorry

context includes lifting_syntax
begin

lemma lookup_fmlookup_transfer[transfer_rule]: 
  "(cr_fmaplist ===> (=) ===> (=)) DAList.lookup fmlookup"
  unfolding rel_fun_def cr_fmaplist_def fmap_of_alist_def_impl
  by (transfer, simp add: fmlookup_of_list)

lemma fmempty_transfer[transfer_rule]: "cr_fmaplist DAList.empty fmempty"
  unfolding rel_fun_def cr_fmaplist_def fmap_of_alist_def_impl
  by (simp add: empty.rep_eq fmap_of_alist_impl.rep_eq)

lemma fmrel_transfer[transfer_rule]:
  "((=) ===> cr_fmaplist ===> cr_fmaplist ===> (=)) perm_alistvalrel fmrel"
  unfolding rel_fun_def cr_fmaplist_def fmap_of_alist_def_impl
  by transfer 
    (auto simp: fmlookup_of_list fmrel_iff perm_listvalrel_eq_rel_option)+

lemma fmupd_transfer[transfer_rule]:
  "((=) ===> (=) ===> cr_fmaplist ===> cr_fmaplist) DAList.update fmupd"
  unfolding rel_fun_def 
  using cr_fmaplist_imp_eq_lookup 
  by (fastforce simp: eq_lookup_imp_cr_fmaplist)

end

lifting_update fmap.lifting
lifting_forget fmap.lifting

end

context 
begin

interpretation fmap_fmaplist .

lemma fmrel_update1:
  assumes "fmlookup xs k = None" 
  shows 
    "fmrel P (fmupd k vx xs) ys = 
    (∃vz zs. 
    fmlookup zs k = None ∧ 
    ys = (fmupd k vz zs) ∧ 
    P vx vz ∧ 
    fmrel P xs zs)"
  including fmap.lifting 
  by (insert assms, transfer) (simp add: perm_alistvalrel_update1)

lemma fmrel_update2:
  assumes "fmlookup ys k = None" 
  shows 
    "fmrel P xs (fmupd k vy ys) = 
    (∃vz zs. 
    fmlookup zs k = None ∧ 
    xs = (fmupd k vz zs) ∧ 
    P vz vy ∧ 
    fmrel P zs ys)"
  including fmap.lifting
  by (insert assms, transfer) (simp add: perm_alistvalrel_update2)

lemma fmrel_induct[consumes 1, case_names Nil update]:
  assumes  P: "fmrel P xs ys"
    and Nil: "R fmempty fmempty"
    and update:
    "⋀k vx xs vy ys.
    ⟦P vx vy; fmrel P xs ys; R xs ys⟧ ⟹ R (fmupd k vx xs) (fmupd k vy ys)"
  shows "R xs ys"
  proof -
    from update have update':
      "⋀k vx xs vy ys ys'. 
      ⟦P vx vy; fmrel P xs ys; R xs ys; (fmupd k vy ys) = ys'⟧ ⟹ 
      R (fmupd k vx xs) ys'"
      by auto
    show "R xs ys" sorry
  qed                    

end

lemma fmrel_to_rtrancl:
  assumes as_r: "(⋀x. r x x)" and rel_rpp_xm_ym: "(fmrel r⇧*⇧*) xm ym" 
  shows "(fmrel r)⇧*⇧* xm ym"
  sorry

The lemma

lemma fmrel_to_rtrancl:
  assumes as_r: "(⋀x. r x x)" and rel_rpp_xm_ym: "(fmrel r⇧*⇧*) xm ym" 
  shows "(fmrel r)⇧*⇧* xm ym"
  sorry

can be transferred following the same methodology.


The code listing below presents the answer using the second methodology.

theory so_htlartfm
imports 
  Complex_Main
  "HOL-Library.Finite_Map"
begin


lemma fmap_eqdom_Cons1:
  assumes as_1: "fmlookup xm i = None"
    and as_2: "fmrel R (fmupd i x xm) ym" 
  shows 
    "(∃z zm. 
    fmlookup zm i = None ∧ ym = (fmupd i z zm) ∧ R x z ∧ fmrel R xm zm)"
proof - 
  from as_2 have eq_dom: "fmdom (fmupd i x xm) = fmdom ym" 
    using fmrel_fmdom_eq by blast
  from as_1 eq_dom as_2 obtain y where y: "fmlookup ym i = Some y"
    by force
  obtain z zm where z_zm: "ym = (fmupd i z zm) ∧ fmlookup zm i = None"
    using y by (smt fmap_ext fmlookup_drop fmupd_lookup)
  {
    assume "¬R x z"
    with as_1 z_zm have "¬fmrel R (fmupd i x xm) ym"
      by (metis fmrel_iff fmupd_lookup option.simps(11))
  }
  with as_2 have c3: "R x z" by auto
  {
    assume "¬fmrel R xm zm"
    with as_1 have "¬fmrel R (fmupd i x xm) ym" 
      by (metis fmrel_iff fmupd_lookup option.rel_sel z_zm)
  }
  with as_2 have c4: "fmrel R xm zm" by auto
  from z_zm c3 c4 show ?thesis by auto
qed

lemma fmap_eqdom_induct [consumes 1, case_names nil step]:
  assumes R: "fmrel R xm ym"
    and nil: "P fmempty fmempty"
    and step: 
    "⋀x xm y ym i. ⟦R x y; fmrel R xm ym; P xm ym⟧ ⟹ 
    P (fmupd i x xm) (fmupd i y ym)"
  shows "P xm ym"
  using R 
proof(induct xm arbitrary: ym)
  case fmempty
  then show ?case
    by (metis fempty_iff fmdom_empty fmfilter_alt_defs(5) 
      fmfilter_false fmrel_fmdom_eq fmrestrict_fset_dom nil)
next
  case (fmupd i x xm) show ?case 
  proof -
    from fmupd.prems(1) obtain y where y: "fmlookup ym i = Some y"
      by (metis fmupd.prems(1) fmrel_cases fmupd_lookup option.discI)
    from fmupd.hyps(2) fmupd.prems(1) fmupd.prems(1) obtain z zm where 
      zm_i_none: "fmlookup zm i = None" and
      ym_eq_z_zm: "ym = (fmupd i z zm)" and 
      R_x_z: "R x z" and
      R_xm_zm: "fmrel R xm zm"
      using fmap_eqdom_Cons1 by metis
    with R_xm_zm fmupd.hyps(1) have P_xm_zm: "P xm zm" by blast
    from R_x_z R_xm_zm P_xm_zm have "P (fmupd i x xm) (fmupd i z zm)" 
      by (rule step)
    then show ?thesis by (simp add: ym_eq_z_zm)
  qed
qed

lemma fmrel_to_rtrancl:
  assumes as_r: "(⋀x. r x x)" 
    and rel_rpp_xm_ym: "(fmrel r⇧*⇧*) xm ym" 
  shows "(fmrel r)⇧*⇧* xm ym"
proof-
  from rel_rpp_xm_ym show "(fmrel r)⇧*⇧* xm ym"
  proof(induct rule: fmap_eqdom_induct)
    case nil then show ?case by auto
  next
    case (step x xm y ym i) show ?case
    proof -
      from as_r have lp_xs_xs: "fmrel r xm xm" by (simp add: fmap.rel_refl)
      from step.hyps(1) have x_xs_y_zs: 
        "(fmrel r)⇧*⇧* (fmupd i x xm) (fmupd i y xm)"
      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: "(fmrel r)⇧*⇧* (fmupd i y xm) (fmupd i z xm)" 
            by (rule r_into_rtranclp, simp add: fmrel_upd lp_xs_xs step.hyps(2))
          from step.IH rt_step_2 show ?thesis by (rule rtranclp_trans) 
        qed      
      qed
      from step.hyps(3) have "(fmrel r)⇧*⇧* (fmupd i y xm) (fmupd i y ym)"
      proof(induction rule: rtranclp_induct)
        case base then show ?case by simp
      next
        case (step ya za) show ?case
        proof -
          have rt_step_2: "(fmrel r)⇧*⇧* (fmupd i y ya) (fmupd i y za)" 
            by (rule r_into_rtranclp, simp add: as_r fmrel_upd step.hyps(2)) 
          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
qed

lemma fmrel_to_trancl:
  assumes as_r: "(⋀x. r x x)" 
    and rel_rpp_xm_ym: "(fmrel r⇧+⇧+) xm ym" 
  shows "(fmrel r)⇧+⇧+ xm ym" 
  by (metis as_r fmrel_to_rtrancl fmap.rel_mono_strong fmap.rel_refl 
      r_into_rtranclp reflclp_tranclp rel_rpp_xm_ym rtranclpD rtranclp_idemp 
      rtranclp_reflclp tranclp.r_into_trancl)

end
  • Thanks a lot for your help! I guess, it was the last blocking lemma in my theory. I'm trying to describe the type system of some programming language. After some code cleanups I will share it somewhere. – Denis Dec 03 '18 at 17:36
  • @Denis Sounds interesting. Indeed, if possible, please provide a link as a comment to one of my answers. – user9716869 - supports Ukraine Dec 03 '18 at 18:28
  • Dear @xanonec, I would like to add these theorems to the Finite_Map theory. Please let me know how you'd like to be attributed. – larsrh Dec 11 '18 at 16:53
  • @xanonec Here is a repository https://github.com/AresEkb/Simple_OCL These lemmas are required to prove that defined type system is a semilattice. I improved some lemmas. For example, I defined a property "a function is bijective on transitive closure" (see `Transitive_Closure_Ext`), it allowed me to generalize https://stackoverflow.com/questions/52565266/what-kind-of-functions-preserve-properties-of-closure – Denis Dec 12 '18 at 05:56
  • @Denis Thank you for your comment. Also, I received your email. Indeed, it may be best to continue our conversation via email, because these comments seem to be starting to deviate from the main topic of the question. – user9716869 - supports Ukraine Dec 12 '18 at 12:21