src/HOL/Algebra/poly/LongDiv.ML
author ballarin
Sat, 10 Feb 2001 08:52:41 +0100
changeset 11093 62c2e0af1d30
parent 10962 cda180b1e2e0
child 11171 8aa53b4591a5
permissions -rw-r--r--
Changes to HOL/Algebra: - Axclasses consolidated (axiom one_not_zero). - Now uses summation operator setsum; SUM has been removed. - Priority of infix assoc changed to 50, in accordance to dvd.
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
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     7
Goal
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     8
  "!!p::('a::ring up). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     9
\    [| deg p <= deg r; deg q <= deg r; \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    10
\       coeff (deg r) p = - (coeff (deg r) q); deg r ~= 0 |] ==> \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    11
\    deg (p + q) < deg r";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    12
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
    13
by (arith_tac 2);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    14
by (rtac deg_aboveI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    15
by (strip_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    16
by (case_tac "deg r = m" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    17
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    18
by (Asm_full_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    19
(* case "deg q ~= m" *)
10959
paulson
parents: 10448
diff changeset
    20
by (subgoal_tac "deg p < m & deg q < m" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    21
by (asm_simp_tac (simpset() addsimps [deg_aboveD]) 1);
10959
paulson
parents: 10448
diff changeset
    22
by (arith_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    23
qed "deg_lcoeff_cancel";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    24
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    25
Goal
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    26
  "!!p::('a::ring up). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    27
\    [| deg p <= deg r; deg q <= deg r; \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    28
\       p ~= -q; coeff (deg r) p = - (coeff (deg r) q) |] ==> \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    29
\    deg (p + q) < deg r";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    30
by (rtac deg_lcoeff_cancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    31
by (REPEAT (atac 1));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    32
by (rtac classical 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    33
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    34
by (etac notE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    35
by (res_inst_tac [("p", "p")] up_repr2D 1 THEN atac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    36
by (res_inst_tac [("p", "q")] up_repr2D 1 THEN atac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    37
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    38
by (asm_full_simp_tac (simpset() addsimps [smult_l_minus]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    39
qed "deg_lcoeff_cancel2";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    40
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    41
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    42
  "!!g::('a::ring up). g ~= 0 ==> \
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    43
\    Ex (% (q, r, k). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    44
\      (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
    45
by (res_inst_tac [("P", "%f. Ex (% (q, r, k). \
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    46
\      (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
    47
  wf_induct 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    48
(* TO DO: replace by measure_induct *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    49
by (res_inst_tac [("f", "eucl_size")] wf_measure 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    50
by (case_tac "eucl_size x < eucl_size g" 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    51
by (res_inst_tac [("x", "(0, x, 0)")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    52
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    53
(* case "eucl_size x >= eucl_size g" *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    54
by (dres_inst_tac [("x", "lcoeff g *s x -- (lcoeff x *s monom (deg x - deg g)) * g")] spec 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    55
by (etac impE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    56
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
    57
by (case_tac "x = 0" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    58
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    59
by (asm_full_simp_tac (simpset() addsimps [eucl_size_def]) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    60
(* case "x ~= 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    61
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    62
by (asm_full_simp_tac (simpset() addsimps [eucl_size_def]) 1);
8006
paulson
parents: 7998
diff changeset
    63
by (Simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    64
by (rtac impI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    65
by (rtac deg_lcoeff_cancel2 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    66
  (* replace by linear arithmetic??? *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    67
  by (rtac le_trans 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    68
  by (rtac deg_smult_ring 1);
8006
paulson
parents: 7998
diff changeset
    69
  by (Simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    70
  by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    71
  by (rtac le_trans 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    72
  by (rtac deg_mult_ring 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    73
  by (rtac le_trans 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    74
  by (rtac add_le_mono1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    75
  by (rtac deg_smult_ring 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    76
(*  by (asm_simp_tac (simpset() addsimps [leI]) 1); *)
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    77
  by (Asm_simp_tac 1);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
    78
  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
    79
  by (arith_tac 1);
8707
paulson
parents: 8006
diff changeset
    80
by (Force_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    81
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    82
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
    83
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    84
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
    85
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
    86
    SUM_extend_below 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    87
by (rtac le_refl 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    88
by (asm_simp_tac (simpset() addsimps [less_not_refl2]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    89
by (asm_simp_tac (simpset() addsimps [diff_diff_right, leI, m_comm]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    90
(* end of subproof deg f1 < deg f *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    91
by (etac exE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    92
by (res_inst_tac [("x", "((%(q,r,k). ((lcoeff g ^ k * lcoeff x) *s monom (deg x - deg g) + q)) xa, (%(q,r,k). r) xa, (%(q,r,k). Suc k) xa)")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    93
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    94
by (rtac conjI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    95
by (dtac sym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    96
by (simp_tac (simpset() addsimps [l_distr, a_assoc]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    97
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    98
by (simp_tac (simpset() addsimps a_ac@[smult_l_distr, smult_r_distr, smult_r_minus, smult_assoc2, smult_assoc1]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    99
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   100
qed "long_div_eucl_size";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   101
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   102
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   103
  "!!g::('a::ring up). g ~= 0 ==> \
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   104
\    Ex (% (q, r, k). \
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   105
\      (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
   106
by (forw_inst_tac [("f", "f")]
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   107
  (simplify (simpset() addsimps [eucl_size_def]) long_div_eucl_size) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   108
by Auto_tac;
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   109
by (case_tac "aa = 0" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   110
by (Blast_tac 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   111
(* case "aa ~= 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   112
by (rotate_tac ~1 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   113
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   114
qed "long_div_ring";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   115
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   116
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   117
  "!!g::('a::ring up). [| g ~= 0; (lcoeff g) dvd <1> |] ==> \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   118
\    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
   119
by (forw_inst_tac [("f", "f")] long_div_ring 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   120
by (etac exE 1);
10448
wenzelm
parents: 8707
diff changeset
   121
by (res_inst_tac [("x", "((%(q,r,k). (inverse(lcoeff g ^k) *s q)) x, \
wenzelm
parents: 8707
diff changeset
   122
\  (%(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
   123
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   124
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   125
by (rtac conjI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   126
by (dtac sym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   127
by (asm_simp_tac (simpset() addsimps [smult_r_distr RS sym, smult_assoc2]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   128
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
   129
(* degree property *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   130
by (etac disjE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   131
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   132
by (rtac disjI2 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   133
by (rtac le_less_trans 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   134
by (rtac deg_smult_ring 1);
8006
paulson
parents: 7998
diff changeset
   135
by (Asm_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   136
qed "long_div_unit";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   137
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   138
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   139
  "!!g::('a::field up). g ~= 0 ==> \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   140
\    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
   141
by (rtac long_div_unit 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   142
by (assume_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   143
by (asm_simp_tac (simpset() addsimps [lcoeff_def, nonzero_lcoeff, field_ax]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   144
qed "long_div_theorem";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   145
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   146
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   147
  "!!g::('a::field up). [| g ~= 0; \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   148
\    f = q1 * g + r1; (r1 = 0 | deg r1 < deg g); \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   149
\    f = q2 * g + r2; (r2 = 0 | deg r2 < deg g) |] ==> q1 = q2";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   150
by (subgoal_tac "(q1 -- q2) * g = r2 -- r1" 1); (* 1 *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   151
by (thin_tac "f = ?x" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   152
by (thin_tac "f = ?x" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   153
by (rtac diff_zero_imp_eq 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   154
by (rtac classical 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   155
by (etac disjE 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   156
(* r1 = 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   157
by (etac disjE 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   158
(* r2 = 0 *)
8707
paulson
parents: 8006
diff changeset
   159
by (force_tac (claset() addDs [integral], simpset()) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   160
(* r2 ~= 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   161
by (dres_inst_tac [("f", "deg"), ("y", "r2 -- r1")] arg_cong 1);
8707
paulson
parents: 8006
diff changeset
   162
by (Force_tac 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   163
(* r1 ~=0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   164
by (etac disjE 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   165
(* r2 = 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   166
by (dres_inst_tac [("f", "deg"), ("y", "r2 -- r1")] arg_cong 1);
8707
paulson
parents: 8006
diff changeset
   167
by (Force_tac 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   168
(* r2 ~= 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   169
by (dres_inst_tac [("f", "deg"), ("y", "r2 -- r1")] arg_cong 1);
8707
paulson
parents: 8006
diff changeset
   170
by (Asm_full_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   171
by (dtac (order_eq_refl RS add_leD2) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   172
by (dtac leD 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   173
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
   174
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   175
(* proof of 1 *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   176
by (rtac diff_zero_imp_eq 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   177
by (Asm_full_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   178
by (dres_inst_tac [("a", "?x+?y")] eq_imp_diff_zero 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   179
by (asm_full_simp_tac (simpset() addsimps (l_distr::minus_add::l_minus::a_ac)) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   180
qed "long_div_quo_unique";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   181
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   182
Goal
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   183
  "!!g::('a::field up). [| g ~= 0; \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   184
\    f = q1 * g + r1; (r1 = 0 | deg r1 < deg g); \
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10962
diff changeset
   185
\    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
   186
by (subgoal_tac "q1 = q2" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   187
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   188
by (asm_full_simp_tac (simpset() addsimps [a_lcancel_eq]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   189
by (rtac long_div_quo_unique 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   190
by (REPEAT (atac 1));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   191
qed "long_div_rem_unique";