CookBook/Solutions.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 11 Mar 2009 17:38:17 +0000
changeset 167 3e30ea95c7aa
parent 166 00d153e32a53
child 168 009ca4807baa
permissions -rw-r--r--
added temporarily some timing test about conversions and simprocs
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
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
     2
imports Base 
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
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
     9
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
    10
let
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    11
  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
    12
    | dest_sum u = [u]
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    13
in
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    14
   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
    15
end *}
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
text {* \solution{fun:makesum} *}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    19
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
    20
      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
    21
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    22
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
    23
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    24
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
    25
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    26
val scan_cmt =
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    27
  let
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    28
    val begin_cmt = Scan.this_string "(*" 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    29
    val end_cmt = Scan.this_string "*)"
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    30
  in
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    31
   begin_cmt |-- Scan.repeat (Scan.unless end_cmt any) --| end_cmt 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    32
    >> (enclose "(**" "**)" o implode)
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    33
  end
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    34
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    35
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
    36
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    37
val scan_all =
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    38
  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
    39
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    40
text {*
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    41
  By using @{text "#> fst"} in the last line, the function 
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    42
  @{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
    43
  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
    44
80
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    45
  @{ML_response [display,gray]
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    46
"let
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    47
  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
    48
  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
    49
in
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    50
  (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
    51
end"
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    52
"(\"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
    53
*}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    54
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    55
text {* \solution{ex:addsimproc} *}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    56
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    57
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
    58
  case term of 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    59
    (@{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
    60
        (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
    61
  | _ => 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
    62
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    63
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
    64
let 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    65
  val sum = HOLogic.mk_number @{typ "nat"} (n1 + n2)
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
    66
  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
    67
in
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
    68
  Goal.prove ctxt [] [] goal (K (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
    69
end
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    70
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    71
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
    72
let 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    73
  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
    74
  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
    75
in
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    76
  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
    77
  handle TERM _ => NONE
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    78
end*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    79
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    80
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
    81
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    82
simproc_setup add_sp ("t1 + t2") = {* K add_sp_aux *}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    83
 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    84
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
    85
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    86
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
    87
  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
    88
txt {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    89
  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
    90
  
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    91
  \begin{minipage}{\textwidth}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    92
  @{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
    93
  \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
    94
*}(*<*)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
    95
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
    96
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
    97
166
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
    98
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
    99
let
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   100
  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
   101
in 
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   102
  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
   103
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
   104
166
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
   105
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
   106
  (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
   107
     @{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
   108
         (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
   109
          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
   110
    | _ $ _ => 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
   111
                 (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
   112
    | 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
   113
    | _ => 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
   114
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
   115
val add_tac = CSUBGOAL (fn (goal, i) =>
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
  let
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
   117
    val ctxt = ProofContext.init (Thm.theory_of_cterm goal)
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
  in
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
    CONVERSION
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
   120
      (Conv.params_conv ~1 (fn ctxt =>
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
   121
        (Conv.prems_conv ~1 (add_conv ctxt) then_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
   122
          Conv.concl_conv ~1 (add_conv ctxt))) ctxt) i
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
  end)*}
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
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
lemma "P (Suc (99 + 1)) ((0 + 0)::nat) (Suc (3 + 3 + 3)) (4 + 1)"
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
  apply(tactic {* add_tac 1 *})?
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
   127
txt {* 
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
   128
  where the simproc produces the goal state
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
   129
  
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
   130
  \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
   131
  @{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
   132
  \end{minipage}\bigskip
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   133
*}(*<*)oops(*>*)
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   134
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   135
subsection {* Tests start here *}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   136
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   137
lemma cheat: "P" sorry
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   138
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   139
ML{*fun timing_wrapper tac st =
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   140
let 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   141
  val t_start = start_timing ();
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   142
  val res = tac st;
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   143
  val t_end = end_timing t_start;
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   144
in
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   145
  (warning (#message t_end); res)
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   146
end*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   147
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   148
ML{* 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   149
fun create_term1 0 = @{term "0::nat"}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   150
  | create_term1 n = 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   151
       Const (@{const_name "plus"}, @{typ "nat\<Rightarrow>nat\<Rightarrow>nat"}) 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   152
         $ (HOLogic.mk_number @{typ "nat"} n) $ (create_term1 (n-1))
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   153
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   154
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   155
ML{* 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   156
fun create_term2 0 = @{term "0::nat"}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   157
  | create_term2 n = 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   158
       Const (@{const_name "plus"}, @{typ "nat\<Rightarrow>nat\<Rightarrow>nat"}) 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   159
         $ (create_term2 (n-1)) $ (HOLogic.mk_number @{typ nat} n)
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   160
*}
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
ML{*
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   163
fun create_term n = 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   164
  HOLogic.mk_Trueprop
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   165
  (@{term "P::nat\<Rightarrow> bool"} $ 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   166
   (Const (@{const_name "plus"}, @{typ "nat\<Rightarrow>nat\<Rightarrow>nat"}) 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   167
    $ (create_term1 n) $ (create_term2 n)))
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   168
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   169
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   170
ML {*
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   171
warning (Syntax.string_of_term @{context} (create_term 4))
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   172
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   173
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   174
ML {* 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   175
val _ = Goal.prove @{context} [] [] (create_term 100) 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   176
   (fn _ => timing_wrapper (EVERY1 [add_tac, rtac @{thm cheat}]))
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   177
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   178
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   179
ML {* 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   180
val _ = Goal.prove @{context} [] [] (create_term 100) 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   181
   (fn _ => timing_wrapper (EVERY1 [simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]), rtac @{thm cheat}]))
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   182
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   183
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   184
ML {* 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   185
val _ = Goal.prove @{context} [] [] (create_term 400) 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   186
  (fn _ => timing_wrapper (EVERY1 [add_tac, rtac @{thm cheat}]))
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   187
*}
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   188
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   189
ML {* 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   190
val _ = Goal.prove @{context} [] [] (create_term 400) 
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   191
   (fn _ => timing_wrapper (EVERY1 [simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]), rtac @{thm cheat}]))
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 166
diff changeset
   192
*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   193
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
end