src/HOL/Import/HOLLightInt.thy
author huffman
Tue, 06 Sep 2011 19:03:41 -0700
changeset 44766 d4d33a4d7548
parent 43785 2bd54d4b5f3d
child 46783 3e89a5cab8d7
permissions -rw-r--r--
avoid using legacy theorem names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
(*  Title:      HOL/Import/HOLLightInt.thy
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
    Author:     Cezary Kaliszyk
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
*)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
header {* Compatibility theorems for HOL Light integers *}
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
theory HOLLightInt imports Main Real GCD begin
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
fun int_coprime where "int_coprime ((a :: int), (b :: int)) = coprime a b"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
lemma DEF_int_coprime:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
  "int_coprime = (\<lambda>u. \<exists>x y. ((fst u) * x) + ((snd u) * y) = int 1)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
  apply (auto simp add: fun_eq_iff)
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
    14
  apply (metis bezout_int mult_commute)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
  by (metis coprime_divisors_nat dvd_triv_left gcd_1_int gcd_add2_int)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
lemma INT_FORALL_POS:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
  "(\<forall>n. P (int n)) = (\<forall>i\<ge>(int 0). P i)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
  by (auto, drule_tac x="nat i" in spec) simp
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
lemma INT_LT_DISCRETE:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
  "(x < y) = (x + int 1 \<le> y)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
  by auto
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
lemma INT_ABS_MUL_1:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  "(abs (x * y) = int 1) = (abs x = int 1 \<and> abs y = int 1)"
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
    27
  by simp (metis dvd_mult_right zdvd1_eq abs_zmult_eq_1 abs_mult mult_1_right)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
lemma dest_int_rep:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
  "\<exists>(n :: nat). real (i :: int) = real n \<or> real i = - real n"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
  by (metis (full_types) of_int_of_nat real_eq_of_int real_of_nat_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
lemma DEF_int_add:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
  "op + = (\<lambda>u ua. floor (real u + real ua))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
  by simp
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
lemma DEF_int_sub:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
  "op - = (\<lambda>u ua. floor (real u - real ua))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  by simp
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
lemma DEF_int_mul:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
  "op * = (\<lambda>u ua. floor (real u * real ua))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
  by (metis floor_number_of number_of_is_id number_of_real_def real_eq_of_int real_of_int_mult)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
lemma DEF_int_abs:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
  "abs = (\<lambda>u. floor (abs (real u)))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
  by (metis floor_real_of_int real_of_int_abs)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
lemma DEF_int_sgn:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
  "sgn = (\<lambda>u. floor (sgn (real u)))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
  by (simp add: sgn_if fun_eq_iff)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
lemma int_sgn_th:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
  "real (sgn (x :: int)) = sgn (real x)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
  by (simp add: sgn_if)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
lemma DEF_int_max:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
  "max = (\<lambda>u ua. floor (max (real u) (real ua)))"
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
    59
  by (metis floor_real_of_int real_of_int_le_iff sup_absorb1 sup_commute sup_max linorder_linear)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
lemma int_max_th:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
  "real (max (x :: int) y) = max (real x) (real y)"
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
    63
  by (metis min_max.le_iff_sup min_max.sup_absorb1 real_of_int_le_iff linorder_linear)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
lemma DEF_int_min:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
  "min = (\<lambda>u ua. floor (min (real u) (real ua)))"
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
    67
  by (metis floor_real_of_int inf_absorb1 inf_absorb2 inf_int_def inf_real_def real_of_int_le_iff linorder_linear)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
lemma int_min_th:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  "real (min (x :: int) y) = min (real x) (real y)"
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
    71
  by (metis inf_absorb1 inf_absorb2 inf_int_def inf_real_def real_of_int_le_iff linorder_linear)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
lemma INT_IMAGE:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
  "(\<exists>n. x = int n) \<or> (\<exists>n. x = - int n)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  by (metis number_of_eq number_of_is_id of_int_of_nat)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
lemma DEF_int_pow:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
  "op ^ = (\<lambda>u ua. floor (real u ^ ua))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
  by (simp add: floor_power)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
lemma DEF_int_divides:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
  "op dvd = (\<lambda>(u :: int) ua. \<exists>x. ua = u * x)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
  by (metis dvdE dvdI)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
lemma DEF_int_divides':
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
  "(a :: int) dvd b = (\<exists>x. b = a * x)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
  by (metis dvdE dvdI)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
definition "int_mod (u :: int) ua ub = (u dvd (ua - ub))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
lemma int_mod_def':
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
  "int_mod = (\<lambda>u ua ub. (u dvd (ua - ub)))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
  by (simp add: int_mod_def_raw)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
lemma int_congruent:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
  "\<forall>x xa xb. int_mod xb x xa = (\<exists>d. x - xa = xb * d)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
  unfolding int_mod_def'
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
  by (auto simp add: DEF_int_divides')
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
lemma int_congruent':
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
  "\<forall>(x :: int) y n. (n dvd x - y) = (\<exists>d. x - y = n * d)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
  using int_congruent[unfolded int_mod_def] .
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
fun int_gcd where
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
  "int_gcd ((a :: int), (b :: int)) = gcd a b"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
definition "hl_mod (k\<Colon>int) (l\<Colon>int) = (if 0 \<le> l then k mod l else k mod - l)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
lemma hl_mod_nonneg:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
  "b \<noteq> 0 \<Longrightarrow> hl_mod a b \<ge> 0"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
  by (simp add: hl_mod_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
lemma hl_mod_lt_abs:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  "b \<noteq> 0 \<Longrightarrow> hl_mod a b < abs b"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
  by (simp add: hl_mod_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
definition "hl_div k l = (if 0 \<le> l then k div l else -(k div (-l)))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
lemma hl_mod_div:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
  "n \<noteq> (0\<Colon>int) \<Longrightarrow> m = hl_div m n * n + hl_mod m n"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
  unfolding hl_div_def hl_mod_def
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
   122
  by auto (metis zmod_zdiv_equality mult_commute mult_minus_left)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
lemma sth:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
  "(\<forall>(x :: int) y z. x + (y + z) = x + y + z) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
   (\<forall>(x :: int) y. x + y = y + x) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
   (\<forall>(x :: int). int 0 + x = x) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
   (\<forall>(x :: int) y z. x * (y * z) = x * y * z) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
   (\<forall>(x :: int) y. x * y = y * x) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
   (\<forall>(x :: int). int 1 * x = x) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
   (\<forall>(x :: int). int 0 * x = int 0) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
   (\<forall>(x :: int) y z. x * (y + z) = x * y + x * z) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
   (\<forall>(x :: int). x ^ 0 = int 1) \<and> (\<forall>(x :: int) n. x ^ Suc n = x * x ^ n)"
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
   134
  by (simp_all add: right_distrib)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
lemma INT_DIVISION:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
  "n ~= int 0 \<Longrightarrow> m = hl_div m n * n + hl_mod m n \<and> int 0 \<le> hl_mod m n \<and> hl_mod m n < abs n"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
  by (auto simp add: hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
lemma INT_DIVMOD_EXIST_0:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
  "\<exists>q r. if n = int 0 then q = int 0 \<and> r = m
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
         else int 0 \<le> r \<and> r < abs n \<and> m = q * n + r"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
  apply (rule_tac x="hl_div m n" in exI)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
  apply (rule_tac x="hl_mod m n" in exI)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
  apply (auto simp add: hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
  unfolding hl_div_def hl_mod_def
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
  by auto
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
lemma DEF_div:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
  "hl_div = (SOME q. \<exists>r. \<forall>m n. if n = int 0 then q m n = int 0 \<and> r m n = m
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
     else (int 0) \<le> (r m n) \<and> (r m n) < (abs n) \<and> m = ((q m n) * n) + (r m n))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
  apply (rule some_equality[symmetric])
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
  apply (rule_tac x="hl_mod" in exI)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
  apply (auto simp add: fun_eq_iff hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
  apply (simp add: hl_div_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
  apply (simp add: hl_mod_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
  apply (drule_tac x="x" in spec)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
  apply (drule_tac x="xa" in spec)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
  apply (case_tac "0 = xa")
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
  apply (simp add: hl_mod_def hl_div_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
  apply (case_tac "xa > 0")
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
  apply (simp add: hl_mod_def hl_div_def)
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
   163
  apply (metis comm_semiring_1_class.normalizing_semiring_rules(24) div_mult_self2 not_less_iff_gr_or_eq order_less_le add_0 zdiv_eq_0_iff mult_commute)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
  apply (simp add: hl_mod_def hl_div_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
  by (metis add.comm_neutral add_pos_nonneg div_mult_self1 less_minus_iff minus_add minus_add_cancel minus_minus mult_zero_right not_square_less_zero zdiv_eq_0_iff zdiv_zminus2)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   166
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
lemma DEF_rem:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
  "hl_mod = (SOME r. \<forall>m n. if n = int 0 then
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
     (if 0 \<le> n then m div n else - (m div - n)) = int 0 \<and> r m n = m
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
     else int 0 \<le> r m n \<and> r m n < abs n \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
            m = (if 0 \<le> n then m div n else - (m div - n)) * n + r m n)"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
  apply (rule some_equality[symmetric])
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
  apply (fold hl_div_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
  apply (auto simp add: fun_eq_iff hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
  apply (simp add: hl_div_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
  apply (simp add: hl_mod_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
  apply (drule_tac x="x" in spec)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
  apply (drule_tac x="xa" in spec)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
  apply (case_tac "0 = xa")
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
  apply (simp add: hl_mod_def hl_div_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
  apply (case_tac "xa > 0")
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
  apply (simp add: hl_mod_def hl_div_def)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   183
  apply (metis add_left_cancel mod_div_equality)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   184
  apply (simp add: hl_mod_def hl_div_def)
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
   185
  by (metis minus_mult_right mod_mult_self2 mod_pos_pos_trivial add_commute zminus_zmod zmod_zminus2 mult_commute)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
lemma DEF_int_gcd:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
  "int_gcd = (SOME d. \<forall>a b. (int 0) \<le> (d (a, b)) \<and> (d (a, b)) dvd a \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
       (d (a, b)) dvd b \<and> (\<exists>x y. d (a, b) = (a * x) + (b * y)))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
  apply (rule some_equality[symmetric])
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
  apply auto
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 43785
diff changeset
   192
  apply (metis bezout_int mult_commute)
43785
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
  apply (auto simp add: fun_eq_iff)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
  apply (drule_tac x="a" in spec)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  apply (drule_tac x="b" in spec)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
  using gcd_greatest_int zdvd_antisym_nonneg
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   197
  by auto
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   199
definition "eqeq x y (r :: 'a \<Rightarrow> 'b \<Rightarrow> bool) = r x y"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   200
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   201
lemma INT_INTEGRAL:
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   202
  "(\<forall>x. int 0 * x = int 0) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   203
   (\<forall>(x :: int) y z. (x + y = x + z) = (y = z)) \<and>
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
   (\<forall>(w :: int) x y z. (w * y + x * z = w * z + x * y) = (w = x \<or> y = z))"
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   205
  by (auto simp add: crossproduct_eq)
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   206
2bd54d4b5f3d HOL/Import for HOLLight revival: Proper theory headers, update generation scripts to SVN version of HOL Light, add some constant maps and compatibility theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   207
end