src/HOL/Algebra/poly/LongDiv.ML
author wenzelm
Mon, 29 Aug 2005 16:18:04 +0200
changeset 17184 3d80209e9a53
parent 14723 b77ce15b625a
child 20282 49c312eaaa11
permissions -rw-r--r--
use AList operations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     1
(*
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     2
    Long division of polynomials
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     3
    $Id$
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     4
    Author: Clemens Ballarin, started 23 June 1999
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     5
*)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     6
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
     7
(* legacy bindings and theorems *)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
     8
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
     9
val deg_aboveI = thm "deg_aboveI";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    10
val smult_l_minus = thm "smult_l_minus";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    11
val deg_monom_ring = thm "deg_monom_ring";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    12
val deg_smult_ring = thm "deg_smult_ring";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    13
val smult_l_distr = thm "smult_l_distr";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    14
val smult_r_distr = thm "smult_r_distr";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    15
val smult_r_minus = thm "smult_r_minus";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    16
val smult_assoc2 = thm "smult_assoc2";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    17
val smult_assoc1 = thm "smult_assoc1";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    18
val monom_mult_smult = thm "monom_mult_smult";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    19
val field_ax = thm "field_ax";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    20
val lcoeff_nonzero = thm "lcoeff_nonzero";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    21
14723
b77ce15b625a moved first lemma in LongDiv.ML to LongDiv.thy
obua
parents: 13783
diff changeset
    22
val lcoeff_def = thm "lcoeff_def";
b77ce15b625a moved first lemma in LongDiv.ML to LongDiv.thy
obua
parents: 13783
diff changeset
    23
val eucl_size_def = thm "eucl_size_def";
b77ce15b625a moved first lemma in LongDiv.ML to LongDiv.thy
obua
parents: 13783
diff changeset
    24
b77ce15b625a moved first lemma in LongDiv.ML to LongDiv.thy
obua
parents: 13783
diff changeset
    25
val SUM_shrink_below_lemma = thm "SUM_shrink_below_lemma";
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    26
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    27
Goal
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    28
  "!! f::(nat=>'a::ring). \
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    29
\    [| m <= n; !!i. i < m ==> f i = 0; P (setsum (%i. f (i+m)) {..n-m}) |] \
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    30
\    ==> P (setsum f {..n})";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    31
by (asm_full_simp_tac 
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    32
    (simpset() addsimps [SUM_shrink_below_lemma, add_diff_inverse, leD]) 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    33
qed "SUM_extend_below";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    34
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    35
Goal
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    36
  "!! p::'a::ring up. \
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    37
\  [| deg p <= n; P (setsum (%i. monom (coeff p i) i) {..n}) |] \
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    38
\    ==> P p";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    39
by (asm_full_simp_tac (simpset() addsimps [thm "up_repr_le"]) 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    40
qed "up_repr2D";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    41
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    42
(* Start of LongDiv *)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    43
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    44
Goal
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    45
  "!!p::('a::ring up). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    46
\    [| deg p <= deg r; deg q <= deg r; \
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    47
\       coeff p (deg r) = - (coeff q (deg r)); deg r ~= 0 |] ==> \
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    48
\    deg (p + q) < deg r";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    49
by (res_inst_tac [("j", "deg r - 1")] le_less_trans 1);
10962
cda180b1e2e0 deleted several obsolete lemmas from NatArith.ML
paulson
parents: 10959
diff changeset
    50
by (arith_tac 2);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    51
by (rtac deg_aboveI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    52
by (strip_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    53
by (case_tac "deg r = m" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    54
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    55
by (Asm_full_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    56
(* case "deg q ~= m" *)
10959
paulson
parents: 10448
diff changeset
    57
by (subgoal_tac "deg p < m & deg q < m" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    58
by (asm_simp_tac (simpset() addsimps [deg_aboveD]) 1);
10959
paulson
parents: 10448
diff changeset
    59
by (arith_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    60
qed "deg_lcoeff_cancel";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    61
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    62
Goal
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    63
  "!!p::('a::ring up). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    64
\    [| deg p <= deg r; deg q <= deg r; \
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    65
\       p ~= -q; coeff p (deg r) = - (coeff q (deg r)) |] ==> \
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    66
\    deg (p + q) < deg r";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    67
by (rtac deg_lcoeff_cancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    68
by (REPEAT (atac 1));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    69
by (rtac classical 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    70
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    71
by (etac notE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    72
by (res_inst_tac [("p", "p")] up_repr2D 1 THEN atac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    73
by (res_inst_tac [("p", "q")] up_repr2D 1 THEN atac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    74
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    75
by (asm_full_simp_tac (simpset() addsimps [smult_l_minus]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    76
qed "deg_lcoeff_cancel2";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    77
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    78
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    79
  "!!g::('a::ring up). g ~= 0 ==> \
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    80
\    Ex (% (q, r, k). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    81
\      (lcoeff g)^k *s f = q * g + r & (eucl_size r < eucl_size g))";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    82
by (res_inst_tac [("P", "%f. Ex (% (q, r, k). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    83
\      (lcoeff g)^k *s f = q * g + r & (eucl_size r < eucl_size g))")]
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    84
  wf_induct 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    85
(* TO DO: replace by measure_induct *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    86
by (res_inst_tac [("f", "eucl_size")] wf_measure 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    87
by (case_tac "eucl_size x < eucl_size g" 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    88
by (res_inst_tac [("x", "(0, x, 0)")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    89
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    90
(* case "eucl_size x >= eucl_size g" *)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
    91
by (dres_inst_tac [("x", "lcoeff g *s x - (monom (lcoeff x) (deg x - deg g)) * g")] spec 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    92
by (etac impE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    93
by (full_simp_tac (simpset() addsimps [inv_image_def, measure_def, lcoeff_def]) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    94
by (case_tac "x = 0" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    95
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    96
by (asm_full_simp_tac (simpset() addsimps [eucl_size_def]) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    97
(* case "x ~= 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    98
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    99
by (asm_full_simp_tac (simpset() addsimps [eucl_size_def]) 1);
13783
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   100
(* by (Simp_tac 1); *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   101
by (rtac impI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   102
by (rtac deg_lcoeff_cancel2 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   103
  (* replace by linear arithmetic??? *)
13783
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   104
  by (rtac le_trans 2);
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   105
  by (rtac deg_smult_ring 2);
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   106
  by (Simp_tac 2);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   107
  by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   108
  by (rtac le_trans 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   109
  by (rtac deg_mult_ring 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   110
  by (rtac le_trans 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   111
(**)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   112
  by (rtac add_le_mono 1); by (rtac le_refl 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   113
    (* term order forces to use this instead of add_le_mono1 *)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   114
  by (rtac deg_monom_ring 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   115
  by (Asm_simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   116
(**)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   117
(*
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   118
  by (rtac add_le_mono1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   119
  by (rtac deg_smult_ring 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   120
(*  by (asm_simp_tac (simpset() addsimps [leI]) 1); *)
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   121
  by (Asm_simp_tac 1);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   122
  by (cut_inst_tac [("m", "deg x - deg g"), ("'a", "'a")] deg_monom_ring 1);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   123
  by (arith_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   124
*)
8707
paulson
parents: 8006
diff changeset
   125
by (Force_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   126
by (Simp_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   127
(**)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   128
  (* This change is probably caused by application of commutativity *)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   129
by (res_inst_tac [("m", "deg g"), ("n", "deg x")] SUM_extend 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   130
by (Simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   131
by (Asm_simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   132
by (arith_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   133
by (res_inst_tac [("m", "deg g"), ("n", "deg g")] SUM_extend_below 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   134
by (rtac le_refl 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   135
by (Asm_simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   136
by (arith_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   137
by (Simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   138
(**)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   139
(*
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   140
by (res_inst_tac [("m", "deg x - deg g"), ("n", "deg x")] SUM_extend 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   141
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   142
by (asm_simp_tac (simpset() addsimps [less_not_refl2 RS not_sym]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   143
by (res_inst_tac [("m", "deg x - deg g"), ("n", "deg x - deg g")]
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   144
    SUM_extend_below 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   145
by (rtac le_refl 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   146
by (asm_simp_tac (simpset() addsimps [less_not_refl2]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   147
by (asm_simp_tac (simpset() addsimps [diff_diff_right, leI, m_comm]) 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   148
*)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   149
(* end of subproof deg f1 < deg f *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   150
by (etac exE 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   151
by (res_inst_tac [("x", "((%(q,r,k). (monom (lcoeff g ^ k * lcoeff x) (deg x - deg g) + q)) xa, (%(q,r,k). r) xa, (%(q,r,k). Suc k) xa)")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   152
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   153
by (dtac sym 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   154
by (simp_tac (simpset() addsimps [l_distr, a_assoc]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   155
  delsimprocs [ring_simproc]) 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   156
by (asm_simp_tac (simpset() delsimprocs [ring_simproc]) 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   157
by (simp_tac (simpset() addsimps [minus_def, smult_r_distr, smult_r_minus,
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   158
    monom_mult_smult, smult_assoc1, smult_assoc2]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   159
  delsimprocs [ring_simproc]) 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   160
by (Simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   161
qed "long_div_eucl_size";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   162
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   163
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   164
  "!!g::('a::ring up). g ~= 0 ==> \
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   165
\    Ex (% (q, r, k). \
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   166
\      (lcoeff g)^k *s f = q * g + r & (r = 0 | deg r < deg g))";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   167
by (forw_inst_tac [("f", "f")]
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   168
  (simplify (simpset() addsimps [eucl_size_def]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   169
    delsimprocs [ring_simproc]) long_div_eucl_size) 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   170
by (auto_tac (claset(), simpset() delsimprocs [ring_simproc]));
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   171
by (case_tac "aa = 0" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   172
by (Blast_tac 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   173
(* case "aa ~= 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   174
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   175
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   176
qed "long_div_ring";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   177
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   178
(* Next one fails *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   179
Goal
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   180
  "!!g::('a::ring up). [| g ~= 0; (lcoeff g) dvd 1 |] ==> \
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   181
\    Ex (% (q, r). f = q * g + r & (r = 0 | deg r < deg g))";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   182
by (forw_inst_tac [("f", "f")] long_div_ring 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   183
by (etac exE 1);
10448
wenzelm
parents: 8707
diff changeset
   184
by (res_inst_tac [("x", "((%(q,r,k). (inverse(lcoeff g ^k) *s q)) x, \
wenzelm
parents: 8707
diff changeset
   185
\  (%(q,r,k). inverse(lcoeff g ^k) *s r) x)")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   186
by (Clarify_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   187
(* by (Simp_tac 1); *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   188
by (rtac conjI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   189
by (dtac sym 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   190
by (asm_simp_tac (simpset() addsimps [smult_r_distr RS sym, smult_assoc2]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   191
  delsimprocs [ring_simproc]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   192
by (asm_simp_tac (simpset() addsimps [l_inverse_ring, unit_power, smult_assoc1 RS sym]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   193
(* degree property *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   194
by (etac disjE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   195
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   196
by (rtac disjI2 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   197
by (rtac le_less_trans 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   198
by (rtac deg_smult_ring 1);
8006
paulson
parents: 7998
diff changeset
   199
by (Asm_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   200
qed "long_div_unit";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   201
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   202
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   203
  "!!g::('a::field up). g ~= 0 ==> \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   204
\    Ex (% (q, r). f = q * g + r & (r = 0 | deg r < deg g))";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   205
by (rtac long_div_unit 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   206
by (assume_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   207
by (asm_simp_tac (simpset() addsimps [lcoeff_def, lcoeff_nonzero, field_ax]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   208
qed "long_div_theorem";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   209
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   210
Goal "- (0::'a::ring) = 0";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   211
by (Simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   212
val uminus_zero = result();
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   213
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   214
Goal "!!a::'a::ring. a - b = 0 ==> a = b";
13783
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   215
by (res_inst_tac [("s", "a - (a - b)")] trans 1);
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   216
by (asm_simp_tac (simpset() delsimprocs [ring_simproc]) 1);
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   217
by (Simp_tac 1);
3294f727e20d Fixed term order for normal form in rings.
ballarin
parents: 13735
diff changeset
   218
by (Simp_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   219
val diff_zero_imp_eq = result();
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   220
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   221
Goal "!!a::'a::ring. a = b ==> a + (-b) = 0";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   222
by (Asm_simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   223
val eq_imp_diff_zero = result();
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   224
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   225
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   226
  "!!g::('a::field up). [| g ~= 0; \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   227
\    f = q1 * g + r1; (r1 = 0 | deg r1 < deg g); \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   228
\    f = q2 * g + r2; (r2 = 0 | deg r2 < deg g) |] ==> q1 = q2";
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   229
by (subgoal_tac "(q1 - q2) * g = r2 - r1" 1); (* 1 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   230
by (thin_tac "f = ?x" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   231
by (thin_tac "f = ?x" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   232
by (rtac diff_zero_imp_eq 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   233
by (rtac classical 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   234
by (etac disjE 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   235
(* r1 = 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   236
by (etac disjE 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   237
(* r2 = 0 *)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   238
by (asm_full_simp_tac (simpset()
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   239
  addsimps [thm "integral_iff", minus_def, l_zero, uminus_zero]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   240
  delsimprocs [ring_simproc]) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   241
(* r2 ~= 0 *)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   242
by (dres_inst_tac [("f", "deg"), ("y", "r2 - r1")] arg_cong 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   243
by (asm_full_simp_tac (simpset() addsimps [minus_def, l_zero, uminus_zero]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   244
  delsimprocs [ring_simproc]) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   245
(* r1 ~=0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   246
by (etac disjE 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   247
(* r2 = 0 *)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   248
by (dres_inst_tac [("f", "deg"), ("y", "r2 - r1")] arg_cong 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   249
by (asm_full_simp_tac (simpset() addsimps [minus_def, l_zero, uminus_zero]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   250
  delsimprocs [ring_simproc]) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   251
(* r2 ~= 0 *)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   252
by (dres_inst_tac [("f", "deg"), ("y", "r2 - r1")] arg_cong 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   253
by (asm_full_simp_tac (simpset() addsimps [minus_def]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   254
  delsimprocs [ring_simproc]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   255
by (dtac (order_eq_refl RS add_leD2) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   256
by (dtac leD 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   257
by (etac notE 1 THEN rtac (deg_add RS le_less_trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   258
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   259
(* proof of 1 *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   260
by (rtac diff_zero_imp_eq 1);
13601
fd3e3d6b37b2 Adapted to new simplifier.
berghofe
parents: 11171
diff changeset
   261
by (hyp_subst_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   262
by (dres_inst_tac [("a", "?x+?y")] eq_imp_diff_zero 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   263
by (Asm_full_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   264
qed "long_div_quo_unique";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   265
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   266
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   267
  "!!g::('a::field up). [| g ~= 0; \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   268
\    f = q1 * g + r1; (r1 = 0 | deg r1 < deg g); \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   269
\    f = q2 * g + r2; (r2 = 0 | deg r2 < deg g) |] ==> r1 = r2";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   270
by (subgoal_tac "q1 = q2" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   271
by (Clarify_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   272
by (res_inst_tac [("a", "q2 * g + r1 - q2 * g"), ("b", "q2 * g + r2 - q2 * g")]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   273
  box_equals 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   274
by (Asm_full_simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   275
by (Simp_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13601
diff changeset
   276
by (Simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   277
by (rtac long_div_quo_unique 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   278
by (REPEAT (atac 1));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   279
qed "long_div_rem_unique";