# HG changeset patch # User huffman # Date 1245351603 25200 # Node ID 32f07b47f9c548c237f7d1ac979ed7496250a90c # Parent 347e9d18f372641cecc5a2dba55e66e90f044c05# Parent 2eb55a82acd948277515e413b89de28b15f0274b merged diff -r 347e9d18f372 -r 32f07b47f9c5 src/HOL/Decision_Procs/Cooper.thy --- a/src/HOL/Decision_Procs/Cooper.thy Thu Jun 18 19:54:21 2009 +0200 +++ b/src/HOL/Decision_Procs/Cooper.thy Thu Jun 18 12:00:03 2009 -0700 @@ -1071,8 +1071,8 @@ "plusinf p = p" recdef \ "measure size" - "\ (And p q) = zlcm (\ p) (\ q)" - "\ (Or p q) = zlcm (\ p) (\ q)" + "\ (And p q) = lcm (\ p) (\ q)" + "\ (Or p q) = lcm (\ p) (\ q)" "\ (Dvd i (CN 0 c e)) = i" "\ (NDvd i (CN 0 c e)) = i" "\ p = 1" @@ -1104,20 +1104,20 @@ proof (induct p rule: iszlfm.induct) case (1 p q) let ?d = "\ (And p q)" - from prems zlcm_pos have dp: "?d >0" by simp + from prems int_lcm_pos have dp: "?d >0" by simp have d1: "\ p dvd \ (And p q)" using prems by simp - hence th: "d\ p ?d" using delta_mono prems(3-4) by(simp del:dvd_zlcm_self1) + hence th: "d\ p ?d" using delta_mono prems(3-4) by(simp del:int_lcm_dvd1) have "\ q dvd \ (And p q)" using prems by simp - hence th': "d\ q ?d" using delta_mono prems by(simp del:dvd_zlcm_self2) + hence th': "d\ q ?d" using delta_mono prems by(simp del:int_lcm_dvd2) from th th' dp show ?case by simp next case (2 p q) let ?d = "\ (And p q)" - from prems zlcm_pos have dp: "?d >0" by simp + from prems int_lcm_pos have dp: "?d >0" by simp have "\ p dvd \ (And p q)" using prems by simp - hence th: "d\ p ?d" using delta_mono prems by(simp del:dvd_zlcm_self1) + hence th: "d\ p ?d" using delta_mono prems by(simp del:int_lcm_dvd1) have "\ q dvd \ (And p q)" using prems by simp - hence th': "d\ q ?d" using delta_mono prems by(simp del:dvd_zlcm_self2) + hence th': "d\ q ?d" using delta_mono prems by(simp del:int_lcm_dvd2) from th th' dp show ?case by simp qed simp_all @@ -1156,8 +1156,8 @@ "d\ p = (\ k. True)" recdef \ "measure size" - "\ (And p q) = zlcm (\ p) (\ q)" - "\ (Or p q) = zlcm (\ p) (\ q)" + "\ (And p q) = lcm (\ p) (\ q)" + "\ (Or p q) = lcm (\ p) (\ q)" "\ (Eq (CN 0 c e)) = c" "\ (NEq (CN 0 c e)) = c" "\ (Lt (CN 0 c e)) = c" @@ -1406,19 +1406,19 @@ using linp proof(induct p rule: iszlfm.induct) case (1 p q) - from prems have dl1: "\ p dvd zlcm (\ p) (\ q)" by simp - from prems have dl2: "\ q dvd zlcm (\ p) (\ q)" by simp - from prems d\_mono[where p = "p" and l="\ p" and l'="zlcm (\ p) (\ q)"] - d\_mono[where p = "q" and l="\ q" and l'="zlcm (\ p) (\ q)"] - dl1 dl2 show ?case by (auto simp add: zlcm_pos) + from prems have dl1: "\ p dvd lcm (\ p) (\ q)" by simp + from prems have dl2: "\ q dvd lcm (\ p) (\ q)" by simp + from prems d\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] + d\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] + dl1 dl2 show ?case by (auto simp add: int_lcm_pos) next case (2 p q) - from prems have dl1: "\ p dvd zlcm (\ p) (\ q)" by simp - from prems have dl2: "\ q dvd zlcm (\ p) (\ q)" by simp - from prems d\_mono[where p = "p" and l="\ p" and l'="zlcm (\ p) (\ q)"] - d\_mono[where p = "q" and l="\ q" and l'="zlcm (\ p) (\ q)"] - dl1 dl2 show ?case by (auto simp add: zlcm_pos) -qed (auto simp add: zlcm_pos) + from prems have dl1: "\ p dvd lcm (\ p) (\ q)" by simp + from prems have dl2: "\ q dvd lcm (\ p) (\ q)" by simp + from prems d\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] + d\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] + dl1 dl2 show ?case by (auto simp add: int_lcm_pos) +qed (auto simp add: int_lcm_pos) lemma a\: assumes linp: "iszlfm p" and d: "d\ p l" and lp: "l > 0" shows "iszlfm (a\ p l) \ d\ (a\ p l) 1 \ (Ifm bbs (l*x #bs) (a\ p l) = Ifm bbs (x#bs) p)"