| author | wenzelm | 
| Tue, 30 Mar 2010 00:12:42 +0200 | |
| changeset 36014 | c51a077680e4 | 
| parent 35419 | d78659d1723e | 
| child 39246 | 9e58f0499f57 | 
| permissions | -rw-r--r-- | 
| 29650 | 1 | (* Title: HOL/ex/ReflectionEx.thy | 
| 20319 | 2 | Author: Amine Chaieb, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 5 | header {* Examples for generic reflection and reification *}
 | |
| 29650 | 6 | |
| 20319 | 7 | theory ReflectionEx | 
| 8 | imports Reflection | |
| 9 | begin | |
| 10 | ||
| 20337 | 11 | text{* This theory presents two methods: reify and reflection *}
 | 
| 20319 | 12 | |
| 20337 | 13 | text{* 
 | 
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 14 | Consider an HOL type 'a, the structure of which is not recongnisable on the theory level. This is the case of bool, arithmetical terms such as int, real etc \dots | 
| 20337 | 15 | In order to implement a simplification on terms of type 'a we often need its structure. | 
| 16 | Traditionnaly such simplifications are written in ML, proofs are synthesized. | |
| 17 | An other strategy is to declare an HOL-datatype tau and an HOL function (the interpretation) that maps elements of tau to elements of 'a. The functionality of @{text reify} is to compute a term s::tau, which is the representant of t. For this it needs equations for the interpretation.
 | |
| 18 | ||
| 19 | NB: All the interpretations supported by @{text reify} must have the type @{text "'b list \<Rightarrow> tau \<Rightarrow> 'a"}.
 | |
| 20 | The method @{text reify} can also be told which subterm of the current subgoal should be reified. The general call for @{text reify} is: @{text "reify eqs (t)"}, where @{text eqs} are the defining equations of the interpretation and @{text "(t)"} is an optional parameter which specifies the subterm to which reification should be applied to. If @{text "(t)"} is abscent, @{text reify} tries to reify the whole subgoal.
 | |
| 21 | ||
| 22 | The method reflection uses @{text reify} and has a very similar signature: @{text "reflection corr_thm eqs (t)"}. Here again @{text eqs} and @{text "(t)"} are as described above and @{text corr_thm} is a thorem proving @{term "I vs (f t) = I vs t"}. We assume that @{text I} is the interpretation and @{text f} is some useful and executable simplification of type @{text "tau \<Rightarrow> tau"}. The method @{text reflection} applies reification and hence the theorem @{term "t = I xs s"} and hence using @{text corr_thm} derives @{term "t = I xs (f s)"}. It then uses normalization by evaluation to prove @{term "f s = s'"} which almost finishes the proof of @{term "t = t'"} where @{term "I xs s' = t'"}.
 | |
| 23 | *} | |
| 24 | ||
| 25 | text{* Example 1 : Propositional formulae and NNF.*}
 | |
| 26 | text{* The type @{text fm} represents simple propositional formulae: *}
 | |
| 27 | ||
| 23547 | 28 | datatype form = TrueF | FalseF | Less nat nat | | 
| 29 | And form form | Or form form | Neg form | ExQ form | |
| 30 | ||
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 31 | fun interp :: "form \<Rightarrow> ('a::ord) list \<Rightarrow> bool" where
 | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 32 | "interp TrueF e = True" | | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 33 | "interp FalseF e = False" | | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 34 | "interp (Less i j) e = (e!i < e!j)" | | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 35 | "interp (And f1 f2) e = (interp f1 e & interp f2 e)" | | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 36 | "interp (Or f1 f2) e = (interp f1 e | interp f2 e)" | | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 37 | "interp (Neg f) e = (~ interp f e)" | | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 38 | "interp (ExQ f) e = (EX x. interp f (x#e))" | 
| 23547 | 39 | |
| 23608 | 40 | lemmas interp_reify_eqs = interp.simps | 
| 41 | declare interp_reify_eqs[reify] | |
| 23547 | 42 | |
| 23608 | 43 | lemma "EX x. x < y & x < z" | 
| 44 | apply (reify ) | |
| 45 | oops | |
| 23547 | 46 | |
| 20319 | 47 | datatype fm = And fm fm | Or fm fm | Imp fm fm | Iff fm fm | NOT fm | At nat | 
| 48 | ||
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 49 | consts Ifm :: "fm \<Rightarrow> bool list \<Rightarrow> bool" | 
| 20319 | 50 | primrec | 
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 51 | "Ifm (At n) vs = vs!n" | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 52 | "Ifm (And p q) vs = (Ifm p vs \<and> Ifm q vs)" | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 53 | "Ifm (Or p q) vs = (Ifm p vs \<or> Ifm q vs)" | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 54 | "Ifm (Imp p q) vs = (Ifm p vs \<longrightarrow> Ifm q vs)" | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 55 | "Ifm (Iff p q) vs = (Ifm p vs = Ifm q vs)" | 
| 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 56 | "Ifm (NOT p) vs = (\<not> (Ifm p vs))" | 
| 20319 | 57 | |
| 23547 | 58 | lemma "Q \<longrightarrow> (D & F & ((~ D) & (~ F)))" | 
| 59 | apply (reify Ifm.simps) | |
| 60 | oops | |
| 61 | ||
| 62 |   text{* Method @{text reify} maps a bool to an fm. For this it needs the 
 | |
| 63 |   semantics of fm, i.e.\ the rewrite rules in @{text Ifm.simps}. *}
 | |
| 64 | ||
| 65 | ||
| 66 | (* You can also just pick up a subterm to reify \<dots> *) | |
| 67 | lemma "Q \<longrightarrow> (D & F & ((~ D) & (~ F)))" | |
| 68 | apply (reify Ifm.simps ("((~ D) & (~ F))"))
 | |
| 69 | oops | |
| 70 | ||
| 71 |   text{* Let's perform NNF. This is a version that tends to generate disjunctions *}
 | |
| 35419 | 72 | primrec fmsize :: "fm \<Rightarrow> nat" where | 
| 20319 | 73 | "fmsize (At n) = 1" | 
| 35419 | 74 | | "fmsize (NOT p) = 1 + fmsize p" | 
| 75 | | "fmsize (And p q) = 1 + fmsize p + fmsize q" | |
| 76 | | "fmsize (Or p q) = 1 + fmsize p + fmsize q" | |
| 77 | | "fmsize (Imp p q) = 2 + fmsize p + fmsize q" | |
| 78 | | "fmsize (Iff p q) = 2 + 2* fmsize p + 2* fmsize q" | |
| 79 | ||
| 80 | lemma [measure_function]: "is_measure fmsize" .. | |
| 20319 | 81 | |
| 35419 | 82 | fun nnf :: "fm \<Rightarrow> fm" | 
| 83 | where | |
| 20319 | 84 | "nnf (At n) = At n" | 
| 35419 | 85 | | "nnf (And p q) = And (nnf p) (nnf q)" | 
| 86 | | "nnf (Or p q) = Or (nnf p) (nnf q)" | |
| 87 | | "nnf (Imp p q) = Or (nnf (NOT p)) (nnf q)" | |
| 88 | | "nnf (Iff p q) = Or (And (nnf p) (nnf q)) (And (nnf (NOT p)) (nnf (NOT q)))" | |
| 89 | | "nnf (NOT (And p q)) = Or (nnf (NOT p)) (nnf (NOT q))" | |
| 90 | | "nnf (NOT (Or p q)) = And (nnf (NOT p)) (nnf (NOT q))" | |
| 91 | | "nnf (NOT (Imp p q)) = And (nnf p) (nnf (NOT q))" | |
| 92 | | "nnf (NOT (Iff p q)) = Or (And (nnf p) (nnf (NOT q))) (And (nnf (NOT p)) (nnf q))" | |
| 93 | | "nnf (NOT (NOT p)) = nnf p" | |
| 94 | | "nnf (NOT p) = NOT p" | |
| 20319 | 95 | |
| 20337 | 96 |   text{* The correctness theorem of nnf: it preserves the semantics of fm *}
 | 
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 97 | lemma nnf[reflection]: "Ifm (nnf p) vs = Ifm p vs" | 
| 20319 | 98 | by (induct p rule: nnf.induct) auto | 
| 99 | ||
| 20337 | 100 |   text{* Now let's perform NNF using our @{term nnf} function defined above. First to the whole subgoal. *}
 | 
| 20319 | 101 | lemma "(\<not> (A = B)) \<and> (B \<longrightarrow> (A \<noteq> (B | C \<and> (B \<longrightarrow> A | D)))) \<longrightarrow> A \<or> B \<and> D" | 
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 102 | apply (reflection Ifm.simps) | 
| 20319 | 103 | oops | 
| 104 | ||
| 20337 | 105 |   text{* Now we specify on which subterm it should be applied*}
 | 
| 20319 | 106 | lemma "(\<not> (A = B)) \<and> (B \<longrightarrow> (A \<noteq> (B | C \<and> (B \<longrightarrow> A | D)))) \<longrightarrow> A \<or> B \<and> D" | 
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 107 | apply (reflection Ifm.simps only: "(B | C \<and> (B \<longrightarrow> A | D))") | 
| 20319 | 108 | oops | 
| 109 | ||
| 110 | ||
| 20337 | 111 | (* Example 2 : Simple arithmetic formulae *) | 
| 20319 | 112 | |
| 20337 | 113 |   text{* The type @{text num} reflects linear expressions over natural number *}
 | 
| 20319 | 114 | datatype num = C nat | Add num num | Mul nat num | Var nat | CN nat nat num | 
| 115 | ||
| 20337 | 116 | text{* This is just technical to make recursive definitions easier. *}
 | 
| 35419 | 117 | primrec num_size :: "num \<Rightarrow> nat" | 
| 118 | where | |
| 20319 | 119 | "num_size (C c) = 1" | 
| 35419 | 120 | | "num_size (Var n) = 1" | 
| 121 | | "num_size (Add a b) = 1 + num_size a + num_size b" | |
| 122 | | "num_size (Mul c a) = 1 + num_size a" | |
| 123 | | "num_size (CN n c a) = 4 + num_size a " | |
| 20319 | 124 | |
| 20337 | 125 |   text{* The semantics of num *}
 | 
| 35419 | 126 | primrec Inum:: "num \<Rightarrow> nat list \<Rightarrow> nat" | 
| 127 | where | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 128 | Inum_C : "Inum (C i) vs = i" | 
| 35419 | 129 | | Inum_Var: "Inum (Var n) vs = vs!n" | 
| 130 | | Inum_Add: "Inum (Add s t) vs = Inum s vs + Inum t vs " | |
| 131 | | Inum_Mul: "Inum (Mul c t) vs = c * Inum t vs " | |
| 132 | | Inum_CN : "Inum (CN n c t) vs = c*(vs!n) + Inum t vs " | |
| 20319 | 133 | |
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 134 | |
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 135 |   text{* Let's reify some nat expressions \dots *}
 | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 136 | lemma "4 * (2*x + (y::nat)) + f a \<noteq> 0" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 137 |   apply (reify Inum.simps ("4 * (2*x + (y::nat)) + f a"))
 | 
| 20319 | 138 | oops | 
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 139 | text{* We're in a bad situation!! x, y and f a have been recongnized as a constants, which is correct but does not correspond to our intuition of the constructor C. It should encapsulate constants, i.e. numbers, i.e. numerals.*}
 | 
| 20337 | 140 | |
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 141 |   text{* So let's leave the @{text "Inum_C"} equation at the end and see what happens \dots*}
 | 
| 20319 | 142 | lemma "4 * (2*x + (y::nat)) \<noteq> 0" | 
| 143 |   apply (reify Inum_Var Inum_Add Inum_Mul Inum_CN Inum_C ("4 * (2*x + (y::nat))"))
 | |
| 144 | oops | |
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 145 | text{* Hmmm let's specialize @{text Inum_C} with numerals.*}
 | 
| 20319 | 146 | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 147 | lemma Inum_number: "Inum (C (number_of t)) vs = number_of t" by simp | 
| 20319 | 148 | lemmas Inum_eqs = Inum_Var Inum_Add Inum_Mul Inum_CN Inum_number | 
| 149 | ||
| 20337 | 150 |   text{* Second attempt *}
 | 
| 20319 | 151 | lemma "1 * (2*x + (y::nat)) \<noteq> 0" | 
| 152 |   apply (reify Inum_eqs ("1 * (2*x + (y::nat))"))
 | |
| 153 | oops | |
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 154 |   text{* That was fine, so let's try another one \dots *}
 | 
| 20319 | 155 | |
| 20337 | 156 | lemma "1 * (2* x + (y::nat) + 0 + 1) \<noteq> 0" | 
| 20319 | 157 |   apply (reify Inum_eqs ("1 * (2*x + (y::nat) + 0 + 1)"))
 | 
| 158 | oops | |
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 159 |   text{* Oh!! 0 is not a variable \dots\ Oh! 0 is not a @{text "number_of"} \dots\ thing. The same for 1. So let's add those equations too *}
 | 
| 20319 | 160 | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 161 | lemma Inum_01: "Inum (C 0) vs = 0" "Inum (C 1) vs = 1" "Inum (C(Suc n)) vs = Suc n" | 
| 20319 | 162 | by simp+ | 
| 20337 | 163 | |
| 20319 | 164 | lemmas Inum_eqs'= Inum_eqs Inum_01 | 
| 20337 | 165 | |
| 166 | text{* Third attempt: *}
 | |
| 20319 | 167 | |
| 168 | lemma "1 * (2*x + (y::nat) + 0 + 1) \<noteq> 0" | |
| 169 |   apply (reify Inum_eqs' ("1 * (2*x + (y::nat) + 0 + 1)"))
 | |
| 170 | oops | |
| 20337 | 171 | text{* Okay, let's try reflection. Some simplifications on num follow. You can skim until the main theorem @{text linum} *}
 | 
| 35419 | 172 | fun lin_add :: "num \<Rightarrow> num \<Rightarrow> num" | 
| 173 | where | |
| 174 | "lin_add (CN n1 c1 r1) (CN n2 c2 r2) = | |
| 20319 | 175 | (if n1=n2 then | 
| 176 | (let c = c1 + c2 | |
| 35419 | 177 | in (if c=0 then lin_add r1 r2 else CN n1 c (lin_add r1 r2))) | 
| 178 | else if n1 \<le> n2 then (CN n1 c1 (lin_add r1 (CN n2 c2 r2))) | |
| 179 | else (CN n2 c2 (lin_add (CN n1 c1 r1) r2)))" | |
| 180 | | "lin_add (CN n1 c1 r1) t = CN n1 c1 (lin_add r1 t)" | |
| 181 | | "lin_add t (CN n2 c2 r2) = CN n2 c2 (lin_add t r2)" | |
| 182 | | "lin_add (C b1) (C b2) = C (b1+b2)" | |
| 183 | | "lin_add a b = Add a b" | |
| 184 | lemma lin_add: "Inum (lin_add t s) bs = Inum (Add t s) bs" | |
| 20319 | 185 | apply (induct t s rule: lin_add.induct, simp_all add: Let_def) | 
| 186 | apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all) | |
| 29667 | 187 | by (case_tac "n1 = n2", simp_all add: algebra_simps) | 
| 20319 | 188 | |
| 35419 | 189 | fun lin_mul :: "num \<Rightarrow> nat \<Rightarrow> num" | 
| 190 | where | |
| 191 | "lin_mul (C j) i = C (i*j)" | |
| 192 | | "lin_mul (CN n c a) i = (if i=0 then (C 0) else CN n (i*c) (lin_mul a i))" | |
| 193 | | "lin_mul t i = (Mul i t)" | |
| 20319 | 194 | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 195 | lemma lin_mul: "Inum (lin_mul t i) bs = Inum (Mul i t) bs" | 
| 35419 | 196 | by (induct t i rule: lin_mul.induct, auto simp add: algebra_simps) | 
| 20319 | 197 | |
| 35419 | 198 | lemma [measure_function]: "is_measure num_size" .. | 
| 199 | ||
| 200 | fun linum:: "num \<Rightarrow> num" | |
| 201 | where | |
| 20319 | 202 | "linum (C b) = C b" | 
| 35419 | 203 | | "linum (Var n) = CN n 1 (C 0)" | 
| 204 | | "linum (Add t s) = lin_add (linum t) (linum s)" | |
| 205 | | "linum (Mul c t) = lin_mul (linum t) c" | |
| 206 | | "linum (CN n c t) = lin_add (linum (Mul c (Var n))) (linum t)" | |
| 20319 | 207 | |
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 208 | lemma linum[reflection] : "Inum (linum t) bs = Inum t bs" | 
| 20319 | 209 | by (induct t rule: linum.induct, simp_all add: lin_mul lin_add) | 
| 210 | ||
| 20337 | 211 |   text{* Now we can use linum to simplify nat terms using reflection *}
 | 
| 212 | lemma "(Suc (Suc 1)) * (x + (Suc 1)*y) = 3*x + 6*y" | |
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 213 | apply (reflection Inum_eqs' only: "(Suc (Suc 1)) * (x + (Suc 1)*y)") | 
| 20319 | 214 | oops | 
| 215 | ||
| 20337 | 216 |   text{* Let's lift this to formulae and see what happens *}
 | 
| 20319 | 217 | |
| 218 | datatype aform = Lt num num | Eq num num | Ge num num | NEq num num | | |
| 219 | Conj aform aform | Disj aform aform | NEG aform | T | F | |
| 35419 | 220 | |
| 221 | primrec linaformsize:: "aform \<Rightarrow> nat" | |
| 222 | where | |
| 20319 | 223 | "linaformsize T = 1" | 
| 35419 | 224 | | "linaformsize F = 1" | 
| 225 | | "linaformsize (Lt a b) = 1" | |
| 226 | | "linaformsize (Ge a b) = 1" | |
| 227 | | "linaformsize (Eq a b) = 1" | |
| 228 | | "linaformsize (NEq a b) = 1" | |
| 229 | | "linaformsize (NEG p) = 2 + linaformsize p" | |
| 230 | | "linaformsize (Conj p q) = 1 + linaformsize p + linaformsize q" | |
| 231 | | "linaformsize (Disj p q) = 1 + linaformsize p + linaformsize q" | |
| 20319 | 232 | |
| 35419 | 233 | lemma [measure_function]: "is_measure linaformsize" .. | 
| 20319 | 234 | |
| 35419 | 235 | primrec is_aform :: "aform => nat list => bool" | 
| 236 | where | |
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 237 | "is_aform T vs = True" | 
| 35419 | 238 | | "is_aform F vs = False" | 
| 239 | | "is_aform (Lt a b) vs = (Inum a vs < Inum b vs)" | |
| 240 | | "is_aform (Eq a b) vs = (Inum a vs = Inum b vs)" | |
| 241 | | "is_aform (Ge a b) vs = (Inum a vs \<ge> Inum b vs)" | |
| 242 | | "is_aform (NEq a b) vs = (Inum a vs \<noteq> Inum b vs)" | |
| 243 | | "is_aform (NEG p) vs = (\<not> (is_aform p vs))" | |
| 244 | | "is_aform (Conj p q) vs = (is_aform p vs \<and> is_aform q vs)" | |
| 245 | | "is_aform (Disj p q) vs = (is_aform p vs \<or> is_aform q vs)" | |
| 20319 | 246 | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 247 |   text{* Let's reify and do reflection *}
 | 
| 20319 | 248 | lemma "(3::nat)*x + t < 0 \<and> (2 * x + y \<noteq> 17)" | 
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 249 | apply (reify Inum_eqs' is_aform.simps) | 
| 20319 | 250 | oops | 
| 251 | ||
| 20337 | 252 | text{* Note that reification handles several interpretations at the same time*}
 | 
| 20319 | 253 | lemma "(3::nat)*x + t < 0 & x*x + t*x + 3 + 1 = z*t*4*z | x + x + 1 < 0" | 
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 254 | apply (reflection Inum_eqs' is_aform.simps only:"x + x + 1") | 
| 20319 | 255 | oops | 
| 256 | ||
| 20337 | 257 |   text{* For reflection we now define a simple transformation on aform: NNF + linum on atoms *}
 | 
| 35419 | 258 | |
| 259 | fun linaform:: "aform \<Rightarrow> aform" | |
| 260 | where | |
| 20319 | 261 | "linaform (Lt s t) = Lt (linum s) (linum t)" | 
| 35419 | 262 | | "linaform (Eq s t) = Eq (linum s) (linum t)" | 
| 263 | | "linaform (Ge s t) = Ge (linum s) (linum t)" | |
| 264 | | "linaform (NEq s t) = NEq (linum s) (linum t)" | |
| 265 | | "linaform (Conj p q) = Conj (linaform p) (linaform q)" | |
| 266 | | "linaform (Disj p q) = Disj (linaform p) (linaform q)" | |
| 267 | | "linaform (NEG T) = F" | |
| 268 | | "linaform (NEG F) = T" | |
| 269 | | "linaform (NEG (Lt a b)) = Ge a b" | |
| 270 | | "linaform (NEG (Ge a b)) = Lt a b" | |
| 271 | | "linaform (NEG (Eq a b)) = NEq a b" | |
| 272 | | "linaform (NEG (NEq a b)) = Eq a b" | |
| 273 | | "linaform (NEG (NEG p)) = linaform p" | |
| 274 | | "linaform (NEG (Conj p q)) = Disj (linaform (NEG p)) (linaform (NEG q))" | |
| 275 | | "linaform (NEG (Disj p q)) = Conj (linaform (NEG p)) (linaform (NEG q))" | |
| 276 | | "linaform p = p" | |
| 20319 | 277 | |
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 278 | lemma linaform: "is_aform (linaform p) vs = is_aform p vs" | 
| 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 279 | by (induct p rule: linaform.induct) (auto simp add: linum) | 
| 20319 | 280 | |
| 281 | lemma "(((Suc(Suc (Suc 0)))*((x::nat) + (Suc (Suc 0)))) + (Suc (Suc (Suc 0))) * ((Suc(Suc (Suc 0)))*((x::nat) + (Suc (Suc 0))))< 0) \<and> (Suc 0 + Suc 0< 0)" | |
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 282 | apply (reflection Inum_eqs' is_aform.simps rules: linaform) | 
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 283 | oops | 
| 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 284 | |
| 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 285 | declare linaform[reflection] | 
| 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 286 | lemma "(((Suc(Suc (Suc 0)))*((x::nat) + (Suc (Suc 0)))) + (Suc (Suc (Suc 0))) * ((Suc(Suc (Suc 0)))*((x::nat) + (Suc (Suc 0))))< 0) \<and> (Suc 0 + Suc 0< 0)" | 
| 28866 
30cd9d89a0fb
reactivated some dead theories (based on hints by Mark Hillebrand);
 wenzelm parents: 
23650diff
changeset | 287 | apply (reflection Inum_eqs' is_aform.simps) | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 288 | oops | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 289 | |
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 290 | text{* We now give an example where Interpretaions have 0 or more than only one envornement of different types and show that automatic reification also deals with binding *}
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 291 | datatype rb = BC bool| BAnd rb rb | BOr rb rb | 
| 35419 | 292 | primrec Irb :: "rb \<Rightarrow> bool" | 
| 293 | where | |
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 294 | "Irb (BC p) = p" | 
| 35419 | 295 | | "Irb (BAnd s t) = (Irb s \<and> Irb t)" | 
| 296 | | "Irb (BOr s t) = (Irb s \<or> Irb t)" | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 297 | |
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 298 | lemma "A \<and> (B \<or> D \<and> B) \<and> A \<and> (B \<or> D \<and> B) \<or> A \<and> (B \<or> D \<and> B) \<or> A \<and> (B \<or> D \<and> B)" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 299 | apply (reify Irb.simps) | 
| 20319 | 300 | oops | 
| 301 | ||
| 23650 
0a6a719d24d5
Show the use of reflection attribute; fixed comments about the order of absorbing equations : f (C x) = x ; now automaically tried as last;
 chaieb parents: 
23624diff
changeset | 302 | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 303 | datatype rint = IC int| IVar nat | IAdd rint rint | IMult rint rint | INeg rint | ISub rint rint | 
| 35419 | 304 | primrec Irint :: "rint \<Rightarrow> int list \<Rightarrow> int" | 
| 305 | where | |
| 306 | Irint_Var: "Irint (IVar n) vs = vs!n" | |
| 307 | | Irint_Neg: "Irint (INeg t) vs = - Irint t vs" | |
| 308 | | Irint_Add: "Irint (IAdd s t) vs = Irint s vs + Irint t vs" | |
| 309 | | Irint_Sub: "Irint (ISub s t) vs = Irint s vs - Irint t vs" | |
| 310 | | Irint_Mult: "Irint (IMult s t) vs = Irint s vs * Irint t vs" | |
| 311 | | Irint_C: "Irint (IC i) vs = i" | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 312 | lemma Irint_C0: "Irint (IC 0) vs = 0" | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 313 | by simp | 
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 314 | lemma Irint_C1: "Irint (IC 1) vs = 1" | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 315 | by simp | 
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 316 | lemma Irint_Cnumberof: "Irint (IC (number_of x)) vs = number_of x" | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 317 | by simp | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 318 | lemmas Irint_simps = Irint_Var Irint_Neg Irint_Add Irint_Sub Irint_Mult Irint_C0 Irint_C1 Irint_Cnumberof | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 319 | lemma "(3::int) * x + y*y - 9 + (- z) = 0" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 320 |   apply (reify Irint_simps ("(3::int) * x + y*y - 9 + (- z)"))
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 321 | oops | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 322 | datatype rlist = LVar nat| LEmpty| LCons rint rlist | LAppend rlist rlist | 
| 35419 | 323 | primrec Irlist :: "rlist \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> (int list)" | 
| 324 | where | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 325 | "Irlist (LEmpty) is vs = []" | 
| 35419 | 326 | | "Irlist (LVar n) is vs = vs!n" | 
| 327 | | "Irlist (LCons i t) is vs = ((Irint i is)#(Irlist t is vs))" | |
| 328 | | "Irlist (LAppend s t) is vs = (Irlist s is vs) @ (Irlist t is vs)" | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 329 | lemma "[(1::int)] = []" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 330 |   apply (reify Irlist.simps Irint_simps ("[1]:: int list"))
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 331 | oops | 
| 20374 | 332 | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 333 | lemma "([(3::int) * x + y*y - 9 + (- z)] @ []) @ xs = [y*y - z - 9 + (3::int) * x]" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 334 |   apply (reify Irlist.simps Irint_simps ("([(3::int) * x + y*y - 9 + (- z)] @ []) @ xs"))
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 335 | oops | 
| 20374 | 336 | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 337 | datatype rnat = NC nat| NVar nat| NSuc rnat | NAdd rnat rnat | NMult rnat rnat | NNeg rnat | NSub rnat rnat | Nlgth rlist | 
| 35419 | 338 | primrec Irnat :: "rnat \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> nat" | 
| 339 | where | |
| 340 | Irnat_Suc: "Irnat (NSuc t) is ls vs = Suc (Irnat t is ls vs)" | |
| 341 | | Irnat_Var: "Irnat (NVar n) is ls vs = vs!n" | |
| 342 | | Irnat_Neg: "Irnat (NNeg t) is ls vs = 0" | |
| 343 | | Irnat_Add: "Irnat (NAdd s t) is ls vs = Irnat s is ls vs + Irnat t is ls vs" | |
| 344 | | Irnat_Sub: "Irnat (NSub s t) is ls vs = Irnat s is ls vs - Irnat t is ls vs" | |
| 345 | | Irnat_Mult: "Irnat (NMult s t) is ls vs = Irnat s is ls vs * Irnat t is ls vs" | |
| 346 | | Irnat_lgth: "Irnat (Nlgth rxs) is ls vs = length (Irlist rxs is ls)" | |
| 347 | | Irnat_C: "Irnat (NC i) is ls vs = i" | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 348 | lemma Irnat_C0: "Irnat (NC 0) is ls vs = 0" | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 349 | by simp | 
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 350 | lemma Irnat_C1: "Irnat (NC 1) is ls vs = 1" | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 351 | by simp | 
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 352 | lemma Irnat_Cnumberof: "Irnat (NC (number_of x)) is ls vs = number_of x" | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 353 | by simp | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 354 | lemmas Irnat_simps = Irnat_Suc Irnat_Var Irnat_Neg Irnat_Add Irnat_Sub Irnat_Mult Irnat_lgth | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 355 | Irnat_C0 Irnat_C1 Irnat_Cnumberof | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 356 | lemma "(Suc n) * length (([(3::int) * x + y*y - 9 + (- z)] @ []) @ xs) = length xs" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 357 |   apply (reify Irnat_simps Irlist.simps Irint_simps ("(Suc n) *length (([(3::int) * x + y*y - 9 + (- z)] @ []) @ xs)"))
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 358 | oops | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 359 | datatype rifm = RT | RF | RVar nat | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 360 | | RNLT rnat rnat | RNILT rnat rint | RNEQ rnat rnat | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 361 | |RAnd rifm rifm | ROr rifm rifm | RImp rifm rifm| RIff rifm rifm | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 362 | | RNEX rifm | RIEX rifm| RLEX rifm | RNALL rifm | RIALL rifm| RLALL rifm | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 363 | | RBEX rifm | RBALL rifm | 
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 364 | |
| 35419 | 365 | primrec Irifm :: "rifm \<Rightarrow> bool list \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> bool" | 
| 366 | where | |
| 367 | "Irifm RT ps is ls ns = True" | |
| 368 | | "Irifm RF ps is ls ns = False" | |
| 369 | | "Irifm (RVar n) ps is ls ns = ps!n" | |
| 370 | | "Irifm (RNLT s t) ps is ls ns = (Irnat s is ls ns < Irnat t is ls ns)" | |
| 371 | | "Irifm (RNILT s t) ps is ls ns = (int (Irnat s is ls ns) < Irint t is)" | |
| 372 | | "Irifm (RNEQ s t) ps is ls ns = (Irnat s is ls ns = Irnat t is ls ns)" | |
| 373 | | "Irifm (RAnd p q) ps is ls ns = (Irifm p ps is ls ns \<and> Irifm q ps is ls ns)" | |
| 374 | | "Irifm (ROr p q) ps is ls ns = (Irifm p ps is ls ns \<or> Irifm q ps is ls ns)" | |
| 375 | | "Irifm (RImp p q) ps is ls ns = (Irifm p ps is ls ns \<longrightarrow> Irifm q ps is ls ns)" | |
| 376 | | "Irifm (RIff p q) ps is ls ns = (Irifm p ps is ls ns = Irifm q ps is ls ns)" | |
| 377 | | "Irifm (RNEX p) ps is ls ns = (\<exists>x. Irifm p ps is ls (x#ns))" | |
| 378 | | "Irifm (RIEX p) ps is ls ns = (\<exists>x. Irifm p ps (x#is) ls ns)" | |
| 379 | | "Irifm (RLEX p) ps is ls ns = (\<exists>x. Irifm p ps is (x#ls) ns)" | |
| 380 | | "Irifm (RBEX p) ps is ls ns = (\<exists>x. Irifm p (x#ps) is ls ns)" | |
| 381 | | "Irifm (RNALL p) ps is ls ns = (\<forall>x. Irifm p ps is ls (x#ns))" | |
| 382 | | "Irifm (RIALL p) ps is ls ns = (\<forall>x. Irifm p ps (x#is) ls ns)" | |
| 383 | | "Irifm (RLALL p) ps is ls ns = (\<forall>x. Irifm p ps is (x#ls) ns)" | |
| 384 | | "Irifm (RBALL p) ps is ls ns = (\<forall>x. Irifm p (x#ps) is ls ns)" | |
| 20374 | 385 | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 386 | lemma " \<forall>x. \<exists>n. ((Suc n) * length (([(3::int) * x + (f t)*y - 9 + (- z)] @ []) @ xs) = length xs) \<and> m < 5*n - length (xs @ [2,3,4,x*z + 8 - y]) \<longrightarrow> (\<exists>p. \<forall>q. p \<and> q \<longrightarrow> r)" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20503diff
changeset | 387 | apply (reify Irifm.simps Irnat_simps Irlist.simps Irint_simps) | 
| 20374 | 388 | oops | 
| 389 | ||
| 22199 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 390 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31021diff
changeset | 391 | (* An example for equations containing type variables *) | 
| 22199 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 392 | datatype prod = Zero | One | Var nat | Mul prod prod | 
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 393 | | Pw prod nat | PNM nat nat prod | 
| 35419 | 394 | primrec Iprod :: " prod \<Rightarrow> ('a::{linordered_idom}) list \<Rightarrow>'a" 
 | 
| 395 | where | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 396 | "Iprod Zero vs = 0" | 
| 35419 | 397 | | "Iprod One vs = 1" | 
| 398 | | "Iprod (Var n) vs = vs!n" | |
| 399 | | "Iprod (Mul a b) vs = (Iprod a vs * Iprod b vs)" | |
| 400 | | "Iprod (Pw a n) vs = ((Iprod a vs) ^ n)" | |
| 401 | | "Iprod (PNM n k t) vs = (vs ! n)^k * Iprod t vs" | |
| 22199 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 402 | consts prodmul:: "prod \<times> prod \<Rightarrow> prod" | 
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 403 | datatype sgn = Pos prod | Neg prod | ZeroEq prod | NZeroEq prod | Tr | F | 
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 404 | | Or sgn sgn | And sgn sgn | 
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 405 | |
| 35419 | 406 | primrec Isgn :: " sgn \<Rightarrow> ('a::{linordered_idom}) list \<Rightarrow>bool"
 | 
| 407 | where | |
| 23624 
82091387f6d7
The order for parameter for interpretation is now inversted:
 chaieb parents: 
23608diff
changeset | 408 | "Isgn Tr vs = True" | 
| 35419 | 409 | | "Isgn F vs = False" | 
| 410 | | "Isgn (ZeroEq t) vs = (Iprod t vs = 0)" | |
| 411 | | "Isgn (NZeroEq t) vs = (Iprod t vs \<noteq> 0)" | |
| 412 | | "Isgn (Pos t) vs = (Iprod t vs > 0)" | |
| 413 | | "Isgn (Neg t) vs = (Iprod t vs < 0)" | |
| 414 | | "Isgn (And p q) vs = (Isgn p vs \<and> Isgn q vs)" | |
| 415 | | "Isgn (Or p q) vs = (Isgn p vs \<or> Isgn q vs)" | |
| 22199 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 416 | |
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 417 | lemmas eqs = Isgn.simps Iprod.simps | 
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 418 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
32960diff
changeset | 419 | lemma "(x::'a::{linordered_idom})^4 * y * z * y^2 * z^23 > 0"
 | 
| 22199 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 420 | apply (reify eqs) | 
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 421 | oops | 
| 
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
 chaieb parents: 
20564diff
changeset | 422 | |
| 20319 | 423 | end |