diff -r a9599d3d7610 -r 5197a2ecb658 src/HOL/Decision_Procs/Cooper.thy --- a/src/HOL/Decision_Procs/Cooper.thy Thu Nov 05 10:39:49 2015 +0100 +++ b/src/HOL/Decision_Procs/Cooper.thy Thu Nov 05 10:39:59 2015 +0100 @@ -18,7 +18,7 @@ datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num | Mul int num -primrec num_size :: "num \ nat" -- \A size for num to make inductive proofs simpler\ +primrec num_size :: "num \ nat" \ \A size for num to make inductive proofs simpler\ where "num_size (C c) = 1" | "num_size (Bound n) = 1" @@ -44,7 +44,7 @@ | Closed nat | NClosed nat -fun fmsize :: "fm \ nat" -- \A size for fm\ +fun fmsize :: "fm \ nat" \ \A size for fm\ where "fmsize (NOT p) = 1 + fmsize p" | "fmsize (And p q) = 1 + fmsize p + fmsize q" @@ -60,7 +60,7 @@ lemma fmsize_pos: "fmsize p > 0" by (induct p rule: fmsize.induct) simp_all -primrec Ifm :: "bool list \ int list \ fm \ bool" -- \Semantics of formulae (fm)\ +primrec Ifm :: "bool list \ int list \ fm \ bool" \ \Semantics of formulae (fm)\ where "Ifm bbs bs T \ True" | "Ifm bbs bs F \ False" @@ -113,7 +113,7 @@ by (induct p arbitrary: bs rule: prep.induct) auto -fun qfree :: "fm \ bool" -- \Quantifier freeness\ +fun qfree :: "fm \ bool" \ \Quantifier freeness\ where "qfree (E p) \ False" | "qfree (A p) \ False" @@ -127,7 +127,7 @@ text \Boundedness and substitution\ -primrec numbound0 :: "num \ bool" -- \a num is INDEPENDENT of Bound 0\ +primrec numbound0 :: "num \ bool" \ \a num is INDEPENDENT of Bound 0\ where "numbound0 (C c) \ True" | "numbound0 (Bound n) \ n > 0" @@ -142,7 +142,7 @@ shows "Inum (b # bs) a = Inum (b' # bs) a" using nb by (induct a rule: num.induct) (auto simp add: gr0_conv_Suc) -primrec bound0 :: "fm \ bool" -- \A Formula is independent of Bound 0\ +primrec bound0 :: "fm \ bool" \ \A Formula is independent of Bound 0\ where "bound0 T \ True" | "bound0 F \ True" @@ -188,7 +188,7 @@ "numbound0 a \ Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t" by (induct t rule: numsubst0.induct) (auto simp: nth_Cons' numbound0_I[where b="b" and b'="b'"]) -primrec subst0:: "num \ fm \ fm" -- \substitue a num into a formula for Bound 0\ +primrec subst0:: "num \ fm \ fm" \ \substitue a num into a formula for Bound 0\ where "subst0 t T = T" | "subst0 t F = F" @@ -254,7 +254,7 @@ lemma decr_qf: "bound0 p \ qfree (decr p)" by (induct p) simp_all -fun isatom :: "fm \ bool" -- \test for atomicity\ +fun isatom :: "fm \ bool" \ \test for atomicity\ where "isatom T \ True" | "isatom F \ True" @@ -856,9 +856,9 @@ (auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf simpfm simpfm_qf simp del: simpfm.simps) -text \Linearity for fm where Bound 0 ranges over @{text "\"}\ +text \Linearity for fm where Bound 0 ranges over \\\\ -fun zsplit0 :: "num \ int \ num" -- \splits the bounded from the unbounded part\ +fun zsplit0 :: "num \ int \ num" \ \splits the bounded from the unbounded part\ where "zsplit0 (C c) = (0, C c)" | "zsplit0 (Bound n) = (if n = 0 then (1, C 0) else (0, Bound n))" @@ -989,7 +989,7 @@ by simp qed -consts iszlfm :: "fm \ bool" -- \Linearity test for fm\ +consts iszlfm :: "fm \ bool" \ \Linearity test for fm\ recdef iszlfm "measure size" "iszlfm (And p q) \ iszlfm p \ iszlfm q" "iszlfm (Or p q) \ iszlfm p \ iszlfm q" @@ -1006,7 +1006,7 @@ lemma zlin_qfree: "iszlfm p \ qfree p" by (induct p rule: iszlfm.induct) auto -consts zlfm :: "fm \ fm" -- \Linearity transformation for fm\ +consts zlfm :: "fm \ fm" \ \Linearity transformation for fm\ recdef zlfm "measure fmsize" "zlfm (And p q) = And (zlfm p) (zlfm q)" "zlfm (Or p q) = Or (zlfm p) (zlfm q)" @@ -1258,7 +1258,7 @@ qed qed auto -consts minusinf :: "fm \ fm" -- \Virtual substitution of @{text "-\"}\ +consts minusinf :: "fm \ fm" \ \Virtual substitution of \-\\\ recdef minusinf "measure size" "minusinf (And p q) = And (minusinf p) (minusinf q)" "minusinf (Or p q) = Or (minusinf p) (minusinf q)" @@ -1273,7 +1273,7 @@ lemma minusinf_qfree: "qfree p \ qfree (minusinf p)" by (induct p rule: minusinf.induct) auto -consts plusinf :: "fm \ fm" -- \Virtual substitution of @{text "+\"}\ +consts plusinf :: "fm \ fm" \ \Virtual substitution of \+\\\ recdef plusinf "measure size" "plusinf (And p q) = And (plusinf p) (plusinf q)" "plusinf (Or p q) = Or (plusinf p) (plusinf q)" @@ -1285,7 +1285,7 @@ "plusinf (Ge (CN 0 c e)) = T" "plusinf p = p" -consts \ :: "fm \ int" -- \Compute @{text "lcm {d| N\<^sup>? Dvd c*x+t \ p}"}\ +consts \ :: "fm \ int" \ \Compute \lcm {d| N\<^sup>? Dvd c*x+t \ p}\\ recdef \ "measure size" "\ (And p q) = lcm (\ p) (\ q)" "\ (Or p q) = lcm (\ p) (\ q)" @@ -1293,7 +1293,7 @@ "\ (NDvd i (CN 0 c e)) = i" "\ p = 1" -consts d_\ :: "fm \ int \ bool" -- \check if a given l divides all the ds above\ +consts d_\ :: "fm \ int \ bool" \ \check if a given l divides all the ds above\ recdef d_\ "measure size" "d_\ (And p q) = (\d. d_\ p d \ d_\ q d)" "d_\ (Or p q) = (\d. d_\ p d \ d_\ q d)" @@ -1354,7 +1354,7 @@ qed simp_all -consts a_\ :: "fm \ int \ fm" -- \adjust the coefficients of a formula\ +consts a_\ :: "fm \ int \ fm" \ \adjust the coefficients of a formula\ recdef a_\ "measure size" "a_\ (And p q) = (\k. And (a_\ p k) (a_\ q k))" "a_\ (Or p q) = (\k. Or (a_\ p k) (a_\ q k))" @@ -1368,7 +1368,7 @@ "a_\ (NDvd i (CN 0 c e))=(\k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" "a_\ p = (\k. p)" -consts d_\ :: "fm \ int \ bool" -- \test if all coeffs c of c divide a given l\ +consts d_\ :: "fm \ int \ bool" \ \test if all coeffs c of c divide a given l\ recdef d_\ "measure size" "d_\ (And p q) = (\k. (d_\ p k) \ (d_\ q k))" "d_\ (Or p q) = (\k. (d_\ p k) \ (d_\ q k))" @@ -1382,7 +1382,7 @@ "d_\ (NDvd i (CN 0 c e))=(\k. c dvd k)" "d_\ p = (\k. True)" -consts \ :: "fm \ int" -- \computes the lcm of all coefficients of x\ +consts \ :: "fm \ int" \ \computes the lcm of all coefficients of x\ recdef \ "measure size" "\ (And p q) = lcm (\ p) (\ q)" "\ (Or p q) = lcm (\ p) (\ q)" @@ -1434,7 +1434,7 @@ "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" "mirror p = p" -text \Lemmas for the correctness of @{text "\_\"}\ +text \Lemmas for the correctness of \\_\\\ lemma dvd1_eq1: fixes x :: int