src/HOL/MicroJava/Comp/NatCanonify.thy
author wenzelm
Sun, 20 Nov 2011 21:05:23 +0100
changeset 45605 a89b4bc311a5
parent 41589 bbd861837ebc
permissions -rw-r--r--
eliminated obsolete "standard";
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/Comp/NatCanonify.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     2
    Author:     Martin Strecker
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     3
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     4
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14981
diff changeset
     5
theory NatCanonify imports Main begin
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     6
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     7
(************************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     8
  (* Canonizer for converting nat to int *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     9
(************************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    10
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    11
lemma nat_add_canonify: "(n1::nat) + n2 = nat ((int n1) + (int n2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    12
by (simp add: nat_add_distrib)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    13
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    14
lemma nat_mult_canonify: "(n1::nat) * n2 = nat ((int n1) * (int n2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    15
by (simp add: nat_mult_distrib)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    16
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    17
lemma nat_diff_canonify: "(n1::nat) - n2 = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    18
  nat (if (int n1) \<le> (int n2) then 0 else (int n1) - (int n2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    19
by (simp add: zdiff_int nat_int)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    20
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    21
lemma nat_le_canonify: "((n1::nat) \<le> n2) = ((int n1) \<le> (int n2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    22
by arith
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    23
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    24
lemma nat_less_canonify: "((n1::nat) < n2) = ((int n1) < (int n2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    25
by arith
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    26
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    27
lemma nat_eq_canonify: "((n1::nat) = n2) = ((int n1) = (int n2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    28
by arith
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    29
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    30
lemma nat_if_canonify: "(if b then (n1::nat) else n2) =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    31
  nat (if b then (int n1) else (int n2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    32
by simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    34
lemma int_nat_canonify: "(int (nat n)) = (if 0 \<le> n then n else 0)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    35
by simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    36
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    37
lemmas nat_canonify = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    38
  nat_add_canonify nat_mult_canonify nat_diff_canonify 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    39
  nat_le_canonify nat_less_canonify nat_eq_canonify nat_if_canonify 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    40
  int_nat_canonify nat_int
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    41
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    42
end