src/HOL/Extraction/QuotRem.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 15244 9473137b3550
child 17145 e623e57b0f44
permissions -rw-r--r--
migrated theory headers to new format
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
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15244
diff changeset
     8
theory QuotRem imports Main begin
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     9
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    10
text {* Derivation of quotient and remainder using program extraction. *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    11
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    12
lemma nat_eq_dec: "\<And>n::nat. m = n \<or> m \<noteq> n"
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    13
  apply (induct m)
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    14
  apply (case_tac n)
15244
9473137b3550 mod becuase of chnage in induct
nipkow
parents: 13931
diff changeset
    15
  apply (case_tac [3] n)
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    16
  apply (simp only: nat.simps, rules?)+
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    17
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    18
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    19
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
    20
proof (induct a)
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    21
  case 0
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    22
  have "0 = Suc b * 0 + 0 \<and> 0 \<le> b" by simp
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    23
  thus ?case by rules
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
  case (Suc a)
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    26
  then obtain r q where I: "a = Suc b * q + r" and "r \<le> b" by rules
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    27
  from nat_eq_dec show ?case
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    28
  proof
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    29
    assume "r = b"
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    30
    with I have "Suc a = Suc b * (Suc q) + 0 \<and> 0 \<le> b" by simp
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    31
    thus ?case by rules
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    32
  next
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    33
    assume "r \<noteq> b"
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    34
    hence "r < b" by (simp add: order_less_le)
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    35
    with I have "Suc a = Suc b * q + (Suc r) \<and> (Suc r) \<le> b" by simp
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    36
    thus ?case by rules
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    37
  qed
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    38
qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    39
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    40
extract division
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    41
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    42
text {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    43
  The program extracted from the above proof looks as follows
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    44
  @{thm [display] division_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    45
  The corresponding correctness theorem is
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    46
  @{thm [display] division_correctness [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    47
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    48
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    49
generate_code
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    50
  test = "division 9 2"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    51
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    52
end