| author | haftmann | 
| Sun, 20 Jan 2019 17:15:49 +0000 | |
| changeset 69699 | 82f57315cade | 
| parent 64632 | 9df24b8b6c0a | 
| permissions | -rw-r--r-- | 
| 60519 | 1 | (* Author: Stefan Berghofer, Lukas Bulwahn, TU Muenchen *) | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 2 | |
| 60500 | 3 | section \<open>A table-based implementation of the reflexive transitive closure\<close> | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 4 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 5 | theory Transitive_Closure_Table | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 6 | imports Main | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 7 | begin | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 8 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 9 | inductive rtrancl_path :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 10 | for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 11 | where | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 12 | base: "rtrancl_path r x [] x" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 13 | | step: "r x y \<Longrightarrow> rtrancl_path r y ys z \<Longrightarrow> rtrancl_path r x (y # ys) z" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 14 | |
| 60519 | 15 | lemma rtranclp_eq_rtrancl_path: "r\<^sup>*\<^sup>* x y \<longleftrightarrow> (\<exists>xs. rtrancl_path r x xs y)" | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 16 | proof | 
| 60519 | 17 | show "\<exists>xs. rtrancl_path r x xs y" if "r\<^sup>*\<^sup>* x y" | 
| 18 | using that | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 19 | proof (induct rule: converse_rtranclp_induct) | 
| 34912 | 20 | case base | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 21 | have "rtrancl_path r y [] y" by (rule rtrancl_path.base) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 22 | then show ?case .. | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 23 | next | 
| 34912 | 24 | case (step x z) | 
| 60500 | 25 | from \<open>\<exists>xs. rtrancl_path r z xs y\<close> | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 26 | obtain xs where "rtrancl_path r z xs y" .. | 
| 60500 | 27 | with \<open>r x z\<close> have "rtrancl_path r x (z # xs) y" | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 28 | by (rule rtrancl_path.step) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 29 | then show ?case .. | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 30 | qed | 
| 60519 | 31 | show "r\<^sup>*\<^sup>* x y" if "\<exists>xs. rtrancl_path r x xs y" | 
| 32 | proof - | |
| 33 | from that obtain xs where "rtrancl_path r x xs y" .. | |
| 34 | then show ?thesis | |
| 35 | proof induct | |
| 36 | case (base x) | |
| 37 | show ?case | |
| 38 | by (rule rtranclp.rtrancl_refl) | |
| 39 | next | |
| 40 | case (step x y ys z) | |
| 41 | from \<open>r x y\<close> \<open>r\<^sup>*\<^sup>* y z\<close> show ?case | |
| 42 | by (rule converse_rtranclp_into_rtranclp) | |
| 43 | qed | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 44 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 45 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 46 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 47 | lemma rtrancl_path_trans: | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 48 | assumes xy: "rtrancl_path r x xs y" | 
| 60519 | 49 | and yz: "rtrancl_path r y ys z" | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 50 | shows "rtrancl_path r x (xs @ ys) z" using xy yz | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 51 | proof (induct arbitrary: z) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 52 | case (base x) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 53 | then show ?case by simp | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 54 | next | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 55 | case (step x y xs) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 56 | then have "rtrancl_path r y (xs @ ys) z" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 57 | by simp | 
| 60500 | 58 | with \<open>r x y\<close> have "rtrancl_path r x (y # (xs @ ys)) z" | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 59 | by (rule rtrancl_path.step) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 60 | then show ?case by simp | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 61 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 62 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 63 | lemma rtrancl_path_appendE: | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 64 | assumes xz: "rtrancl_path r x (xs @ y # ys) z" | 
| 60519 | 65 | obtains "rtrancl_path r x (xs @ [y]) y" and "rtrancl_path r y ys z" | 
| 66 | using xz | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 67 | proof (induct xs arbitrary: x) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 68 | case Nil | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 69 | then have "rtrancl_path r x (y # ys) z" by simp | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 70 | then obtain xy: "r x y" and yz: "rtrancl_path r y ys z" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 71 | by cases auto | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 72 | from xy have "rtrancl_path r x [y] y" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 73 | by (rule rtrancl_path.step [OF _ rtrancl_path.base]) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 74 | then have "rtrancl_path r x ([] @ [y]) y" by simp | 
| 60519 | 75 | then show thesis using yz by (rule Nil) | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 76 | next | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 77 | case (Cons a as) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 78 | then have "rtrancl_path r x (a # (as @ y # ys)) z" by simp | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 79 | then obtain xa: "r x a" and az: "rtrancl_path r a (as @ y # ys) z" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 80 | by cases auto | 
| 60519 | 81 | show thesis | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 82 | proof (rule Cons(1) [OF _ az]) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 83 | assume "rtrancl_path r y ys z" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 84 | assume "rtrancl_path r a (as @ [y]) y" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 85 | with xa have "rtrancl_path r x (a # (as @ [y])) y" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 86 | by (rule rtrancl_path.step) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 87 | then have "rtrancl_path r x ((a # as) @ [y]) y" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 88 | by simp | 
| 60519 | 89 | then show thesis using \<open>rtrancl_path r y ys z\<close> | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 90 | by (rule Cons(2)) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 91 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 92 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 93 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 94 | lemma rtrancl_path_distinct: | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 95 | assumes xy: "rtrancl_path r x xs y" | 
| 61764 | 96 | obtains xs' where "rtrancl_path r x xs' y" and "distinct (x # xs')" and "set xs' \<subseteq> set xs" | 
| 60519 | 97 | using xy | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 98 | proof (induct xs rule: measure_induct_rule [of length]) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 99 | case (less xs) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 100 | show ?case | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 101 | proof (cases "distinct (x # xs)") | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 102 | case True | 
| 61764 | 103 | with \<open>rtrancl_path r x xs y\<close> show ?thesis by (rule less) simp | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 104 | next | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 105 | case False | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 106 | then have "\<exists>as bs cs a. x # xs = as @ [a] @ bs @ [a] @ cs" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 107 | by (rule not_distinct_decomp) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 108 | then obtain as bs cs a where xxs: "x # xs = as @ [a] @ bs @ [a] @ cs" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 109 | by iprover | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 110 | show ?thesis | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 111 | proof (cases as) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 112 | case Nil | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 113 | with xxs have x: "x = a" and xs: "xs = bs @ a # cs" | 
| 35423 | 114 | by auto | 
| 61764 | 115 | from x xs \<open>rtrancl_path r x xs y\<close> have cs: "rtrancl_path r x cs y" "set cs \<subseteq> set xs" | 
| 35423 | 116 | by (auto elim: rtrancl_path_appendE) | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 117 | from xs have "length cs < length xs" by simp | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 118 | then show ?thesis | 
| 61764 | 119 | by (rule less(1))(blast intro: cs less(2) order_trans del: subsetI)+ | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 120 | next | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 121 | case (Cons d ds) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 122 | with xxs have xs: "xs = ds @ a # (bs @ [a] @ cs)" | 
| 35423 | 123 | by auto | 
| 60500 | 124 | with \<open>rtrancl_path r x xs y\<close> obtain xa: "rtrancl_path r x (ds @ [a]) a" | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 125 | and ay: "rtrancl_path r a (bs @ a # cs) y" | 
| 35423 | 126 | by (auto elim: rtrancl_path_appendE) | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 127 | from ay have "rtrancl_path r a cs y" by (auto elim: rtrancl_path_appendE) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 128 | with xa have xy: "rtrancl_path r x ((ds @ [a]) @ cs) y" | 
| 35423 | 129 | by (rule rtrancl_path_trans) | 
| 61764 | 130 | from xs have set: "set ((ds @ [a]) @ cs) \<subseteq> set xs" by auto | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 131 | from xs have "length ((ds @ [a]) @ cs) < length xs" by simp | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 132 | then show ?thesis | 
| 61764 | 133 | by (rule less(1))(blast intro: xy less(2) set[THEN subsetD])+ | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 134 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 135 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 136 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 137 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 138 | inductive rtrancl_tab :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 139 | for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 140 | where | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 141 | base: "rtrancl_tab r xs x x" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 142 | | step: "x \<notin> set xs \<Longrightarrow> r x y \<Longrightarrow> rtrancl_tab r (x # xs) y z \<Longrightarrow> rtrancl_tab r xs x z" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 143 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 144 | lemma rtrancl_path_imp_rtrancl_tab: | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 145 | assumes path: "rtrancl_path r x xs y" | 
| 60519 | 146 | and x: "distinct (x # xs)" | 
| 147 |     and ys: "({x} \<union> set xs) \<inter> set ys = {}"
 | |
| 148 | shows "rtrancl_tab r ys x y" | |
| 149 | using path x ys | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 150 | proof (induct arbitrary: ys) | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 151 | case base | 
| 60519 | 152 | show ?case | 
| 153 | by (rule rtrancl_tab.base) | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 154 | next | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 155 | case (step x y zs z) | 
| 60519 | 156 | then have "x \<notin> set ys" | 
| 157 | by auto | |
| 158 | from step have "distinct (y # zs)" | |
| 159 | by simp | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 160 |   moreover from step have "({y} \<union> set zs) \<inter> set (x # ys) = {}"
 | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 161 | by auto | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 162 | ultimately have "rtrancl_tab r (x # ys) y z" | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 163 | by (rule step) | 
| 60519 | 164 | with \<open>x \<notin> set ys\<close> \<open>r x y\<close> show ?case | 
| 165 | by (rule rtrancl_tab.step) | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 166 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 167 | |
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 168 | lemma rtrancl_tab_imp_rtrancl_path: | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 169 | assumes tab: "rtrancl_tab r ys x y" | 
| 60519 | 170 | obtains xs where "rtrancl_path r x xs y" | 
| 171 | using tab | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 172 | proof induct | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 173 | case base | 
| 60519 | 174 | from rtrancl_path.base show ?case | 
| 175 | by (rule base) | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 176 | next | 
| 60519 | 177 | case step | 
| 178 | show ?case | |
| 179 | by (iprover intro: step rtrancl_path.step) | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 180 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 181 | |
| 60519 | 182 | lemma rtranclp_eq_rtrancl_tab_nil: "r\<^sup>*\<^sup>* x y \<longleftrightarrow> rtrancl_tab r [] x y" | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 183 | proof | 
| 60519 | 184 | show "rtrancl_tab r [] x y" if "r\<^sup>*\<^sup>* x y" | 
| 185 | proof - | |
| 186 | from that obtain xs where "rtrancl_path r x xs y" | |
| 187 | by (auto simp add: rtranclp_eq_rtrancl_path) | |
| 188 | then obtain xs' where xs': "rtrancl_path r x xs' y" and distinct: "distinct (x # xs')" | |
| 189 | by (rule rtrancl_path_distinct) | |
| 190 |     have "({x} \<union> set xs') \<inter> set [] = {}"
 | |
| 191 | by simp | |
| 192 | with xs' distinct show ?thesis | |
| 193 | by (rule rtrancl_path_imp_rtrancl_tab) | |
| 194 | qed | |
| 195 | show "r\<^sup>*\<^sup>* x y" if "rtrancl_tab r [] x y" | |
| 196 | proof - | |
| 197 | from that obtain xs where "rtrancl_path r x xs y" | |
| 198 | by (rule rtrancl_tab_imp_rtrancl_path) | |
| 199 | then show ?thesis | |
| 200 | by (auto simp add: rtranclp_eq_rtrancl_path) | |
| 201 | qed | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 202 | qed | 
| 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 203 | |
| 60519 | 204 | declare rtranclp_rtrancl_eq [code del] | 
| 205 | declare rtranclp_eq_rtrancl_tab_nil [THEN iffD2, code_pred_intro] | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 206 | |
| 60519 | 207 | code_pred rtranclp | 
| 208 | using rtranclp_eq_rtrancl_tab_nil [THEN iffD1] by fastforce | |
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: diff
changeset | 209 | |
| 64632 | 210 | lemma rtrancl_path_Range: "\<lbrakk> rtrancl_path R x xs y; z \<in> set xs \<rbrakk> \<Longrightarrow> Rangep R z" | 
| 61765 | 211 | by(induction rule: rtrancl_path.induct) auto | 
| 212 | ||
| 64632 | 213 | lemma rtrancl_path_Range_end: "\<lbrakk> rtrancl_path R x xs y; xs \<noteq> [] \<rbrakk> \<Longrightarrow> Rangep R y" | 
| 61765 | 214 | by(induction rule: rtrancl_path.induct)(auto elim: rtrancl_path.cases) | 
| 215 | ||
| 216 | lemma rtrancl_path_nth: | |
| 217 | "\<lbrakk> rtrancl_path R x xs y; i < length xs \<rbrakk> \<Longrightarrow> R ((x # xs) ! i) (xs ! i)" | |
| 218 | proof(induction arbitrary: i rule: rtrancl_path.induct) | |
| 219 | case step thus ?case by(cases i) simp_all | |
| 220 | qed simp | |
| 221 | ||
| 222 | lemma rtrancl_path_last: "\<lbrakk> rtrancl_path R x xs y; xs \<noteq> [] \<rbrakk> \<Longrightarrow> last xs = y" | |
| 223 | by(induction rule: rtrancl_path.induct)(auto elim: rtrancl_path.cases) | |
| 224 | ||
| 225 | lemma rtrancl_path_mono: | |
| 226 | "\<lbrakk> rtrancl_path R x p y; \<And>x y. R x y \<Longrightarrow> S x y \<rbrakk> \<Longrightarrow> rtrancl_path S x p y" | |
| 227 | by(induction rule: rtrancl_path.induct)(auto intro: rtrancl_path.intros) | |
| 228 | ||
| 62390 | 229 | end |