src/HOL/NatArith.thy
author nipkow
Fri, 04 Feb 2005 17:14:42 +0100
changeset 15497 53bca254719a
parent 15439 71c0f98e31f1
child 15539 333a88244569
permissions -rw-r--r--
Added semi-lattice locales and reorganized fold1 lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10214
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/NatArith.thy
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
     2
    ID:         $Id$
13297
wenzelm
parents: 11655
diff changeset
     3
    Author:     Tobias Nipkow and Markus Wenzel
wenzelm
parents: 11655
diff changeset
     4
*)
10214
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
     5
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
     6
header {*Further Arithmetic Facts Concerning the Natural Numbers*}
10214
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15048
diff changeset
     8
theory NatArith
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     9
imports Nat
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15048
diff changeset
    10
files "arith_data.ML"
c69542757a4d New theory header syntax.
nipkow
parents: 15048
diff changeset
    11
begin
10214
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
    12
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
    13
setup arith_setup
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
    14
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    15
text{*The following proofs may rely on the arithmetic proof procedures.*}
13297
wenzelm
parents: 11655
diff changeset
    16
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    17
lemma pred_nat_trancl_eq_le: "((m, n) : pred_nat^*) = (m \<le> n)"
14208
144f45277d5a misc tidying
paulson
parents: 13499
diff changeset
    18
by (simp add: less_eq reflcl_trancl [symmetric]
144f45277d5a misc tidying
paulson
parents: 13499
diff changeset
    19
            del: reflcl_trancl, arith)
11454
7514e5e21cb8 Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents: 11324
diff changeset
    20
10214
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
    21
lemma nat_diff_split:
10599
2df753cf86e9 miniscoping of nat_diff_split
paulson
parents: 10214
diff changeset
    22
    "P(a - b::nat) = ((a<b --> P 0) & (ALL d. a = b + d --> P d))"
13297
wenzelm
parents: 11655
diff changeset
    23
    -- {* elimination of @{text -} on @{text nat} *}
wenzelm
parents: 11655
diff changeset
    24
  by (cases "a<b" rule: case_split)
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    25
     (auto simp add: diff_is_0_eq [THEN iffD2])
11324
82406bd816a5 nat_diff_split_asm, for the assumptions
paulson
parents: 11181
diff changeset
    26
82406bd816a5 nat_diff_split_asm, for the assumptions
paulson
parents: 11181
diff changeset
    27
lemma nat_diff_split_asm:
82406bd816a5 nat_diff_split_asm, for the assumptions
paulson
parents: 11181
diff changeset
    28
    "P(a - b::nat) = (~ (a < b & ~ P 0 | (EX d. a = b + d & ~ P d)))"
13297
wenzelm
parents: 11655
diff changeset
    29
    -- {* elimination of @{text -} on @{text nat} in assumptions *}
11324
82406bd816a5 nat_diff_split_asm, for the assumptions
paulson
parents: 11181
diff changeset
    30
  by (simp split: nat_diff_split)
10214
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
    31
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
    32
lemmas [arith_split] = nat_diff_split split_min split_max
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
    33
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    34
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    35
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    36
lemma le_square: "m \<le> m*(m::nat)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    37
by (induct_tac "m", auto)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    38
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    39
lemma le_cube: "(m::nat) \<le> m*(m*m)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    40
by (induct_tac "m", auto)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    41
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    42
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    43
text{*Subtraction laws, mostly by Clemens Ballarin*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    44
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    45
lemma diff_less_mono: "[| a < (b::nat); c \<le> a |] ==> a-c < b-c"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    46
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    47
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    48
lemma less_diff_conv: "(i < j-k) = (i+k < (j::nat))"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    49
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    50
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    51
lemma le_diff_conv: "(j-k \<le> (i::nat)) = (j \<le> i+k)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    52
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    53
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    54
lemma le_diff_conv2: "k \<le> j ==> (i \<le> j-k) = (i+k \<le> (j::nat))"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    55
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    56
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    57
lemma diff_diff_cancel [simp]: "i \<le> (n::nat) ==> n - (n - i) = i"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    58
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    59
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    60
lemma le_add_diff: "k \<le> (n::nat) ==> m \<le> n + m - k"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    61
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    62
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    63
(*Replaces the previous diff_less and le_diff_less, which had the stronger
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    64
  second premise n\<le>m*)
15439
71c0f98e31f1 made diff_less a simp rule
nipkow
parents: 15404
diff changeset
    65
lemma diff_less[simp]: "!!m::nat. [| 0<n; 0<m |] ==> m - n < m"
15404
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    66
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    67
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    68
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    69
(** Simplification of relational expressions involving subtraction **)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    70
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    71
lemma diff_diff_eq: "[| k \<le> m;  k \<le> (n::nat) |] ==> ((m-k) - (n-k)) = (m-n)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    72
by (simp split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    73
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    74
lemma eq_diff_iff: "[| k \<le> m;  k \<le> (n::nat) |] ==> (m-k = n-k) = (m=n)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    75
by (auto split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    76
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    77
lemma less_diff_iff: "[| k \<le> m;  k \<le> (n::nat) |] ==> (m-k < n-k) = (m<n)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    78
by (auto split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    79
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    80
lemma le_diff_iff: "[| k \<le> m;  k \<le> (n::nat) |] ==> (m-k \<le> n-k) = (m\<le>n)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    81
by (auto split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    82
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    83
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    84
text{*(Anti)Monotonicity of subtraction -- by Stephan Merz*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    85
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    86
(* Monotonicity of subtraction in first argument *)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    87
lemma diff_le_mono: "m \<le> (n::nat) ==> (m-l) \<le> (n-l)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    88
by (simp split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    89
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    90
lemma diff_le_mono2: "m \<le> (n::nat) ==> (l-n) \<le> (l-m)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    91
by (simp split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    92
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    93
lemma diff_less_mono2: "[| m < (n::nat); m<l |] ==> (l-n) < (l-m)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    94
by (simp split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    95
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    96
lemma diffs0_imp_equal: "!!m::nat. [| m-n = 0; n-m = 0 |] ==>  m=n"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    97
by (simp split add: nat_diff_split)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    98
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
    99
text{*Lemmas for ex/Factorization*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   100
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   101
lemma one_less_mult: "[| Suc 0 < n; Suc 0 < m |] ==> Suc 0 < m*n"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   102
by (case_tac "m", auto)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   103
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   104
lemma n_less_m_mult_n: "[| Suc 0 < n; Suc 0 < m |] ==> n<m*n"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   105
by (case_tac "m", auto)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   106
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   107
lemma n_less_n_mult_m: "[| Suc 0 < n; Suc 0 < m |] ==> n<n*m"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   108
by (case_tac "m", auto)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   109
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   110
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   111
text{*Rewriting to pull differences out*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   112
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   113
lemma diff_diff_right [simp]: "k\<le>j --> i - (j - k) = i + (k::nat) - j"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   114
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   115
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   116
lemma diff_Suc_diff_eq1 [simp]: "k \<le> j ==> m - Suc (j - k) = m + k - Suc j"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   117
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   118
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   119
lemma diff_Suc_diff_eq2 [simp]: "k \<le> j ==> Suc (j - k) - m = Suc j - (k + m)"
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   120
by arith
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   121
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   122
(*The others are
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   123
      i - j - k = i - (j + k),
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   124
      k \<le> j ==> j - k + i = j + i - k,
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   125
      k \<le> j ==> i + (j - k) = i + j - k *)
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   126
declare diff_diff_left [simp] 
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   127
        diff_add_assoc [symmetric, simp]
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   128
        diff_add_assoc2[symmetric, simp]
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   129
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   130
text{*At present we prove no analogue of @{text not_less_Least} or @{text
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   131
Least_Suc}, since there appears to be no need.*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   132
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   133
ML
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   134
{*
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   135
val nat_diff_split = thm "nat_diff_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   136
val nat_diff_split_asm = thm "nat_diff_split_asm";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   137
val pred_nat_trancl_eq_le = thm "pred_nat_trancl_eq_le";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   138
val nat_diff_split = thm "nat_diff_split";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   139
val nat_diff_split_asm = thm "nat_diff_split_asm";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   140
val le_square = thm "le_square";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   141
val le_cube = thm "le_cube";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   142
val diff_less_mono = thm "diff_less_mono";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   143
val less_diff_conv = thm "less_diff_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   144
val le_diff_conv = thm "le_diff_conv";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   145
val le_diff_conv2 = thm "le_diff_conv2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   146
val diff_diff_cancel = thm "diff_diff_cancel";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   147
val le_add_diff = thm "le_add_diff";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   148
val diff_less = thm "diff_less";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   149
val diff_diff_eq = thm "diff_diff_eq";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   150
val eq_diff_iff = thm "eq_diff_iff";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   151
val less_diff_iff = thm "less_diff_iff";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   152
val le_diff_iff = thm "le_diff_iff";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   153
val diff_le_mono = thm "diff_le_mono";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   154
val diff_le_mono2 = thm "diff_le_mono2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   155
val diff_less_mono2 = thm "diff_less_mono2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   156
val diffs0_imp_equal = thm "diffs0_imp_equal";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   157
val one_less_mult = thm "one_less_mult";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   158
val n_less_m_mult_n = thm "n_less_m_mult_n";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   159
val n_less_n_mult_m = thm "n_less_n_mult_m";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   160
val diff_diff_right = thm "diff_diff_right";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   161
val diff_Suc_diff_eq1 = thm "diff_Suc_diff_eq1";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   162
val diff_Suc_diff_eq2 = thm "diff_Suc_diff_eq2";
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   163
*}
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   164
a9a762f586b5 removal of NatArith.ML and Product_Type.ML
paulson
parents: 15140
diff changeset
   165
10214
77349ed89f45 *** empty log message ***
nipkow
parents:
diff changeset
   166
end