src/HOL/ex/Reflection_Examples.thy
author haftmann
Mon, 02 Aug 2021 10:01:06 +0000
changeset 74101 d804e93ae9ff
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
moved theory Bit_Operations into Main corpus
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51240
a7a04b449e8b updated headers;
wenzelm
parents: 51144
diff changeset
     1
(*  Title:      HOL/ex/Reflection_Examples.thy
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     3
*)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     4
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     5
section \<open>Examples for generic reflection and reification\<close>
29650
cc3958d31b1d Reflection.thy now in HOL/Library
haftmann
parents: 28866
diff changeset
     6
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
     7
theory Reflection_Examples
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 61933
diff changeset
     8
imports Complex_Main "HOL-Library.Reflection"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     9
begin
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    10
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    11
text \<open>This theory presents two methods: reify and reflection\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    12
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    13
text \<open>
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    14
Consider an HOL type \<open>\<sigma>\<close>, the structure of which is not recongnisable
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    15
on the theory level.  This is the case of \<^typ>\<open>bool\<close>, arithmetical terms such as \<^typ>\<open>int\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    16
\<^typ>\<open>real\<close> etc \dots  In order to implement a simplification on terms of type \<open>\<sigma>\<close> we
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    17
often need its structure.  Traditionnaly such simplifications are written in ML,
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    18
proofs are synthesized.
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    19
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    20
An other strategy is to declare an HOL datatype \<open>\<tau>\<close> and an HOL function (the
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    21
interpretation) that maps elements of \<open>\<tau>\<close> to elements of \<open>\<sigma>\<close>.
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    22
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    23
The functionality of \<open>reify\<close> then is, given a term \<open>t\<close> of type \<open>\<sigma>\<close>,
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    24
to compute a term \<open>s\<close> of type \<open>\<tau>\<close>.  For this it needs equations for the
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    25
interpretation.
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    26
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    27
N.B: All the interpretations supported by \<open>reify\<close> must have the type
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    28
\<open>'a list \<Rightarrow> \<tau> \<Rightarrow> \<sigma>\<close>.  The method \<open>reify\<close> can also be told which subterm
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    29
of the current subgoal should be reified.  The general call for \<open>reify\<close> is
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    30
\<open>reify eqs (t)\<close>, where \<open>eqs\<close> are the defining equations of the interpretation
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    31
and \<open>(t)\<close> is an optional parameter which specifies the subterm to which reification
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    32
should be applied to.  If \<open>(t)\<close> is abscent, \<open>reify\<close> tries to reify the whole
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    33
subgoal.
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    34
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    35
The method \<open>reflection\<close> uses \<open>reify\<close> and has a very similar signature:
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    36
\<open>reflection corr_thm eqs (t)\<close>.  Here again \<open>eqs\<close> and \<open>(t)\<close>
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    37
are as described above and \<open>corr_thm\<close> is a theorem proving
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    38
\<^prop>\<open>I vs (f t) = I vs t\<close>.  We assume that \<open>I\<close> is the interpretation
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    39
and \<open>f\<close> is some useful and executable simplification of type \<open>\<tau> \<Rightarrow> \<tau>\<close>.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    40
The method \<open>reflection\<close> applies reification and hence the theorem \<^prop>\<open>t = I xs s\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    41
and hence using \<open>corr_thm\<close> derives \<^prop>\<open>t = I xs (f s)\<close>.  It then uses
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    42
normalization by equational rewriting to prove \<^prop>\<open>f s = s'\<close> which almost finishes
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    43
the proof of \<^prop>\<open>t = t'\<close> where \<^prop>\<open>I xs s' = t'\<close>.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    44
\<close>
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    45
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    46
text \<open>Example 1 : Propositional formulae and NNF.\<close>
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    47
text \<open>The type \<open>fm\<close> represents simple propositional formulae:\<close>
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
    48
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
    49
datatype form = TrueF | FalseF | Less nat nat
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    50
  | And form form | Or form form | Neg form | ExQ form
23547
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    51
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    52
primrec interp :: "form \<Rightarrow> ('a::ord) list \<Rightarrow> bool"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    53
where
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    54
  "interp TrueF vs \<longleftrightarrow> True"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    55
| "interp FalseF vs \<longleftrightarrow> False"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    56
| "interp (Less i j) vs \<longleftrightarrow> vs ! i < vs ! j"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    57
| "interp (And f1 f2) vs \<longleftrightarrow> interp f1 vs \<and> interp f2 vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    58
| "interp (Or f1 f2) vs \<longleftrightarrow> interp f1 vs \<or> interp f2 vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    59
| "interp (Neg f) vs \<longleftrightarrow> \<not> interp f vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    60
| "interp (ExQ f) vs \<longleftrightarrow> (\<exists>v. interp f (v # vs))"
23547
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    61
23608
e65e36dbe0d2 Some examples for reifying type variables
chaieb
parents: 23547
diff changeset
    62
lemmas interp_reify_eqs = interp.simps
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    63
declare interp_reify_eqs [reify]
23547
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    64
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    65
lemma "\<exists>x. x < y \<and> x < z"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    66
  apply reify
23608
e65e36dbe0d2 Some examples for reifying type variables
chaieb
parents: 23547
diff changeset
    67
  oops
23547
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    68
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
    69
datatype fm = And fm fm | Or fm fm | Imp fm fm | Iff fm fm | Not fm | At nat
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    70
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    71
primrec Ifm :: "fm \<Rightarrow> bool list \<Rightarrow> bool"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    72
where
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    73
  "Ifm (At n) vs \<longleftrightarrow> vs ! n"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    74
| "Ifm (And p q) vs \<longleftrightarrow> Ifm p vs \<and> Ifm q vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    75
| "Ifm (Or p q) vs \<longleftrightarrow> Ifm p vs \<or> Ifm q vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    76
| "Ifm (Imp p q) vs \<longleftrightarrow> Ifm p vs \<longrightarrow> Ifm q vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    77
| "Ifm (Iff p q) vs \<longleftrightarrow> Ifm p vs = Ifm q vs"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
    78
| "Ifm (Not p) vs \<longleftrightarrow> \<not> Ifm p vs"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    79
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    80
lemma "Q \<longrightarrow> (D \<and> F \<and> ((\<not> D) \<and> (\<not> F)))"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    81
  apply (reify Ifm.simps)
23547
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    82
oops
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    83
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    84
text \<open>Method \<open>reify\<close> maps a \<^typ>\<open>bool\<close> to an \<^typ>\<open>fm\<close>.  For this it needs the 
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    85
semantics of \<open>fm\<close>, i.e.\ the rewrite rules in \<open>Ifm.simps\<close>.\<close>
23547
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    86
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    87
text \<open>You can also just pick up a subterm to reify.\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    88
lemma "Q \<longrightarrow> (D \<and> F \<and> ((\<not> D) \<and> (\<not> F)))"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    89
  apply (reify Ifm.simps ("((\<not> D) \<and> (\<not> F))"))
23547
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    90
oops
cb1203d8897c More examples
chaieb
parents: 23477
diff changeset
    91
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    92
text \<open>Let's perform NNF. This is a version that tends to generate disjunctions\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    93
primrec fmsize :: "fm \<Rightarrow> nat"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
    94
where
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    95
  "fmsize (At n) = 1"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
    96
| "fmsize (Not p) = 1 + fmsize p"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
    97
| "fmsize (And p q) = 1 + fmsize p + fmsize q"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
    98
| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
    99
| "fmsize (Imp p q) = 2 + fmsize p + fmsize q"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   100
| "fmsize (Iff p q) = 2 + 2* fmsize p + 2* fmsize q"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   101
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   102
lemma [measure_function]: "is_measure fmsize" ..
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   103
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   104
fun nnf :: "fm \<Rightarrow> fm"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   105
where
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   106
  "nnf (At n) = At n"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   107
| "nnf (And p q) = And (nnf p) (nnf q)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   108
| "nnf (Or p q) = Or (nnf p) (nnf q)"
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   109
| "nnf (Imp p q) = Or (nnf (Not p)) (nnf q)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   110
| "nnf (Iff p q) = Or (And (nnf p) (nnf q)) (And (nnf (Not p)) (nnf (Not q)))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   111
| "nnf (Not (And p q)) = Or (nnf (Not p)) (nnf (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   112
| "nnf (Not (Or p q)) = And (nnf (Not p)) (nnf (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   113
| "nnf (Not (Imp p q)) = And (nnf p) (nnf (Not q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   114
| "nnf (Not (Iff p q)) = Or (And (nnf p) (nnf (Not q))) (And (nnf (Not p)) (nnf q))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   115
| "nnf (Not (Not p)) = nnf p"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents: 69597
diff changeset
   116
| "nnf (Not p) = Not p"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   117
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   118
text \<open>The correctness theorem of \<^const>\<open>nnf\<close>: it preserves the semantics of \<^typ>\<open>fm\<close>\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   119
lemma nnf [reflection]:
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   120
  "Ifm (nnf p) vs = Ifm p vs"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   121
  by (induct p rule: nnf.induct) auto
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   122
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   123
text \<open>Now let's perform NNF using our \<^const>\<open>nnf\<close> function defined above.  First to the
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   124
  whole subgoal.\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   125
lemma "A \<noteq> B \<and> (B \<longrightarrow> A \<noteq> (B \<or> C \<and> (B \<longrightarrow> A \<or> D))) \<longrightarrow> A \<or> B \<and> D"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   126
  apply (reflection Ifm.simps)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   127
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   128
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   129
text \<open>Now we specify on which subterm it should be applied\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   130
lemma "A \<noteq> B \<and> (B \<longrightarrow> A \<noteq> (B \<or> C \<and> (B \<longrightarrow> A \<or> D))) \<longrightarrow> A \<or> B \<and> D"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   131
  apply (reflection Ifm.simps only: "B \<or> C \<and> (B \<longrightarrow> A \<or> D)")
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   132
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   133
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   134
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   135
text \<open>Example 2: Simple arithmetic formulae\<close>
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   136
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   137
text \<open>The type \<open>num\<close> reflects linear expressions over natural number\<close>
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   138
datatype num = C nat | Add num num | Mul nat num | Var nat | CN nat nat num
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   139
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   140
text \<open>This is just technical to make recursive definitions easier.\<close>
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   141
primrec num_size :: "num \<Rightarrow> nat" 
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   142
where
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   143
  "num_size (C c) = 1"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   144
| "num_size (Var n) = 1"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   145
| "num_size (Add a b) = 1 + num_size a + num_size b"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   146
| "num_size (Mul c a) = 1 + num_size a"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   147
| "num_size (CN n c a) = 4 + num_size a "
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   148
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   149
lemma [measure_function]: "is_measure num_size" ..
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   150
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   151
text \<open>The semantics of num\<close>
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   152
primrec Inum:: "num \<Rightarrow> nat list \<Rightarrow> nat"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   153
where
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   154
  Inum_C  : "Inum (C i) vs = i"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   155
| Inum_Var: "Inum (Var n) vs = vs!n"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   156
| Inum_Add: "Inum (Add s t) vs = Inum s vs + Inum t vs "
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   157
| Inum_Mul: "Inum (Mul c t) vs = c * Inum t vs "
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   158
| Inum_CN : "Inum (CN n c t) vs = c*(vs!n) + Inum t vs "
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   159
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   160
text \<open>Let's reify some nat expressions \dots\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   161
lemma "4 * (2 * x + (y::nat)) + f a \<noteq> 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   162
  apply (reify Inum.simps ("4 * (2 * x + (y::nat)) + f a"))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   163
oops
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   164
text \<open>We're in a bad situation! \<open>x\<close>, \<open>y\<close> and \<open>f\<close> have been recongnized
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   165
as constants, which is correct but does not correspond to our intuition of the constructor C.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   166
It should encapsulate constants, i.e. numbers, i.e. numerals.\<close>
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   167
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   168
text \<open>So let's leave the \<open>Inum_C\<close> equation at the end and see what happens \dots\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   169
lemma "4 * (2 * x + (y::nat)) \<noteq> 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   170
  apply (reify Inum_Var Inum_Add Inum_Mul Inum_CN Inum_C ("4 * (2 * x + (y::nat))"))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   171
oops
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   172
text \<open>Hm, let's specialize \<open>Inum_C\<close> with numerals.\<close>
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   173
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 41413
diff changeset
   174
lemma Inum_number: "Inum (C (numeral t)) vs = numeral t" by simp
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   175
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
   176
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   177
text \<open>Second attempt\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   178
lemma "1 * (2 * x + (y::nat)) \<noteq> 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   179
  apply (reify Inum_eqs ("1 * (2 * x + (y::nat))"))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   180
oops
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   181
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   182
text\<open>That was fine, so let's try another one \dots\<close>
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   183
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   184
lemma "1 * (2 * x + (y::nat) + 0 + 1) \<noteq> 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   185
  apply (reify Inum_eqs ("1 * (2 * x + (y::nat) + 0 + 1)"))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   186
oops
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   187
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   188
text \<open>Oh!! 0 is not a variable \dots\ Oh! 0 is not a \<open>numeral\<close> \dots\ thing.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   189
The same for 1. So let's add those equations, too.\<close>
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   190
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   191
lemma Inum_01: "Inum (C 0) vs = 0" "Inum (C 1) vs = 1" "Inum (C(Suc n)) vs = Suc n"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   192
  by simp_all
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   193
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   194
lemmas Inum_eqs'= Inum_eqs Inum_01
20337
36e2fae2c68a *** empty log message ***
chaieb
parents: 20319
diff changeset
   195
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   196
text\<open>Third attempt:\<close>
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   197
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   198
lemma "1 * (2 * x + (y::nat) + 0 + 1) \<noteq> 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   199
  apply (reify Inum_eqs' ("1 * (2 * x + (y::nat) + 0 + 1)"))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   200
oops
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   201
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   202
text \<open>Okay, let's try reflection. Some simplifications on \<^typ>\<open>num\<close> follow. You can
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   203
  skim until the main theorem \<open>linum\<close>.\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   204
  
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   205
fun lin_add :: "num \<Rightarrow> num \<Rightarrow> num"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   206
where
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   207
  "lin_add (CN n1 c1 r1) (CN n2 c2 r2) =
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   208
    (if n1 = n2 then 
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   209
      (let c = c1 + c2
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   210
       in (if c = 0 then lin_add r1 r2 else CN n1 c (lin_add r1 r2)))
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   211
     else if n1 \<le> n2 then (CN n1 c1 (lin_add r1 (CN n2 c2 r2))) 
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   212
     else (CN n2 c2 (lin_add (CN n1 c1 r1) r2)))"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   213
| "lin_add (CN n1 c1 r1) t = CN n1 c1 (lin_add r1 t)"  
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   214
| "lin_add t (CN n2 c2 r2) = CN n2 c2 (lin_add t r2)" 
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   215
| "lin_add (C b1) (C b2) = C (b1 + b2)"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   216
| "lin_add a b = Add a b"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   217
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   218
lemma lin_add:
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   219
  "Inum (lin_add t s) bs = Inum (Add t s) bs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   220
  apply (induct t s rule: lin_add.induct, simp_all add: Let_def)
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   221
  apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   222
  apply (case_tac "n1 = n2", simp_all add: algebra_simps)
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   223
  done
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   224
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   225
fun lin_mul :: "num \<Rightarrow> nat \<Rightarrow> num"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   226
where
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   227
  "lin_mul (C j) i = C (i * j)"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   228
| "lin_mul (CN n c a) i = (if i=0 then (C 0) else CN n (i * c) (lin_mul a i))"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   229
| "lin_mul t i = (Mul i t)"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   230
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   231
lemma lin_mul:
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   232
  "Inum (lin_mul t i) bs = Inum (Mul i t) bs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   233
  by (induct t i rule: lin_mul.induct) (auto simp add: algebra_simps)
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   234
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   235
fun linum:: "num \<Rightarrow> num"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   236
where
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   237
  "linum (C b) = C b"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   238
| "linum (Var n) = CN n 1 (C 0)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   239
| "linum (Add t s) = lin_add (linum t) (linum s)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   240
| "linum (Mul c t) = lin_mul (linum t) c"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   241
| "linum (CN n c t) = lin_add (linum (Mul c (Var n))) (linum t)"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   242
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   243
lemma linum [reflection]:
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   244
  "Inum (linum t) bs = Inum t bs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   245
  by (induct t rule: linum.induct) (simp_all add: lin_mul lin_add)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   246
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   247
text \<open>Now we can use linum to simplify nat terms using reflection\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   248
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   249
lemma "Suc (Suc 1) * (x + Suc 1 * y) = 3 * x + 6 * y"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   250
  apply (reflection Inum_eqs' only: "Suc (Suc 1) * (x + Suc 1 * y)")
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   251
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   252
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   253
text \<open>Let's lift this to formulae and see what happens\<close>
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   254
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   255
datatype aform = Lt num num  | Eq num num | Ge num num | NEq num num | 
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   256
  Conj aform aform | Disj aform aform | NEG aform | T | F
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   257
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   258
primrec linaformsize:: "aform \<Rightarrow> nat"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   259
where
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   260
  "linaformsize T = 1"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   261
| "linaformsize F = 1"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   262
| "linaformsize (Lt a b) = 1"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   263
| "linaformsize (Ge a b) = 1"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   264
| "linaformsize (Eq a b) = 1"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   265
| "linaformsize (NEq a b) = 1"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   266
| "linaformsize (NEG p) = 2 + linaformsize p"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   267
| "linaformsize (Conj p q) = 1 + linaformsize p + linaformsize q"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   268
| "linaformsize (Disj p q) = 1 + linaformsize p + linaformsize q"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   269
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   270
lemma [measure_function]: "is_measure linaformsize" ..
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   271
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   272
primrec is_aform :: "aform => nat list => bool"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   273
where
28866
30cd9d89a0fb reactivated some dead theories (based on hints by Mark Hillebrand);
wenzelm
parents: 23650
diff changeset
   274
  "is_aform T vs = True"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   275
| "is_aform F vs = False"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   276
| "is_aform (Lt a b) vs = (Inum a vs < Inum b vs)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   277
| "is_aform (Eq a b) vs = (Inum a vs = Inum b vs)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   278
| "is_aform (Ge a b) vs = (Inum a vs \<ge> Inum b vs)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   279
| "is_aform (NEq a b) vs = (Inum a vs \<noteq> Inum b vs)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   280
| "is_aform (NEG p) vs = (\<not> (is_aform p vs))"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   281
| "is_aform (Conj p q) vs = (is_aform p vs \<and> is_aform q vs)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   282
| "is_aform (Disj p q) vs = (is_aform p vs \<or> is_aform q vs)"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   283
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   284
text\<open>Let's reify and do reflection\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   285
lemma "(3::nat) * x + t < 0 \<and> (2 * x + y \<noteq> 17)"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   286
  apply (reify Inum_eqs' is_aform.simps) 
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   287
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   288
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   289
text \<open>Note that reification handles several interpretations at the same time\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   290
lemma "(3::nat) * x + t < 0 \<and> x * x + t * x + 3 + 1 = z * t * 4 * z \<or> x + x + 1 < 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   291
  apply (reflection Inum_eqs' is_aform.simps only: "x + x + 1") 
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   292
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   293
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   294
text \<open>For reflection we now define a simple transformation on aform: NNF + linum on atoms\<close>
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   295
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   296
fun linaform:: "aform \<Rightarrow> aform"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   297
where
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   298
  "linaform (Lt s t) = Lt (linum s) (linum t)"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   299
| "linaform (Eq s t) = Eq (linum s) (linum t)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   300
| "linaform (Ge s t) = Ge (linum s) (linum t)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   301
| "linaform (NEq s t) = NEq (linum s) (linum t)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   302
| "linaform (Conj p q) = Conj (linaform p) (linaform q)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   303
| "linaform (Disj p q) = Disj (linaform p) (linaform q)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   304
| "linaform (NEG T) = F"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   305
| "linaform (NEG F) = T"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   306
| "linaform (NEG (Lt a b)) = Ge a b"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   307
| "linaform (NEG (Ge a b)) = Lt a b"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   308
| "linaform (NEG (Eq a b)) = NEq a b"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   309
| "linaform (NEG (NEq a b)) = Eq a b"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   310
| "linaform (NEG (NEG p)) = linaform p"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   311
| "linaform (NEG (Conj p q)) = Disj (linaform (NEG p)) (linaform (NEG q))"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   312
| "linaform (NEG (Disj p q)) = Conj (linaform (NEG p)) (linaform (NEG q))"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   313
| "linaform p = p"
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   314
28866
30cd9d89a0fb reactivated some dead theories (based on hints by Mark Hillebrand);
wenzelm
parents: 23650
diff changeset
   315
lemma linaform: "is_aform (linaform p) vs = is_aform p vs"
30cd9d89a0fb reactivated some dead theories (based on hints by Mark Hillebrand);
wenzelm
parents: 23650
diff changeset
   316
  by (induct p rule: linaform.induct) (auto simp add: linum)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   317
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   318
lemma "(Suc (Suc (Suc 0)) * ((x::nat) + Suc (Suc 0)) + Suc (Suc (Suc 0)) *
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   319
  (Suc (Suc (Suc 0))) * ((x::nat) + Suc (Suc 0))) < 0 \<and> Suc 0 + Suc 0 < 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   320
  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: 23624
diff changeset
   321
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: 23624
diff changeset
   322
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   323
declare linaform [reflection]
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   324
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   325
lemma "(Suc (Suc (Suc 0)) * ((x::nat) + Suc (Suc 0)) + Suc (Suc (Suc 0)) *
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   326
  (Suc (Suc (Suc 0))) * ((x::nat) + Suc (Suc 0))) < 0 \<and> Suc 0 + Suc 0 < 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   327
  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: 20503
diff changeset
   328
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
   329
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   330
text \<open>We now give an example where interpretaions have zero or more than only
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   331
  one envornement of different types and show that automatic reification also deals with
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   332
  bindings\<close>
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   333
  
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   334
datatype rb = BC bool | BAnd rb rb | BOr rb rb
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   335
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   336
primrec Irb :: "rb \<Rightarrow> bool"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   337
where
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   338
  "Irb (BC p) \<longleftrightarrow> p"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   339
| "Irb (BAnd s t) \<longleftrightarrow> Irb s \<and> Irb t"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   340
| "Irb (BOr s t) \<longleftrightarrow> 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: 20503
diff changeset
   341
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
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)"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   343
  apply (reify Irb.simps)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   344
oops
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   345
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   346
datatype rint = IC int | IVar nat | IAdd rint rint | IMult rint rint
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   347
  | INeg rint | ISub rint rint
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: 23624
diff changeset
   348
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   349
primrec Irint :: "rint \<Rightarrow> int list \<Rightarrow> int"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   350
where
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   351
  Irint_Var: "Irint (IVar n) vs = vs ! n"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   352
| Irint_Neg: "Irint (INeg t) vs = - Irint t vs"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   353
| Irint_Add: "Irint (IAdd s t) vs = Irint s vs + Irint t vs"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   354
| Irint_Sub: "Irint (ISub s t) vs = Irint s vs - Irint t vs"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   355
| Irint_Mult: "Irint (IMult s t) vs = Irint s vs * Irint t vs"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   356
| Irint_C: "Irint (IC i) vs = i"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   357
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   358
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: 20503
diff changeset
   359
  by simp
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   360
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   361
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: 20503
diff changeset
   362
  by simp
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   363
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 41413
diff changeset
   364
lemma Irint_Cnumeral: "Irint (IC (numeral x)) vs = numeral 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: 20503
diff changeset
   365
  by simp
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   366
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 41413
diff changeset
   367
lemmas Irint_simps = Irint_Var Irint_Neg Irint_Add Irint_Sub Irint_Mult Irint_C0 Irint_C1 Irint_Cnumeral
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   368
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   369
lemma "(3::int) * x + y * y - 9 + (- z) = 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   370
  apply (reify Irint_simps ("(3::int) * x + y * y - 9 + (- z)"))
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
   371
  oops
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   372
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   373
datatype rlist = LVar nat | LEmpty | LCons rint rlist | LAppend rlist rlist
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   374
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   375
primrec Irlist :: "rlist \<Rightarrow> int list \<Rightarrow> int list list \<Rightarrow> int list"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   376
where
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   377
  "Irlist (LEmpty) is vs = []"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   378
| "Irlist (LVar n) is vs = vs ! n"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   379
| "Irlist (LCons i t) is vs = Irint i is # Irlist t is vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   380
| "Irlist (LAppend s t) is vs = Irlist s is vs @ Irlist t is vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   381
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
   382
lemma "[(1::int)] = []"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   383
  apply (reify Irlist.simps Irint_simps ("[1] :: int list"))
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
   384
  oops
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   385
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   386
lemma "([(3::int) * x + y * y - 9 + (- z)] @ []) @ xs = [y * y - z - 9 + (3::int) * x]"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   387
  apply (reify Irlist.simps Irint_simps ("([(3::int) * x + y * y - 9 + (- z)] @ []) @ xs"))
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
   388
  oops
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   389
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   390
datatype rnat = NC nat| NVar nat| NSuc rnat | NAdd rnat rnat | NMult rnat rnat
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   391
  | NNeg rnat | NSub rnat rnat | Nlgth rlist
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   392
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   393
primrec Irnat :: "rnat \<Rightarrow> int list \<Rightarrow> int list list \<Rightarrow> nat list \<Rightarrow> nat"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   394
where
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   395
  Irnat_Suc: "Irnat (NSuc t) is ls vs = Suc (Irnat t is ls vs)"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   396
| Irnat_Var: "Irnat (NVar n) is ls vs = vs ! n"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   397
| Irnat_Neg: "Irnat (NNeg t) is ls vs = 0"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   398
| Irnat_Add: "Irnat (NAdd s t) is ls vs = Irnat s is ls vs + Irnat t is ls vs"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   399
| Irnat_Sub: "Irnat (NSub s t) is ls vs = Irnat s is ls vs - Irnat t is ls vs"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   400
| Irnat_Mult: "Irnat (NMult s t) is ls vs = Irnat s is ls vs * Irnat t is ls vs"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   401
| Irnat_lgth: "Irnat (Nlgth rxs) is ls vs = length (Irlist rxs is ls)"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   402
| Irnat_C: "Irnat (NC i) is ls vs = i"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   403
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   404
lemma Irnat_C0: "Irnat (NC 0) is ls vs = 0"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   405
  by simp
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   406
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   407
lemma Irnat_C1: "Irnat (NC 1) is ls vs = 1"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   408
  by simp
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   409
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 41413
diff changeset
   410
lemma Irnat_Cnumeral: "Irnat (NC (numeral x)) is ls vs = numeral x"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   411
  by simp
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   412
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
   413
lemmas Irnat_simps = Irnat_Suc Irnat_Var Irnat_Neg Irnat_Add Irnat_Sub Irnat_Mult Irnat_lgth
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 41413
diff changeset
   414
  Irnat_C0 Irnat_C1 Irnat_Cnumeral
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   415
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   416
lemma "Suc n * length (([(3::int) * x + y * y - 9 + (- z)] @ []) @ xs) = length xs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   417
  apply (reify Irnat_simps Irlist.simps Irint_simps
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   418
     ("Suc n * length (([(3::int) * x + y * y - 9 + (- z)] @ []) @ xs)"))
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
   419
  oops
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   420
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   421
datatype rifm = RT | RF | RVar nat
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
   422
  | RNLT rnat rnat | RNILT rnat rint | RNEQ rnat rnat
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   423
  | RAnd rifm rifm | ROr rifm rifm | RImp rifm rifm| RIff rifm rifm
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   424
  | RNEX rifm | RIEX rifm | RLEX rifm | RNALL rifm | RIALL rifm | RLALL rifm
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
   425
  | RBEX rifm | RBALL rifm
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   426
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   427
primrec Irifm :: "rifm \<Rightarrow> bool list \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> bool"
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   428
where
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   429
  "Irifm RT ps is ls ns \<longleftrightarrow> True"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   430
| "Irifm RF ps is ls ns \<longleftrightarrow> False"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   431
| "Irifm (RVar n) ps is ls ns \<longleftrightarrow> ps ! n"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   432
| "Irifm (RNLT s t) ps is ls ns \<longleftrightarrow> Irnat s is ls ns < Irnat t is ls ns"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   433
| "Irifm (RNILT s t) ps is ls ns \<longleftrightarrow> int (Irnat s is ls ns) < Irint t is"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   434
| "Irifm (RNEQ s t) ps is ls ns \<longleftrightarrow> Irnat s is ls ns = Irnat t is ls ns"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   435
| "Irifm (RAnd p q) ps is ls ns \<longleftrightarrow> Irifm p ps is ls ns \<and> Irifm q ps is ls ns"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   436
| "Irifm (ROr p q) ps is ls ns \<longleftrightarrow> Irifm p ps is ls ns \<or> Irifm q ps is ls ns"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   437
| "Irifm (RImp p q) ps is ls ns \<longleftrightarrow> Irifm p ps is ls ns \<longrightarrow> Irifm q ps is ls ns"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   438
| "Irifm (RIff p q) ps is ls ns \<longleftrightarrow> Irifm p ps is ls ns = Irifm q ps is ls ns"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   439
| "Irifm (RNEX p) ps is ls ns \<longleftrightarrow> (\<exists>x. Irifm p ps is ls (x # ns))"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   440
| "Irifm (RIEX p) ps is ls ns \<longleftrightarrow> (\<exists>x. Irifm p ps (x # is) ls ns)"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   441
| "Irifm (RLEX p) ps is ls ns \<longleftrightarrow> (\<exists>x. Irifm p ps is (x # ls) ns)"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   442
| "Irifm (RBEX p) ps is ls ns \<longleftrightarrow> (\<exists>x. Irifm p (x # ps) is ls ns)"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   443
| "Irifm (RNALL p) ps is ls ns \<longleftrightarrow> (\<forall>x. Irifm p ps is ls (x#ns))"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   444
| "Irifm (RIALL p) ps is ls ns \<longleftrightarrow> (\<forall>x. Irifm p ps (x # is) ls ns)"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   445
| "Irifm (RLALL p) ps is ls ns \<longleftrightarrow> (\<forall>x. Irifm p ps is (x#ls) ns)"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   446
| "Irifm (RBALL p) ps is ls ns \<longleftrightarrow> (\<forall>x. Irifm p (x # ps) is ls ns)"
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   447
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   448
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)"
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
   449
  apply (reify Irifm.simps Irnat_simps Irlist.simps Irint_simps)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   450
oops
01b711328990 Reification now handels binders.
chaieb
parents: 20337
diff changeset
   451
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   452
text \<open>An example for equations containing type variables\<close>
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   453
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   454
datatype prod = Zero | One | Var nat | Mul prod prod 
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   455
  | Pw prod nat | PNM nat nat prod
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   456
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   457
primrec Iprod :: " prod \<Rightarrow> ('a::linordered_idom) list \<Rightarrow>'a" 
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   458
where
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23608
diff changeset
   459
  "Iprod Zero vs = 0"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   460
| "Iprod One vs = 1"
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   461
| "Iprod (Var n) vs = vs ! n"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   462
| "Iprod (Mul a b) vs = Iprod a vs * Iprod b vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   463
| "Iprod (Pw a n) vs = Iprod a vs ^ n"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   464
| "Iprod (PNM n k t) vs = (vs ! n) ^ k * Iprod t vs"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35419
diff changeset
   465
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   466
datatype sgn = Pos prod | Neg prod | ZeroEq prod | NZeroEq prod | Tr | F 
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   467
  | Or sgn sgn | And sgn sgn
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   468
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   469
primrec Isgn :: "sgn \<Rightarrow> ('a::linordered_idom) list \<Rightarrow> bool"
35419
d78659d1723e more recdef (and old primrec) hunting
krauss
parents: 35028
diff changeset
   470
where 
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   471
  "Isgn Tr vs \<longleftrightarrow> True"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   472
| "Isgn F vs \<longleftrightarrow> False"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   473
| "Isgn (ZeroEq t) vs \<longleftrightarrow> Iprod t vs = 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   474
| "Isgn (NZeroEq t) vs \<longleftrightarrow> Iprod t vs \<noteq> 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   475
| "Isgn (Pos t) vs \<longleftrightarrow> Iprod t vs > 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   476
| "Isgn (Neg t) vs \<longleftrightarrow> Iprod t vs < 0"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   477
| "Isgn (And p q) vs \<longleftrightarrow> Isgn p vs \<and> Isgn q vs"
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   478
| "Isgn (Or p q) vs \<longleftrightarrow> Isgn p vs \<or> Isgn q vs"
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   479
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   480
lemmas eqs = Isgn.simps Iprod.simps
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   481
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   482
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: 20564
diff changeset
   483
  apply (reify eqs)
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   484
  oops
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 20564
diff changeset
   485
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   486
end
51093
9d7aa2bb097b tuned, particulary name
haftmann
parents: 47108
diff changeset
   487