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