CookBook/Solutions.thy
author berghofe
Mon, 13 Oct 2008 17:51:59 +0200
changeset 35 d5c090b9a2b1
parent 25 e2f9f94b26d4
child 47 4daf913fdbe1
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
25
e2f9f94b26d4 Antiquotation setup is now contained in theory Base.
berghofe
parents: 15
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
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
chapter {* Solutions to Most Exercises *}
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
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
ML {* 
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  fun rev_sum t =
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  let
35
berghofe
parents: 25
diff changeset
    12
   fun dest_sum (Const (@{const_name plus}, _) $ u $ u') = u' :: dest_sum u
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
     | dest_sum u = [u]
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
   in
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
     foldl1 (HOLogic.mk_binop @{const_name plus}) (dest_sum t)
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
   end;
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
*}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
text {* \solution{fun:makesum} *}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
ML {*
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  fun make_sum t1 t2 =
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
      HOLogic.mk_nat (HOLogic.dest_nat t1 + HOLogic.dest_nat t2)
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
*}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
end