| author | wenzelm | 
| Wed, 03 Sep 2008 11:44:52 +0200 | |
| changeset 28108 | 1b08ed83b79e | 
| parent 25972 | 94b15338da8d | 
| child 28455 | a79701b14a30 | 
| permissions | -rw-r--r-- | 
| 1476 | 1 | (* Title: HOL/Lambda/Commutation.thy | 
| 1278 | 2 | ID: $Id$ | 
| 1476 | 3 | Author: Tobias Nipkow | 
| 1278 | 4 | Copyright 1995 TU Muenchen | 
| 5 | *) | |
| 6 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 7 | header {* Abstract commutation and confluence notions *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 8 | |
| 16417 | 9 | theory Commutation imports Main begin | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 10 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 11 | subsection {* Basic definitions *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 12 | |
| 19086 | 13 | definition | 
| 22270 | 14 | square :: "['a => 'a => bool, 'a => 'a => bool, 'a => 'a => bool, 'a => 'a => bool] => bool" where | 
| 19086 | 15 | "square R S T U = | 
| 22270 | 16 | (\<forall>x y. R x y --> (\<forall>z. S x z --> (\<exists>u. T y u \<and> U z u)))" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 17 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19363diff
changeset | 18 | definition | 
| 22270 | 19 | commute :: "['a => 'a => bool, 'a => 'a => bool] => bool" where | 
| 19086 | 20 | "commute R S = square R S S R" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 21 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19363diff
changeset | 22 | definition | 
| 22270 | 23 |   diamond :: "('a => 'a => bool) => bool" where
 | 
| 19086 | 24 | "diamond R = commute R R" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 25 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19363diff
changeset | 26 | definition | 
| 22270 | 27 |   Church_Rosser :: "('a => 'a => bool) => bool" where
 | 
| 19086 | 28 | "Church_Rosser R = | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22270diff
changeset | 29 | (\<forall>x y. (sup R (R^--1))^** x y --> (\<exists>z. R^** x z \<and> R^** y z))" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 30 | |
| 19363 | 31 | abbreviation | 
| 22270 | 32 |   confluent :: "('a => 'a => bool) => bool" where
 | 
| 33 | "confluent R == diamond (R^**)" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 34 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 35 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 36 | subsection {* Basic lemmas *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 37 | |
| 25972 | 38 | subsubsection {* @{text "square"} *}
 | 
| 1278 | 39 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 40 | lemma square_sym: "square R S T U ==> square S R U T" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 41 | apply (unfold square_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 42 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 43 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 44 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 45 | lemma square_subset: | 
| 22270 | 46 | "[| square R S T U; T \<le> T' |] ==> square R S T' U" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 47 | apply (unfold square_def) | 
| 22270 | 48 | apply (blast dest: predicate2D) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 49 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 50 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 51 | lemma square_reflcl: | 
| 22270 | 52 | "[| square R S T (R^==); S \<le> T |] ==> square (R^==) S T (R^==)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 53 | apply (unfold square_def) | 
| 22270 | 54 | apply (blast dest: predicate2D) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 55 | done | 
| 1278 | 56 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 57 | lemma square_rtrancl: | 
| 22270 | 58 | "square R S S T ==> square (R^**) S S (T^**)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 59 | apply (unfold square_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 60 | apply (intro strip) | 
| 23750 | 61 | apply (erule rtranclp_induct) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 62 | apply blast | 
| 23750 | 63 | apply (blast intro: rtranclp.rtrancl_into_rtrancl) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 64 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 65 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 66 | lemma square_rtrancl_reflcl_commute: | 
| 22270 | 67 | "square R S (S^**) (R^==) ==> commute (R^**) (S^**)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 68 | apply (unfold commute_def) | 
| 22270 | 69 | apply (fastsimp dest: square_reflcl square_sym [THEN square_rtrancl]) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 70 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 71 | |
| 1278 | 72 | |
| 25972 | 73 | subsubsection {* @{text "commute"} *}
 | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 74 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 75 | lemma commute_sym: "commute R S ==> commute S R" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 76 | apply (unfold commute_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 77 | apply (blast intro: square_sym) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 78 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 79 | |
| 22270 | 80 | lemma commute_rtrancl: "commute R S ==> commute (R^**) (S^**)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 81 | apply (unfold commute_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 82 | apply (blast intro: square_rtrancl square_sym) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 83 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 84 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 85 | lemma commute_Un: | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22270diff
changeset | 86 | "[| commute R T; commute S T |] ==> commute (sup R S) T" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 87 | apply (unfold commute_def square_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 88 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 89 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 90 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 91 | |
| 25972 | 92 | subsubsection {* @{text "diamond"}, @{text "confluence"}, and @{text "union"} *}
 | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 93 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 94 | lemma diamond_Un: | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22270diff
changeset | 95 | "[| diamond R; diamond S; commute R S |] ==> diamond (sup R S)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 96 | apply (unfold diamond_def) | 
| 23815 | 97 | apply (blast intro: commute_Un commute_sym) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 98 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 99 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 100 | lemma diamond_confluent: "diamond R ==> confluent R" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 101 | apply (unfold diamond_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 102 | apply (erule commute_rtrancl) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 103 | done | 
| 1278 | 104 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 105 | lemma square_reflcl_confluent: | 
| 22270 | 106 | "square R R (R^==) (R^==) ==> confluent R" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 107 | apply (unfold diamond_def) | 
| 22270 | 108 | apply (fast intro: square_rtrancl_reflcl_commute elim: square_subset) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 109 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 110 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 111 | lemma confluent_Un: | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22270diff
changeset | 112 | "[| confluent R; confluent S; commute (R^**) (S^**) |] ==> confluent (sup R S)" | 
| 23750 | 113 | apply (rule rtranclp_sup_rtranclp [THEN subst]) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 114 | apply (blast dest: diamond_Un intro: diamond_confluent) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 115 | done | 
| 1278 | 116 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 117 | lemma diamond_to_confluence: | 
| 22270 | 118 | "[| diamond R; T \<le> R; R \<le> T^** |] ==> confluent T" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 119 | apply (force intro: diamond_confluent | 
| 23750 | 120 | dest: rtranclp_subset [symmetric]) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 121 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 122 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 123 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 124 | subsection {* Church-Rosser *}
 | 
| 1278 | 125 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 126 | lemma Church_Rosser_confluent: "Church_Rosser R = confluent R" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 127 | apply (unfold square_def commute_def diamond_def Church_Rosser_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 128 |   apply (tactic {* safe_tac HOL_cs *})
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 129 |    apply (tactic {*
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 130 | blast_tac (HOL_cs addIs | 
| 23750 | 131 | [thm "sup_ge2" RS thm "rtranclp_mono" RS thm "predicate2D" RS thm "rtranclp_trans", | 
| 132 | thm "rtranclp_converseI", thm "conversepI", | |
| 133 | thm "sup_ge1" RS thm "rtranclp_mono" RS thm "predicate2D"]) 1 *}) | |
| 134 | apply (erule rtranclp_induct) | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 135 | apply blast | 
| 23750 | 136 | apply (blast del: rtranclp.rtrancl_refl intro: rtranclp_trans) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 137 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 138 | |
| 13089 | 139 | |
| 140 | subsection {* Newman's lemma *}
 | |
| 141 | ||
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 142 | text {* Proof by Stefan Berghofer *}
 | 
| 13346 | 143 | |
| 13343 | 144 | theorem newman: | 
| 22270 | 145 | assumes wf: "wfP (R\<inverse>\<inverse>)" | 
| 146 | and lc: "\<And>a b c. R a b \<Longrightarrow> R a c \<Longrightarrow> | |
| 147 | \<exists>d. R\<^sup>*\<^sup>* b d \<and> R\<^sup>*\<^sup>* c d" | |
| 148 | shows "\<And>b c. R\<^sup>*\<^sup>* a b \<Longrightarrow> R\<^sup>*\<^sup>* a c \<Longrightarrow> | |
| 149 | \<exists>d. R\<^sup>*\<^sup>* b d \<and> R\<^sup>*\<^sup>* c d" | |
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 150 | using wf | 
| 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 151 | proof induct | 
| 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 152 | case (less x b c) | 
| 23464 | 153 | have xc: "R\<^sup>*\<^sup>* x c" by fact | 
| 154 | have xb: "R\<^sup>*\<^sup>* x b" by fact thus ?case | |
| 23750 | 155 | proof (rule converse_rtranclpE) | 
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 156 | assume "x = b" | 
| 22270 | 157 | with xc have "R\<^sup>*\<^sup>* b c" by simp | 
| 17589 | 158 | thus ?thesis by iprover | 
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 159 | next | 
| 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 160 | fix y | 
| 22270 | 161 | assume xy: "R x y" | 
| 162 | assume yb: "R\<^sup>*\<^sup>* y b" | |
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 163 | from xc show ?thesis | 
| 23750 | 164 | proof (rule converse_rtranclpE) | 
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 165 | assume "x = c" | 
| 22270 | 166 | with xb have "R\<^sup>*\<^sup>* c b" by simp | 
| 17589 | 167 | thus ?thesis by iprover | 
| 13089 | 168 | next | 
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 169 | fix y' | 
| 22270 | 170 | assume y'c: "R\<^sup>*\<^sup>* y' c" | 
| 171 | assume xy': "R x y'" | |
| 172 | with xy have "\<exists>u. R\<^sup>*\<^sup>* y u \<and> R\<^sup>*\<^sup>* y' u" by (rule lc) | |
| 173 | then obtain u where yu: "R\<^sup>*\<^sup>* y u" and y'u: "R\<^sup>*\<^sup>* y' u" by iprover | |
| 174 | from xy have "R\<inverse>\<inverse> y x" .. | |
| 175 | from this and yb yu have "\<exists>d. R\<^sup>*\<^sup>* b d \<and> R\<^sup>*\<^sup>* u d" by (rule less) | |
| 176 | then obtain v where bv: "R\<^sup>*\<^sup>* b v" and uv: "R\<^sup>*\<^sup>* u v" by iprover | |
| 177 | from xy' have "R\<inverse>\<inverse> y' x" .. | |
| 23750 | 178 | moreover from y'u and uv have "R\<^sup>*\<^sup>* y' v" by (rule rtranclp_trans) | 
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 179 | moreover note y'c | 
| 22270 | 180 | ultimately have "\<exists>d. R\<^sup>*\<^sup>* v d \<and> R\<^sup>*\<^sup>* c d" by (rule less) | 
| 181 | then obtain w where vw: "R\<^sup>*\<^sup>* v w" and cw: "R\<^sup>*\<^sup>* c w" by iprover | |
| 23750 | 182 | from bv vw have "R\<^sup>*\<^sup>* b w" by (rule rtranclp_trans) | 
| 17589 | 183 | with cw show ?thesis by iprover | 
| 13089 | 184 | qed | 
| 185 | qed | |
| 186 | qed | |
| 187 | ||
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 188 | text {*
 | 
| 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 189 | \medskip Alternative version. Partly automated by Tobias | 
| 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 190 | Nipkow. Takes 2 minutes (2002). | 
| 13346 | 191 | |
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 192 | This is the maximal amount of automation possible at the moment. | 
| 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 193 | *} | 
| 13346 | 194 | |
| 13349 
7d4441c8c46a
Added "using" to the beginning of original newman proof again, because
 berghofe parents: 
13346diff
changeset | 195 | theorem newman': | 
| 22270 | 196 | assumes wf: "wfP (R\<inverse>\<inverse>)" | 
| 197 | and lc: "\<And>a b c. R a b \<Longrightarrow> R a c \<Longrightarrow> | |
| 198 | \<exists>d. R\<^sup>*\<^sup>* b d \<and> R\<^sup>*\<^sup>* c d" | |
| 199 | shows "\<And>b c. R\<^sup>*\<^sup>* a b \<Longrightarrow> R\<^sup>*\<^sup>* a c \<Longrightarrow> | |
| 200 | \<exists>d. R\<^sup>*\<^sup>* b d \<and> R\<^sup>*\<^sup>* c d" | |
| 18241 | 201 | using wf | 
| 13346 | 202 | proof induct | 
| 203 | case (less x b c) | |
| 22270 | 204 | note IH = `\<And>y b c. \<lbrakk>R\<inverse>\<inverse> y x; R\<^sup>*\<^sup>* y b; R\<^sup>*\<^sup>* y c\<rbrakk> | 
| 205 | \<Longrightarrow> \<exists>d. R\<^sup>*\<^sup>* b d \<and> R\<^sup>*\<^sup>* c d` | |
| 23464 | 206 | have xc: "R\<^sup>*\<^sup>* x c" by fact | 
| 207 | have xb: "R\<^sup>*\<^sup>* x b" by fact | |
| 13346 | 208 | thus ?case | 
| 23750 | 209 | proof (rule converse_rtranclpE) | 
| 13346 | 210 | assume "x = b" | 
| 22270 | 211 | with xc have "R\<^sup>*\<^sup>* b c" by simp | 
| 17589 | 212 | thus ?thesis by iprover | 
| 13346 | 213 | next | 
| 214 | fix y | |
| 22270 | 215 | assume xy: "R x y" | 
| 216 | assume yb: "R\<^sup>*\<^sup>* y b" | |
| 13346 | 217 | from xc show ?thesis | 
| 23750 | 218 | proof (rule converse_rtranclpE) | 
| 13346 | 219 | assume "x = c" | 
| 22270 | 220 | with xb have "R\<^sup>*\<^sup>* c b" by simp | 
| 17589 | 221 | thus ?thesis by iprover | 
| 13346 | 222 | next | 
| 223 | fix y' | |
| 22270 | 224 | assume y'c: "R\<^sup>*\<^sup>* y' c" | 
| 225 | assume xy': "R x y'" | |
| 226 | with xy obtain u where u: "R\<^sup>*\<^sup>* y u" "R\<^sup>*\<^sup>* y' u" | |
| 18241 | 227 | by (blast dest: lc) | 
| 13346 | 228 | from yb u y'c show ?thesis | 
| 23750 | 229 | by (blast del: rtranclp.rtrancl_refl | 
| 230 | intro: rtranclp_trans | |
| 22270 | 231 | dest: IH [OF conversepI, OF xy] IH [OF conversepI, OF xy']) | 
| 13346 | 232 | qed | 
| 233 | qed | |
| 234 | qed | |
| 235 | ||
| 10179 | 236 | end |