ProgTutorial/Solutions.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 22 Jul 2009 01:10:14 +0200
changeset 278 c6b64fa9f301
parent 272 2ff4867c00cf
child 293 0a567f923b42
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Solutions
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
     2
imports Base "Recipes/Timing"
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
     5
chapter {* Solutions to Most Exercises\label{ch:solutions} *}
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
text {* \solution{fun:revsum} *}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
271
949957531f63 modified solution(s) for "revsum" example
griff
parents: 270
diff changeset
     9
ML{*fun rev_sum((p as Const(@{const_name plus},_)) $ t $ u) = p $ u $ rev_sum t
949957531f63 modified solution(s) for "revsum" example
griff
parents: 270
diff changeset
    10
  | rev_sum t = t *}
949957531f63 modified solution(s) for "revsum" example
griff
parents: 270
diff changeset
    11
949957531f63 modified solution(s) for "revsum" example
griff
parents: 270
diff changeset
    12
text {* \solution{fun:revsum typed} *}
949957531f63 modified solution(s) for "revsum" example
griff
parents: 270
diff changeset
    13
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    14
ML{*fun rev_sum t =
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    15
let
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    16
  fun dest_sum (Const (@{const_name plus}, _) $ u $ u') = u' :: dest_sum u
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    17
    | dest_sum u = [u]
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    18
in
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    19
   foldl1 (HOLogic.mk_binop @{const_name plus}) (dest_sum t)
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    20
end *}
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
271
949957531f63 modified solution(s) for "revsum" example
griff
parents: 270
diff changeset
    22
270
bfcabed9079e included an alternative solution for the "rev_sum" example as a comment
griff
parents: 225
diff changeset
    23
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
text {* \solution{fun:makesum} *}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    26
ML{*fun make_sum t1 t2 =
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    27
      HOLogic.mk_nat (HOLogic.dest_nat t1 + HOLogic.dest_nat t2) *}
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    29
text {* \solution{ex:scancmts} *}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    30
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    31
ML{*val any = Scan.one (Symbol.not_eof)
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    32
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    33
val scan_cmt =
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    34
let
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    35
  val begin_cmt = Scan.this_string "(*" 
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    36
  val end_cmt = Scan.this_string "*)"
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    37
in
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    38
  begin_cmt |-- Scan.repeat (Scan.unless end_cmt any) --| end_cmt 
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    39
  >> (enclose "(**" "**)" o implode)
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    40
end
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    41
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    42
val parser = Scan.repeat (scan_cmt || any)
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    43
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    44
val scan_all =
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
    45
      Scan.finite Symbol.stopper parser >> implode #> fst *}
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    46
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    47
text {*
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    48
  By using @{text "#> fst"} in the last line, the function 
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    49
  @{ML scan_all} retruns a string, instead of the pair a parser would
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    50
  normally return. For example:
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    51
80
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    52
  @{ML_response [display,gray]
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    53
"let
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    54
  val input1 = (explode \"foo bar\")
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    55
  val input2 = (explode \"foo (*test*) bar (*test*)\")
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    56
in
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    57
  (scan_all input1, scan_all input2)
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    58
end"
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    59
"(\"foo bar\", \"foo (**test**) bar (**test**)\")"}
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    60
*}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    61
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    62
text {* \solution{ex:addsimproc} *}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    63
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    64
ML{*fun dest_sum term =
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    65
  case term of 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    66
    (@{term "(op +):: nat \<Rightarrow> nat \<Rightarrow> nat"} $ t1 $ t2) =>
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    67
        (snd (HOLogic.dest_number t1), snd (HOLogic.dest_number t2))
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    68
  | _ => raise TERM ("dest_sum", [term])
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    69
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    70
fun get_sum_thm ctxt t (n1, n2) =  
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    71
let 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    72
  val sum = HOLogic.mk_number @{typ "nat"} (n1 + n2)
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
    73
  val goal = Logic.mk_equals (t, sum)
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    74
in
214
7e04dc2368b0 updated to latest Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    75
  Goal.prove ctxt [] [] goal (K (Arith_Data.arith_tac ctxt 1))
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    76
end
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    77
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    78
fun add_sp_aux ss t =
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    79
let 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    80
  val ctxt = Simplifier.the_context ss
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    81
  val t' = term_of t
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    82
in
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    83
  SOME (get_sum_thm ctxt t' (dest_sum t'))
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    84
  handle TERM _ => NONE
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    85
end*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    86
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    87
text {* The setup for the simproc is *}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    88
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 175
diff changeset
    89
simproc_setup %gray add_sp ("t1 + t2") = {* K add_sp_aux *}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    90
 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    91
text {* and a test case is the lemma *}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    92
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    93
lemma "P (Suc (99 + 1)) ((0 + 0)::nat) (Suc (3 + 3 + 3)) (4 + 1)"
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
    94
  apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]) 1 *})
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    95
txt {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    96
  where the simproc produces the goal state
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    97
  
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    98
  \begin{minipage}{\textwidth}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    99
  @{subgoals [display]}
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   100
  \end{minipage}\bigskip
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   101
*}(*<*)oops(*>*)
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   102
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   103
text {* \solution{ex:addconversion} *}
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   104
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   105
text {*
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   106
  The following code assumes the function @{ML dest_sum} from the previous
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   107
  exercise.
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   108
*}
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   109
166
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   110
ML{*fun add_simple_conv ctxt ctrm =
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   111
let
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   112
  val trm =  Thm.term_of ctrm
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   113
in 
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   114
  get_sum_thm ctxt trm (dest_sum trm)
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   115
end
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   116
166
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   117
fun add_conv ctxt ctrm =
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   118
  (case Thm.term_of ctrm of
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   119
     @{term "(op +)::nat \<Rightarrow> nat \<Rightarrow> nat"} $ _ $ _ => 
166
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   120
         (Conv.binop_conv (add_conv ctxt)
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   121
          then_conv (Conv.try_conv (add_simple_conv ctxt))) ctrm         
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   122
    | _ $ _ => Conv.combination_conv 
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   123
                 (add_conv ctxt) (add_conv ctxt) ctrm
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   124
    | Abs _ => Conv.abs_conv (fn (_, ctxt) => add_conv ctxt) ctxt ctrm
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   125
    | _ => Conv.all_conv ctrm)
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   126
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   127
fun add_tac ctxt = CSUBGOAL (fn (goal, i) =>
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   128
  CONVERSION
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   129
    (Conv.params_conv ~1 (fn ctxt =>
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   130
       (Conv.prems_conv ~1 (add_conv ctxt) then_conv
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   131
          Conv.concl_conv ~1 (add_conv ctxt))) ctxt) i)*}
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   132
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   133
text {*
175
7c09bd3227c5 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
   134
  A test case for this conversion is as follows
168
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   135
*}
009ca4807baa polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   136
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   137
lemma "P (Suc (99 + 1)) ((0 + 0)::nat) (Suc (3 + 3 + 3)) (4 + 1)"
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   138
  apply(tactic {* add_tac @{context} 1 *})?
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   139
txt {* 
175
7c09bd3227c5 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
   140
  where it produces the goal state
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   141
  
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   142
  \begin{minipage}{\textwidth}
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   143
  @{subgoals [display]}
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   144
  \end{minipage}\bigskip
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   145
*}(*<*)oops(*>*)
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   146
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   147
text {* \solution{ex:addconversion} *}
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   148
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   149
text {* 
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   150
  We use the timing function @{ML timing_wrapper} from Recipe~\ref{rec:timing}.
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   151
  To measure any difference between the simproc and conversion, we will create 
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   152
  mechanically terms involving additions and then set up a goal to be 
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   153
  simplified. We have to be careful to set up the goal so that
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   154
  other parts of the simplifier do not interfere. For this we construct an
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   155
  unprovable goal which, after simplification, we are going to ``prove'' with
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   156
  the help of ``\isacommand{sorry}'', that is the method @{ML SkipProof.cheat_tac}
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   157
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   158
  For constructing test cases, we first define a function that returns a 
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   159
  complete binary tree whose leaves are numbers and the nodes are 
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   160
  additions.
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   161
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   162
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   163
ML{*fun term_tree n =
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   164
let
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   165
  val count = ref 0; 
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   166
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   167
  fun term_tree_aux n =
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   168
    case n of
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   169
      0 => (count := !count + 1; HOLogic.mk_number @{typ nat} (!count))
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   170
    | _ => Const (@{const_name "plus"}, @{typ "nat\<Rightarrow>nat\<Rightarrow>nat"}) 
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   171
             $ (term_tree_aux (n - 1)) $ (term_tree_aux (n - 1))
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   172
in
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   173
  term_tree_aux n
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   174
end*}
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   175
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   176
text {*
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   177
  This function generates for example:
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   178
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   179
  @{ML_response_fake [display,gray] 
239
b63c72776f03 replaced "warning" with "writeln"
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   180
  "writeln (Syntax.string_of_term @{context} (term_tree 2))" 
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   181
  "(1 + 2) + (3 + 4)"} 
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   182
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   183
  The next function constructs a goal of the form @{text "P \<dots>"} with a term 
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   184
  produced by @{ML term_tree} filled in.
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   185
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   186
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   187
ML{*fun goal n = HOLogic.mk_Trueprop (@{term "P::nat\<Rightarrow> bool"} $ (term_tree n))*}
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   188
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   189
text {*
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   190
  Note that the goal needs to be wrapped in a @{term "Trueprop"}. Next we define
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   191
  two tactics, @{text "c_tac"} and @{text "s_tac"}, for the conversion and simproc,
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   192
  respectively. The idea is to first apply the conversion (respectively simproc) and 
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   193
  then prove the remaining goal using @{ML "cheat_tac" in SkipProof}.
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   194
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   195
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   196
ML{*local
225
7859fc59249a section for further material about simple inductive
Christian Urban <urbanc@in.tum.de>
parents: 214
diff changeset
   197
  fun mk_tac tac = 
7859fc59249a section for further material about simple inductive
Christian Urban <urbanc@in.tum.de>
parents: 214
diff changeset
   198
        timing_wrapper (EVERY1 [tac, K (SkipProof.cheat_tac @{theory})])
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   199
in
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   200
val c_tac = mk_tac (add_tac @{context}) 
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   201
val s_tac = mk_tac (simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]))
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   202
end*}
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   203
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   204
text {*
175
7c09bd3227c5 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
   205
  This is all we need to let the conversion run against the simproc:
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   206
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   207
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   208
ML{*val _ = Goal.prove @{context} [] [] (goal 8) (K c_tac)
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   209
val _ = Goal.prove @{context} [] [] (goal 8) (K s_tac)*}
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   210
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   211
text {*
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   212
  If you do the exercise, you can see that both ways of simplifying additions
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   213
  perform relatively similar with perhaps some advantages for the
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   214
  simproc. That means the simplifier, even if much more complicated than
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   215
  conversions, is quite efficient for tasks it is designed for. It usually does not
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   216
  make sense to implement general-purpose rewriting using
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   217
  conversions. Conversions only have clear advantages in special situations:
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 168
diff changeset
   218
  for example if you need to have control over innermost or outermost
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   219
  rewriting, or when rewriting rules are lead to non-termination.
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   220
*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   221
270
bfcabed9079e included an alternative solution for the "rev_sum" example as a comment
griff
parents: 225
diff changeset
   222
end