# HG changeset patch # User paulson # Date 1250259616 -3600 # Node ID 532c11b965b56d6dcfe3b82b747523bd359c9b68 # Parent 37d87022cad37733426b6445e82cb841c7975eaa Put back characters that had been erased by Emacs diff -r 37d87022cad3 -r 532c11b965b5 src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Thu Aug 13 17:19:54 2009 +0100 +++ b/src/HOL/Bali/Basis.thy Fri Aug 14 15:20:16 2009 +0100 @@ -63,36 +63,36 @@ by (auto intro: r_into_rtrancl rtrancl_trans) lemma triangle_lemma: - "\ \ a b c. \(a,b)\r; (a,c)\r\ \ b=c; (a,x)\r*; (a,y)\r*\ - \ (x,y)\r* \ (y,x)\r*" + "\ \ a b c. \(a,b)\r; (a,c)\r\ \ b=c; (a,x)\r\<^sup>*; (a,y)\r\<^sup>*\ + \ (x,y)\r\<^sup>* \ (y,x)\r\<^sup>*" proof - note converse_rtrancl_induct = converse_rtrancl_induct [consumes 1] note converse_rtranclE = converse_rtranclE [consumes 1] assume unique: "\ a b c. \(a,b)\r; (a,c)\r\ \ b=c" - assume "(a,x)\r*" - then show "(a,y)\r* \ (x,y)\r* \ (y,x)\r*" + assume "(a,x)\r\<^sup>*" + then show "(a,y)\r\<^sup>* \ (x,y)\r\<^sup>* \ (y,x)\r\<^sup>*" proof (induct rule: converse_rtrancl_induct) - assume "(x,y)\r*" + assume "(x,y)\r\<^sup>*" then show ?thesis by blast next fix a v assume a_v_r: "(a, v) \ r" and - v_x_rt: "(v, x) \ r*" and - a_y_rt: "(a, y) \ r*" and - hyp: "(v, y) \ r* \ (x, y) \ r* \ (y, x) \ r*" + v_x_rt: "(v, x) \ r\<^sup>*" and + a_y_rt: "(a, y) \ r\<^sup>*" and + hyp: "(v, y) \ r\<^sup>* \ (x, y) \ r\<^sup>* \ (y, x) \ r\<^sup>*" from a_y_rt - show "(x, y) \ r* \ (y, x) \ r*" + show "(x, y) \ r\<^sup>* \ (y, x) \ r\<^sup>*" proof (cases rule: converse_rtranclE) assume "a=y" - with a_v_r v_x_rt have "(y,x) \ r*" + with a_v_r v_x_rt have "(y,x) \ r\<^sup>*" by (auto intro: r_into_rtrancl rtrancl_trans) then show ?thesis by blast next fix w assume a_w_r: "(a, w) \ r" and - w_y_rt: "(w, y) \ r*" + w_y_rt: "(w, y) \ r\<^sup>*" from a_v_r a_w_r unique have "v=w" by auto @@ -105,7 +105,7 @@ lemma rtrancl_cases [consumes 1, case_names Refl Trancl]: - "\(a,b)\r*; a = b \ P; (a,b)\r+ \ P\ \ P" + "\(a,b)\r\<^sup>*; a = b \ P; (a,b)\r\<^sup>+ \ P\ \ P" apply (erule rtranclE) apply (auto dest: rtrancl_into_trancl1) done