src/HOL/Extraction/QuotRem.thy
author berghofe
Tue, 13 Nov 2007 10:55:08 +0100
changeset 25419 e6a56be0ccaa
parent 24348 c708ea5b109a
child 27436 9581777503e9
permissions -rw-r--r--
Moved nat_eq_dec to Util.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Extraction/QuotRem.thy
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     4
*)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     5
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     6
header {* Quotient and remainder *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     7
25419
e6a56be0ccaa Moved nat_eq_dec to Util.thy
berghofe
parents: 24348
diff changeset
     8
theory QuotRem imports Util begin
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     9
text {* Derivation of quotient and remainder using program extraction. *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    10
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    11
theorem division: "\<exists>r q. a = Suc b * q + r \<and> r \<le> b"
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    12
proof (induct a)
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    13
  case 0
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    14
  have "0 = Suc b * 0 + 0 \<and> 0 \<le> b" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    15
  thus ?case by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    16
next
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    17
  case (Suc a)
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    18
  then obtain r q where I: "a = Suc b * q + r" and "r \<le> b" by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    19
  from nat_eq_dec show ?case
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    20
  proof
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    21
    assume "r = b"
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    22
    with I have "Suc a = Suc b * (Suc q) + 0 \<and> 0 \<le> b" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    23
    thus ?case by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    24
  next
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    25
    assume "r \<noteq> b"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 22845
diff changeset
    26
    with `r \<le> b` have "r < b" by (simp add: order_less_le)
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    27
    with I have "Suc a = Suc b * q + (Suc r) \<and> (Suc r) \<le> b" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    28
    thus ?case by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    29
  qed
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    30
qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    31
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    32
extract division
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    33
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    34
text {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    35
  The program extracted from the above proof looks as follows
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    36
  @{thm [display] division_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    37
  The corresponding correctness theorem is
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    38
  @{thm [display] division_correctness [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    39
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    40
17145
e623e57b0f44 Adapted to new code generator syntax.
berghofe
parents: 16417
diff changeset
    41
code_module Div
e623e57b0f44 Adapted to new code generator syntax.
berghofe
parents: 16417
diff changeset
    42
contains
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    43
  test = "division 9 2"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    44
24348
c708ea5b109a renamed code_gen to export_code
haftmann
parents: 23373
diff changeset
    45
export_code division in SML
20593
5af400cc64d5 added some stuff for code generation 2
haftmann
parents: 17604
diff changeset
    46
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    47
end