CookBook/Solutions.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 28 Nov 2008 05:56:28 +0100
changeset 52 a04bdee4fb1e
parent 47 4daf913fdbe1
child 56 126646f2aa88
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 {* 
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    10
fun rev_sum t =
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    11
let
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    12
 fun dest_sum (Const (@{const_name plus}, _) $ u $ u') = u' :: dest_sum u
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    13
   | dest_sum u = [u]
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    14
 in
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    15
   foldl1 (HOLogic.mk_binop @{const_name plus}) (dest_sum t)
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    16
 end;
15
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 {*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    22
fun make_sum t1 t2 =
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
    23
    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
    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