| 7998 |      1 | (*
 | 
|  |      2 |     Experimental theory: long division of polynomials
 | 
|  |      3 |     $Id$
 | 
|  |      4 |     Author: Clemens Ballarin, started 23 June 1999
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 16417 |      7 | theory LongDiv imports PolyHomo begin
 | 
| 7998 |      8 | 
 | 
| 21423 |      9 | definition
 | 
|  |     10 |   lcoeff :: "'a::ring up => 'a" where
 | 
|  |     11 |   "lcoeff p = coeff p (deg p)"
 | 
| 7998 |     12 | 
 | 
| 21423 |     13 | definition
 | 
|  |     14 |   eucl_size :: "'a::zero up => nat" where
 | 
|  |     15 |   "eucl_size p = (if p = 0 then 0 else deg p + 1)"
 | 
| 14723 |     16 | 
 | 
|  |     17 | lemma SUM_shrink_below_lemma:
 | 
|  |     18 |   "!! f::(nat=>'a::ring). (ALL i. i < m --> f i = 0) --> 
 | 
|  |     19 |   setsum (%i. f (i+m)) {..d} = setsum f {..m+d}"
 | 
|  |     20 |   apply (induct_tac d)
 | 
| 15481 |     21 |    apply (induct_tac m)
 | 
| 21423 |     22 |     apply simp
 | 
|  |     23 |    apply force
 | 
|  |     24 |   apply (simp add: ab_semigroup_add_class.add_commute [of m]) 
 | 
|  |     25 |   done
 | 
|  |     26 | 
 | 
|  |     27 | lemma SUM_extend_below: 
 | 
|  |     28 |   "!! f::(nat=>'a::ring).  
 | 
|  |     29 |      [| m <= n; !!i. i < m ==> f i = 0; P (setsum (%i. f (i+m)) {..n-m}) |]  
 | 
|  |     30 |      ==> P (setsum f {..n})"
 | 
|  |     31 |   by (simp add: SUM_shrink_below_lemma add_diff_inverse leD)
 | 
|  |     32 | 
 | 
|  |     33 | lemma up_repr2D: 
 | 
|  |     34 |   "!! p::'a::ring up.  
 | 
|  |     35 |    [| deg p <= n; P (setsum (%i. monom (coeff p i) i) {..n}) |]  
 | 
|  |     36 |      ==> P p"
 | 
|  |     37 |   by (simp add: up_repr_le)
 | 
|  |     38 | 
 | 
|  |     39 | 
 | 
|  |     40 | (* Start of LongDiv *)
 | 
|  |     41 | 
 | 
|  |     42 | lemma deg_lcoeff_cancel: 
 | 
|  |     43 |   "!!p::('a::ring up).  
 | 
|  |     44 |      [| deg p <= deg r; deg q <= deg r;  
 | 
|  |     45 |         coeff p (deg r) = - (coeff q (deg r)); deg r ~= 0 |] ==>  
 | 
|  |     46 |      deg (p + q) < deg r"
 | 
|  |     47 |   apply (rule_tac j = "deg r - 1" in le_less_trans)
 | 
|  |     48 |    prefer 2
 | 
|  |     49 |    apply arith
 | 
|  |     50 |   apply (rule deg_aboveI)
 | 
|  |     51 |   apply (case_tac "deg r = m")
 | 
|  |     52 |    apply clarify
 | 
|  |     53 |    apply simp
 | 
|  |     54 |   (* case "deg q ~= m" *)
 | 
|  |     55 |    apply (subgoal_tac "deg p < m & deg q < m")
 | 
|  |     56 |     apply (simp (no_asm_simp) add: deg_aboveD)
 | 
|  |     57 |   apply arith
 | 
|  |     58 |   done
 | 
|  |     59 | 
 | 
|  |     60 | lemma deg_lcoeff_cancel2: 
 | 
|  |     61 |   "!!p::('a::ring up).  
 | 
|  |     62 |      [| deg p <= deg r; deg q <= deg r;  
 | 
|  |     63 |         p ~= -q; coeff p (deg r) = - (coeff q (deg r)) |] ==>  
 | 
|  |     64 |      deg (p + q) < deg r"
 | 
|  |     65 |   apply (rule deg_lcoeff_cancel)
 | 
|  |     66 |      apply assumption+
 | 
|  |     67 |   apply (rule classical)
 | 
|  |     68 |   apply clarify
 | 
|  |     69 |   apply (erule notE)
 | 
|  |     70 |   apply (rule_tac p = p in up_repr2D, assumption)
 | 
|  |     71 |   apply (rule_tac p = q in up_repr2D, assumption)
 | 
|  |     72 |   apply (rotate_tac -1)
 | 
|  |     73 |   apply (simp add: smult_l_minus)
 | 
|  |     74 |   done
 | 
|  |     75 | 
 | 
|  |     76 | lemma long_div_eucl_size: 
 | 
|  |     77 |   "!!g::('a::ring up). g ~= 0 ==>  
 | 
|  |     78 |      Ex (% (q, r, k).  
 | 
|  |     79 |        (lcoeff g)^k *s f = q * g + r & (eucl_size r < eucl_size g))"
 | 
|  |     80 |   apply (rule_tac P = "%f. Ex (% (q, r, k) . (lcoeff g) ^k *s f = q * g + r & (eucl_size r < eucl_size g))" in wf_induct)
 | 
|  |     81 |   (* TO DO: replace by measure_induct *)
 | 
|  |     82 |   apply (rule_tac f = eucl_size in wf_measure)
 | 
|  |     83 |   apply (case_tac "eucl_size x < eucl_size g")
 | 
|  |     84 |    apply (rule_tac x = "(0, x, 0)" in exI)
 | 
|  |     85 |    apply (simp (no_asm_simp))
 | 
|  |     86 |   (* case "eucl_size x >= eucl_size g" *)
 | 
|  |     87 |   apply (drule_tac x = "lcoeff g *s x - (monom (lcoeff x) (deg x - deg g)) * g" in spec)
 | 
|  |     88 |   apply (erule impE)
 | 
|  |     89 |    apply (simp (no_asm_use) add: inv_image_def measure_def lcoeff_def)
 | 
|  |     90 |    apply (case_tac "x = 0")
 | 
|  |     91 |     apply (rotate_tac -1)
 | 
|  |     92 |     apply (simp add: eucl_size_def)
 | 
|  |     93 |     (* case "x ~= 0 *)
 | 
|  |     94 |     apply (rotate_tac -1)
 | 
|  |     95 |    apply (simp add: eucl_size_def)
 | 
|  |     96 |    apply (rule impI)
 | 
|  |     97 |    apply (rule deg_lcoeff_cancel2)
 | 
|  |     98 |   (* replace by linear arithmetic??? *)
 | 
|  |     99 |       apply (rule_tac [2] le_trans)
 | 
|  |    100 |        apply (rule_tac [2] deg_smult_ring)
 | 
|  |    101 |       prefer 2
 | 
|  |    102 |       apply simp
 | 
|  |    103 |      apply (simp (no_asm))
 | 
|  |    104 |      apply (rule le_trans)
 | 
|  |    105 |       apply (rule deg_mult_ring)
 | 
|  |    106 |      apply (rule le_trans)
 | 
|  |    107 | (**)
 | 
|  |    108 |       apply (rule add_le_mono)
 | 
|  |    109 |        apply (rule le_refl)
 | 
|  |    110 |     (* term order forces to use this instead of add_le_mono1 *)
 | 
|  |    111 |       apply (rule deg_monom_ring)
 | 
|  |    112 |      apply (simp (no_asm_simp))
 | 
|  |    113 |     apply force
 | 
|  |    114 |    apply (simp (no_asm))
 | 
|  |    115 | (**)
 | 
|  |    116 |    (* This change is probably caused by application of commutativity *)
 | 
|  |    117 |    apply (rule_tac m = "deg g" and n = "deg x" in SUM_extend)
 | 
|  |    118 |      apply (simp (no_asm))
 | 
|  |    119 |     apply (simp (no_asm_simp))
 | 
|  |    120 |     apply arith
 | 
|  |    121 |    apply (rule_tac m = "deg g" and n = "deg g" in SUM_extend_below)
 | 
|  |    122 |      apply (rule le_refl)
 | 
|  |    123 |     apply (simp (no_asm_simp))
 | 
|  |    124 |     apply arith
 | 
|  |    125 |    apply (simp (no_asm))
 | 
|  |    126 | (**)
 | 
|  |    127 | (* end of subproof deg f1 < deg f *)
 | 
|  |    128 |   apply (erule exE)
 | 
|  |    129 |   apply (rule_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) " in exI)
 | 
|  |    130 |   apply clarify
 | 
|  |    131 |   apply (drule sym)
 | 
|  |    132 |   apply (tactic {* simp_tac (simpset() addsimps [thm "l_distr", thm "a_assoc"]
 | 
|  |    133 |     delsimprocs [ring_simproc]) 1 *})
 | 
|  |    134 |   apply (tactic {* asm_simp_tac (simpset() delsimprocs [ring_simproc]) 1 *})
 | 
|  |    135 |   apply (tactic {* simp_tac (simpset () addsimps [thm "minus_def", thm "smult_r_distr",
 | 
|  |    136 |     thm "smult_r_minus", thm "monom_mult_smult", thm "smult_assoc1", thm "smult_assoc2"]
 | 
|  |    137 |     delsimprocs [ring_simproc]) 1 *})
 | 
|  |    138 |   apply simp
 | 
|  |    139 |   done
 | 
|  |    140 | 
 | 
|  |    141 | ML {* simplify (simpset() addsimps [thm "eucl_size_def"]
 | 
|  |    142 |     delsimprocs [ring_simproc]) (thm "long_div_eucl_size") *}
 | 
|  |    143 | 
 | 
|  |    144 | thm long_div_eucl_size [simplified]
 | 
|  |    145 | 
 | 
|  |    146 | lemma long_div_ring: 
 | 
|  |    147 |   "!!g::('a::ring up). g ~= 0 ==>  
 | 
|  |    148 |      Ex (% (q, r, k).  
 | 
|  |    149 |        (lcoeff g)^k *s f = q * g + r & (r = 0 | deg r < deg g))"
 | 
|  |    150 |   apply (tactic {* forw_inst_tac [("f", "f")]
 | 
|  |    151 |     (simplify (simpset() addsimps [thm "eucl_size_def"]
 | 
|  |    152 |       delsimprocs [ring_simproc]) (thm "long_div_eucl_size")) 1 *})
 | 
|  |    153 |   apply (tactic {* auto_tac (claset(), simpset() delsimprocs [ring_simproc]) *})
 | 
|  |    154 |   apply (case_tac "aa = 0")
 | 
|  |    155 |    apply blast
 | 
|  |    156 |   (* case "aa ~= 0 *)
 | 
|  |    157 |   apply (rotate_tac -1)
 | 
|  |    158 |   apply auto
 | 
|  |    159 |   done
 | 
|  |    160 | 
 | 
|  |    161 | (* Next one fails *)
 | 
|  |    162 | lemma long_div_unit: 
 | 
|  |    163 |   "!!g::('a::ring up). [| g ~= 0; (lcoeff g) dvd 1 |] ==>  
 | 
|  |    164 |      Ex (% (q, r). f = q * g + r & (r = 0 | deg r < deg g))"
 | 
|  |    165 |   apply (frule_tac f = "f" in long_div_ring)
 | 
|  |    166 |   apply (erule exE)
 | 
|  |    167 |   apply (rule_tac x = "((% (q,r,k) . (inverse (lcoeff g ^k) *s q)) x, (% (q,r,k) . inverse (lcoeff g ^k) *s r) x) " in exI)
 | 
|  |    168 |   apply clarify
 | 
|  |    169 |   apply (rule conjI)
 | 
|  |    170 |    apply (drule sym)
 | 
|  |    171 |    apply (tactic {* asm_simp_tac
 | 
|  |    172 |      (simpset() addsimps [thm "smult_r_distr" RS sym, thm "smult_assoc2"]
 | 
|  |    173 |      delsimprocs [ring_simproc]) 1 *})
 | 
|  |    174 |    apply (simp (no_asm_simp) add: l_inverse_ring unit_power smult_assoc1 [symmetric])
 | 
|  |    175 |   (* degree property *)
 | 
|  |    176 |    apply (erule disjE)
 | 
|  |    177 |     apply (simp (no_asm_simp))
 | 
|  |    178 |   apply (rule disjI2)
 | 
|  |    179 |   apply (rule le_less_trans)
 | 
|  |    180 |    apply (rule deg_smult_ring)
 | 
|  |    181 |   apply (simp (no_asm_simp))
 | 
|  |    182 |   done
 | 
|  |    183 | 
 | 
|  |    184 | lemma long_div_theorem: 
 | 
|  |    185 |   "!!g::('a::field up). g ~= 0 ==>  
 | 
|  |    186 |      Ex (% (q, r). f = q * g + r & (r = 0 | deg r < deg g))"
 | 
|  |    187 |   apply (rule long_div_unit)
 | 
|  |    188 |    apply assumption
 | 
|  |    189 |   apply (simp (no_asm_simp) add: lcoeff_def lcoeff_nonzero field_ax)
 | 
|  |    190 |   done
 | 
|  |    191 | 
 | 
|  |    192 | lemma uminus_zero: "- (0::'a::ring) = 0"
 | 
|  |    193 |   by simp
 | 
|  |    194 | 
 | 
|  |    195 | lemma diff_zero_imp_eq: "!!a::'a::ring. a - b = 0 ==> a = b"
 | 
|  |    196 |   apply (rule_tac s = "a - (a - b) " in trans)
 | 
|  |    197 |    apply (tactic {* asm_simp_tac (simpset() delsimprocs [ring_simproc]) 1 *})
 | 
|  |    198 |    apply simp
 | 
|  |    199 |   apply (simp (no_asm))
 | 
|  |    200 |   done
 | 
|  |    201 | 
 | 
|  |    202 | lemma eq_imp_diff_zero: "!!a::'a::ring. a = b ==> a + (-b) = 0"
 | 
|  |    203 |   by simp
 | 
|  |    204 | 
 | 
|  |    205 | lemma long_div_quo_unique: 
 | 
|  |    206 |   "!!g::('a::field up). [| g ~= 0;  
 | 
|  |    207 |      f = q1 * g + r1; (r1 = 0 | deg r1 < deg g);  
 | 
|  |    208 |      f = q2 * g + r2; (r2 = 0 | deg r2 < deg g) |] ==> q1 = q2"
 | 
|  |    209 |   apply (subgoal_tac "(q1 - q2) * g = r2 - r1") (* 1 *)
 | 
|  |    210 |    apply (erule_tac V = "f = ?x" in thin_rl)
 | 
|  |    211 |   apply (erule_tac V = "f = ?x" in thin_rl)
 | 
|  |    212 |   apply (rule diff_zero_imp_eq)
 | 
|  |    213 |   apply (rule classical)
 | 
|  |    214 |   apply (erule disjE)
 | 
|  |    215 |   (* r1 = 0 *)
 | 
|  |    216 |     apply (erule disjE)
 | 
|  |    217 |   (* r2 = 0 *)
 | 
|  |    218 |      apply (tactic {* asm_full_simp_tac (simpset()
 | 
|  |    219 |        addsimps [thm "integral_iff", thm "minus_def", thm "l_zero", thm "uminus_zero"]
 | 
|  |    220 |        delsimprocs [ring_simproc]) 1 *})
 | 
|  |    221 |   (* r2 ~= 0 *)
 | 
|  |    222 |     apply (drule_tac f = "deg" and y = "r2 - r1" in arg_cong)
 | 
|  |    223 |     apply (tactic {* asm_full_simp_tac (simpset() addsimps
 | 
|  |    224 |       [thm "minus_def", thm "l_zero", thm "uminus_zero"] delsimprocs [ring_simproc]) 1 *})
 | 
|  |    225 |   (* r1 ~=0 *)
 | 
|  |    226 |    apply (erule disjE)
 | 
|  |    227 |   (* r2 = 0 *)
 | 
|  |    228 |     apply (drule_tac f = "deg" and y = "r2 - r1" in arg_cong)
 | 
|  |    229 |     apply (tactic {* asm_full_simp_tac (simpset() addsimps
 | 
|  |    230 |       [thm "minus_def", thm "l_zero", thm "uminus_zero"] delsimprocs [ring_simproc]) 1 *})
 | 
|  |    231 |   (* r2 ~= 0 *)
 | 
|  |    232 |    apply (drule_tac f = "deg" and y = "r2 - r1" in arg_cong)
 | 
|  |    233 |    apply (tactic {* asm_full_simp_tac (simpset() addsimps [thm "minus_def"]
 | 
|  |    234 |      delsimprocs [ring_simproc]) 1 *})
 | 
|  |    235 |    apply (drule order_eq_refl [THEN add_leD2])
 | 
|  |    236 |    apply (drule leD)
 | 
|  |    237 |    apply (erule notE, rule deg_add [THEN le_less_trans])
 | 
|  |    238 |    apply (simp (no_asm_simp))
 | 
|  |    239 |   (* proof of 1 *)
 | 
|  |    240 |    apply (rule diff_zero_imp_eq)
 | 
|  |    241 |   apply hypsubst
 | 
|  |    242 |   apply (drule_tac a = "?x+?y" in eq_imp_diff_zero)
 | 
|  |    243 |   apply simp
 | 
|  |    244 |   done
 | 
|  |    245 | 
 | 
|  |    246 | lemma long_div_rem_unique: 
 | 
|  |    247 |   "!!g::('a::field up). [| g ~= 0;  
 | 
|  |    248 |      f = q1 * g + r1; (r1 = 0 | deg r1 < deg g);  
 | 
|  |    249 |      f = q2 * g + r2; (r2 = 0 | deg r2 < deg g) |] ==> r1 = r2"
 | 
|  |    250 |   apply (subgoal_tac "q1 = q2")
 | 
|  |    251 |    apply clarify
 | 
|  |    252 |    apply (rule_tac a = "q2 * g + r1 - q2 * g" and b = "q2 * g + r2 - q2 * g" in box_equals)
 | 
|  |    253 |      apply simp
 | 
|  |    254 |     apply (simp (no_asm))
 | 
|  |    255 |    apply (simp (no_asm))
 | 
|  |    256 |   apply (rule long_div_quo_unique)
 | 
|  |    257 |   apply assumption+
 | 
| 14723 |    258 |   done
 | 
| 7998 |    259 | 
 | 
|  |    260 | end
 |