diff -r 4dc65845eab3 -r d8d7d1b785af src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Wed Feb 24 11:55:52 2010 +0100 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Mar 01 13:40:23 2010 +0100 @@ -273,10 +273,10 @@ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" shows "allpolys isnpoly t \ isnpoly c \ allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm) -constdefs tmneg :: "tm \ tm" +definition tmneg :: "tm \ tm" where "tmneg t \ tmmul t (C (- 1,1))" -constdefs tmsub :: "tm \ tm \ tm" +definition tmsub :: "tm \ tm \ tm" where "tmsub s t \ (if s = t then CP 0\<^sub>p else tmadd (s,tmneg t))" lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)" @@ -477,26 +477,26 @@ lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)" by (induct p rule: not.induct) auto -constdefs conj :: "fm \ fm \ fm" +definition conj :: "fm \ fm \ fm" where "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else if p = q then p else And p q)" lemma conj[simp]: "Ifm vs bs (conj p q) = Ifm vs bs (And p q)" by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) -constdefs disj :: "fm \ fm \ fm" +definition disj :: "fm \ fm \ fm" where "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p else if p=q then p else Or p q)" lemma disj[simp]: "Ifm vs bs (disj p q) = Ifm vs bs (Or p q)" by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) -constdefs imp :: "fm \ fm \ fm" +definition imp :: "fm \ fm \ fm" where "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp[simp]: "Ifm vs bs (imp p q) = Ifm vs bs (Imp p q)" by (cases "p=F \ q=T",simp_all add: imp_def) -constdefs iff :: "fm \ fm \ fm" +definition iff :: "fm \ fm \ fm" where "iff p q \ (if (p = q) then T else if (p = NOT q \ NOT p = q) then F else if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else Iff p q)" @@ -776,10 +776,10 @@ lemma bound0_qf: "bound0 p \ qfree p" by (induct p, simp_all) -constdefs djf:: "('a \ fm) \ 'a \ fm \ fm" +definition djf :: "('a \ fm) \ 'a \ fm \ fm" where "djf f p q \ (if q=T then T else if q=F then f p else (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q))" -constdefs evaldjf:: "('a \ fm) \ 'a list \ fm" +definition evaldjf :: "('a \ fm) \ 'a list \ fm" where "evaldjf f ps \ foldr (djf f) ps F" lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)" @@ -823,7 +823,7 @@ thus ?thesis by (simp only: list_all_iff) qed -constdefs DJ :: "(fm \ fm) \ fm \ fm" +definition DJ :: "(fm \ fm) \ fm \ fm" where "DJ f p \ evaldjf f (disjuncts p)" lemma DJ: assumes fdj: "\ p q. Ifm vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))" @@ -869,10 +869,10 @@ "conjuncts T = []" "conjuncts p = [p]" -constdefs list_conj :: "fm list \ fm" +definition list_conj :: "fm list \ fm" where "list_conj ps \ foldr conj ps T" -constdefs CJNB:: "(fm \ fm) \ fm \ fm" +definition CJNB :: "(fm \ fm) \ fm \ fm" where "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs in conj (decr0 (list_conj yes)) (f (list_conj no)))" @@ -1158,7 +1158,7 @@ "conjs p = [p]" lemma conjs_ci: "(\ q \ set (conjs p). Ifm vs bs q) = Ifm vs bs p" by (induct p rule: conjs.induct, auto) -constdefs list_disj :: "fm list \ fm" +definition list_disj :: "fm list \ fm" where "list_disj ps \ foldr disj ps F" lemma list_conj: "Ifm vs bs (list_conj ps) = (\p\ set ps. Ifm vs bs p)"