src/HOL/ex/ReflectionEx.thy
author wenzelm
Sun, 03 Jun 2007 23:16:47 +0200
changeset 23219 87ad6e8a5f2c
parent 22199 b617ddd200eb
child 23477 f4b83f03cac9
permissions -rw-r--r--
tuned document;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     1
(*
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     3
    Author:     Amine Chaieb, TU Muenchen
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     4
*)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     5
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     6
header {* Examples for generic reflection and reification *}
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     7
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     8
theory ReflectionEx
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     9
imports Reflection
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    10
begin
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    11
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    12
text{* This theory presents two methods: reify and reflection *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    13
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    14
text{* 
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    15
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> 
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    16
In order to implement a simplification on terms of type 'a we often need its structure.
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    17
Traditionnaly such simplifications are written in ML, proofs are synthesized.
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    18
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.
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    19
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    20
NB: All the interpretations supported by @{text reify} must have the type @{text "'b list \<Rightarrow> tau \<Rightarrow> 'a"}.
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    21
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.
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    22
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    23
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'"}.
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    24
*}
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    25
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    26
text{* Example 1 : Propositional formulae and NNF.*}
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    27
text{* The type @{text fm} represents simple propositional formulae: *}
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    28
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    29
datatype fm = And fm fm | Or fm fm | Imp fm fm | Iff fm fm | NOT fm | At nat
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    30
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    31
consts Ifm :: "bool list \<Rightarrow> fm \<Rightarrow> bool"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    32
primrec
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    33
  "Ifm vs (At n) = vs!n"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    34
  "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    35
  "Ifm vs (Or p q) = (Ifm vs p \<or> Ifm vs q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    36
  "Ifm vs (Imp p q) = (Ifm vs p \<longrightarrow> Ifm vs q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    37
  "Ifm vs (Iff p q) = (Ifm vs p = Ifm vs q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    38
  "Ifm vs (NOT p) = (\<not> (Ifm vs p))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    39
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    40
consts fmsize :: "fm \<Rightarrow> nat"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    41
primrec
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    42
  "fmsize (At n) = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    43
  "fmsize (NOT p) = 1 + fmsize p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    44
  "fmsize (And p q) = 1 + fmsize p + fmsize q"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    45
  "fmsize (Or p q) = 1 + fmsize p + fmsize q"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    46
  "fmsize (Imp p q) = 2 + fmsize p + fmsize q"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    47
  "fmsize (Iff p q) = 2 + 2* fmsize p + 2* fmsize q"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    48
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    49
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    50
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    51
  text{* Method @{text reify} maps a bool to an fm. For this it needs the 
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    52
  semantics of fm, i.e.\ the rewrite rules in @{text Ifm.simps}. *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    53
lemma "Q \<longrightarrow> (D & F & ((~ D) & (~ F)))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    54
apply (reify Ifm.simps)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    55
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    56
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    57
  (* You can also just pick up a subterm to reify \<dots> *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    58
lemma "Q \<longrightarrow> (D & F & ((~ D) & (~ F)))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    59
apply (reify Ifm.simps ("((~ D) & (~ F))"))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    60
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    61
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    62
  text{* Let's perform NNF. This is a version that tends to generate disjunctions *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    63
consts nnf :: "fm \<Rightarrow> fm"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    64
recdef nnf "measure fmsize"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    65
  "nnf (At n) = At n"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    66
  "nnf (And p q) = And (nnf p) (nnf q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    67
  "nnf (Or p q) = Or (nnf p) (nnf q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    68
  "nnf (Imp p q) = Or (nnf (NOT p)) (nnf q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    69
  "nnf (Iff p q) = Or (And (nnf p) (nnf q)) (And (nnf (NOT p)) (nnf (NOT q)))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    70
  "nnf (NOT (And p q)) = Or (nnf (NOT p)) (nnf (NOT q))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    71
  "nnf (NOT (Or p q)) = And (nnf (NOT p)) (nnf (NOT q))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    72
  "nnf (NOT (Imp p q)) = And (nnf p) (nnf (NOT q))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    73
  "nnf (NOT (Iff p q)) = Or (And (nnf p) (nnf (NOT q))) (And (nnf (NOT p)) (nnf q))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    74
  "nnf (NOT (NOT p)) = nnf p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    75
  "nnf (NOT p) = NOT p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    76
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    77
  text{* The correctness theorem of nnf: it preserves the semantics of fm *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    78
lemma nnf: "Ifm vs (nnf p) = Ifm vs p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    79
  by (induct p rule: nnf.induct) auto
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    80
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    81
  text{* Now let's perform NNF using our @{term nnf} function defined above. First to the whole subgoal. *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    82
lemma "(\<not> (A = B)) \<and> (B \<longrightarrow> (A \<noteq> (B | C \<and> (B \<longrightarrow> A | D)))) \<longrightarrow> A \<or> B \<and> D"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    83
apply (reflection nnf Ifm.simps)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    84
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    85
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    86
  text{* Now we specify on which subterm it should be applied*}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    87
lemma "(\<not> (A = B)) \<and> (B \<longrightarrow> (A \<noteq> (B | C \<and> (B \<longrightarrow> A | D)))) \<longrightarrow> A \<or> B \<and> D"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    88
apply (reflection nnf Ifm.simps ("(B | C \<and> (B \<longrightarrow> A | D))"))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    89
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    90
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    91
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    92
  (* Example 2 : Simple arithmetic formulae *)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    93
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    94
  text{* The type @{text num} reflects linear expressions over natural number *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    95
datatype num = C nat | Add num num | Mul nat num | Var nat | CN nat nat num
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    96
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    97
text{* This is just technical to make recursive definitions easier. *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    98
consts num_size :: "num \<Rightarrow> nat" 
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    99
primrec 
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   100
  "num_size (C c) = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   101
  "num_size (Var n) = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   102
  "num_size (Add a b) = 1 + num_size a + num_size b"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   103
  "num_size (Mul c a) = 1 + num_size a"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   104
  "num_size (CN n c a) = 4 + num_size a "
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   105
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   106
  text{* The semantics of num *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   107
consts Inum:: "nat list \<Rightarrow> num \<Rightarrow> nat"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   108
primrec 
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   109
  Inum_C  : "Inum vs (C i) = i"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   110
  Inum_Var: "Inum vs (Var n) = vs!n"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   111
  Inum_Add: "Inum vs (Add s t) = Inum vs s + Inum vs t"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   112
  Inum_Mul: "Inum vs (Mul c t) = c * Inum vs t"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   113
  Inum_CN : "Inum vs (CN n c t) = c*(vs!n) + Inum vs t"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   114
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   115
  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: 20503
diff changeset
   116
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: 20503
diff changeset
   117
  apply (reify Inum.simps ("4 * (2*x + (y::nat)) + f a"))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   118
oops
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   119
text{* We're in a bad situation!! The term above has been recongnized as a constant, which is correct but does not correspond to our intuition of the constructor C. It should encapsulate constants, i.e. numbers, i.e. numerals.*}
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   120
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   121
  text{* So let's leave the Inum_C equation at the end and see what happens \<dots>*}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   122
lemma "4 * (2*x + (y::nat)) \<noteq> 0"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   123
  apply (reify Inum_Var Inum_Add Inum_Mul Inum_CN Inum_C ("4 * (2*x + (y::nat))"))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   124
oops
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   125
text{* Better, but it still reifies @{term x} to @{term "C x"}. Note that the reification depends on the order of the equations. The problem is that the right hand side of @{thm Inum_C} matches any term of type nat, which makes things bad. We want only numerals to match\<dots> So let's specialize @{text Inum_C} with numerals.*}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   126
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   127
lemma Inum_number: "Inum vs (C (number_of t)) = number_of t" by simp
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   128
lemmas Inum_eqs = Inum_Var Inum_Add Inum_Mul Inum_CN Inum_number
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   129
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   130
  text{* Second attempt *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   131
lemma "1 * (2*x + (y::nat)) \<noteq> 0"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   132
  apply (reify Inum_eqs ("1 * (2*x + (y::nat))"))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   133
oops
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   134
  text{* That was fine, so let's try an other one\<dots> *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   135
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   136
lemma "1 * (2* x + (y::nat) + 0 + 1) \<noteq> 0"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   137
  apply (reify Inum_eqs ("1 * (2*x + (y::nat) + 0 + 1)"))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   138
oops
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   139
  text{* Oh!! 0 is not a variable \<dots> Oh! 0 is not a number_of .. thing. The same for 1. So let's add those equations too *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   140
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   141
lemma Inum_01: "Inum vs (C 0) = 0" "Inum vs (C 1) = 1" "Inum vs (C(Suc n)) = Suc n"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   142
  by simp+
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   143
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   144
lemmas Inum_eqs'= Inum_eqs Inum_01
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   145
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   146
text{* Third attempt: *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   147
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   148
lemma "1 * (2*x + (y::nat) + 0 + 1) \<noteq> 0"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   149
  apply (reify Inum_eqs' ("1 * (2*x + (y::nat) + 0 + 1)"))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   150
oops
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   151
text{* Okay, let's try reflection. Some simplifications on num follow. You can skim until the main theorem @{text linum} *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   152
consts lin_add :: "num \<times> num \<Rightarrow> num"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   153
recdef lin_add "measure (\<lambda>(x,y). ((size x) + (size y)))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   154
  "lin_add (CN n1 c1 r1,CN n2 c2 r2) =
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   155
  (if n1=n2 then 
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   156
  (let c = c1 + c2
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   157
  in (if c=0 then lin_add(r1,r2) else CN n1 c (lin_add (r1,r2))))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   158
  else if n1 \<le> n2 then (CN n1 c1 (lin_add (r1,CN n2 c2 r2))) 
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   159
  else (CN n2 c2 (lin_add (CN n1 c1 r1,r2))))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   160
  "lin_add (CN n1 c1 r1,t) = CN n1 c1 (lin_add (r1, t))"  
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   161
  "lin_add (t,CN n2 c2 r2) = CN n2 c2 (lin_add (t,r2))" 
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   162
  "lin_add (C b1, C b2) = C (b1+b2)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   163
  "lin_add (a,b) = Add a b"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   164
lemma lin_add: "Inum bs (lin_add (t,s)) = Inum bs (Add t s)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   165
apply (induct t s rule: lin_add.induct, simp_all add: Let_def)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   166
apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   167
by (case_tac "n1 = n2", simp_all add: ring_eq_simps)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   168
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   169
consts lin_mul :: "num \<Rightarrow> nat \<Rightarrow> num"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   170
recdef lin_mul "measure size "
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   171
  "lin_mul (C j) = (\<lambda> i. C (i*j))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   172
  "lin_mul (CN n c a) = (\<lambda> i. if i=0 then (C 0) else CN n (i*c) (lin_mul a i))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   173
  "lin_mul t = (\<lambda> i. Mul i t)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   174
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   175
lemma lin_mul: "Inum bs (lin_mul t i) = Inum bs (Mul i 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: 20503
diff changeset
   176
by (induct t arbitrary: i rule: lin_mul.induct, auto simp add: ring_eq_simps)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   177
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   178
consts linum:: "num \<Rightarrow> num"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   179
recdef linum "measure num_size"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   180
  "linum (C b) = C b"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   181
  "linum (Var n) = CN n 1 (C 0)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   182
  "linum (Add t s) = lin_add (linum t, linum s)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   183
  "linum (Mul c t) = lin_mul (linum t) c"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   184
  "linum (CN n c t) = lin_add (linum (Mul c (Var n)),linum t)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   185
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   186
lemma linum : "Inum vs (linum t) = Inum vs t"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   187
by (induct t rule: linum.induct, simp_all add: lin_mul lin_add)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   188
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   189
  text{* Now we can use linum to simplify nat terms using reflection *}
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   190
lemma "(Suc (Suc 1)) * (x + (Suc 1)*y) = 3*x + 6*y"
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: 20503
diff changeset
   191
(* apply (reflection linum Inum_eqs' ("(Suc (Suc 1)) * (x + (Suc 1)*y)")) *)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   192
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   193
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   194
  text{* Let's lift this to formulae and see what happens *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   195
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   196
datatype aform = Lt num num  | Eq num num | Ge num num | NEq num num | 
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   197
  Conj aform aform | Disj aform aform | NEG aform | T | F
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   198
consts linaformsize:: "aform \<Rightarrow> nat"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   199
recdef linaformsize "measure size"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   200
  "linaformsize T = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   201
  "linaformsize F = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   202
  "linaformsize (Lt a b) = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   203
  "linaformsize (Ge a b) = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   204
  "linaformsize (Eq a b) = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   205
  "linaformsize (NEq a b) = 1"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   206
  "linaformsize (NEG p) = 2 + linaformsize p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   207
  "linaformsize (Conj p q) = 1 + linaformsize p + linaformsize q"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   208
  "linaformsize (Disj p q) = 1 + linaformsize p + linaformsize q"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   209
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   210
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   211
consts aform :: "nat list => aform => bool"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   212
primrec
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   213
  "aform vs T = True"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   214
  "aform vs F = False"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   215
  "aform vs (Lt a b) = (Inum vs a < Inum vs b)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   216
  "aform vs (Eq a b) = (Inum vs a = Inum vs b)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   217
  "aform vs (Ge a b) = (Inum vs a \<ge> Inum vs b)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   218
  "aform vs (NEq a b) = (Inum vs a \<noteq> Inum vs b)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   219
  "aform vs (NEG p) = (\<not> (aform vs p))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   220
  "aform vs (Conj p q) = (aform vs p \<and> aform vs q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   221
  "aform vs (Disj p q) = (aform vs p \<or> aform vs q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   222
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: 20503
diff changeset
   223
  text{* Let's reify and do reflection *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   224
lemma "(3::nat)*x + t < 0 \<and> (2 * x + y \<noteq> 17)"
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: 20503
diff changeset
   225
(* apply (reify Inum_eqs' aform.simps) *)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   226
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   227
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   228
text{* Note that reification handles several interpretations at the same time*}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   229
lemma "(3::nat)*x + t < 0 & x*x + t*x + 3 + 1 = z*t*4*z | x + x + 1 < 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: 20503
diff changeset
   230
(* apply (reflection linum Inum_eqs' aform.simps ("x + x + 1")) *)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   231
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   232
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   233
  text{* For reflection we now define a simple transformation on aform: NNF + linum on atoms *}
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   234
consts linaform:: "aform \<Rightarrow> aform"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   235
recdef linaform "measure linaformsize"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   236
  "linaform (Lt s t) = Lt (linum s) (linum t)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   237
  "linaform (Eq s t) = Eq (linum s) (linum t)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   238
  "linaform (Ge s t) = Ge (linum s) (linum t)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   239
  "linaform (NEq s t) = NEq (linum s) (linum t)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   240
  "linaform (Conj p q) = Conj (linaform p) (linaform q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   241
  "linaform (Disj p q) = Disj (linaform p) (linaform q)"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   242
  "linaform (NEG T) = F"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   243
  "linaform (NEG F) = T"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   244
  "linaform (NEG (Lt a b)) = Ge a b"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   245
  "linaform (NEG (Ge a b)) = Lt a b"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   246
  "linaform (NEG (Eq a b)) = NEq a b"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   247
  "linaform (NEG (NEq a b)) = Eq a b"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   248
  "linaform (NEG (NEG p)) = linaform p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   249
  "linaform (NEG (Conj p q)) = Disj (linaform (NEG p)) (linaform (NEG q))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   250
  "linaform (NEG (Disj p q)) = Conj (linaform (NEG p)) (linaform (NEG q))"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   251
  "linaform p = p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   252
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   253
lemma linaform: "aform vs (linaform p) = aform vs p"
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   254
  by (induct p rule: linaform.induct, auto simp add: linum)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   255
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   256
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)"
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: 20503
diff changeset
   257
(*   apply (reflection linaform Inum_eqs' aform.simps)  *)
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: 20503
diff changeset
   258
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: 20503
diff changeset
   259
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: 20503
diff changeset
   260
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: 20503
diff changeset
   261
datatype rb = BC bool| BAnd rb rb | BOr rb rb
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: 20503
diff changeset
   262
consts Irb :: "rb \<Rightarrow> bool"
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: 20503
diff changeset
   263
primrec
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: 20503
diff changeset
   264
  "Irb (BAnd s t) = (Irb s \<and> Irb t)"
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: 20503
diff changeset
   265
  "Irb (BOr s t) = (Irb s \<or> Irb t)"
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: 20503
diff changeset
   266
  "Irb (BC p) = p"
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: 20503
diff changeset
   267
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: 20503
diff changeset
   268
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: 20503
diff changeset
   269
apply (reify Irb.simps)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   270
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   271
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: 20503
diff changeset
   272
datatype rint = IC int| IVar nat | IAdd rint rint | IMult rint rint | INeg rint | ISub rint rint
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: 20503
diff changeset
   273
consts Irint :: "int list \<Rightarrow> rint \<Rightarrow> 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: 20503
diff changeset
   274
primrec
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: 20503
diff changeset
   275
Irint_Var: "Irint vs (IVar n) = vs!n"
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: 20503
diff changeset
   276
Irint_Neg: "Irint vs (INeg t) = - Irint vs t"
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: 20503
diff changeset
   277
Irint_Add: "Irint vs (IAdd s t) = Irint vs s + Irint vs t"
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: 20503
diff changeset
   278
Irint_Sub: "Irint vs (ISub s t) = Irint vs s - Irint vs t"
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: 20503
diff changeset
   279
Irint_Mult: "Irint vs (IMult s t) = Irint vs s * Irint vs t"
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: 20503
diff changeset
   280
Irint_C: "Irint vs (IC i) = i"
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: 20503
diff changeset
   281
lemma Irint_C0: "Irint vs (IC 0) = 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: 20503
diff changeset
   282
  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: 20503
diff changeset
   283
lemma Irint_C1: "Irint vs (IC 1) = 1"
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: 20503
diff changeset
   284
  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: 20503
diff changeset
   285
lemma Irint_Cnumberof: "Irint vs (IC (number_of x)) = number_of 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: 20503
diff changeset
   286
  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: 20503
diff changeset
   287
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: 20503
diff changeset
   288
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: 20503
diff changeset
   289
  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: 20503
diff changeset
   290
  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: 20503
diff changeset
   291
datatype rlist = LVar nat| LEmpty| LCons rint rlist | LAppend rlist rlist
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: 20503
diff changeset
   292
consts Irlist :: "int list \<Rightarrow> (int list) list \<Rightarrow> rlist \<Rightarrow> (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: 20503
diff changeset
   293
primrec
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: 20503
diff changeset
   294
  "Irlist is vs (LEmpty) = []"
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: 20503
diff changeset
   295
  "Irlist is vs (LVar n) = vs!n"
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: 20503
diff changeset
   296
  "Irlist is vs (LCons i t) = ((Irint is i)#(Irlist is vs t))"
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: 20503
diff changeset
   297
  "Irlist is vs (LAppend s t) = (Irlist is vs s) @ (Irlist is vs t)"
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: 20503
diff changeset
   298
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: 20503
diff changeset
   299
  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: 20503
diff changeset
   300
  oops
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   301
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: 20503
diff changeset
   302
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: 20503
diff changeset
   303
  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: 20503
diff changeset
   304
  oops
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   305
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: 20503
diff changeset
   306
datatype rnat = NC nat| NVar nat| NSuc rnat | NAdd rnat rnat | NMult rnat rnat | NNeg rnat | NSub rnat rnat | Nlgth rlist
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: 20503
diff changeset
   307
consts Irnat :: "int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> rnat \<Rightarrow> nat"
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   308
primrec
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: 20503
diff changeset
   309
Irnat_Suc: "Irnat is ls vs (NSuc t) = Suc (Irnat is ls vs t)"
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: 20503
diff changeset
   310
Irnat_Var: "Irnat is ls vs (NVar n) = vs!n"
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: 20503
diff changeset
   311
Irnat_Neg: "Irnat is ls vs (NNeg t) = - Irnat is ls vs t"
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: 20503
diff changeset
   312
Irnat_Add: "Irnat is ls vs (NAdd s t) = Irnat is ls vs s + Irnat is ls vs t"
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: 20503
diff changeset
   313
Irnat_Sub: "Irnat is ls vs (NSub s t) = Irnat is ls vs s - Irnat is ls vs t"
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: 20503
diff changeset
   314
Irnat_Mult: "Irnat is ls vs (NMult s t) = Irnat is ls vs s * Irnat is ls vs t"
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: 20503
diff changeset
   315
Irnat_lgth: "Irnat is ls vs (Nlgth rxs) = length (Irlist is ls rxs)"
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: 20503
diff changeset
   316
Irnat_C: "Irnat is ls vs (NC i) = i"
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: 20503
diff changeset
   317
lemma Irnat_C0: "Irnat is ls vs (NC 0) = 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: 20503
diff changeset
   318
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: 20503
diff changeset
   319
lemma Irnat_C1: "Irnat is ls vs (NC 1) = 1"
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: 20503
diff changeset
   320
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: 20503
diff changeset
   321
lemma Irnat_Cnumberof: "Irnat is ls vs (NC (number_of x)) = number_of 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: 20503
diff changeset
   322
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: 20503
diff changeset
   323
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: 20503
diff changeset
   324
  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: 20503
diff changeset
   325
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: 20503
diff changeset
   326
  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: 20503
diff changeset
   327
  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: 20503
diff changeset
   328
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: 20503
diff changeset
   329
  | 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: 20503
diff changeset
   330
  |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: 20503
diff changeset
   331
  | 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: 20503
diff changeset
   332
  | RBEX rifm | RBALL 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: 20503
diff changeset
   333
consts Irifm :: "bool list \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> rifm \<Rightarrow> bool"
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: 20503
diff changeset
   334
primrec
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: 20503
diff changeset
   335
"Irifm ps is ls ns RT = True"
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: 20503
diff changeset
   336
"Irifm ps is ls ns RF = False"
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: 20503
diff changeset
   337
  "Irifm ps is ls ns (RVar n) = ps!n"
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: 20503
diff changeset
   338
"Irifm ps is ls ns (RNLT s t) = (Irnat is ls ns s < Irnat is ls ns t)"
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: 20503
diff changeset
   339
"Irifm ps is ls ns (RNILT s t) = (int (Irnat is ls ns s) < Irint is t)"
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: 20503
diff changeset
   340
"Irifm ps is ls ns (RNEQ s t) = (Irnat is ls ns s = Irnat is ls ns t)"
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: 20503
diff changeset
   341
"Irifm ps is ls ns (RAnd p q) = (Irifm ps is ls ns p \<and> Irifm ps is ls ns q)"
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: 20503
diff changeset
   342
"Irifm ps is ls ns (ROr p q) = (Irifm ps is ls ns p \<or> Irifm ps is ls ns q)"
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: 20503
diff changeset
   343
"Irifm ps is ls ns (RImp p q) = (Irifm ps is ls ns p \<longrightarrow> Irifm ps is ls ns q)"
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: 20503
diff changeset
   344
"Irifm ps is ls ns (RIff p q) = (Irifm ps is ls ns p = Irifm ps is ls ns q)"
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: 20503
diff changeset
   345
"Irifm ps is ls ns (RNEX p) =  (\<exists>x. Irifm ps is ls (x#ns) p)"
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: 20503
diff changeset
   346
"Irifm ps is ls ns (RIEX p) =  (\<exists>x. Irifm ps (x#is) ls ns p)"
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: 20503
diff changeset
   347
"Irifm ps is ls ns (RLEX p) =  (\<exists>x. Irifm ps is (x#ls) ns p)"
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: 20503
diff changeset
   348
"Irifm ps is ls ns (RBEX p) =  (\<exists>x. Irifm (x#ps) is ls ns p)"
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: 20503
diff changeset
   349
"Irifm ps is ls ns (RNALL p) = (\<forall>x. Irifm ps is ls (x#ns) p)"
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: 20503
diff changeset
   350
"Irifm ps is ls ns (RIALL p) = (\<forall>x. Irifm ps (x#is) ls ns p)"
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: 20503
diff changeset
   351
"Irifm ps is ls ns (RLALL p) = (\<forall>x. Irifm ps is (x#ls) ns p)"
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: 20503
diff changeset
   352
"Irifm ps is ls ns (RBALL p) = (\<forall>x. Irifm (x#ps) is ls ns p)"
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   353
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: 20503
diff changeset
   354
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: 20503
diff changeset
   355
  apply (reify Irifm.simps Irnat_simps Irlist.simps Irint_simps)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   356
oops
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   357
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   358
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   359
  (* An example for equations containing type variables *)	
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   360
datatype prod = Zero | One | Var nat | Mul prod prod 
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   361
  | Pw prod nat | PNM nat nat prod
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   362
consts Iprod :: "('a::{ordered_idom,recpower}) list \<Rightarrow> prod \<Rightarrow> 'a" 
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   363
primrec
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   364
  "Iprod vs Zero = 0"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   365
  "Iprod vs One = 1"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   366
  "Iprod vs (Var n) = vs!n"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   367
  "Iprod vs (Mul a b) = (Iprod vs a * Iprod vs b)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   368
  "Iprod vs (Pw a n) = ((Iprod vs a) ^ n)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   369
  "Iprod vs (PNM n k t) = (vs ! n)^k * Iprod vs t"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   370
consts prodmul:: "prod \<times> prod \<Rightarrow> prod"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   371
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: 20564
diff changeset
   372
  | Or sgn sgn | And sgn sgn
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   373
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   374
consts Isgn :: "('a::{ordered_idom, recpower}) list \<Rightarrow> sgn \<Rightarrow> bool"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   375
primrec 
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   376
  "Isgn vs Tr = True"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   377
  "Isgn vs F = False"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   378
  "Isgn vs (ZeroEq t) = (Iprod vs t = 0)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   379
  "Isgn vs (NZeroEq t) = (Iprod vs t \<noteq> 0)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   380
  "Isgn vs (Pos t) = (Iprod vs t > 0)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   381
  "Isgn vs (Neg t) = (Iprod vs t < 0)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   382
  "Isgn vs (And p q) = (Isgn vs p \<and> Isgn vs q)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   383
  "Isgn vs (Or p q) = (Isgn vs p \<or> Isgn vs q)"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   384
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   385
lemmas eqs = Isgn.simps Iprod.simps
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   386
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   387
lemma "(x::int)^4 * y * z * y^2 * z^23 > 0"
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   388
  apply (reify eqs)
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   389
  oops
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   390
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   391
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   392
end