diff --git a/thys/Completeness/Completeness.thy b/thys/Completeness/Completeness.thy --- a/thys/Completeness/Completeness.thy +++ b/thys/Completeness/Completeness.thy @@ -1,933 +1,935 @@ section "Completeness" theory Completeness imports Tree Sequents begin subsection "pseq: type represents a processed sequent" type_synonym "atom" = "(signs * predicate * vbl list)" type_synonym nform = "(nat * formula)" type_synonym pseq = "(atom list * nform list)" definition sequent :: "pseq => formula list" where "sequent = (%(atoms,nforms) . map snd nforms @ map (% (z,p,vs) . FAtom z p vs) atoms)" definition pseq :: "formula list => pseq" where "pseq fs = ([],map (%f.(0,f)) fs)" definition atoms :: "pseq => atom list" where "atoms = fst" definition nforms :: "pseq => nform list" where "nforms = snd" subsection "subs: SATAxiom" definition SATAxiom :: "formula list => bool" where "SATAxiom fs \ (? n vs . FAtom Pos n vs : set fs & FAtom Neg n vs : set fs)" subsection "subs: a CutFreePC justifiable backwards proof step" definition subsFAtom :: "[atom list,(nat * formula) list,signs,predicate,vbl list] => pseq set" where "subsFAtom atms nAs z P vs = { ((z,P,vs)#atms,nAs) }" definition subsFConj :: "[atom list,(nat * formula) list,signs,formula,formula] => pseq set" where "subsFConj atms nAs z A0 A1 = (case z of Pos => { (atms,(0,A0)#nAs),(atms,(0,A1)#nAs) } | Neg => { (atms,(0,A0)#(0,A1)#nAs) })" definition subsFAll :: "[atom list,(nat * formula) list,nat,signs,formula,vbl set] => pseq set" where "subsFAll atms nAs n z A frees = (case z of Pos => { let v = freshVar frees in (atms,(0,instanceF v A)#nAs) } | Neg => { (atms,(0,instanceF (X n) A)#nAs @ [(Suc n,FAll Neg A)]) })" definition subs :: "pseq => pseq set" where "subs = (% pseq . if SATAxiom (sequent pseq) then {} else let (atms,nforms) = pseq in case nforms of [] => {} | nA#nAs => let (n,A) = nA in (case A of FAtom z P vs => subsFAtom atms nAs z P vs | FConj z A0 A1 => subsFConj atms nAs z A0 A1 | FAll z A => subsFAll atms nAs n z A (freeVarsFL (sequent pseq))))" subsection "proofTree(Gamma) says whether tree(Gamma) is a proof" definition proofTree :: "(nat * pseq) set => bool" where "proofTree A \ bounded A & founded subs (SATAxiom o sequent) A" subsection "path: considers, contains, costBarrier" definition considers :: "[nat => pseq,nat * formula,nat] => bool" where "considers f nA n = (case (snd (f n)) of [] => False | x#xs => x=nA)" definition "contains" :: "[nat => pseq,nat * formula,nat] => bool" where "contains f nA n \ nA : set (snd (f n))" definition costBarrier :: "[nat * formula,pseq] => nat" where (****** costBarrier justifies: eventually contains ==> eventually considers with a termination thm, (barrier strictly decreases in |N). ******) "costBarrier nA = (%(atms,nAs). let barrier = takeWhile (%x. nA ~= x) nAs in let costs = map (exp 3 o size o snd) barrier in sumList costs)" subsection "path: eventually" (****** Could do this with composable temporal operators, but following paper proof first. ******) definition EV :: "[nat => bool] => bool" where "EV f == (? n . f n)" (****** This allows blast_tac like searching to be done without hassle of removing exists rules. (hopefully). ******) subsection "path: counter model" definition counterM :: "(nat => pseq) => object set" where "counterM f = range obj" definition counterEvalP :: "(nat => pseq) => predicate => object list => bool" where "counterEvalP f = (%p args . ! i . ~(EV (contains f (i,FAtom Pos p (map (X o inv obj) args)))))" definition counterModel :: "(nat => pseq) => model" where "counterModel f = Abs_model (counterM f, counterEvalP f)" primrec counterAssign :: "vbl => object" where "counterAssign (X n) = obj n" (* just deX *) subsection "subs: finite" lemma finite_subs: "finite (subs gamma)" apply(simp add: subs_def subsFAtom_def subsFConj_def subsFAll_def Let_def split_beta split: if_splits list.split formula.split signs.split) done lemma fansSubs: "fans subs" apply(rule fansI) apply(rule, rule finite_subs) done lemma subs_def2: "!!gamma. ~ SATAxiom (sequent gamma) ==> subs gamma = (case nforms gamma of [] => {} | nA#nAs => let (n,A) = nA in (case A of FAtom z P vs => subsFAtom (atoms gamma) nAs z P vs | FConj z A0 A1 => subsFConj (atoms gamma) nAs z A0 A1 | FAll z A => subsFAll (atoms gamma) nAs n z A (freeVarsFL (sequent gamma))))" apply(simp add: subs_def Let_def nforms_def atoms_def split_beta split: list.split) done subsection "inherited: proofTree" lemma proofTree_def2: "proofTree = ( % x . bounded x & founded subs (SATAxiom o sequent) x)" apply(rule ext) apply(simp add: proofTree_def) done lemma inheritedProofTree: "inherited subs proofTree" apply(simp add: proofTree_def2) apply(auto intro: inheritedJoinI inheritedBounded inheritedFounded) done lemma proofTreeI: "[| bounded A; founded subs (SATAxiom o sequent) A |] ==> proofTree A" apply(simp add: proofTree_def) done lemma proofTreeBounded: "proofTree A ==> bounded A" apply(simp add: proofTree_def) done lemma proofTreeTerminal: "proofTree A ==> (n,delta) : A ==> terminal subs delta ==> SATAxiom (sequent delta)" apply(simp add: proofTree_def founded_def) apply blast done subsection "pseq: lemma" lemma snd_o_Pair: "(snd o (Pair x)) = (% x. x)" apply(rule ext) by simp lemma sequent_pseq: "sequent (pseq fs) = fs" by (simp add: pseq_def sequent_def snd_o_Pair) subsection "SATAxiom: proofTree" lemma SATAxiomTerminal[rule_format]: "SATAxiom (sequent gamma) --> terminal subs gamma" apply(simp add: subs_def proofTree_def terminal_def founded_def bounded_def) done lemma SATAxiomBounded:"SATAxiom (sequent gamma) ==> bounded (tree subs gamma)" apply(frule SATAxiomTerminal) apply(subst treeEquation) apply(simp add: subs_def proofTree_def terminal_def founded_def bounded_def) apply(force simp add: boundedByInsert boundedByEmpty) done lemma SATAxiomFounded: "SATAxiom (sequent gamma) ==> founded subs (SATAxiom o sequent) (tree subs gamma)" apply(frule SATAxiomTerminal) apply(subst treeEquation) apply(simp add: subs_def proofTree_def terminal_def founded_def bounded_def) done lemma SATAxiomProofTree[rule_format]: "SATAxiom (sequent gamma) --> proofTree (tree subs gamma)" apply(blast intro: proofTreeI SATAxiomBounded SATAxiomFounded) done lemma SATAxiomEq: "(proofTree (tree subs gamma) & terminal subs gamma) = SATAxiom (sequent gamma)" apply(blast intro: SATAxiomProofTree proofTreeTerminal tree.tree0 SATAxiomTerminal) done \ \FIXME tjr blast sensitive to obj imp vs meta imp for pTT\ subsection "SATAxioms are deductions: - needed" lemma SAT_deduction: "SATAxiom x ==> x : deductions CutFreePC" apply(simp add: SATAxiom_def) apply(elim exE) apply(rule_tac x="[FAtom Pos n vs,FAtom Neg n vs]" in inDed_mono) apply (blast intro!: SATAxiomI rulesInPCs, force) done subsection "proofTrees are deductions: subs respects rules - messy start and end" lemma subsJustified': notes ss = subs_def2 nforms_def Let_def atoms_def sequent_def subsFAtom_def subsFConj_def subsFAll_def shows "\ SATAxiom (sequent (ats, (n, f) # list)) --> \ terminal subs (ats, (n, f) # list) --> (\sigma\subs (ats, (n, f) # list). sequent sigma \ deductions CutFreePC) --> sequent (ats, (n, f) # list) \ deductions CutFreePC" apply (rule_tac A=f in formula_signs_cases, clarify) apply(simp add: ss) - apply(rule PermI) apply assumption apply(rule perm_append_Cons) apply (rule rulesInPCs, clarify) apply(simp add: ss) - apply(rule PermI) apply assumption apply(rule perm_append_Cons) apply (rule rulesInPCs, clarify) apply(simp add: ss) apply(elim conjE) + apply(rule PermI) apply assumption apply simp_all + apply (rule rulesInPCs, clarify) apply(simp add: ss) + apply(rule PermI) apply assumption apply simp_all + apply (rule rulesInPCs, clarify) apply(simp add: ss) apply(elim conjE) apply(rule ConjI') apply assumption apply assumption apply(rule rulesInPCs)+ apply clarify apply(simp add: ss) apply(rule DisjI) apply assumption apply(rule rulesInPCs)+ apply clarify apply(simp add: ss) apply(rule AllI) apply assumption apply(rule finiteFreshVar) apply(rule finite_freeVarsFL) apply (rule rulesInPCs, clarify) apply(simp add: ss) apply(rule ContrI) \ \!\ apply(rule_tac w = "X n" in ExI) apply(rule inDed_mono) apply assumption apply force apply(rule rulesInPCs)+ done lemma subsJustified: "!! gamma. ~ terminal subs gamma ==> ! sigma : subs gamma . sequent sigma : deductions (CutFreePC) ==> sequent gamma : deductions (CutFreePC)" apply(case_tac "SATAxiom (sequent gamma)") apply(erule SAT_deduction) apply(case_tac gamma) apply(rename_tac ats nfs) apply(case_tac nfs) apply(simp add: terminal_def subs_def sequent_def Let_def) apply(case_tac a) apply(blast intro: subsJustified'[rule_format]) done subsection "proofTrees are deductions: instance of boundedTreeInduction" lemmas proofTreeD = proofTree_def [THEN iffD1] lemma proofTreeDeductionD[rule_format]: "proofTree(tree subs gamma) \ sequent gamma : deductions (CutFreePC)" apply(rule boundedTreeInduction[OF fansSubs]) apply(erule proofTreeBounded) apply(rule foundedMono) apply (force dest: proofTreeD, simp) apply(blast intro: SAT_deduction foundedMono subsJustified) apply(blast intro: subsJustified) done subsection "contains, considers:" lemma contains_def2: "contains f iA n = (iA : set (nforms (f n)))" apply(simp add: contains_def nforms_def) done lemma considers_def2: "considers f iA n = ( ? nAs . nforms (f n) = iA#nAs)" apply(simp add: considers_def nforms_def split: list.split) done lemmas containsI = contains_def2[THEN iffD2] subsection "path: nforms = [] implies" lemma nformsNoContains: "[| branch subs gamma f; !n . ~proofTree (tree subs (f n)); nforms (f n) = [] |] ==> ~ contains f iA n" apply(simp add: contains_def2) done \ \FIXME tjr assumptions not required\ lemma nformsTerminal: "nforms (f n) = [] ==> terminal subs (f n)" apply(simp add: subs_def Let_def terminal_def nforms_def split_beta) done lemma nformsStops: "!!f. [| branch subs gamma f; !n . ~proofTree (tree subs (f n)); nforms (f n) = [] |] ==> nforms (f (Suc n)) = [] & atoms (f (Suc n)) = atoms (f n)" apply(subgoal_tac "f (Suc n) = f n") apply simp apply(blast intro: branchStops nformsTerminal) done subsection "path: cases" lemma terminalNFormCases: "!!f. terminal subs (f n) | (? i A nAs . nforms (f n) = (i,A)#nAs)" apply (rule disjCI, simp) apply(rule nformsTerminal) apply(case_tac "nforms (f n)") apply simp apply force done lemma cases[elim_format]: "terminal subs (f n) | (\ (terminal subs (f n) \ (? i A nAs . nforms (f n) = (i,A)#nAs)))" apply(auto elim: terminalNFormCases[elim_format]) done subsection "path: contains not terminal and propagate condition" lemma containsNotTerminal: "[| branch subs gamma f; !n . ~proofTree (tree subs (f n)); contains f iA n |] ==> ~ (terminal subs (f n))" apply(case_tac "SATAxiom (sequent (f n))") apply(blast dest: SATAxiomEq[THEN iffD2]) apply(drule_tac x=n in spec) apply (simp add: subs_def subs_def subsFAtom_def subsFConj_def subsFAll_def Let_def contains_def terminal_def nforms_def split_beta branch_def split: list.split signs.split expand_case_formula, force) done lemma containsPropagates: "!!f. [| branch subs gamma f; !n . ~proofTree (tree subs (f n)); contains f iA n |] ==> contains f iA (Suc n) | considers f iA n" apply(frule_tac containsNotTerminal) apply force apply force apply(frule_tac branchSubs) apply assumption apply(case_tac "considers f iA n") apply simp apply simp apply(simp add: contains_def) apply(case_tac "f n") apply simp apply(drule split_list) apply(elim exE conjE) apply simp apply(case_tac ys) apply (simp add: considers_def, simp) apply(case_tac "SATAxiom (sequent (f n))") apply(blast dest: iffD2[OF SATAxiomEq]) apply(simp add: subs_def2 nforms_def Let_def) apply (case_tac aa, simp) apply(case_tac ba) apply(simp add: subsFAtom_def) apply(rename_tac signs a b) apply(case_tac signs) apply(simp add: subsFConj_def) apply force apply(simp add: subsFConj_def) apply(rename_tac signs a) apply(case_tac signs) apply(simp add: subsFAll_def Let_def) apply(simp add: subsFAll_def Let_def) done subsection "path: no consider lemmas" lemma noConsidersD: "!!f. ~considers f iA n ==> nforms (f n) = x#xs ==> iA ~= x" by(simp add: considers_def2) lemma considersD: "!!f. considers f iA n ==> ? xs . nforms (f n) = iA#xs" by(simp add: considers_def2) subsection "path: contains initially" lemma contains_initially: "branch subs (pseq gamma) f \ A : set gamma \ (contains f (0,A) 0)" apply(drule branch0) apply(simp add: contains_def pseq_def) done lemma contains_initialEVs: "branch subs (pseq gamma) f \ A : set gamma \ EV (contains f (0,A))" apply(simp add: EV_def) apply(fast dest: contains_initially) done subsection "termination: (for EV contains implies EV considers)" lemmas r = wf_induct[of "measure msrFn", OF wf_measure] for msrFn lemmas r' = r[simplified measure_def inv_image_def less_than_def less_eq mem_Collect_eq] lemma r'': "(\ x. (\ y. ( ((msrFn::'a \ nat) y) < ((msrFn :: 'a \ nat) x)) \ P y) \ P x) \ P a" by (blast intro: r' [of _ P]) lemma terminationRule [rule_format]: "! n. P n --> (~(P (Suc n)) | (P (Suc n) & msrFn (Suc n) < (msrFn::nat => nat) n)) ==> P m --> (? n . P n & ~(P (Suc n)))" (is "_ \ ?P m") apply (rule r''[of msrFn "?P" m], blast) done \ \FIXME ugly\ subsection "costBarrier: lemmas" subsection "costBarrier: exp3 lemmas - bit specific..." lemma exp3Min: "exp 3 a > 0" by (induct a, simp, simp) lemma exp1: "exp 3 (A) + exp 3 (B) < 3 * ((exp 3 A) * (exp 3 B))" using exp3Min[of A] exp3Min[of B] apply(case_tac "exp 3 A") apply(simp add: exp3Min) apply(case_tac "exp 3 B") apply (simp add: exp3Min, simp) done lemma exp1': "exp 3 (A) < 3 * ((exp 3 A) * (exp 3 B)) + C" apply(subgoal_tac "exp 3 (A) < 3 * ((exp 3 A) * (exp 3 B))") apply arith apply(case_tac "exp 3 A") using exp3Min[of A] apply arith apply(case_tac "exp 3 B") using exp3Min[of B] apply arith apply simp done lemma exp2: "Suc 0 < 3 * exp 3 (B)" using exp3Min[of B] apply arith done lemma expSum: "exp x (a+b) = (exp x a) * (exp x b)" apply(induct a, auto) done subsection "costBarrier: decreases whilst contains and unconsiders" lemma costBarrierDecreases': notes ss = subs_def2 nforms_def Let_def subsFAtom_def subsFConj_def subsFAll_def costBarrier_def atoms_def exp3Min expSum shows "~SATAxiom (sequent (a, (num,fm) # list)) --> iA ~= (num, fm) --> \ proofTree (tree subs (a, (num, fm) # list)) --> fSucn : subs (a, (num, fm) # list) --> iA \ set list --> costBarrier iA (fSucn) < costBarrier iA (a, (num, fm) # list)" apply(rule_tac A=fm in formula_signs_cases) \ \atoms\ apply(simp add: ss) apply(simp add: ss) \ \conj\ apply clarify apply(simp add: ss) apply(erule disjE) apply(simp add: ss exp2) apply(simp add: ss exp2) \ \disj\ apply clarify apply(simp add: ss exp1 exp1') \ \all\ apply clarify apply(simp add: ss size_instance) \ \ex\ apply clarify apply(simp add: ss size_instance) done lemma costBarrierDecreases: "[| branch subs gamma f; !n . ~proofTree (tree subs (f n)); contains f iA n; ~(considers f iA) n |] ==> costBarrier iA (f (Suc n)) < costBarrier iA (f n)" apply(subgoal_tac "\ terminal subs (f n)") apply(subgoal_tac "\ SATAxiom (sequent (f n))") apply(subgoal_tac "f (Suc n) \ subs (f n)") apply(frule_tac x=n in spec) apply(case_tac "f n", simp) apply(case_tac b, simp) apply(simp add: contains_def) apply(case_tac aa) apply(rename_tac num fm) apply simp apply(simp add: contains_def considers_def) apply(rule costBarrierDecreases'[rule_format]) apply force+ apply(rule branchSubs) apply assumption apply assumption apply(blast dest: SATAxiomTerminal) apply(blast dest: containsNotTerminal) done \ \FIXME boring splitting etc.\ subsection "path: EV contains implies EV considers" lemma considersContains: "considers f iA n \ contains f iA n" apply(simp add: considers_def contains_def) apply(cases "snd (f n)", auto) done lemma containsConsiders: "[| branch subs gamma f; !n . ~ proofTree (tree subs (f n)); EV (contains f iA) |] ==> EV (considers f iA)" apply(simp add: EV_def) apply(erule exE) apply(case_tac "considers f iA n") apply force apply(subgoal_tac "\n. (contains f iA n \ \ considers f iA n) \ \ (contains f iA (Suc n) \ \ considers f iA (Suc n))") apply(erule exE) apply(simp, clarify) apply(case_tac "contains f iA (Suc na)") apply force apply(force dest!: containsPropagates) apply(rule_tac msrFn = "%n. costBarrier iA (f n)" and P = "%n. contains f iA n & ~ considers f iA n" in terminationRule) prefer 2 apply force apply(case_tac "(contains f iA (Suc na) \ \ considers f iA (Suc na))") apply simp apply(erule costBarrierDecreases) apply simp_all done subsection "EV contains: common lemma" lemma lemmA: "[| branch subs gamma f; ! n. ~ proofTree (tree subs (f n)); EV (contains f (i,A)) |] ==> ? n nAs. ~SATAxiom (sequent (f n)) & (nforms (f n) = (i,A) # nAs & f (Suc n) : subs (f n))" apply(frule containsConsiders) apply(assumption+) apply(unfold EV_def) apply(erule exE, frule considersContains) apply(unfold considers_def) apply(case_tac "snd (f n)") apply force apply simp apply(rule_tac x=n in exI) apply (intro conjI, rule) apply(blast dest!: SATAxiomEq[THEN iffD2]) apply(rule_tac x=list in exI) apply(simp add: nforms_def) apply(frule containsNotTerminal) apply force apply(assumption+) apply(blast dest!: branchSubs) done subsection "EV contains: FConj,FDisj,FAll" lemma EV_disj: "(EV P | EV Q) = EV (\n. P n | Q n)" apply (unfold EV_def, force) done lemma evContainsConj: "[| EV (contains f (i,FConj Pos A0 A1)); branch subs gamma f; !n . ~ proofTree (tree subs (f n)) |] ==> EV (contains f (0,A0)) | EV (contains f (0,A1))" apply(drule lemmA) apply(assumption+) apply(subgoal_tac "EV (\ n. contains f (0,A0) n | contains f (0,A1) n)") apply(simp add: EV_disj) apply(unfold EV_def) apply clarify apply(rename_tac n n' nAs, rule_tac x="Suc n'" in exI) apply(simp add: subs_def2 Let_def) apply(simp add: subsFConj_def) apply (simp add: contains_def nforms_def, auto) done lemma evContainsDisj: "[| EV (contains f (i,FConj Neg A0 A1)); branch subs gamma f; !n . ~ proofTree (tree subs (f n)) |] ==> EV (contains f (0,A0)) & EV (contains f (0,A1))" apply(drule lemmA) apply(assumption+) apply(rule conjI) apply(unfold EV_def) apply(erule exE) back apply(rule_tac x="Suc n" in exI) apply(clarify, simp add: subs_def2 Let_def) apply(simp add: subsFConj_def) apply(simp add: contains_def nforms_def) apply(erule exE) back apply(rule_tac x="Suc n" in exI) apply(clarify, simp add: subs_def2 Let_def) apply(simp add: subsFConj_def) apply(simp add: contains_def nforms_def) done lemma evContainsAll: "[| EV (contains f (i,FAll Pos body)); branch subs gamma f; !n . ~ proofTree (tree subs (f n)) |] ==> ? v . EV (contains f (0,instanceF v body))" apply(drule lemmA) apply(assumption+) apply(erule exE) apply(rule_tac x="freshVar(freeVarsFL (sequent (f n)))" in exI) apply(unfold EV_def) apply(rule_tac x="Suc n" in exI) apply(clarify, simp add: subs_def2 Let_def) apply(simp add: subsFAll_def Let_def) apply(simp add: contains_def nforms_def) done lemma evContainsEx_instance: "[| EV (contains f (i,FAll Neg body)); branch subs gamma f; !n . ~ proofTree (tree subs (f n)) |] ==> EV (contains f (0,instanceF (X i) body))" apply(drule lemmA) apply(assumption+) apply(erule exE) apply(unfold EV_def) apply(rule_tac x="Suc n" in exI) apply(clarify, simp add: subs_def2 Let_def) apply(simp add: subsFAll_def Let_def) apply(simp add: contains_def nforms_def) done lemma evContainsEx_repeat: " [| branch subs gamma f; !n . ~ proofTree (tree subs (f n)); EV (contains f (i,FAll Neg body)) |] ==> EV (contains f (Suc i,FAll Neg body))" apply(drule lemmA) apply(assumption+) apply(erule exE) apply(unfold EV_def) apply(rule_tac x="Suc n" in exI) apply(clarify, simp add: subs_def2 Let_def) apply(simp add: subsFAll_def Let_def) apply(simp add: contains_def nforms_def) done subsection "EV contains: lemmas (temporal related)" (****** Should have abstracted to have temporal ops: EV : (nat -> bool) -> nat -> bool AW : (nat -> bool) -> nat -> bool NEXT : (nat -> bool) -> nat -> bool then require: EV P and EV B imp (\n. A n & B n) already think have done similar proofs to this one elsewhere, Q: where? share proofs? ******) lemma lemma1: "[| P A n ; !A n. P A n --> P A (Suc n) |] ==> P A (n + m)" apply (induct m, simp, simp) done lemma lemma2: "[| P A n ; P B m ; ! A n. P A n --> P A (Suc n) |] ==> ? n . P A n & P B n" apply (rule exI[of _ "n+m"], rule) apply(blast intro!: lemma1) apply(rule subst[OF add.commute]) apply(blast intro!: lemma1) done subsection "EV contains: FAtoms" lemma notTerminalNotSATAxiom: "\ terminal subs gamma \ \ SATAxiom (sequent gamma)" apply(erule contrapos_nn) apply(erule SATAxiomTerminal) done lemma notTerminalNforms: "\ terminal subs (f n) \ nforms (f n) \ []" apply(erule contrapos_nn) apply(erule nformsTerminal) done lemma atomsPropagate: "[| branch subs gamma f |] ==> x : set (atoms (f n)) --> x : set (atoms (f (Suc n)))" apply(cases "terminal subs (f n)") apply(drule branchStops) apply assumption apply simp apply(drule branchSubs) apply assumption apply rule apply(frule notTerminalNotSATAxiom) apply(frule notTerminalNforms) apply(simp add: subs_def2) apply(cases "nforms (f n)") apply simp apply(simp add: Let_def) apply(case_tac a, auto) apply(case_tac ba, auto) apply(simp add: subsFAtom_def atoms_def) apply(simp add: subsFConj_def atoms_def) apply(rename_tac signs a b) apply(case_tac signs) apply force apply force apply(simp add: subsFAll_def atoms_def) apply(rename_tac signs a) apply(case_tac signs) apply(force simp: Let_def) apply force done subsection "EV contains: FEx cases" lemma evContainsEx0_allRepeats: "[| branch subs gamma f; !n . ~ proofTree (tree subs (f n)); EV (contains f (0,FAll Neg A)) |] ==> EV (contains f (i,FAll Neg A))" apply (induct i, simp) apply(blast dest!: evContainsEx_repeat) done lemma evContainsEx0_allInstances: "[| branch subs gamma f; !n . ~ proofTree (tree subs (f n)); EV (contains f (0,FAll Neg A)) |] ==> EV (contains f (0,instanceF (X i) A))" apply(blast dest!: evContainsEx0_allRepeats intro!: evContainsEx_instance) done subsection "pseq: creates initial pseq" lemma containsPSeq0D: "branch subs (pseq fs) f \ contains f (i,A) 0 \ i=0" apply(drule branch0) apply (simp add: pseq_def contains_def, blast) done subsection "EV contains: contain any (i,FEx y) means contain all (i,FEx y)" (****** Now, show (Suc i,FEx v A) present means (i,FEx v A) present (or at start). Assumes initial pseq contains only (0,form) pairs. ------ Show that only way of introducing a (Suc i,FEx_) was from (i,FEx_). contains n == considers n or contains n+. Want to find the point where it was introduced. Have, P true all n or fails at 0 or for a (first) successor. (!n. P n) | (~P 0) | (P n & ~ P (Suc n)) Instance this with P n = ~(contains ..FEx.. n). ******) lemma claim: "(A | B | C) = (~C --> ~B --> A)" by auto lemma natPredCases: "(!n. P n) | (~P 0) | (? n . P n & ~ P (Suc n))" apply(rule claim[THEN iffD2]) apply(intro impI) apply simp apply rule apply(induct_tac n) apply auto done lemma containsNotTerminal': "\ branch subs gamma f; !n . ~proofTree (tree subs (f n)); contains f iA n \ \ ~ (terminal subs (f n))" apply (rule containsNotTerminal, auto) done lemma notTerminalSucNotTerminal: "\ \ terminal subs (f (Suc n)); branch subs gamma f \ \ \ terminal subs (f n)" apply(erule contrapos_nn) apply(rule_tac branchTerminalPropagates[of _ _ _ _ 1, simplified]) apply(assumption+) done \ \FIXME move to Tree?\ lemma evContainsExSuc_containsEx: "[| branch subs (pseq fs) f; !n . ~ proofTree (tree subs (f n)); EV (contains f (Suc i,FAll Neg body)) |] ==> EV (contains f (i,FAll Neg body))" apply(cut_tac P="%n. ~ contains f (Suc i,FAll Neg body) n" in natPredCases) apply simp apply(erule disjE) apply(simp add: EV_def) apply(erule disjE) apply(blast dest!: containsPSeq0D) apply(thin_tac "EV _") apply(erule exE) apply(erule conjE) apply(unfold EV_def) apply(rule_tac x=n in exI) apply(rule considersContains) apply(frule containsNotTerminal') apply(assumption+) apply(frule notTerminalSucNotTerminal) apply assumption apply(thin_tac "\ terminal x y" for x y) apply(frule branchSubs) apply assumption apply(frule notTerminalNforms) apply(case_tac "SATAxiom (sequent (f n))") apply(drule SATAxiomTerminal) apply simp apply(subgoal_tac "(\i A nAs. nforms (f n) = (i, A) # nAs)") prefer 2 apply(rule_tac f=f and n=n in cases) apply simp apply(case_tac "nforms (f n)") apply simp apply(case_tac a, force) apply(erule exE)+ apply(unfold considers_def) apply(simp add: nforms_def) \ \shift A into succedent\ apply(rule_tac P="snd (f n) = (ia, A) # nAs" in rev_mp) apply assumption apply(thin_tac "snd (f n) = (ia, A) # nAs") apply(rule_tac A=A in formula_signs_cases) apply(auto simp add: subs_def2 nforms_def Let_def subsFAtom_def subsFConj_def subsFAll_def contains_def2 Let_def) done \ \phew, bit precarious, but not too much going on besides unfolding defns. and computing a bit. Need to set simps up\ subsection "EV contains: contain any (i,FEx y) means contain all (i,FEx y)" lemma evContainsEx_containsEx0: "[| branch subs (pseq fs) f; !n . ~ proofTree (tree subs (f n)) |] ==> EV (contains f (i,FAll Neg A)) --> EV (contains f (0,FAll Neg A))" apply (induct i, simp) apply(blast dest!: evContainsExSuc_containsEx) done lemma evContainsExval: "[| EV (contains f (i,FAll Neg body)); branch subs (pseq fs) f; !n . ~ proofTree (tree subs (f n)) |] ==> ! v . EV (contains f (0,instanceF v body))" apply rule apply(induct_tac v) apply(blast intro!: evContainsEx0_allInstances dest!: evContainsEx_containsEx0) done subsection "EV contains: atoms" lemma atomsInSequentI[rule_format]: " (z,P,vs) : set (fst ps) --> FAtom z P vs : set (sequent ps)" apply(simp add: sequent_def) apply(cases ps, force) done lemma evContainsAtom1: "[| branch subs (pseq fs) f; !n . ~ proofTree (tree subs (f n)); EV (contains f (i,FAtom z P vs)) |] ==> ? n . (z,P,vs) : set (fst (f n))" apply(drule lemmA) apply(assumption+) apply(erule exE) apply(rule_tac x="Suc n" in exI) apply(simp add: subs_def2) apply clarify apply(simp add: subs_def2 Let_def) apply(simp add: subsFAtom_def) done lemmas atomsPropagate'' = atomsPropagate[rule_format] lemmas atomsPropagate' = atomsPropagate''[simplified atoms_def, simplified] lemma evContainsAtom: "[| branch subs (pseq fs) f; !n . ~ proofTree (tree subs (f n)); EV (contains f (i,FAtom z P vs)) |] ==> ? n . (! m . FAtom z P vs : set (sequent (f (n + m))))" apply(frule evContainsAtom1) apply(assumption+) apply(erule exE) apply (rule_tac x=n in exI, rule) apply(rule atomsInSequentI) apply (induct_tac m, simp, simp) apply(rule atomsPropagate') apply(assumption+) done lemma notEvContainsBothAtoms: "[| branch subs (pseq fs) f; !n . ~ proofTree (tree subs (f n)) |] ==> ~ EV (contains f (i,FAtom Pos p vs)) | ~ EV (contains f (j,FAtom Neg p vs))" apply clarify apply(frule evContainsAtom) apply(assumption+) apply(thin_tac "EV (contains f (j, FAtom Neg p vs))") apply(frule evContainsAtom) apply(assumption+) apply(thin_tac "EV (contains f (i, FAtom Pos p vs))") apply(erule_tac exE)+ apply(drule_tac x=na in spec) back apply(drule_tac x=n in spec) back apply(simp add: ac_simps) apply(subgoal_tac "SATAxiom (sequent (f (n+na)))") apply(force dest: SATAxiomProofTree) apply(force simp add: SATAxiom_def) done subsection "counterModel: lemmas" lemma counterModelInRepn: "(counterM f,counterEvalP f) : model" apply(simp add: model_def counterM_def) done lemmas Abs_counterModel_inverse = counterModelInRepn[THEN Abs_model_inverse] lemma inv_obj_obj: "inv obj (obj n) = n" using inj_obj apply simp done lemma map_X_map_counterAssign: "map X (map (inv obj) (map counterAssign xs)) = xs" apply(simp) apply(subgoal_tac "(X \ (inv obj \ counterAssign)) = (% x . x)") apply simp apply(rule ext) apply(case_tac x) apply(simp add: inv_obj_obj) done lemma objectsCounterModel: "objects (counterModel f) = { z . ? i . z = obj i }" apply(simp add: objects_def counterModel_def) apply(simp add: Abs_counterModel_inverse) apply(simp add: counterM_def) by force lemma inCounterM: "counterAssign v : objects (counterModel f)" apply(induct v) apply(simp add: objectsCounterModel) by blast lemma counterAssign_eqI[rule_format]: "x : objects (counterModel f) --> z = X (inv obj x) --> counterAssign z = x" apply(force simp: objectsCounterModel inj_obj) done lemma evalPCounterModel: "M = counterModel f ==> evalP M = counterEvalP f" apply(simp add: evalP_def counterModel_def Abs_counterModel_inverse) done subsection "counterModel: all path formula value false - step by step" lemma path_evalF': notes ss = evalPCounterModel counterEvalP_def map_X_map_counterAssign map_map[symmetric] and ss1 = instanceF_def evalF_subF_eq comp_vblcase id_def[symmetric] shows "[| branch subs (pseq fs) f; !n . ~ proofTree (tree subs (f n)) |] ==> (? i . EV (contains f (i,A))) \ ~(evalF (counterModel f) counterAssign A)" apply (rule_tac strong_formula_induct, rule) apply(rule formula_signs_cases) \ \atom\ apply(simp add: ss del: map_map) apply(rule, rule) apply(simp add: ss del: map_map) apply(force dest: notEvContainsBothAtoms) \ \conj\ apply(force dest: evContainsConj) \ \disj\ apply(force dest: evContainsDisj) \ \all\ apply(rule, rule) apply(erule exE) apply(drule_tac evContainsAll) apply assumption apply assumption apply(erule exE) apply(drule_tac x="(instanceF v f1)" in spec) apply(erule impE, force simp: size_instance)+ apply simp apply(simp add: ss1) apply(rule_tac x="counterAssign v" in bexI) apply simp apply(simp add: inCounterM) \ \ex\ apply(rule, rule) apply(erule exE) apply(drule_tac evContainsExval) apply assumption apply assumption apply simp apply rule apply(simp add: objectsCounterModel) apply(erule exE) apply(drule_tac x="X i" in spec) apply(drule_tac x="(instanceF (X i) f1)" in spec) apply(erule impE, force simp: size_instance)+ apply(simp add: ss1) done lemmas path_evalF'' = mp[OF path_evalF'] subsection "adequacy" lemma counterAssignModelAssign: "counterAssign : modelAssigns (counterModel gamma)" apply (simp add: modelAssigns_def, rule) apply (erule rangeE, simp) apply(rule inCounterM) done lemma branch_contains_initially: "branch subs (pseq fs) f \ x : set fs \ contains f (0,x) 0" apply(simp add: contains_def branch0 pseq_def) done lemma path_evalF: "[| branch subs (pseq fs) f; \n. \ proofTree (tree subs (f n)); x \ set fs |] ==> \ evalF (counterModel f) counterAssign x" apply (rule path_evalF'', assumption, assumption) apply(rule_tac x=0 in exI) apply(simp add: EV_def) apply(rule_tac x=0 in exI) apply(simp add: branch_contains_initially) done lemma validProofTree: "~proofTree (tree subs (pseq fs)) ==> ~(validS fs)" apply(simp add: validS_def evalS_def) apply(subgoal_tac "\f. branch subs (pseq fs) f \ (\n. \ proofTree (tree subs (f n)))") apply(elim exE conjE) apply(rule_tac x="counterModel f" in exI) apply(rule_tac x=counterAssign in bexI) apply(force dest!: path_evalF) apply(rule counterAssignModelAssign) apply(rule failingBranchExistence) apply(rule inheritedProofTree) apply (rule fansSubs, assumption) done lemma adequacy[simplified sequent_pseq]: "validS fs ==> (sequent (pseq fs)) : deductions CutFreePC" apply(rule proofTreeDeductionD) apply(rule ccontr) apply(force dest!: validProofTree) done end diff --git a/thys/Completeness/PermutationLemmas.thy b/thys/Completeness/PermutationLemmas.thy --- a/thys/Completeness/PermutationLemmas.thy +++ b/thys/Completeness/PermutationLemmas.thy @@ -1,184 +1,172 @@ section "Permutation Lemmas" theory PermutationLemmas -imports "HOL-Library.List_Permutation" "HOL-Library.Multiset" +imports "HOL-Library.Multiset" begin \ \following function is very close to that in multisets- now we can make the connection that x <~~> y iff the multiset of x is the same as that of y\ subsection "perm, count equivalence" primrec count :: "'a \ 'a list \ nat" where "count x [] = 0" | "count x (y#ys) = (if x=y then 1 else 0) + count x ys" -lemma perm_count: "A <~~> B \ (\ x. count x A = count x B)" - by(induct set: perm) auto +lemma count_eq: + \count x xs = Multiset.count (mset xs) x\ + by (induction xs) simp_all + +lemma perm_count: "mset A = mset B \ (\ x. count x A = count x B)" + by (simp add: count_eq) lemma count_0: "(\x. count x B = 0) = (B = [])" by(induct B) auto lemma count_Suc: "count a B = Suc m \ a : set B" apply(induct B) apply auto apply(case_tac "a = aa") apply auto done lemma count_append: "count a (xs@ys) = count a xs + count a ys" by(induct xs) auto -lemma count_perm: "!! B. (\ x. count x A = count x B) \ A <~~> B" - apply(induct A) - apply(simp add: count_0) -proof - - fix a list B - assume a: "\B. \x. count x list = count x B \ list <~~> B" - and b: "\x. count x (a # list) = count x B" - from b have "a : set B" - apply auto - apply (drule_tac x=a in spec, simp) apply(metis count_Suc) done - from split_list[OF this] obtain xs ys where B: "B = xs@a#ys" by blast - let ?B' = "xs@ys" - from b have "\x. count x list = count x ?B'" by(simp add: count_append B) - from a[OF this] have c: "list <~~> xs@ys" . - hence "a#list <~~> a#(xs@ys)" by rule - also have "a#(xs@ys) <~~> xs@a#ys" by(rule perm_append_Cons) - also (perm.trans) note B[symmetric] - finally show "a # list <~~> B" . -qed +lemma count_perm: "!! B. (\ x. count x A = count x B) \ mset A = mset B" + by (simp add: count_eq multiset_eq_iff) -lemma perm_count_conv: "A <~~> B = (\ x. count x A = count x B)" - apply(blast intro!: perm_count count_perm) done +lemma perm_count_conv: "mset A = mset B \ (\ x. count x A = count x B)" + by (simp add: count_eq multiset_eq_iff) subsection "Properties closed under Perm and Contr hold for x iff hold for remdups x" lemma remdups_append: "y : set ys --> remdups (ws@y#ys) = remdups (ws@ys)" apply (induct ws, simp) apply (case_tac "y = a", simp, simp) done -lemma perm_contr': assumes perm[rule_format]: "! xs ys. xs <~~> ys --> (P xs = P ys)" +lemma perm_contr': assumes perm[rule_format]: "! xs ys. mset xs = mset ys --> (P xs = P ys)" and contr'[rule_format]: "! x xs. P(x#x#xs) = P (x#xs)" shows "! xs. length xs = n --> (P xs = P (remdups xs))" apply(induct n rule: nat_less_induct) proof (safe) fix xs :: "'a list" assume a[rule_format]: "\mys. length ys = m \ P ys = P (remdups ys)" show "P xs = P (remdups xs)" proof (cases "distinct xs") case True thus ?thesis by(simp add:distinct_remdups_id) next case False from not_distinct_decomp[OF this] obtain ws ys zs y where xs: "xs = ws@[y]@ys@[y]@zs" by force have "P xs = P (ws@[y]@ys@[y]@zs)" by (simp add: xs) also have "... = P ([y,y]@ws@ys@zs)" apply(rule perm) apply(rule iffD2[OF perm_count_conv]) apply rule apply(simp add: count_append) done also have "... = P ([y]@ws@ys@zs)" apply simp apply(rule contr') done also have "... = P (ws@ys@[y]@zs)" apply(rule perm) apply(rule iffD2[OF perm_count_conv]) apply rule apply(simp add: count_append) done also have "... = P (remdups (ws@ys@[y]@zs))" apply(rule a) by(auto simp: xs) also have "(remdups (ws@ys@[y]@zs)) = (remdups xs)" apply(simp add: xs remdups_append) done finally show "P xs = P (remdups xs)" . qed qed -lemma perm_contr: assumes perm: "! xs ys. xs <~~> ys --> (P xs = P ys)" +lemma perm_contr: assumes perm: "! xs ys. mset xs = mset ys --> (P xs = P ys)" and contr': "! x xs. P(x#x#xs) = P (x#xs)" shows "(P xs = P (remdups xs))" apply(rule perm_contr'[OF perm contr', rule_format]) by force subsection "List properties closed under Perm, Weak and Contr are monotonic in the set of the list" definition rem :: "'a => 'a list => 'a list" where "rem x xs = filter (%y. y ~= x) xs" lemma rem: "x ~: set (rem x xs)" by(simp add: rem_def) lemma length_rem: "length (rem x xs) <= length xs" by(simp add: rem_def) lemma rem_notin: "x ~: set xs ==> rem x xs = xs" apply(simp add: rem_def) apply(rule filter_True) apply force done -lemma perm_weak_filter': assumes perm[rule_format]: "! xs ys. xs <~~> ys --> (P xs = P ys)" +lemma perm_weak_filter': assumes perm[rule_format]: "! xs ys. mset xs = mset ys --> (P xs = P ys)" and weak[rule_format]: "! x xs. P xs --> P (x#xs)" - shows "! ys. P (ys@filter Q xs) --> P (ys@xs)" - apply (induct xs, simp, rule) - apply rule - apply simp - apply (case_tac "Q a", simp) - apply(drule_tac x="ys@[a]" in spec) apply simp - apply simp - apply(drule_tac x="ys@[a]" in spec) apply simp - apply(erule impE) - apply(subgoal_tac "(ys @ a # filter Q xs) <~~> a#ys@filter Q xs") - apply(simp add: perm) - apply(rule weak) apply simp - apply(rule perm_sym) apply(rule perm_append_Cons) - . +shows "! ys. P (ys@filter Q xs) --> P (ys@xs)" +proof (rule allI, rule impI) + fix ys + define zs where \zs = filter (Not \ Q) xs\ + assume \P (ys @ filter Q xs)\ + then have \P (filter Q xs @ ys)\ + apply (subst perm) defer apply assumption apply simp done + then have \P (zs @ filter Q xs @ ys)\ + apply (induction zs) + apply (simp_all add: weak) + done + with zs_def show \P (ys @ xs)\ + apply (subst perm) defer apply assumption apply simp done +qed -lemma perm_weak_filter: assumes perm: "! xs ys. xs <~~> ys --> (P xs = P ys)" +lemma perm_weak_filter: assumes perm: "! xs ys. mset xs = mset ys --> (P xs = P ys)" and weak: "! x xs. P xs --> P (x#xs)" shows "P (filter Q xs) ==> P xs" using perm_weak_filter'[OF perm weak, rule_format, of "[]", simplified] by blast \ \right, now in a position to prove that in presence of perm, contr and weak, set x leq set y and x : ded implies y : ded\ lemma perm_weak_contr_mono: - assumes perm: "! xs ys. xs <~~> ys --> (P xs = P ys)" + assumes perm: "! xs ys. mset xs = mset ys --> (P xs = P ys)" and contr: "! x xs. P (x#x#xs) --> P (x#xs)" and weak: "! x xs. P xs --> P (x#xs)" and xy: "set x <= set y" and Px : "P x" shows "P y" proof - from contr weak have contr': "! x xs. P(x#x#xs) = P (x#xs)" by blast define y' where "y' = filter (% z. z : set x) y" from xy have "set x = set y'" apply(simp add: y'_def) apply blast done - hence rxry': "remdups x <~~> remdups y'" by(simp add: perm_remdups_iff_eq_set) + hence rxry': "mset (remdups x) = mset (remdups y')" + using set_eq_iff_mset_remdups_eq by auto from Px perm_contr[OF perm contr'] have Prx: "P (remdups x)" by simp - with rxry' have "P (remdups y')" by(simp add: perm) + with rxry' have "P (remdups y')" apply (subst perm) defer apply assumption apply simp done with perm_contr[OF perm contr'] have "P y'" by simp thus "P y" apply(simp add: y'_def) apply(rule perm_weak_filter[OF perm weak]) . qed (* No, not used subsection "Following used in Soundness" primrec multiset_of_list :: "'a list \ 'a multiset" where "multiset_of_list [] = {#}" | "multiset_of_list (x#xs) = {#x#} + multiset_of_list xs" lemma count_count[symmetric]: "count x A = Multiset.count (multiset_of_list A) x" by (induct A) simp_all lemma perm_multiset: "A <~~> B = (multiset_of_list A = multiset_of_list B)" apply(simp add: perm_count_conv) apply(simp add: multiset_eq_iff) apply(simp add: count_count) done lemma set_of_multiset_of_list: "set_of (multiset_of_list A) = set A" by (induct A) auto *) end diff --git a/thys/Completeness/Sequents.thy b/thys/Completeness/Sequents.thy --- a/thys/Completeness/Sequents.thy +++ b/thys/Completeness/Sequents.thy @@ -1,313 +1,310 @@ section "Sequents" theory Sequents imports Formula begin type_synonym sequent = "formula list" definition evalS :: "[model,vbl => object,formula list] => bool" where "evalS M phi fs \ (? f : set fs . evalF M phi f = True)" lemma evalS_nil[simp]: "evalS M phi [] = False" by(simp add: evalS_def) lemma evalS_cons[simp]: "evalS M phi (A # Gamma) = (evalF M phi A | evalS M phi Gamma)" by(simp add: evalS_def) lemma evalS_append: "evalS M phi (Gamma @ Delta) = (evalS M phi Gamma | evalS M phi Delta)" by(force simp add: evalS_def) lemma evalS_equiv[rule_format]: "(equalOn (freeVarsFL Gamma) f g) --> (evalS M f Gamma = evalS M g Gamma)" apply (induct Gamma, simp, rule) apply(simp add: freeVarsFL_cons) apply(drule_tac equalOn_UnD) apply(blast dest: evalF_equiv) done definition modelAssigns :: "[model] => (vbl => object) set" where "modelAssigns M = { phi . range phi <= objects M }" lemma modelAssignsI: "range f <= objects M \ f : modelAssigns M" by(simp add: modelAssigns_def) lemma modelAssignsD: "f : modelAssigns M \ range f <= objects M" by(simp add: modelAssigns_def) definition validS :: "formula list => bool" where "validS fs \ (! M . ! phi : modelAssigns M . evalS M phi fs = True)" subsection "Rules" type_synonym rule = "sequent * (sequent set)" definition concR :: "rule => sequent" where "concR = (%(conc,prems). conc)" definition premsR :: "rule => sequent set" where "premsR = (%(conc,prems). prems)" definition mapRule :: "(formula => formula) => rule => rule" where "mapRule = (%f (conc,prems) . (map f conc,(map f) ` prems))" lemma mapRuleI: "[| A = map f a; B = (map f) ` b |] ==> (A,B) = mapRule f (a,b)" by(simp add: mapRule_def) \ \FIXME tjr would like symmetric\ subsection "Deductions" (*FIXME. I don't see why plain Pow_mono is rejected.*) lemmas Powp_mono [mono] = Pow_mono [to_pred pred_subset_eq] inductive_set deductions :: "rule set => formula list set" for rules :: "rule set" (****** * Given a set of rules, * 1. Given a rule conc/prem(i) in rules, * and the prem(i) are deductions from rules, * then conc is a deduction from rules. * 2. can derive permutation of any deducible formula list. * (supposed to be multisets not lists). ******) where inferI: "[| (conc,prems) : rules; prems : Pow(deductions(rules)) |] ==> conc : deductions(rules)" (* perms "[| permutation conc' conc; conc' : deductions(rules) |] ==> conc : deductions(rules)" *) lemma mono_deductions: "[| A <= B |] ==> deductions(A) <= deductions(B)" apply(best intro: deductions.inferI elim: deductions.induct) done (*lemmas deductionsMono = mono_deductions*) (* -- "tjr following should be subsetD?" lemmas deductionSubsetI = mono_deductions[THEN subsetD] thm deductionSubsetI *) (****** * (f : formula -> formula) extended structurally over rules, deductions etc... * (((If f maps rules into themselves then can consider mapping derivation trees.))) * (((Is the asm necessary - think not?))) * The mapped deductions from the rules are same as * the deductions from the mapped rules. * * WHY: * * map f `` deductions rules <= deductions (mapRule f `` rules) (this thm) * <= deductions rules (closed) * * If rules are closed under f then so are deductions. * Can take f = (subst u v) and have application to exercise #1. * * Q: maybe also make f dual mapping, (what about quantifier side conditions...?). ******) (* lemma map_deductions: "map f ` deductions rules <= deductions (mapRule f ` rules)" apply(rule subsetI) apply (erule_tac imageE, simp) apply(erule deductions.induct) apply(blast intro: deductions.inferI mapRuleI) done lemma deductionsCloseRules: "! (conc,prems) : S . prems <= deductions R --> conc : deductions R ==> deductions (R Un S) = deductions R" apply(rule equalityI) prefer 2 apply(rule mono_deductions) apply blast apply(rule subsetI) apply (erule_tac deductions.induct, simp) apply(erule conjE) apply(thin_tac "prems \ deductions (R \ S)") apply(erule disjE) apply(rule inferI) apply assumption apply force apply blast done *) subsection "Basic Rule sets" definition "Axioms = { z. ? p vs. z = ([FAtom Pos p vs,FAtom Neg p vs],{}) }" definition "Conjs = { z. ? A0 A1 Delta Gamma. z = (FConj Pos A0 A1#Gamma @ Delta,{A0#Gamma,A1#Delta}) }" definition "Disjs = { z. ? A0 A1 Gamma. z = (FConj Neg A0 A1#Gamma,{A0#A1#Gamma}) }" definition "Alls = { z. ? A x Gamma. z = (FAll Pos A#Gamma,{instanceF x A#Gamma}) & x ~: freeVarsFL (FAll Pos A#Gamma) }" definition "Exs = { z. ? A x Gamma. z = (FAll Neg A#Gamma,{instanceF x A#Gamma})}" definition "Weaks = { z. ? A Gamma. z = (A#Gamma,{Gamma})}" definition "Contrs = { z. ? A Gamma. z = (A#Gamma,{A#A#Gamma})}" definition "Cuts = { z. ? C Delta Gamma. z = (Gamma @ Delta,{C#Gamma,FNot C#Delta})}" definition - "Perms = { z. ? Gamma Gamma' . z = (Gamma,{Gamma'}) & Gamma <~~> Gamma'}" + "Perms = { z. ? Gamma Gamma' . z = (Gamma,{Gamma'}) & mset Gamma = mset Gamma'}" definition "DAxioms = { z. ? p vs. z = ([FAtom Neg p vs,FAtom Pos p vs],{}) }" lemma AxiomI: "[| Axioms <= A |] ==> [FAtom Pos p vs,FAtom Neg p vs] : deductions(A)" apply(rule deductions.inferI) apply(auto simp add: Axioms_def) done lemma DAxiomsI: "[| DAxioms <= A |] ==> [FAtom Neg p vs,FAtom Pos p vs] : deductions(A)" apply(rule deductions.inferI) apply(auto simp add: DAxioms_def) done lemma DisjI: "[| A0#A1#Gamma : deductions(A); Disjs <= A |] ==> (FConj Neg A0 A1#Gamma) : deductions(A)" apply(rule deductions.inferI) apply(auto simp add: Disjs_def) done lemma ConjI: "[| (A0#Gamma) : deductions(A); (A1#Delta) : deductions(A); Conjs <= A |] ==> FConj Pos A0 A1#Gamma @ Delta : deductions(A)" apply(rule_tac prems="{A0#Gamma,A1#Delta}" in deductions.inferI) apply(auto simp add: Conjs_def) apply force done lemma AllI: "[| instanceF w A#Gamma : deductions(R); w ~: freeVarsFL (FAll Pos A#Gamma); Alls <= R |] ==> (FAll Pos A#Gamma) : deductions(R)" apply(rule_tac prems="{instanceF w A#Gamma}" in deductions.inferI) apply(auto simp add: Alls_def) done lemma ExI: "[| instanceF w A#Gamma : deductions(R); Exs <= R |] ==> (FAll Neg A#Gamma) : deductions(R)" apply(rule_tac prems = "{instanceF w A#Gamma}" in deductions.inferI) apply(auto simp add: Exs_def) done lemma WeakI: "[| Gamma : deductions R; Weaks <= R |] ==> A#Gamma : deductions(R)" apply(rule_tac prems="{Gamma}" in deductions.inferI) apply(auto simp add: Weaks_def) done lemma ContrI: "[| A#A#Gamma : deductions R; Contrs <= R |] ==> A#Gamma : deductions(R)" apply(rule_tac prems="{A#A#Gamma}" in deductions.inferI) apply(auto simp add: Contrs_def) done -lemma PermI: "[| Gamma' : deductions R; Gamma <~~> Gamma'; Perms <= R |] ==> Gamma : deductions(R)" +lemma PermI: "[| Gamma' : deductions R; mset Gamma = mset Gamma'; Perms <= R |] ==> Gamma : deductions(R)" apply(rule_tac prems="{Gamma'}" in deductions.inferI) apply(auto simp add: Perms_def) done subsection "Derived Rules" lemma WeakI1: "[| Gamma : deductions(A); Weaks <= A |] ==> (Delta @ Gamma) : deductions(A)" apply (induct Delta, simp) apply(auto intro: WeakI) done lemma WeakI2: "[| Gamma : deductions(A); Perms <= A; Weaks <= A |] ==> (Gamma @ Delta) : deductions(A)" - apply(blast intro: PermI perm_append_swap WeakI1) done + apply (auto intro: PermI [of \Delta @ Gamma\] WeakI1) + done lemma SATAxiomI: "[| Axioms <= A; Weaks <= A; Perms <= A; forms = [FAtom Pos n vs,FAtom Neg n vs] @ Gamma |] ==> forms : deductions(A)" apply(simp only:) apply(blast intro: WeakI2 AxiomI) done lemma DisjI1: "[| (A1#Gamma) : deductions(A); Disjs <= A; Weaks <= A |] ==> FConj Neg A0 A1#Gamma : deductions(A)" apply(blast intro: DisjI WeakI) done lemma DisjI2: "!!A. [| (A0#Gamma) : deductions(A); Disjs <= A; Weaks <= A; Perms <= A |] ==> FConj Neg A0 A1#Gamma : deductions(A)" apply(rule DisjI) - apply(rule PermI[OF _ perm.swap]) + apply(rule PermI [of \A1 # A0 # Gamma\]) + apply simp_all apply(rule WeakI) . \ \FIXME the following 4 lemmas could all be proved for the standard rule sets using monotonicity as below\ \ \we keep proofs as in original, but they are slightly ugly, and do not state what is intuitively happening\ lemma perm_tmp4: "Perms \ R \ A @ (a # list) @ (a # list) : deductions R \ (a # a # A) @ list @ list : deductions R" apply (rule PermI, auto) - apply(simp add: perm_count_conv count_append) done + done lemma weaken_append[rule_format]: "Contrs <= R ==> Perms <= R ==> !A. A @ Gamma @ Gamma : deductions(R) --> A @ Gamma : deductions(R)" apply (induct_tac Gamma, simp, rule) apply rule apply(drule_tac x="a#a#A" in spec) apply(erule_tac impE) apply(rule perm_tmp4) apply(assumption, assumption) apply(thin_tac "A @ (a # list) @ a # list \ deductions R") apply simp apply(frule_tac ContrI) apply assumption apply(thin_tac "a # a # A @ list \ deductions R") apply(rule PermI) apply assumption apply(simp add: perm_count_conv count_append) by assumption \ \FIXME horrible\ lemma ListWeakI: "Perms <= R ==> Contrs <= R ==> x # Gamma @ Gamma : deductions(R) ==> x # Gamma : deductions(R)" by(rule weaken_append[of R "[x]" Gamma, simplified]) lemma ConjI': "[| (A0#Gamma) : deductions(A); (A1#Gamma) : deductions(A); Contrs <= A; Conjs <= A; Perms <= A |] ==> FConj Pos A0 A1#Gamma : deductions(A)" apply(rule ListWeakI, assumption, assumption) apply(rule ConjI) . subsection "Standard Rule Sets For Predicate Calculus" definition PC :: "rule set" where "PC = Union {Perms,Axioms,Conjs,Disjs,Alls,Exs,Weaks,Contrs,Cuts}" definition CutFreePC :: "rule set" where "CutFreePC = Union {Perms,Axioms,Conjs,Disjs,Alls,Exs,Weaks,Contrs}" lemma rulesInPCs: "Axioms <= PC" "Axioms <= CutFreePC" "Conjs <= PC" "Conjs <= CutFreePC" "Disjs <= PC" "Disjs <= CutFreePC" "Alls <= PC" "Alls <= CutFreePC" "Exs <= PC" "Exs <= CutFreePC" "Weaks <= PC" "Weaks <= CutFreePC" "Contrs <= PC" "Contrs <= CutFreePC" "Perms <= PC" "Perms <= CutFreePC" "Cuts <= PC" "CutFreePC <= PC" by(auto simp: PC_def CutFreePC_def) subsection "Monotonicity for CutFreePC deductions" \ \these lemmas can be used to replace complicated permutation reasoning above\ \ \essentially if x is a deduction, and set x subset set y, then y is a deduction\ definition inDed :: "formula list => bool" where "inDed xs \ xs : deductions CutFreePC" -lemma perm: "! xs ys. xs <~~> ys --> (inDed xs = inDed ys)" - apply(subgoal_tac "! xs ys. xs <~~> ys --> inDed xs --> inDed ys") - apply (blast intro: perm_sym, clarify) - apply(simp add: inDed_def) - apply (rule PermI, assumption) - apply(rule perm_sym) apply assumption - by(blast intro!: rulesInPCs) +lemma perm: "! xs ys. mset xs = mset ys --> (inDed xs = inDed ys)" + by (metis PermI inDed_def rulesInPCs(16)) lemma contr: "! x xs. inDed (x#x#xs) --> inDed (x#xs)" apply(simp add: inDed_def) apply(blast intro!: ContrI rulesInPCs) done lemma weak: "! x xs. inDed xs --> inDed (x#xs)" apply(simp add: inDed_def) apply(blast intro!: WeakI rulesInPCs) done lemma inDed_mono'[simplified inDed_def]: "set x <= set y ==> inDed x ==> inDed y" using perm_weak_contr_mono[OF perm contr weak] . lemma inDed_mono[simplified inDed_def]: "inDed x ==> set x <= set y ==> inDed y" using perm_weak_contr_mono[OF perm contr weak] . end diff --git a/thys/Completeness/Soundness.thy b/thys/Completeness/Soundness.thy --- a/thys/Completeness/Soundness.thy +++ b/thys/Completeness/Soundness.thy @@ -1,98 +1,102 @@ section "Soundness" theory Soundness imports Completeness begin -lemma permutation_validS: "fs <~~> gs --> (validS fs = validS gs)" +lemma permutation_validS: "mset fs = mset gs --> (validS fs = validS gs)" apply(simp add: validS_def) apply(simp add: evalS_def) - apply(simp add: perm_set_eq) + using mset_eq_setD apply blast done lemma modelAssigns_vblcase: "phi \ modelAssigns M \ x \ objects M \ vblcase x phi \ modelAssigns M" apply (simp add: modelAssigns_def, rule) apply(erule_tac rangeE) apply(case_tac xaa rule: vbl_casesE, auto) done lemma tmp: "(!x : A. P x | Q) ==> (! x : A. P x) | Q " by blast lemma soundnessFAll: "!!Gamma. [| u ~: freeVarsFL (FAll Pos A # Gamma); validS (instanceF u A # Gamma) |] ==> validS (FAll Pos A # Gamma)" apply (simp add: validS_def, rule) apply (drule_tac x=M in spec, rule) apply(simp add: evalF_instance) apply (rule tmp, rule) apply(drule_tac x="% y. if y = u then x else phi y" in bspec) apply(simp add: modelAssigns_def) apply force apply(erule disjE) apply (rule disjI1, simp) apply(subgoal_tac "evalF M (vblcase x (\y. if y = u then x else phi y)) A = evalF M (vblcase x phi) A") apply force apply(rule evalF_equiv) apply(rule equalOn_vblcaseI) apply(rule,rule) apply(simp add: freeVarsFL_cons) apply (rule equalOnI, force) apply(rule disjI2) apply(subgoal_tac "evalS M (\y. if y = u then x else phi y) Gamma = evalS M phi Gamma") apply force apply(rule evalS_equiv) apply(rule equalOnI) apply(force simp: freeVarsFL_cons) done lemma soundnessFEx: "validS (instanceF x A # Gamma) ==> validS (FAll Neg A # Gamma)" apply(simp add: validS_def) apply (simp add: evalF_instance, rule, rule) apply(drule_tac x=M in spec) apply (drule_tac x=phi in bspec, assumption) apply(erule disjE) apply(rule disjI1) apply (rule_tac x="phi x" in bexI, assumption) apply(force dest: modelAssignsD subsetD) apply (rule disjI2, assumption) done lemma soundnessFCut: "[| validS (C # Gamma); validS (FNot C # Delta) |] ==> validS (Gamma @ Delta)" (* apply(force simp: validS_def evalS_append evalS_cons evalF_FNot)*) apply (simp add: validS_def, rule, rule) apply(drule_tac x=M in spec) apply(drule_tac x=M in spec) apply(drule_tac x=phi in bspec) apply assumption apply(drule_tac x=phi in bspec) apply assumption apply (simp add: evalS_append evalF_FNot, blast) done lemma soundness: "fs : deductions(PC) ==> (validS fs)" apply(erule_tac deductions.induct) apply(drule_tac PowD) apply(subgoal_tac "prems \ {x. validS x}") prefer 2 apply force apply(thin_tac "prems \ deductions PC \ {x. validS x}") apply(simp add: subset_eq) apply(simp add: PC_def) apply(elim disjE) - apply(force simp add: Perms_def permutation_validS) + apply (auto simp add: Perms_def) + apply (subst permutation_validS) + defer + apply assumption + apply simp_all apply(force simp: Axioms_def validS_def evalS_def) apply(force simp: Conjs_def validS_def evalS_def) apply(force simp: Disjs_def validS_def evalS_def) apply(simp add: Alls_def) apply(force intro: soundnessFAll) apply(simp add: Exs_def) apply(force intro: soundnessFEx) apply(force simp: Weaks_def validS_def evalS_def) apply(force simp: Contrs_def validS_def evalS_def) apply(force simp: Cuts_def intro: soundnessFCut) done lemma completeness: "fs : deductions (PC) = validS fs" apply rule apply(rule soundness) apply assumption apply(subgoal_tac "fs : deductions CutFreePC") apply(rule subsetD) prefer 2 apply assumption apply(rule mono_deductions) apply(simp add: PC_def CutFreePC_def) apply blast apply(rule adequacy) by assumption end diff --git a/thys/Groebner_Macaulay/Cone_Decomposition.thy b/thys/Groebner_Macaulay/Cone_Decomposition.thy --- a/thys/Groebner_Macaulay/Cone_Decomposition.thy +++ b/thys/Groebner_Macaulay/Cone_Decomposition.thy @@ -1,4893 +1,4888 @@ (* Author: Alexander Maletzky *) section \Cone Decompositions\ theory Cone_Decomposition imports Groebner_Bases.Groebner_PM Monomial_Module Hilbert_Function begin subsection \More Properties of Reduced Gr\"obner Bases\ context pm_powerprod begin lemmas reduced_GB_subset_monic_Polys = punit.reduced_GB_subset_monic_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum] lemmas reduced_GB_is_monomial_set_Polys = punit.reduced_GB_is_monomial_set_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum] lemmas is_red_reduced_GB_monomial_lt_GB_Polys = punit.is_red_reduced_GB_monomial_lt_GB_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum] lemmas reduced_GB_monomial_lt_reduced_GB_Polys = punit.reduced_GB_monomial_lt_reduced_GB_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum] end subsection \Quotient Ideals\ definition quot_set :: "'a set \ 'a \ 'a::semigroup_mult set" (infixl "\
" 55) where "quot_set A x = (*) x -` A" lemma quot_set_iff: "a \ A \
x \ x * a \ A" by (simp add: quot_set_def) lemma quot_setI: "x * a \ A \ a \ A \
x" by (simp only: quot_set_iff) lemma quot_setD: "a \ A \
x \ x * a \ A" by (simp only: quot_set_iff) lemma quot_set_quot_set [simp]: "A \
x \
y = A \
x * y" by (rule set_eqI) (simp add: quot_set_iff mult.assoc) lemma quot_set_one [simp]: "A \
(1::_::monoid_mult) = A" by (rule set_eqI) (simp add: quot_set_iff) lemma ideal_quot_set_ideal [simp]: "ideal (ideal B \
x) = (ideal B) \
(x::_::comm_ring)" proof show "ideal (ideal B \
x) \ ideal B \
x" proof fix b assume "b \ ideal (ideal B \
x)" thus "b \ ideal B \
x" proof (induct b rule: ideal.span_induct') case base show ?case by (simp add: quot_set_iff ideal.span_zero) next case (step b q p) hence "x * b \ ideal B" and "x * p \ ideal B" by (simp_all add: quot_set_iff) hence "x * b + q * (x * p) \ ideal B" by (intro ideal.span_add ideal.span_scale[where c=q]) thus ?case by (simp only: quot_set_iff algebra_simps) qed qed qed (fact ideal.span_superset) lemma quot_set_image_times: "inj ((*) x) \ ((*) x ` A) \
x = A" by (simp add: quot_set_def inj_vimage_image_eq) subsection \Direct Decompositions of Polynomial Rings\ context pm_powerprod begin definition normal_form :: "(('x \\<^sub>0 nat) \\<^sub>0 'a) set \ (('x \\<^sub>0 nat) \\<^sub>0 'a::field) \ (('x \\<^sub>0 nat) \\<^sub>0 'a::field)" where "normal_form F p = (SOME q. (punit.red (punit.reduced_GB F))\<^sup>*\<^sup>* p q \ \ punit.is_red (punit.reduced_GB F) q)" text \Of course, @{const normal_form} could be defined in a much more general context.\ context fixes X :: "'x set" assumes fin_X: "finite X" begin context fixes F :: "(('x \\<^sub>0 nat) \\<^sub>0 'a::field) set" assumes F_sub: "F \ P[X]" begin lemma normal_form: shows "(punit.red (punit.reduced_GB F))\<^sup>*\<^sup>* p (normal_form F p)" (is ?thesis1) and "\ punit.is_red (punit.reduced_GB F) (normal_form F p)" (is ?thesis2) proof - from fin_X F_sub have "finite (punit.reduced_GB F)" by (rule finite_reduced_GB_Polys) hence "wfP (punit.red (punit.reduced_GB F))\\" by (rule punit.red_wf_finite) then obtain q where "(punit.red (punit.reduced_GB F))\<^sup>*\<^sup>* p q" and "\ punit.is_red (punit.reduced_GB F) q" unfolding punit.is_red_def not_not by (rule relation.wf_imp_nf_ex) hence "(punit.red (punit.reduced_GB F))\<^sup>*\<^sup>* p q \ \ punit.is_red (punit.reduced_GB F) q" .. hence "?thesis1 \ ?thesis2" unfolding normal_form_def by (rule someI) thus ?thesis1 and ?thesis2 by simp_all qed lemma normal_form_unique: assumes "(punit.red (punit.reduced_GB F))\<^sup>*\<^sup>* p q" and "\ punit.is_red (punit.reduced_GB F) q" shows "normal_form F p = q" proof (rule relation.ChurchRosser_unique_final) from fin_X F_sub have "punit.is_Groebner_basis (punit.reduced_GB F)" by (rule reduced_GB_is_GB_Polys) thus "relation.is_ChurchRosser (punit.red (punit.reduced_GB F))" by (simp only: punit.is_Groebner_basis_def) next show "(punit.red (punit.reduced_GB F))\<^sup>*\<^sup>* p (normal_form F p)" by (rule normal_form) next have "\ punit.is_red (punit.reduced_GB F) (normal_form F p)" by (rule normal_form) thus "relation.is_final (punit.red (punit.reduced_GB F)) (normal_form F p)" by (simp add: punit.is_red_def) next from assms(2) show "relation.is_final (punit.red (punit.reduced_GB F)) q" by (simp add: punit.is_red_def) qed fact lemma normal_form_id_iff: "normal_form F p = p \ (\ punit.is_red (punit.reduced_GB F) p)" proof assume "normal_form F p = p" with normal_form(2)[of p] show "\ punit.is_red (punit.reduced_GB F) p" by simp next assume "\ punit.is_red (punit.reduced_GB F) p" with rtranclp.rtrancl_refl show "normal_form F p = p" by (rule normal_form_unique) qed lemma normal_form_normal_form: "normal_form F (normal_form F p) = normal_form F p" by (simp add: normal_form_id_iff normal_form) lemma normal_form_zero: "normal_form F 0 = 0" by (simp add: normal_form_id_iff punit.irred_0) lemma normal_form_map_scale: "normal_form F (c \ p) = c \ (normal_form F p)" by (intro normal_form_unique punit.is_irred_map_scale normal_form) (simp add: punit.map_scale_eq_monom_mult punit.red_rtrancl_mult normal_form) lemma normal_form_uminus: "normal_form F (- p) = - normal_form F p" by (intro normal_form_unique punit.red_rtrancl_uminus normal_form) (simp add: punit.is_red_uminus normal_form) lemma normal_form_plus_normal_form: "normal_form F (normal_form F p + normal_form F q) = normal_form F p + normal_form F q" by (intro normal_form_unique rtranclp.rtrancl_refl punit.is_irred_plus normal_form) lemma normal_form_minus_normal_form: "normal_form F (normal_form F p - normal_form F q) = normal_form F p - normal_form F q" by (intro normal_form_unique rtranclp.rtrancl_refl punit.is_irred_minus normal_form) lemma normal_form_ideal_Polys: "normal_form (ideal F \ P[X]) = normal_form F" proof - let ?F = "ideal F \ P[X]" from fin_X have eq: "punit.reduced_GB ?F = punit.reduced_GB F" proof (rule reduced_GB_unique_Polys) from fin_X F_sub show "punit.is_reduced_GB (punit.reduced_GB F)" by (rule reduced_GB_is_reduced_GB_Polys) next from fin_X F_sub have "ideal (punit.reduced_GB F) = ideal F" by (rule reduced_GB_ideal_Polys) also have "\ = ideal (ideal F \ P[X])" proof (intro subset_antisym ideal.span_subset_spanI) from ideal.span_superset[of F] F_sub have "F \ ideal F \ P[X]" by simp thus "F \ ideal (ideal F \ P[X])" using ideal.span_superset by (rule subset_trans) qed blast finally show "ideal (punit.reduced_GB F) = ideal (ideal F \ P[X])" . qed blast show ?thesis by (rule ext) (simp only: normal_form_def eq) qed lemma normal_form_diff_in_ideal: "p - normal_form F p \ ideal F" proof - from normal_form(1) have "p - normal_form F p \ ideal (punit.reduced_GB F)" by (rule punit.red_rtranclp_diff_in_pmdl[simplified]) also from fin_X F_sub have "\ = ideal F" by (rule reduced_GB_ideal_Polys) finally show ?thesis . qed lemma normal_form_zero_iff: "normal_form F p = 0 \ p \ ideal F" proof assume "normal_form F p = 0" with normal_form_diff_in_ideal[of p] show "p \ ideal F" by simp next assume "p \ ideal F" hence "p - (p - normal_form F p) \ ideal F" using normal_form_diff_in_ideal by (rule ideal.span_diff) also from fin_X F_sub have "\ = ideal (punit.reduced_GB F)" by (rule reduced_GB_ideal_Polys[symmetric]) finally have *: "normal_form F p \ ideal (punit.reduced_GB F)" by simp show "normal_form F p = 0" proof (rule ccontr) from fin_X F_sub have "punit.is_Groebner_basis (punit.reduced_GB F)" by (rule reduced_GB_is_GB_Polys) moreover note * moreover assume "normal_form F p \ 0" ultimately obtain g where "g \ punit.reduced_GB F" and "g \ 0" and a: "lpp g adds lpp (normal_form F p)" by (rule punit.GB_adds_lt[simplified]) note this(1, 2) moreover from \normal_form F p \ 0\ have "lpp (normal_form F p) \ keys (normal_form F p)" by (rule punit.lt_in_keys) ultimately have "punit.is_red (punit.reduced_GB F) (normal_form F p)" using a by (rule punit.is_red_addsI[simplified]) with normal_form(2) show False .. qed qed lemma normal_form_eq_iff: "normal_form F p = normal_form F q \ p - q \ ideal F" proof - have "p - q - (normal_form F p - normal_form F q) = (p - normal_form F p) - (q - normal_form F q)" by simp also from normal_form_diff_in_ideal normal_form_diff_in_ideal have "\ \ ideal F" by (rule ideal.span_diff) finally have *: "p - q - (normal_form F p - normal_form F q) \ ideal F" . show ?thesis proof assume "normal_form F p = normal_form F q" with * show "p - q \ ideal F" by simp next assume "p - q \ ideal F" hence "p - q - (p - q - (normal_form F p - normal_form F q)) \ ideal F" using * by (rule ideal.span_diff) hence "normal_form F (normal_form F p - normal_form F q) = 0" by (simp add: normal_form_zero_iff) thus "normal_form F p = normal_form F q" by (simp add: normal_form_minus_normal_form) qed qed lemma Polys_closed_normal_form: assumes "p \ P[X]" shows "normal_form F p \ P[X]" proof - from fin_X F_sub have "punit.reduced_GB F \ P[X]" by (rule reduced_GB_Polys) with fin_X show ?thesis using assms normal_form(1) by (rule punit.dgrad_p_set_closed_red_rtrancl[OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]) qed lemma image_normal_form_iff: "p \ normal_form F ` P[X] \ (p \ P[X] \ \ punit.is_red (punit.reduced_GB F) p)" proof assume "p \ normal_form F ` P[X]" then obtain q where "q \ P[X]" and p: "p = normal_form F q" .. from this(1) show "p \ P[X] \ \ punit.is_red (punit.reduced_GB F) p" unfolding p by (intro conjI Polys_closed_normal_form normal_form) next assume "p \ P[X] \ \ punit.is_red (punit.reduced_GB F) p" hence "p \ P[X]" and "\ punit.is_red (punit.reduced_GB F) p" by simp_all from this(2) have "normal_form F p = p" by (simp add: normal_form_id_iff) from this[symmetric] \p \ P[X]\ show "p \ normal_form F ` P[X]" by (rule image_eqI) qed end lemma direct_decomp_ideal_insert: fixes F and f defines "I \ ideal (insert f F)" defines "L \ (ideal F \
f) \ P[X]" assumes "F \ P[X]" and "f \ P[X]" shows "direct_decomp (I \ P[X]) [ideal F \ P[X], (*) f ` normal_form L ` P[X]]" (is "direct_decomp _ ?ss") proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain x y where x: "x \ ideal F \ P[X]" and y: "y \ (*) f ` normal_form L ` P[X]" and qs: "qs = [x, y]" by (rule listset_doubletonE) have "sum_list qs = x + y" by (simp add: qs) also have "\ \ I \ P[X]" unfolding I_def proof (intro IntI ideal.span_add Polys_closed_plus) have "ideal F \ ideal (insert f F)" by (rule ideal.span_mono) blast with x show "x \ ideal (insert f F)" and "x \ P[X]" by blast+ next from y obtain p where "p \ P[X]" and y: "y = f * normal_form L p" by blast have "f \ ideal (insert f F)" by (rule ideal.span_base) simp hence "normal_form L p * f \ ideal (insert f F)" by (rule ideal.span_scale) thus "y \ ideal (insert f F)" by (simp only: mult.commute y) have "L \ P[X]" by (simp add: L_def) hence "normal_form L p \ P[X]" using \p \ P[X]\ by (rule Polys_closed_normal_form) with assms(4) show "y \ P[X]" unfolding y by (rule Polys_closed_times) qed finally show "sum_list qs \ I \ P[X]" . next fix a assume "a \ I \ P[X]" hence "a \ I" and "a \ P[X]" by simp_all from assms(3, 4) have "insert f F \ P[X]" by simp then obtain F0 q0 where "F0 \ insert f F" and "finite F0" and q0: "\f0. q0 f0 \ P[X]" and a: "a = (\f0\F0. q0 f0 * f0)" using \a \ P[X]\ \a \ I\ unfolding I_def by (rule in_idealE_Polys) blast obtain q a' where a': "a' \ ideal F" and "a' \ P[X]" and "q \ P[X]" and a: "a = q * f + a'" proof (cases "f \ F0") case True with \F0 \ insert f F\ have "F0 - {f} \ F" by blast show ?thesis proof have "(\f0\F0 - {f}. q0 f0 * f0) \ ideal (F0 - {f})" by (rule ideal.sum_in_spanI) also from \F0 - {f} \ F\ have "\ \ ideal F" by (rule ideal.span_mono) finally show "(\f0\F0 - {f}. q0 f0 * f0) \ ideal F" . next show "(\f0\F0 - {f}. q0 f0 * f0) \ P[X]" proof (intro Polys_closed_sum Polys_closed_times q0) fix f0 assume "f0 \ F0 - {f}" also have "\ \ F0" by blast also have "\ \ insert f F" by fact also have "\ \ P[X]" by fact finally show "f0 \ P[X]" . qed next from \finite F0\ True show "a = q0 f * f + (\f0\F0 - {f}. q0 f0 * f0)" by (simp only: a sum.remove) qed fact next case False with \F0 \ insert f F\ have "F0 \ F" by blast show ?thesis proof have "a \ ideal F0" unfolding a by (rule ideal.sum_in_spanI) also from \F0 \ F\ have "\ \ ideal F" by (rule ideal.span_mono) finally show "a \ ideal F" . next show "a = 0 * f + a" by simp qed (fact \a \ P[X]\, fact zero_in_Polys) qed let ?a = "f * (normal_form L q)" have "L \ P[X]" by (simp add: L_def) hence "normal_form L q \ P[X]" using \q \ P[X]\ by (rule Polys_closed_normal_form) with assms(4) have "?a \ P[X]" by (rule Polys_closed_times) from \L \ P[X]\ have "q - normal_form L q \ ideal L" by (rule normal_form_diff_in_ideal) also have "\ \ ideal (ideal F \
f)" unfolding L_def by (rule ideal.span_mono) blast finally have "f * (q - normal_form L q) \ ideal F" by (simp add: quot_set_iff) with \a' \ ideal F\ have "a' + f * (q - normal_form L q) \ ideal F" by (rule ideal.span_add) hence "a - ?a \ ideal F" by (simp add: a algebra_simps) define qs where "qs = [a - ?a, ?a]" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI allI impI) have "a - ?a \ ideal F \ P[X]" proof from assms(4) \a \ P[X]\ \normal_form L q \ P[X]\ show "a - ?a \ P[X]" by (intro Polys_closed_minus Polys_closed_times) qed fact moreover from \q \ P[X]\ have "?a \ (*) f ` normal_form L ` P[X]" by (intro imageI) ultimately show "qs \ listset ?ss" using qs_def by (rule listset_doubletonI) next fix qs0 assume "qs0 \ listset ?ss \ a = sum_list qs0" hence "qs0 \ listset ?ss" and "a = sum_list qs0" by simp_all from this(1) obtain x y where "x \ ideal F \ P[X]" and "y \ (*) f ` normal_form L ` P[X]" and qs0: "qs0 = [x, y]" by (rule listset_doubletonE) from this(2) obtain a0 where "a0 \ P[X]" and y: "y = f * normal_form L a0" by blast from \x \ ideal F \ P[X]\ have "x \ ideal F" by simp have x: "x = a - y" by (simp add: \a = sum_list qs0\ qs0) have "f * (normal_form L q - normal_form L a0) = x - (a - ?a)" by (simp add: x y a algebra_simps) also from \x \ ideal F\ \a - ?a \ ideal F\ have "\ \ ideal F" by (rule ideal.span_diff) finally have "normal_form L q - normal_form L a0 \ ideal F \
f" by (rule quot_setI) moreover from \L \ P[X]\ \q \ P[X]\ \a0 \ P[X]\ have "normal_form L q - normal_form L a0 \ P[X]" by (intro Polys_closed_minus Polys_closed_normal_form) ultimately have "normal_form L q - normal_form L a0 \ L" by (simp add: L_def) also have "\ \ ideal L" by (fact ideal.span_superset) finally have "normal_form L q - normal_form L a0 = 0" using \L \ P[X]\ by (simp only: normal_form_minus_normal_form flip: normal_form_zero_iff) thus "qs0 = qs" by (simp add: qs0 qs_def x y) qed (simp_all add: qs_def) qed corollary direct_decomp_ideal_normal_form: assumes "F \ P[X]" shows "direct_decomp P[X] [ideal F \ P[X], normal_form F ` P[X]]" proof - from assms one_in_Polys have "direct_decomp (ideal (insert 1 F) \ P[X]) [ideal F \ P[X], (*) 1 ` normal_form ((ideal F \
1) \ P[X]) ` P[X]]" by (rule direct_decomp_ideal_insert) moreover have "ideal (insert 1 F) = UNIV" by (simp add: ideal_eq_UNIV_iff_contains_one ideal.span_base) moreover from refl have "((*) 1 \ normal_form F) ` P[X] = normal_form F ` P[X]" by (rule image_cong) simp ultimately show ?thesis using assms by (simp add: image_comp normal_form_ideal_Polys) qed end subsection \Basic Cone Decompositions\ definition cone :: "((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) \ (('x \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_0) set" where "cone hU = (*) (fst hU) ` P[snd hU]" lemma coneI: "p = a * h \ a \ P[U] \ p \ cone (h, U)" by (auto simp: cone_def mult.commute[of a]) lemma coneE: assumes "p \ cone (h, U)" obtains a where "a \ P[U]" and "p = a * h" using assms by (auto simp: cone_def mult.commute) lemma cone_empty: "cone (h, {}) = range (\c. c \ h)" by (auto simp: Polys_empty map_scale_eq_times intro: coneI elim!: coneE) lemma cone_zero [simp]: "cone (0, U) = {0}" by (auto simp: cone_def intro: zero_in_Polys) lemma cone_one [simp]: "cone (1::_ \\<^sub>0 'a::comm_semiring_1, U) = P[U]" by (auto simp: cone_def) lemma zero_in_cone: "0 \ cone hU" by (auto simp: cone_def intro!: image_eqI zero_in_Polys) corollary empty_not_in_map_cone: "{} \ set (map cone ps)" using zero_in_cone by fastforce lemma tip_in_cone: "h \ cone (h::_ \\<^sub>0 _::comm_semiring_1, U)" using _ one_in_Polys by (rule coneI) simp lemma cone_closed_plus: assumes "a \ cone hU" and "b \ cone hU" shows "a + b \ cone hU" proof - obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast with assms have "a \ cone (h, U)" and "b \ cone (h, U)" by simp_all from this(1) obtain a' where "a' \ P[U]" and a: "a = a' * h" by (rule coneE) from \b \ cone (h, U)\ obtain b' where "b' \ P[U]" and b: "b = b' * h" by (rule coneE) have "a + b = (a' + b') * h" by (simp only: a b algebra_simps) moreover from \a' \ P[U]\ \b' \ P[U]\ have "a' + b' \ P[U]" by (rule Polys_closed_plus) ultimately show ?thesis unfolding hU by (rule coneI) qed lemma cone_closed_uminus: assumes "(a::_ \\<^sub>0 _::comm_ring) \ cone hU" shows "- a \ cone hU" proof - obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast with assms have "a \ cone (h, U)" by simp from this(1) obtain a' where "a' \ P[U]" and a: "a = a' * h" by (rule coneE) have "- a = (- a') * h" by (simp add: a) moreover from \a' \ P[U]\ have "- a' \ P[U]" by (rule Polys_closed_uminus) ultimately show ?thesis unfolding hU by (rule coneI) qed lemma cone_closed_minus: assumes "(a::_ \\<^sub>0 _::comm_ring) \ cone hU" and "b \ cone hU" shows "a - b \ cone hU" proof - from assms(2) have "- b \ cone hU" by (rule cone_closed_uminus) with assms(1) have "a + (- b) \ cone hU" by (rule cone_closed_plus) thus ?thesis by simp qed lemma cone_closed_times: assumes "a \ cone (h, U)" and "q \ P[U]" shows "q * a \ cone (h, U)" proof - from assms(1) obtain a' where "a' \ P[U]" and a: "a = a' * h" by (rule coneE) have "q * a = (q * a') * h" by (simp only: a ac_simps) moreover from assms(2) \a' \ P[U]\ have "q * a' \ P[U]" by (rule Polys_closed_times) ultimately show ?thesis by (rule coneI) qed corollary cone_closed_monom_mult: assumes "a \ cone (h, U)" and "t \ .[U]" shows "punit.monom_mult c t a \ cone (h, U)" proof - from assms(2) have "monomial c t \ P[U]" by (rule Polys_closed_monomial) with assms(1) have "monomial c t * a \ cone (h, U)" by (rule cone_closed_times) thus ?thesis by (simp only: times_monomial_left) qed lemma coneD: assumes "p \ cone (h, U)" and "p \ 0" shows "lpp h adds lpp (p::_ \\<^sub>0 _::{comm_semiring_0,semiring_no_zero_divisors})" proof - from assms(1) obtain a where p: "p = a * h" by (rule coneE) with assms(2) have "a \ 0" and "h \ 0" by auto hence "lpp p = lpp a + lpp h" unfolding p by (rule lp_times) also have "\ = lpp h + lpp a" by (rule add.commute) finally show ?thesis by (rule addsI) qed lemma cone_mono_1: assumes "h' \ P[U]" shows "cone (h' * h, U) \ cone (h, U)" proof fix p assume "p \ cone (h' * h, U)" then obtain a' where "a' \ P[U]" and "p = a' * (h' * h)" by (rule coneE) from this(2) have "p = a' * h' * h" by (simp only: mult.assoc) moreover from \a' \ P[U]\ assms have "a' * h' \ P[U]" by (rule Polys_closed_times) ultimately show "p \ cone (h, U)" by (rule coneI) qed lemma cone_mono_2: assumes "U1 \ U2" shows "cone (h, U1) \ cone (h, U2)" proof from assms have "P[U1] \ P[U2]" by (rule Polys_mono) fix p assume "p \ cone (h, U1)" then obtain a where "a \ P[U1]" and "p = a * h" by (rule coneE) note this(2) moreover from \a \ P[U1]\ \P[U1] \ P[U2]\ have "a \ P[U2]" .. ultimately show "p \ cone (h, U2)" by (rule coneI) qed lemma cone_subsetD: assumes "cone (h1, U1) \ cone (h2::_ \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}, U2)" shows "h2 dvd h1" and "h1 \ 0 \ U1 \ U2" proof - from tip_in_cone assms have "h1 \ cone (h2, U2)" .. then obtain a1' where "a1' \ P[U2]" and h1: "h1 = a1' * h2" by (rule coneE) from this(2) have "h1 = h2 * a1'" by (simp only: mult.commute) thus "h2 dvd h1" .. assume "h1 \ 0" with h1 have "a1' \ 0" and "h2 \ 0" by auto show "U1 \ U2" proof fix x assume "x \ U1" hence "monomial (1::'a) (Poly_Mapping.single x 1) \ P[U1]" (is "?p \ _") by (intro Polys_closed_monomial PPs_closed_single) with refl have "?p * h1 \ cone (h1, U1)" by (rule coneI) hence "?p * h1 \ cone (h2, U2)" using assms .. then obtain a where "a \ P[U2]" and "?p * h1 = a * h2" by (rule coneE) from this(2) have "(?p * a1') * h2 = a * h2" by (simp only: h1 ac_simps) hence "?p * a1' = a" using \h2 \ 0\ by (rule times_canc_right) with \a \ P[U2]\ have "a1' * ?p \ P[U2]" by (simp add: mult.commute) hence "?p \ P[U2]" using \a1' \ P[U2]\ \a1' \ 0\ by (rule times_in_PolysD) thus "x \ U2" by (simp add: Polys_def PPs_def) qed qed lemma cone_subset_PolysD: assumes "cone (h::_ \\<^sub>0 'a::{comm_semiring_1,semiring_no_zero_divisors}, U) \ P[X]" shows "h \ P[X]" and "h \ 0 \ U \ X" proof - from tip_in_cone assms show "h \ P[X]" .. assume "h \ 0" show "U \ X" proof fix x assume "x \ U" hence "monomial (1::'a) (Poly_Mapping.single x 1) \ P[U]" (is "?p \ _") by (intro Polys_closed_monomial PPs_closed_single) with refl have "?p * h \ cone (h, U)" by (rule coneI) hence "?p * h \ P[X]" using assms .. hence "h * ?p \ P[X]" by (simp only: mult.commute) hence "?p \ P[X]" using \h \ P[X]\ \h \ 0\ by (rule times_in_PolysD) thus "x \ X" by (simp add: Polys_def PPs_def) qed qed lemma cone_subset_PolysI: assumes "h \ P[X]" and "h \ 0 \ U \ X" shows "cone (h, U) \ P[X]" proof (cases "h = 0") case True thus ?thesis by (simp add: zero_in_Polys) next case False hence "U \ X" by (rule assms(2)) hence "P[U] \ P[X]" by (rule Polys_mono) show ?thesis proof fix a assume "a \ cone (h, U)" then obtain q where "q \ P[U]" and a: "a = q * h" by (rule coneE) from this(1) \P[U] \ P[X]\ have "q \ P[X]" .. from this assms(1) show "a \ P[X]" unfolding a by (rule Polys_closed_times) qed qed lemma cone_image_times: "(*) a ` cone (h, U) = cone (a * h, U)" by (auto simp: ac_simps image_image intro!: image_eqI coneI elim!: coneE) lemma cone_image_times': "(*) a ` cone hU = cone (apfst ((*) a) hU)" proof - obtain h U where "hU = (h, U)" using prod.exhaust by blast thus ?thesis by (simp add: cone_image_times) qed lemma homogeneous_set_coneI: assumes "homogeneous h" shows "homogeneous_set (cone (h, U))" proof (rule homogeneous_setI) fix a n assume "a \ cone (h, U)" then obtain q where "q \ P[U]" and a: "a = q * h" by (rule coneE) from this(1) show "hom_component a n \ cone (h, U)" unfolding a proof (induct q rule: poly_mapping_plus_induct) case 1 show ?case by (simp add: zero_in_cone) next case (2 p c t) have "p \ P[U]" proof (intro PolysI subsetI) fix s assume "s \ keys p" moreover from 2(2) this have "s \ keys (monomial c t)" by auto ultimately have "s \ keys (monomial c t + p)" by (rule in_keys_plusI2) also from 2(4) have "\ \ .[U]" by (rule PolysD) finally show "s \ .[U]" . qed hence *: "hom_component (p * h) n \ cone (h, U)" by (rule 2(3)) from 2(1) have "t \ keys (monomial c t)" by simp hence "t \ keys (monomial c t + p)" using 2(2) by (rule in_keys_plusI1) also from 2(4) have "\ \ .[U]" by (rule PolysD) finally have "monomial c t \ P[U]" by (rule Polys_closed_monomial) with refl have "monomial c t * h \ cone (h, U)" (is "?h \ _") by (rule coneI) from assms have "homogeneous ?h" by (simp add: homogeneous_times) hence "hom_component ?h n = (?h when n = poly_deg ?h)" by (rule hom_component_of_homogeneous) with \?h \ cone (h, U)\ have **: "hom_component ?h n \ cone (h, U)" by (simp add: when_def zero_in_cone) have "hom_component ((monomial c t + p) * h) n = hom_component ?h n + hom_component (p * h) n" by (simp only: distrib_right hom_component_plus) also from ** * have "\ \ cone (h, U)" by (rule cone_closed_plus) finally show ?case . qed qed lemma subspace_cone: "phull.subspace (cone hU)" using zero_in_cone cone_closed_plus proof (rule phull.subspaceI) fix c a assume "a \ cone hU" moreover obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast ultimately have "a \ cone (h, U)" by simp thus "c \ a \ cone hU" unfolding hU punit.map_scale_eq_monom_mult using zero_in_PPs by (rule cone_closed_monom_mult) qed lemma direct_decomp_cone_insert: fixes h :: "_ \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}" assumes "x \ U" shows "direct_decomp (cone (h, insert x U)) [cone (h, U), cone (monomial 1 (Poly_Mapping.single x (Suc 0)) * h, insert x U)]" proof - let ?x = "Poly_Mapping.single x (Suc 0)" define xx where "xx = monomial (1::'a) ?x" show "direct_decomp (cone (h, insert x U)) [cone (h, U), cone (xx * h, insert x U)]" (is "direct_decomp _ ?ss") proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain a b where "a \ cone (h, U)" and b: "b \ cone (xx * h, insert x U)" and qs: "qs = [a, b]" by (rule listset_doubletonE) note this(1) also have "cone (h, U) \ cone (h, insert x U)" by (rule cone_mono_2) blast finally have a: "a \ cone (h, insert x U)" . have "cone (xx * h, insert x U) \ cone (h, insert x U)" by (rule cone_mono_1) (simp add: xx_def Polys_def PPs_closed_single) with b have "b \ cone (h, insert x U)" .. with a have "a + b \ cone (h, insert x U)" by (rule cone_closed_plus) thus "sum_list qs \ cone (h, insert x U)" by (simp add: qs) next fix a assume "a \ cone (h, insert x U)" then obtain q where "q \ P[insert x U]" and a: "a = q * h" by (rule coneE) define qU where "qU = except q (- .[U])" define qx where "qx = except q .[U]" have q: "q = qU + qx" by (simp only: qU_def qx_def add.commute flip: except_decomp) have "qU \ P[U]" by (rule PolysI) (simp add: qU_def keys_except) have x_adds: "?x adds t" if "t \ keys qx" for t unfolding adds_poly_mapping le_fun_def proof fix y show "lookup ?x y \ lookup t y" proof (cases "y = x") case True from that have "t \ keys q" and "t \ .[U]" by (simp_all add: qx_def keys_except) from \q \ P[insert x U]\ have "keys q \ .[insert x U]" by (rule PolysD) with \t \ keys q\ have "t \ .[insert x U]" .. hence "keys t \ insert x U" by (rule PPsD) moreover from \t \ .[U]\ have "\ keys t \ U" by (simp add: PPs_def) ultimately have "x \ keys t" by blast thus ?thesis by (simp add: lookup_single True in_keys_iff) next case False thus ?thesis by (simp add: lookup_single) qed qed define qx' where "qx' = Poly_Mapping.map_key ((+) ?x) qx" have lookup_qx': "lookup qx' = (\t. lookup qx (?x + t))" by (rule ext) (simp add: qx'_def map_key.rep_eq) have "qx' * xx = punit.monom_mult 1 ?x qx'" by (simp only: xx_def mult.commute flip: times_monomial_left) also have "\ = qx" by (auto simp: punit.lookup_monom_mult lookup_qx' add.commute[of ?x] adds_minus simp flip: not_in_keys_iff_lookup_eq_zero dest: x_adds intro!: poly_mapping_eqI) finally have qx: "qx = qx' * xx" by (rule sym) have "qx' \ P[insert x U]" proof (intro PolysI subsetI) fix t assume "t \ keys qx'" hence "t + ?x \ keys qx" by (simp only: lookup_qx' in_keys_iff not_False_eq_True add.commute) also have "\ \ keys q" by (auto simp: qx_def keys_except) also from \q \ P[insert x U]\ have "\ \ .[insert x U]" by (rule PolysD) finally have "(t + ?x) - ?x \ .[insert x U]" by (rule PPs_closed_minus) thus "t \ .[insert x U]" by simp qed define qs where "qs = [qU * h, qx' * (xx * h)]" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI allI impI) from refl \qU \ P[U]\ have "qU * h \ cone (h, U)" by (rule coneI) moreover from refl \qx' \ P[insert x U]\ have "qx' * (xx * h) \ cone (xx * h, insert x U)" by (rule coneI) ultimately show "qs \ listset ?ss" using qs_def by (rule listset_doubletonI) next fix qs0 assume "qs0 \ listset ?ss \ a = sum_list qs0" hence "qs0 \ listset ?ss" and a0: "a = sum_list qs0" by simp_all from this(1) obtain p1 p2 where "p1 \ cone (h, U)" and p2: "p2 \ cone (xx * h, insert x U)" and qs0: "qs0 = [p1, p2]" by (rule listset_doubletonE) from this(1) obtain qU0 where "qU0 \ P[U]" and p1: "p1 = qU0 * h" by (rule coneE) from p2 obtain qx0 where p2: "p2 = qx0 * (xx * h)" by (rule coneE) show "qs0 = qs" proof (cases "h = 0") case True thus ?thesis by (simp add: qs_def qs0 p1 p2) next case False from a0 have "(qU - qU0) * h = (qx0 - qx') * xx * h" by (simp add: a qs0 p1 p2 q qx algebra_simps) hence eq: "qU - qU0 = (qx0 - qx') * xx" using False by (rule times_canc_right) have "qx0 = qx'" proof (rule ccontr) assume "qx0 \ qx'" hence "qx0 - qx' \ 0" by simp moreover have "xx \ 0" by (simp add: xx_def monomial_0_iff) ultimately have "lpp ((qx0 - qx') * xx) = lpp (qx0 - qx') + lpp xx" by (rule lp_times) also have "lpp xx = ?x" by (simp add: xx_def punit.lt_monomial) finally have "?x adds lpp (qU - qU0)" by (simp add: eq) hence "lookup ?x x \ lookup (lpp (qU - qU0)) x" by (simp only: adds_poly_mapping le_fun_def) hence "x \ keys (lpp (qU - qU0))" by (simp add: in_keys_iff lookup_single) moreover have "lpp (qU - qU0) \ keys (qU - qU0)" proof (rule punit.lt_in_keys) from \qx0 - qx' \ 0\ \xx \ 0\ show "qU - qU0 \ 0" unfolding eq by (rule times_not_zero) qed ultimately have "x \ indets (qU - qU0)" by (rule in_indetsI) from \qU \ P[U]\ \qU0 \ P[U]\ have "qU - qU0 \ P[U]" by (rule Polys_closed_minus) hence "indets (qU - qU0) \ U" by (rule PolysD) with \x \ indets (qU - qU0)\ have "x \ U" .. with assms show False .. qed moreover from this eq have "qU0 = qU" by simp ultimately show ?thesis by (simp only: qs_def qs0 p1 p2) qed qed (simp_all add: qs_def a q qx, simp only: algebra_simps) qed qed definition valid_decomp :: "'x set \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ 'x set) list \ bool" where "valid_decomp X ps \ ((\(h, U)\set ps. h \ P[X] \ h \ 0 \ U \ X))" definition monomial_decomp :: "((('x \\<^sub>0 nat) \\<^sub>0 'a::{one,zero}) \ 'x set) list \ bool" where "monomial_decomp ps \ (\hU\set ps. is_monomial (fst hU) \ punit.lc (fst hU) = 1)" definition hom_decomp :: "((('x \\<^sub>0 nat) \\<^sub>0 'a::{one,zero}) \ 'x set) list \ bool" where "hom_decomp ps \ (\hU\set ps. homogeneous (fst hU))" definition cone_decomp :: "(('x \\<^sub>0 nat) \\<^sub>0 'a) set \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::comm_semiring_0) \ 'x set) list \ bool" where "cone_decomp T ps \ direct_decomp T (map cone ps)" lemma valid_decompI: "(\h U. (h, U) \ set ps \ h \ P[X]) \ (\h U. (h, U) \ set ps \ h \ 0) \ (\h U. (h, U) \ set ps \ U \ X) \ valid_decomp X ps" unfolding valid_decomp_def by blast lemma valid_decompD: assumes "valid_decomp X ps" and "(h, U) \ set ps" shows "h \ P[X]" and "h \ 0" and "U \ X" using assms unfolding valid_decomp_def by blast+ lemma valid_decompD_finite: assumes "finite X" and "valid_decomp X ps" and "(h, U) \ set ps" shows "finite U" proof - from assms(2, 3) have "U \ X" by (rule valid_decompD) thus ?thesis using assms(1) by (rule finite_subset) qed lemma valid_decomp_Nil: "valid_decomp X []" by (simp add: valid_decomp_def) lemma valid_decomp_concat: assumes "\ps. ps \ set pss \ valid_decomp X ps" shows "valid_decomp X (concat pss)" proof (rule valid_decompI) fix h U assume "(h, U) \ set (concat pss)" then obtain ps where "ps \ set pss" and "(h, U) \ set ps" unfolding set_concat .. from this(1) have "valid_decomp X ps" by (rule assms) thus "h \ P[X]" and "h \ 0" and "U \ X" using \(h, U) \ set ps\ by (rule valid_decompD)+ qed corollary valid_decomp_append: assumes "valid_decomp X ps" and "valid_decomp X qs" shows "valid_decomp X (ps @ qs)" proof - have "valid_decomp X (concat [ps, qs])" by (rule valid_decomp_concat) (auto simp: assms) thus ?thesis by simp qed lemma valid_decomp_map_times: assumes "valid_decomp X ps" and "s \ P[X]" and "s \ (0::_ \\<^sub>0 _::semiring_no_zero_divisors)" shows "valid_decomp X (map (apfst ((*) s)) ps)" proof (rule valid_decompI) fix h U assume "(h, U) \ set (map (apfst ((*) s)) ps)" then obtain x where "x \ set ps" and "(h, U) = apfst ((*) s) x" unfolding set_map .. moreover obtain a b where "x = (a, b)" using prod.exhaust by blast ultimately have h: "h = s * a" and "(a, U) \ set ps" by simp_all from assms(1) this(2) have "a \ P[X]" and "a \ 0" and "U \ X" by (rule valid_decompD)+ from assms(2) this(1) show "h \ P[X]" unfolding h by (rule Polys_closed_times) from assms(3) \a \ 0\ show "h \ 0" unfolding h by (rule times_not_zero) from \U \ X\ show "U \ X" . qed lemma monomial_decompI: "(\h U. (h, U) \ set ps \ is_monomial h) \ (\h U. (h, U) \ set ps \ punit.lc h = 1) \ monomial_decomp ps" by (auto simp: monomial_decomp_def) lemma monomial_decompD: assumes "monomial_decomp ps" and "(h, U) \ set ps" shows "is_monomial h" and "punit.lc h = 1" using assms by (auto simp: monomial_decomp_def) lemma monomial_decomp_append_iff: "monomial_decomp (ps @ qs) \ monomial_decomp ps \ monomial_decomp qs" by (auto simp: monomial_decomp_def) lemma monomial_decomp_concat: "(\ps. ps \ set pss \ monomial_decomp ps) \ monomial_decomp (concat pss)" by (induct pss) (auto simp: monomial_decomp_def) lemma monomial_decomp_map_times: assumes "monomial_decomp ps" and "is_monomial f" and "punit.lc f = (1::'a::semiring_1)" shows "monomial_decomp (map (apfst ((*) f)) ps)" proof (rule monomial_decompI) fix h U assume "(h, U) \ set (map (apfst ((*) f)) ps)" then obtain x where "x \ set ps" and "(h, U) = apfst ((*) f) x" unfolding set_map .. moreover obtain a b where "x = (a, b)" using prod.exhaust by blast ultimately have h: "h = f * a" and "(a, U) \ set ps" by simp_all from assms(1) this(2) have "is_monomial a" and "punit.lc a = 1" by (rule monomial_decompD)+ from this(1) have "monomial (punit.lc a) (lpp a) = a" by (rule punit.monomial_eq_itself) moreover define t where "t = lpp a" ultimately have a: "a = monomial 1 t" by (simp only: \punit.lc a = 1\) from assms(2) have "monomial (punit.lc f) (lpp f) = f" by (rule punit.monomial_eq_itself) moreover define s where "s = lpp f" ultimately have f: "f = monomial 1 s" by (simp only: assms(3)) show "is_monomial h" by (simp add: h a f times_monomial_monomial monomial_is_monomial) show "punit.lc h = 1" by (simp add: h a f times_monomial_monomial) qed lemma monomial_decomp_monomial_in_cone: assumes "monomial_decomp ps" and "hU \ set ps" and "a \ cone hU" shows "monomial (lookup a t) t \ cone hU" proof (cases "t \ keys a") case True obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast with assms(2) have "(h, U) \ set ps" by simp with assms(1) have "is_monomial h" by (rule monomial_decompD) then obtain c s where h: "h = monomial c s" by (rule is_monomial_monomial) from assms(3) obtain q where "q \ P[U]" and "a = q * h" unfolding hU by (rule coneE) from this(2) have "a = h * q" by (simp only: mult.commute) also have "\ = punit.monom_mult c s q" by (simp only: h times_monomial_left) finally have a: "a = punit.monom_mult c s q" . with True have "t \ keys (punit.monom_mult c s q)" by simp hence "t \ (+) s ` keys q" using punit.keys_monom_mult_subset[simplified] .. then obtain u where "u \ keys q" and t: "t = s + u" .. note this(1) also from \q \ P[U]\ have "keys q \ .[U]" by (rule PolysD) finally have "u \ .[U]" . have "monomial (lookup a t) t = monomial (lookup q u) u * h" by (simp add: a t punit.lookup_monom_mult h times_monomial_monomial mult.commute) moreover from \u \ .[U]\ have "monomial (lookup q u) u \ P[U]" by (rule Polys_closed_monomial) ultimately show ?thesis unfolding hU by (rule coneI) next case False thus ?thesis by (simp add: zero_in_cone in_keys_iff) qed lemma monomial_decomp_sum_list_monomial_in_cone: assumes "monomial_decomp ps" and "a \ sum_list ` listset (map cone ps)" and "t \ keys a" obtains c h U where "(h, U) \ set ps" and "c \ 0" and "monomial c t \ cone (h, U)" proof - from assms(2) obtain qs where qs_in: "qs \ listset (map cone ps)" and a: "a = sum_list qs" .. from assms(3) keys_sum_list_subset have "t \ Keys (set qs)" unfolding a .. then obtain q where "q \ set qs" and "t \ keys q" by (rule in_KeysE) from this(1) obtain i where "i < length qs" and q: "q = qs ! i" by (metis in_set_conv_nth) moreover from qs_in have "length qs = length (map cone ps)" by (rule listsetD) ultimately have "i < length (map cone ps)" by simp moreover from qs_in this have "qs ! i \ (map cone ps) ! i" by (rule listsetD) ultimately have "ps ! i \ set ps" and "q \ cone (ps ! i)" by (simp_all add: q) with assms(1) have *: "monomial (lookup q t) t \ cone (ps ! i)" by (rule monomial_decomp_monomial_in_cone) obtain h U where psi: "ps ! i = (h, U)" using prod.exhaust by blast show ?thesis proof from \ps ! i \ set ps\ show "(h, U) \ set ps" by (simp only: psi) next from \t \ keys q\ show "lookup q t \ 0" by (simp add: in_keys_iff) next from * show "monomial (lookup q t) t \ cone (h, U)" by (simp only: psi) qed qed lemma hom_decompI: "(\h U. (h, U) \ set ps \ homogeneous h) \ hom_decomp ps" by (auto simp: hom_decomp_def) lemma hom_decompD: "hom_decomp ps \ (h, U) \ set ps \ homogeneous h" by (auto simp: hom_decomp_def) lemma hom_decomp_append_iff: "hom_decomp (ps @ qs) \ hom_decomp ps \ hom_decomp qs" by (auto simp: hom_decomp_def) lemma hom_decomp_concat: "(\ps. ps \ set pss \ hom_decomp ps) \ hom_decomp (concat pss)" by (induct pss) (auto simp: hom_decomp_def) lemma hom_decomp_map_times: assumes "hom_decomp ps" and "homogeneous f" shows "hom_decomp (map (apfst ((*) f)) ps)" proof (rule hom_decompI) fix h U assume "(h, U) \ set (map (apfst ((*) f)) ps)" then obtain x where "x \ set ps" and "(h, U) = apfst ((*) f) x" unfolding set_map .. moreover obtain a b where "x = (a, b)" using prod.exhaust by blast ultimately have h: "h = f * a" and "(a, U) \ set ps" by simp_all from assms(1) this(2) have "homogeneous a" by (rule hom_decompD) with assms(2) show "homogeneous h" unfolding h by (rule homogeneous_times) qed lemma monomial_decomp_imp_hom_decomp: assumes "monomial_decomp ps" shows "hom_decomp ps" proof (rule hom_decompI) fix h U assume "(h, U) \ set ps" with assms have "is_monomial h" by (rule monomial_decompD) then obtain c t where h: "h = monomial c t" by (rule is_monomial_monomial) show "homogeneous h" unfolding h by (fact homogeneous_monomial) qed lemma cone_decompI: "direct_decomp T (map cone ps) \ cone_decomp T ps" unfolding cone_decomp_def by blast lemma cone_decompD: "cone_decomp T ps \ direct_decomp T (map cone ps)" unfolding cone_decomp_def by blast lemma cone_decomp_cone_subset: assumes "cone_decomp T ps" and "hU \ set ps" shows "cone hU \ T" proof fix p assume "p \ cone hU" from assms(2) obtain i where "i < length ps" and hU: "hU = ps ! i" by (metis in_set_conv_nth) define qs where "qs = (map 0 ps)[i := p]" have "sum_list qs \ T" proof (intro direct_decompD listsetI) from assms(1) show "direct_decomp T (map cone ps)" by (rule cone_decompD) next fix j assume "j < length (map cone ps)" with \i < length ps\ \p \ cone hU\ show "qs ! j \ map cone ps ! j" by (auto simp: qs_def nth_list_update zero_in_cone hU) qed (simp add: qs_def) also have "sum_list qs = qs ! i" by (rule sum_list_eq_nthI) (simp_all add: qs_def \i < length ps\) also from \i < length ps\ have "\ = p" by (simp add: qs_def) finally show "p \ T" . qed lemma cone_decomp_indets: assumes "cone_decomp T ps" and "T \ P[X]" and "(h, U) \ set ps" shows "h \ P[X]" and "h \ (0::_ \\<^sub>0 _::{comm_semiring_1,semiring_no_zero_divisors}) \ U \ X" proof - from assms(1, 3) have "cone (h, U) \ T" by (rule cone_decomp_cone_subset) hence "cone (h, U) \ P[X]" using assms(2) by (rule subset_trans) thus "h \ P[X]" and "h \ 0 \ U \ X" by (rule cone_subset_PolysD)+ qed lemma cone_decomp_closed_plus: assumes "cone_decomp T ps" and "a \ T" and "b \ T" shows "a + b \ T" proof - from assms(1) have dd: "direct_decomp T (map cone ps)" by (rule cone_decompD) then obtain qsa where qsa: "qsa \ listset (map cone ps)" and a: "a = sum_list qsa" using assms(2) by (rule direct_decompE) from dd assms(3) obtain qsb where qsb: "qsb \ listset (map cone ps)" and b: "b = sum_list qsb" by (rule direct_decompE) from qsa have "length qsa = length (map cone ps)" by (rule listsetD) moreover from qsb have "length qsb = length (map cone ps)" by (rule listsetD) ultimately have "a + b = sum_list (map2 (+) qsa qsb)" by (simp only: sum_list_map2_plus a b) also from dd have "sum_list (map2 (+) qsa qsb) \ T" proof (rule direct_decompD) from qsa qsb show "map2 (+) qsa qsb \ listset (map cone ps)" proof (rule listset_closed_map2) fix c p1 p2 assume "c \ set (map cone ps)" then obtain hU where c: "c = cone hU" by auto assume "p1 \ c" and "p2 \ c" thus "p1 + p2 \ c" unfolding c by (rule cone_closed_plus) qed qed finally show ?thesis . qed lemma cone_decomp_closed_uminus: assumes "cone_decomp T ps" and "(a::_ \\<^sub>0 _::comm_ring) \ T" shows "- a \ T" proof - from assms(1) have dd: "direct_decomp T (map cone ps)" by (rule cone_decompD) then obtain qsa where qsa: "qsa \ listset (map cone ps)" and a: "a = sum_list qsa" using assms(2) by (rule direct_decompE) from qsa have "length qsa = length (map cone ps)" by (rule listsetD) have "- a = sum_list (map uminus qsa)" unfolding a by (induct qsa, simp_all) also from dd have "\ \ T" proof (rule direct_decompD) from qsa show "map uminus qsa \ listset (map cone ps)" proof (rule listset_closed_map) fix c p assume "c \ set (map cone ps)" then obtain hU where c: "c = cone hU" by auto assume "p \ c" thus "- p \ c" unfolding c by (rule cone_closed_uminus) qed qed finally show ?thesis . qed corollary cone_decomp_closed_minus: assumes "cone_decomp T ps" and "(a::_ \\<^sub>0 _::comm_ring) \ T" and "b \ T" shows "a - b \ T" proof - from assms(1, 3) have "- b \ T" by (rule cone_decomp_closed_uminus) with assms(1, 2) have "a + (- b) \ T" by (rule cone_decomp_closed_plus) thus ?thesis by simp qed lemma cone_decomp_Nil: "cone_decomp {0} []" by (auto simp: cone_decomp_def intro: direct_decompI_alt) lemma cone_decomp_singleton: "cone_decomp (cone (t, U)) [(t, U)]" by (simp add: cone_decomp_def direct_decomp_singleton) lemma cone_decomp_append: assumes "direct_decomp T [S1, S2]" and "cone_decomp S1 ps" and "cone_decomp S2 qs" shows "cone_decomp T (ps @ qs)" proof (rule cone_decompI) from assms(2) have "direct_decomp S1 (map cone ps)" by (rule cone_decompD) with assms(1) have "direct_decomp T ([S2] @ map cone ps)" by (rule direct_decomp_direct_decomp) hence "direct_decomp T (S2 # map cone ps)" by simp moreover from assms(3) have "direct_decomp S2 (map cone qs)" by (rule cone_decompD) ultimately have "direct_decomp T (map cone ps @ map cone qs)" by (intro direct_decomp_direct_decomp) thus "direct_decomp T (map cone (ps @ qs))" by simp qed lemma cone_decomp_concat: assumes "direct_decomp T ss" and "length pss = length ss" and "\i. i < length ss \ cone_decomp (ss ! i) (pss ! i)" shows "cone_decomp T (concat pss)" using assms(2, 1, 3) proof (induct pss ss arbitrary: T rule: list_induct2) case Nil from Nil(1) show ?case by (simp add: cone_decomp_def) next case (Cons ps pss s ss) have "0 < length (s # ss)" by simp hence "cone_decomp ((s # ss) ! 0) ((ps # pss) ! 0)" by (rule Cons.prems) hence "cone_decomp s ps" by simp hence *: "direct_decomp s (map cone ps)" by (rule cone_decompD) with Cons.prems(1) have "direct_decomp T (ss @ map cone ps)" by (rule direct_decomp_direct_decomp) hence 1: "direct_decomp T [sum_list ` listset ss, sum_list ` listset (map cone ps)]" and 2: "direct_decomp (sum_list ` listset ss) ss" by (auto dest: direct_decomp_appendD intro!: empty_not_in_map_cone) note 1 moreover from 2 have "cone_decomp (sum_list ` listset ss) (concat pss)" proof (rule Cons.hyps) fix i assume "i < length ss" hence "Suc i < length (s # ss)" by simp hence "cone_decomp ((s # ss) ! Suc i) ((ps # pss) ! Suc i)" by (rule Cons.prems) thus "cone_decomp (ss ! i) (pss ! i)" by simp qed moreover have "cone_decomp (sum_list ` listset (map cone ps)) ps" proof (intro cone_decompI direct_decompI refl) from * show "inj_on sum_list (listset (map cone ps))" by (simp only: direct_decomp_def bij_betw_def) qed ultimately have "cone_decomp T (concat pss @ ps)" by (rule cone_decomp_append) hence "direct_decomp T (map cone (concat pss) @ map cone ps)" by (simp add: cone_decomp_def) hence "direct_decomp T (map cone ps @ map cone (concat pss))" - using perm_append_swap by (rule direct_decomp_perm) + by (auto intro: direct_decomp_perm) thus ?case by (simp add: cone_decomp_def) qed lemma cone_decomp_map_times: assumes "cone_decomp T ps" shows "cone_decomp ((*) s ` T) (map (apfst ((*) (s::_ \\<^sub>0 _::{comm_ring_1,ring_no_zero_divisors}))) ps)" proof (rule cone_decompI) from assms have "direct_decomp T (map cone ps)" by (rule cone_decompD) hence "direct_decomp ((*) s ` T) (map ((`) ((*) s)) (map cone ps))" by (rule direct_decomp_image_times) (rule times_canc_left) also have "map ((`) ((*) s)) (map cone ps) = map cone (map (apfst ((*) s)) ps)" by (simp add: cone_image_times') finally show "direct_decomp ((*) s ` T) (map cone (map (apfst ((*) s)) ps))" . qed lemma cone_decomp_perm: - assumes "cone_decomp T ps" and "perm ps qs" + assumes "cone_decomp T ps" and "mset ps = mset qs" shows "cone_decomp T qs" using assms(1) unfolding cone_decomp_def proof (rule direct_decomp_perm) - from \perm ps qs\ show \perm (map cone ps) (map cone qs)\ - by (simp add: perm_iff_eq_mset) + from \mset ps = mset qs\ show \mset (map cone ps) = mset (map cone qs)\ + by simp qed lemma valid_cone_decomp_subset_Polys: assumes "valid_decomp X ps" and "cone_decomp T ps" shows "T \ P[X]" proof fix p assume "p \ T" from assms(2) have "direct_decomp T (map cone ps)" by (rule cone_decompD) then obtain qs where "qs \ listset (map cone ps)" and p: "p = sum_list qs" using \p \ T\ by (rule direct_decompE) from assms(1) this(1) show "p \ P[X]" unfolding p proof (induct ps arbitrary: qs) case Nil from Nil(2) show ?case by (simp add: zero_in_Polys) next case (Cons a ps) obtain h U where a: "a = (h, U)" using prod.exhaust by blast hence "(h, U) \ set (a # ps)" by simp with Cons.prems(1) have "h \ P[X]" and "U \ X" by (rule valid_decompD)+ hence "cone a \ P[X]" unfolding a by (rule cone_subset_PolysI) from Cons.prems(1) have "valid_decomp X ps" by (simp add: valid_decomp_def) from Cons.prems(2) have "qs \ listset (cone a # map cone ps)" by simp then obtain q qs' where "q \ cone a" and qs': "qs' \ listset (map cone ps)" and qs: "qs = q # qs'" by (rule listset_ConsE) from this(1) \cone a \ P[X]\ have "q \ P[X]" .. moreover from \valid_decomp X ps\ qs' have "sum_list qs' \ P[X]" by (rule Cons.hyps) ultimately have "q + sum_list qs' \ P[X]" by (rule Polys_closed_plus) thus ?case by (simp add: qs) qed qed lemma homogeneous_set_cone_decomp: assumes "cone_decomp T ps" and "hom_decomp ps" shows "homogeneous_set T" proof (rule homogeneous_set_direct_decomp) from assms(1) show "direct_decomp T (map cone ps)" by (rule cone_decompD) next fix cn assume "cn \ set (map cone ps)" then obtain hU where "hU \ set ps" and cn: "cn = cone hU" unfolding set_map .. moreover obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast ultimately have "(h, U) \ set ps" by simp with assms(2) have "homogeneous h" by (rule hom_decompD) thus "homogeneous_set cn" unfolding cn hU by (rule homogeneous_set_coneI) qed lemma subspace_cone_decomp: assumes "cone_decomp T ps" shows "phull.subspace (T::(_ \\<^sub>0 _::field) set)" proof (rule phull.subspace_direct_decomp) from assms show "direct_decomp T (map cone ps)" by (rule cone_decompD) next fix cn assume "cn \ set (map cone ps)" then obtain hU where "hU \ set ps" and cn: "cn = cone hU" unfolding set_map .. show "phull.subspace cn" unfolding cn by (rule subspace_cone) qed definition pos_decomp :: "((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list" ("(_\<^sub>+)" [1000] 999) where "pos_decomp ps = filter (\p. snd p \ {}) ps" definition standard_decomp :: "nat \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ 'x set) list \ bool" where "standard_decomp k ps \ (\(h, U)\set (ps\<^sub>+). k \ poly_deg h \ (\d. k \ d \ d \ poly_deg h \ (\(h', U')\set ps. poly_deg h' = d \ card U \ card U')))" lemma pos_decomp_Nil [simp]: "[]\<^sub>+ = []" by (simp add: pos_decomp_def) lemma pos_decomp_subset: "set (ps\<^sub>+) \ set ps" by (simp add: pos_decomp_def) lemma pos_decomp_append: "(ps @ qs)\<^sub>+ = ps\<^sub>+ @ qs\<^sub>+" by (simp add: pos_decomp_def) lemma pos_decomp_concat: "(concat pss)\<^sub>+ = concat (map pos_decomp pss)" by (metis (mono_tags, lifting) filter_concat map_eq_conv pos_decomp_def) lemma pos_decomp_map: "(map (apfst f) ps)\<^sub>+ = map (apfst f) (ps\<^sub>+)" by (metis (mono_tags, lifting) pos_decomp_def filter_cong filter_map o_apply snd_apfst) lemma card_Diff_pos_decomp: "card {(h, U) \ set qs - set (qs\<^sub>+). P h} = card {h. (h, {}) \ set qs \ P h}" proof - have "{h. (h, {}) \ set qs \ P h} = fst ` {(h, U) \ set qs - set (qs\<^sub>+). P h}" by (auto simp: pos_decomp_def image_Collect) also have "card \ = card {(h, U) \ set qs - set (qs\<^sub>+). P h}" by (rule card_image, auto simp: pos_decomp_def intro: inj_onI) finally show ?thesis by (rule sym) qed lemma standard_decompI: assumes "\h U. (h, U) \ set (ps\<^sub>+) \ k \ poly_deg h" and "\h U d. (h, U) \ set (ps\<^sub>+) \ k \ d \ d \ poly_deg h \ (\h' U'. (h', U') \ set ps \ poly_deg h' = d \ card U \ card U')" shows "standard_decomp k ps" unfolding standard_decomp_def using assms by blast lemma standard_decompD: "standard_decomp k ps \ (h, U) \ set (ps\<^sub>+) \ k \ poly_deg h" unfolding standard_decomp_def by blast lemma standard_decompE: assumes "standard_decomp k ps" and "(h, U) \ set (ps\<^sub>+)" and "k \ d" and "d \ poly_deg h" obtains h' U' where "(h', U') \ set ps" and "poly_deg h' = d" and "card U \ card U'" using assms unfolding standard_decomp_def by blast lemma standard_decomp_Nil: "ps\<^sub>+ = [] \ standard_decomp k ps" by (simp add: standard_decomp_def) lemma standard_decomp_singleton: "standard_decomp (poly_deg h) [(h, U)]" by (simp add: standard_decomp_def pos_decomp_def) lemma standard_decomp_concat: assumes "\ps. ps \ set pss \ standard_decomp k ps" shows "standard_decomp k (concat pss)" proof (rule standard_decompI) fix h U assume "(h, U) \ set ((concat pss)\<^sub>+)" then obtain ps where "ps \ set pss" and *: "(h, U) \ set (ps\<^sub>+)" by (auto simp: pos_decomp_concat) from this(1) have "standard_decomp k ps" by (rule assms) thus "k \ poly_deg h" using * by (rule standard_decompD) fix d assume "k \ d" and "d \ poly_deg h" with \standard_decomp k ps\ * obtain h' U' where "(h', U') \ set ps" and "poly_deg h' = d" and "card U \ card U'" by (rule standard_decompE) note this(2, 3) moreover from \(h', U') \ set ps\ \ps \ set pss\ have "(h', U') \ set (concat pss)" by auto ultimately show "\h' U'. (h', U') \ set (concat pss) \ poly_deg h' = d \ card U \ card U'" by blast qed corollary standard_decomp_append: assumes "standard_decomp k ps" and "standard_decomp k qs" shows "standard_decomp k (ps @ qs)" proof - have "standard_decomp k (concat [ps, qs])" by (rule standard_decomp_concat) (auto simp: assms) thus ?thesis by simp qed lemma standard_decomp_map_times: assumes "standard_decomp k ps" and "valid_decomp X ps" and "s \ (0::_ \\<^sub>0 'a::semiring_no_zero_divisors)" shows "standard_decomp (k + poly_deg s) (map (apfst ((*) s)) ps)" proof (rule standard_decompI) fix h U assume "(h, U) \ set ((map (apfst ((*) s)) ps)\<^sub>+)" then obtain h0 where 1: "(h0, U) \ set (ps\<^sub>+)" and h: "h = s * h0" by (fastforce simp: pos_decomp_map) from this(1) pos_decomp_subset have "(h0, U) \ set ps" .. with assms(2) have "h0 \ 0" by (rule valid_decompD) with assms(3) have deg_h: "poly_deg h = poly_deg s + poly_deg h0" unfolding h by (rule poly_deg_times) moreover from assms(1) 1 have "k \ poly_deg h0" by (rule standard_decompD) ultimately show "k + poly_deg s \ poly_deg h" by simp fix d assume "k + poly_deg s \ d" and "d \ poly_deg h" hence "k \ d - poly_deg s" and "d - poly_deg s \ poly_deg h0" by (simp_all add: deg_h) with assms(1) 1 obtain h' U' where 2: "(h', U') \ set ps" and "poly_deg h' = d - poly_deg s" and "card U \ card U'" by (rule standard_decompE) from assms(2) this(1) have "h' \ 0" by (rule valid_decompD) with assms(3) have deg_h': "poly_deg (s * h') = poly_deg s + poly_deg h'" by (rule poly_deg_times) from 2 have "(s * h', U') \ set (map (apfst ((*) s)) ps)" by force moreover from \k + poly_deg s \ d\ \poly_deg h' = d - poly_deg s\ have "poly_deg (s * h') = d" by (simp add: deg_h') ultimately show "\h' U'. (h', U') \ set (map (apfst ((*) s)) ps) \ poly_deg h' = d \ card U \ card U'" using \card U \ card U'\ by fastforce qed lemma standard_decomp_nonempty_unique: assumes "finite X" and "valid_decomp X ps" and "standard_decomp k ps" and "ps\<^sub>+ \ []" shows "k = Min (poly_deg ` fst ` set (ps\<^sub>+))" proof - let ?A = "poly_deg ` fst ` set (ps\<^sub>+)" define m where "m = Min ?A" have "finite ?A" by simp moreover from assms(4) have "?A \ {}" by simp ultimately have "m \ ?A" unfolding m_def by (rule Min_in) then obtain h U where "(h, U) \ set (ps\<^sub>+)" and m: "m = poly_deg h" by fastforce have m_min: "m \ poly_deg h'" if "(h', U') \ set (ps\<^sub>+)" for h' U' proof - from that have "poly_deg (fst (h', U')) \ ?A" by (intro imageI) with \finite ?A\ have "m \ poly_deg (fst (h', U'))" unfolding m_def by (rule Min_le) thus ?thesis by simp qed show ?thesis proof (rule linorder_cases) assume "k < m" hence "k \ poly_deg h" by (simp add: m) with assms(3) \(h, U) \ set (ps\<^sub>+)\ le_refl obtain h' U' where "(h', U') \ set ps" and "poly_deg h' = k" and "card U \ card U'" by (rule standard_decompE) from this(2) \k < m\ have "\ m \ poly_deg h'" by simp with m_min have "(h', U') \ set (ps\<^sub>+)" by blast with \(h', U') \ set ps\ have "U' = {}" by (simp add: pos_decomp_def) with \card U \ card U'\ have "U = {} \ infinite U" by (simp add: card_eq_0_iff) thus ?thesis proof assume "U = {}" with \(h, U) \ set (ps\<^sub>+)\ show ?thesis by (simp add: pos_decomp_def) next assume "infinite U" moreover from assms(1, 2) have "finite U" proof (rule valid_decompD_finite) from \(h, U) \ set (ps\<^sub>+)\ show "(h, U) \ set ps" by (simp add: pos_decomp_def) qed ultimately show ?thesis .. qed next assume "m < k" hence "\ k \ m" by simp moreover from assms(3) \(h, U) \ set (ps\<^sub>+)\ have "k \ m" unfolding m by (rule standard_decompD) ultimately show ?thesis .. qed (simp only: m_def) qed lemma standard_decomp_SucE: assumes "finite X" and "U \ X" and "h \ P[X]" and "h \ (0::_ \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors})" obtains ps where "valid_decomp X ps" and "cone_decomp (cone (h, U)) ps" and "standard_decomp (Suc (poly_deg h)) ps" and "is_monomial h \ punit.lc h = 1 \ monomial_decomp ps" and "homogeneous h \ hom_decomp ps" proof - from assms(2, 1) have "finite U" by (rule finite_subset) thus ?thesis using assms(2) that proof (induct U arbitrary: thesis rule: finite_induct) case empty from assms(3, 4) have "valid_decomp X [(h, {})]" by (simp add: valid_decomp_def) moreover note cone_decomp_singleton moreover have "standard_decomp (Suc (poly_deg h)) [(h, {})]" by (rule standard_decomp_Nil) (simp add: pos_decomp_def) ultimately show ?case by (rule empty) (simp_all add: monomial_decomp_def hom_decomp_def) next case (insert x U) from insert.prems(1) have "x \ X" and "U \ X" by simp_all from this(2) obtain ps where 0: "valid_decomp X ps" and 1: "cone_decomp (cone (h, U)) ps" and 2: "standard_decomp (Suc (poly_deg h)) ps" and 3: "is_monomial h \ punit.lc h = 1 \ monomial_decomp ps" and 4: "homogeneous h \ hom_decomp ps" by (rule insert.hyps) blast let ?x = "monomial (1::'a) (Poly_Mapping.single x (Suc 0))" have "?x \ 0" by (simp add: monomial_0_iff) with assms(4) have deg: "poly_deg (?x * h) = Suc (poly_deg h)" by (simp add: poly_deg_times poly_deg_monomial deg_pm_single) define qs where "qs = [(?x * h, insert x U)]" show ?case proof (rule insert.prems) from \x \ X\ have "?x \ P[X]" by (intro Polys_closed_monomial PPs_closed_single) hence "?x * h \ P[X]" using assms(3) by (rule Polys_closed_times) moreover from \?x \ 0\ assms(4) have "?x * h \ 0" by (rule times_not_zero) ultimately have "valid_decomp X qs" using insert.hyps(1) \x \ X\ \U \ X\ by (simp add: qs_def valid_decomp_def) with 0 show "valid_decomp X (ps @ qs)" by (rule valid_decomp_append) next show "cone_decomp (cone (h, insert x U)) (ps @ qs)" proof (rule cone_decomp_append) show "direct_decomp (cone (h, insert x U)) [cone (h, U), cone (?x * h, insert x U)]" using insert.hyps(2) by (rule direct_decomp_cone_insert) next show "cone_decomp (cone (?x * h, insert x U)) qs" by (simp add: qs_def cone_decomp_singleton) qed (fact 1) next from standard_decomp_singleton[of "?x * h" "insert x U"] have "standard_decomp (Suc (poly_deg h)) qs" by (simp add: deg qs_def) with 2 show "standard_decomp (Suc (poly_deg h)) (ps @ qs)" by (rule standard_decomp_append) next assume "is_monomial h" and "punit.lc h = 1" hence "monomial_decomp ps" by (rule 3) moreover have "monomial_decomp qs" proof - have "is_monomial (?x * h)" by (metis \is_monomial h\ is_monomial_monomial monomial_is_monomial mult.commute mult.right_neutral mult_single) thus ?thesis by (simp add: monomial_decomp_def qs_def lc_times \punit.lc h = 1\) qed ultimately show "monomial_decomp (ps @ qs)" by (simp only: monomial_decomp_append_iff) next assume "homogeneous h" hence "hom_decomp ps" by (rule 4) moreover from \homogeneous h\ have "hom_decomp qs" by (simp add: hom_decomp_def qs_def homogeneous_times) ultimately show "hom_decomp (ps @ qs)" by (simp only: hom_decomp_append_iff) qed qed qed lemma standard_decomp_geE: assumes "finite X" and "valid_decomp X ps" and "cone_decomp (T::(('x \\<^sub>0 nat) \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}) set) ps" and "standard_decomp k ps" and "k \ d" obtains qs where "valid_decomp X qs" and "cone_decomp T qs" and "standard_decomp d qs" and "monomial_decomp ps \ monomial_decomp qs" and "hom_decomp ps \ hom_decomp qs" proof - have "\qs. valid_decomp X qs \ cone_decomp T qs \ standard_decomp (k + i) qs \ (monomial_decomp ps \ monomial_decomp qs) \ (hom_decomp ps \ hom_decomp qs)" for i proof (induct i) case 0 from assms(2, 3, 4) show ?case unfolding add_0_right by blast next case (Suc i) then obtain qs where 0: "valid_decomp X qs" and 1: "cone_decomp T qs" and 2: "standard_decomp (k + i) qs" and 3: "monomial_decomp ps \ monomial_decomp qs" and 4: "hom_decomp ps \ hom_decomp qs" by blast let ?P = "\hU. poly_deg (fst hU) \ k + i" define rs where "rs = filter (- ?P) qs" define ss where "ss = filter ?P qs" have "set rs \ set qs" by (auto simp: rs_def) have "set ss \ set qs" by (auto simp: ss_def) define f where "f = (\hU. SOME ps'. valid_decomp X ps' \ cone_decomp (cone hU) ps' \ standard_decomp (Suc (poly_deg ((fst hU)::('x \\<^sub>0 _) \\<^sub>0 'a))) ps' \ (monomial_decomp ps \ monomial_decomp ps') \ (hom_decomp ps \ hom_decomp ps'))" have "valid_decomp X (f hU) \ cone_decomp (cone hU) (f hU) \ standard_decomp (Suc (k + i)) (f hU) \ (monomial_decomp ps \ monomial_decomp (f hU)) \ (hom_decomp ps \ hom_decomp (f hU))" if "hU \ set rs" for hU proof - obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast with that have eq: "poly_deg (fst hU) = k + i" by (simp add: rs_def) from that \set rs \ set qs\ have "(h, U) \ set qs" unfolding hU .. with 0 have "U \ X" and "h \ P[X]" and "h \ 0" by (rule valid_decompD)+ with assms(1) obtain ps' where "valid_decomp X ps'" and "cone_decomp (cone (h, U)) ps'" and "standard_decomp (Suc (poly_deg h)) ps'" and md: "is_monomial h \ punit.lc h = 1 \ monomial_decomp ps'" and hd: "homogeneous h \ hom_decomp ps'" by (rule standard_decomp_SucE) blast note this(1-3) moreover have "monomial_decomp ps'" if "monomial_decomp ps" proof - from that have "monomial_decomp qs" by (rule 3) hence "is_monomial h" and "punit.lc h = 1" using \(h, U) \ set qs\ by (rule monomial_decompD)+ thus ?thesis by (rule md) qed moreover have "hom_decomp ps'" if "hom_decomp ps" proof - from that have "hom_decomp qs" by (rule 4) hence "homogeneous h" using \(h, U) \ set qs\ by (rule hom_decompD) thus ?thesis by (rule hd) qed ultimately have "valid_decomp X ps' \ cone_decomp (cone hU) ps' \ standard_decomp (Suc (poly_deg (fst hU))) ps' \ (monomial_decomp ps \ monomial_decomp ps') \ (hom_decomp ps \ hom_decomp ps')" by (simp add: hU) thus ?thesis unfolding f_def eq by (rule someI) qed hence f1: "\ps. ps \ set (map f rs) \ valid_decomp X ps" and f2: "\hU. hU \ set rs \ cone_decomp (cone hU) (f hU)" and f3: "\ps. ps \ set (map f rs) \ standard_decomp (Suc (k + i)) ps" and f4: "\ps'. monomial_decomp ps \ ps' \ set (map f rs) \ monomial_decomp ps'" and f5: "\ps'. hom_decomp ps \ ps' \ set (map f rs) \ hom_decomp ps'" by auto define rs' where "rs' = concat (map f rs)" show ?case unfolding add_Suc_right proof (intro exI conjI impI) have "valid_decomp X ss" proof (rule valid_decompI) fix h U assume "(h, U) \ set ss" hence "(h, U) \ set qs" using \set ss \ set qs\ .. with 0 show "h \ P[X]" and "h \ 0" and "U \ X" by (rule valid_decompD)+ qed moreover have "valid_decomp X rs'" unfolding rs'_def using f1 by (rule valid_decomp_concat) ultimately show "valid_decomp X (ss @ rs')" by (rule valid_decomp_append) next from 1 have "direct_decomp T (map cone qs)" by (rule cone_decompD) hence "direct_decomp T ((map cone ss) @ (map cone rs))" unfolding ss_def rs_def by (rule direct_decomp_split_map) hence ss: "cone_decomp (sum_list ` listset (map cone ss)) ss" and "cone_decomp (sum_list ` listset (map cone rs)) rs" and T: "direct_decomp T [sum_list ` listset (map cone ss), sum_list ` listset (map cone rs)]" by (auto simp: cone_decomp_def dest: direct_decomp_appendD intro!: empty_not_in_map_cone) from this(2) have "direct_decomp (sum_list ` listset (map cone rs)) (map cone rs)" by (rule cone_decompD) hence "cone_decomp (sum_list ` listset (map cone rs)) rs'" unfolding rs'_def proof (rule cone_decomp_concat) fix i assume *: "i < length (map cone rs)" hence "rs ! i \ set rs" by simp hence "cone_decomp (cone (rs ! i)) (f (rs ! i))" by (rule f2) with * show "cone_decomp (map cone rs ! i) (map f rs ! i)" by simp qed simp with T ss show "cone_decomp T (ss @ rs')" by (rule cone_decomp_append) next have "standard_decomp (Suc (k + i)) ss" proof (rule standard_decompI) fix h U assume "(h, U) \ set (ss\<^sub>+)" hence "(h, U) \ set (qs\<^sub>+)" and "poly_deg h \ k + i" by (simp_all add: pos_decomp_def ss_def) from 2 this(1) have "k + i \ poly_deg h" by (rule standard_decompD) with \poly_deg h \ k + i\ show "Suc (k + i) \ poly_deg h" by simp fix d' assume "Suc (k + i) \ d'" and "d' \ poly_deg h" from this(1) have "k + i \ d'" and "d' \ k + i" by simp_all from 2 \(h, U) \ set (qs\<^sub>+)\ this(1) obtain h' U' where "(h', U') \ set qs" and "poly_deg h' = d'" and "card U \ card U'" using \d' \ poly_deg h\ by (rule standard_decompE) moreover from \d' \ k + i\ this(1, 2) have "(h', U') \ set ss" by (simp add: ss_def) ultimately show "\h' U'. (h', U') \ set ss \ poly_deg h' = d' \ card U \ card U'" by blast qed moreover have "standard_decomp (Suc (k + i)) rs'" unfolding rs'_def using f3 by (rule standard_decomp_concat) ultimately show "standard_decomp (Suc (k + i)) (ss @ rs')" by (rule standard_decomp_append) next assume *: "monomial_decomp ps" hence "monomial_decomp qs" by (rule 3) hence "monomial_decomp ss" by (simp add: monomial_decomp_def ss_def) moreover have "monomial_decomp rs'" unfolding rs'_def using f4[OF *] by (rule monomial_decomp_concat) ultimately show "monomial_decomp (ss @ rs')" by (simp only: monomial_decomp_append_iff) next assume *: "hom_decomp ps" hence "hom_decomp qs" by (rule 4) hence "hom_decomp ss" by (simp add: hom_decomp_def ss_def) moreover have "hom_decomp rs'" unfolding rs'_def using f5[OF *] by (rule hom_decomp_concat) ultimately show "hom_decomp (ss @ rs')" by (simp only: hom_decomp_append_iff) qed qed then obtain qs where 1: "valid_decomp X qs" and 2: "cone_decomp T qs" and "standard_decomp (k + (d - k)) qs" and 4: "monomial_decomp ps \ monomial_decomp qs" and 5: "hom_decomp ps \ hom_decomp qs" by blast from this(3) assms(5) have "standard_decomp d qs" by simp with 1 2 show ?thesis using 4 5 .. qed subsection \Splitting w.r.t. Ideals\ context fixes X :: "'x set" begin definition splits_wrt :: "(((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list) \ (('x \\<^sub>0 nat) \\<^sub>0 'a::comm_ring_1) set \ (('x \\<^sub>0 nat) \\<^sub>0 'a) set \ bool" where "splits_wrt pqs T F \ cone_decomp T (fst pqs @ snd pqs) \ (\hU\set (fst pqs). cone hU \ ideal F \ P[X]) \ (\(h, U)\set (snd pqs). cone (h, U) \ P[X] \ cone (h, U) \ ideal F = {0})" lemma splits_wrtI: assumes "cone_decomp T (ps @ qs)" and "\h U. (h, U) \ set ps \ cone (h, U) \ P[X]" and "\h U. (h, U) \ set ps \ h \ ideal F" and "\h U. (h, U) \ set qs \ cone (h, U) \ P[X]" and "\h U a. (h, U) \ set qs \ a \ cone (h, U) \ a \ ideal F \ a = 0" shows "splits_wrt (ps, qs) T F" unfolding splits_wrt_def fst_conv snd_conv proof (intro conjI ballI) fix hU assume "hU \ set ps" moreover obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast ultimately have "(h, U) \ set ps" by simp hence "cone (h, U) \ P[X]" and "h \ ideal F" by (rule assms)+ from _ this(1) show "cone hU \ ideal F \ P[X]" unfolding hU proof (rule Int_greatest) show "cone (h, U) \ ideal F" proof fix a assume "a \ cone (h, U)" then obtain a' where "a' \ P[U]" and a: "a = a' * h" by (rule coneE) from \h \ ideal F\ show "a \ ideal F" unfolding a by (rule ideal.span_scale) qed qed next fix hU assume "hU \ set qs" moreover obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast ultimately have "(h, U) \ set qs" by simp hence "cone (h, U) \ P[X]" and "\a. a \ cone (h, U) \ a \ ideal F \ a = 0" by (rule assms)+ moreover have "0 \ cone (h, U) \ ideal F" by (simp add: zero_in_cone ideal.span_zero) ultimately show "case hU of (h, U) \ cone (h, U) \ P[X] \ cone (h, U) \ ideal F = {0}" by (fastforce simp: hU) qed (fact assms)+ lemma splits_wrtI_valid_decomp: assumes "valid_decomp X ps" and "valid_decomp X qs" and "cone_decomp T (ps @ qs)" and "\h U. (h, U) \ set ps \ h \ ideal F" and "\h U a. (h, U) \ set qs \ a \ cone (h, U) \ a \ ideal F \ a = 0" shows "splits_wrt (ps, qs) T F" using assms(3) _ _ _ assms(5) proof (rule splits_wrtI) fix h U assume "(h, U) \ set ps" thus "h \ ideal F" by (rule assms(4)) from assms(1) \(h, U) \ set ps\ have "h \ P[X]" and "U \ X" by (rule valid_decompD)+ thus "cone (h, U) \ P[X]" by (rule cone_subset_PolysI) next fix h U assume "(h, U) \ set qs" with assms(2) have "h \ P[X]" by (rule valid_decompD) moreover from assms(2) \(h, U) \ set qs\ have "U \ X" by (rule valid_decompD) ultimately show "cone (h, U) \ P[X]" by (rule cone_subset_PolysI) qed lemma splits_wrtD: assumes "splits_wrt (ps, qs) T F" shows "cone_decomp T (ps @ qs)" and "hU \ set ps \ cone hU \ ideal F \ P[X]" and "hU \ set qs \ cone hU \ P[X]" and "hU \ set qs \ cone hU \ ideal F = {0}" using assms by (fastforce simp: splits_wrt_def)+ lemma splits_wrt_image_sum_list_fst_subset: assumes "splits_wrt (ps, qs) T F" shows "sum_list ` listset (map cone ps) \ ideal F \ P[X]" proof fix x assume x_in: "x \ sum_list ` listset (map cone ps)" have "listset (map cone ps) \ listset (map (\_. ideal F \ P[X]) ps)" proof (rule listset_mono) fix i assume i: "i < length (map (\_. ideal F \ P[X]) ps)" hence "ps ! i \ set ps" by simp with assms(1) have "cone (ps ! i) \ ideal F \ P[X]" by (rule splits_wrtD) with i show "map cone ps ! i \ map (\_. ideal F \ P[X]) ps ! i" by simp qed simp hence "sum_list ` listset (map cone ps) \ sum_list ` listset (map (\_. ideal F \ P[X]) ps)" by (rule image_mono) with x_in have "x \ sum_list ` listset (map (\_. ideal F \ P[X]) ps)" .. then obtain xs where xs: "xs \ listset (map (\_. ideal F \ P[X]) ps)" and x: "x = sum_list xs" .. have 1: "y \ ideal F \ P[X]" if "y \ set xs" for y proof - from that obtain i where "i < length xs" and y: "y = xs ! i" by (metis in_set_conv_nth) moreover from xs have "length xs = length (map (\_. ideal F \ P[X]) ps)" by (rule listsetD) ultimately have "i < length (map (\_. ideal F \ P[X]) ps)" by simp moreover from xs this have "xs ! i \ (map (\_. ideal F \ P[X]) ps) ! i" by (rule listsetD) ultimately show "y \ ideal F \ P[X]" by (simp add: y) qed show "x \ ideal F \ P[X]" unfolding x proof show "sum_list xs \ ideal F" proof (rule ideal.span_closed_sum_list[simplified]) fix y assume "y \ set xs" hence "y \ ideal F \ P[X]" by (rule 1) thus "y \ ideal F" by simp qed next show "sum_list xs \ P[X]" proof (rule Polys_closed_sum_list) fix y assume "y \ set xs" hence "y \ ideal F \ P[X]" by (rule 1) thus "y \ P[X]" by simp qed qed qed lemma splits_wrt_image_sum_list_snd_subset: assumes "splits_wrt (ps, qs) T F" shows "sum_list ` listset (map cone qs) \ P[X]" proof fix x assume x_in: "x \ sum_list ` listset (map cone qs)" have "listset (map cone qs) \ listset (map (\_. P[X]) qs)" proof (rule listset_mono) fix i assume i: "i < length (map (\_. P[X]) qs)" hence "qs ! i \ set qs" by simp with assms(1) have "cone (qs ! i) \ P[X]" by (rule splits_wrtD) with i show "map cone qs ! i \ map (\_. P[X]) qs ! i" by simp qed simp hence "sum_list ` listset (map cone qs) \ sum_list ` listset (map (\_. P[X]) qs)" by (rule image_mono) with x_in have "x \ sum_list ` listset (map (\_. P[X]) qs)" .. then obtain xs where xs: "xs \ listset (map (\_. P[X]) qs)" and x: "x = sum_list xs" .. show "x \ P[X]" unfolding x proof (rule Polys_closed_sum_list) fix y assume "y \ set xs" then obtain i where "i < length xs" and y: "y = xs ! i" by (metis in_set_conv_nth) moreover from xs have "length xs = length (map (\_. P[X]::(_ \\<^sub>0 'a) set) qs)" by (rule listsetD) ultimately have "i < length (map (\_. P[X]) qs)" by simp moreover from xs this have "xs ! i \ (map (\_. P[X]) qs) ! i" by (rule listsetD) ultimately show "y \ P[X]" by (simp add: y) qed qed lemma splits_wrt_cone_decomp_1: assumes "splits_wrt (ps, qs) T F" and "monomial_decomp qs" and "is_monomial_set (F::(_ \\<^sub>0 'a::field) set)" \\The last two assumptions are missing in the paper.\ shows "cone_decomp (T \ ideal F) ps" proof - from assms(1) have *: "cone_decomp T (ps @ qs)" by (rule splits_wrtD) hence "direct_decomp T (map cone ps @ map cone qs)" by (simp add: cone_decomp_def) hence 1: "direct_decomp (sum_list ` listset (map cone ps)) (map cone ps)" and 2: "direct_decomp T [sum_list ` listset (map cone ps), sum_list ` listset (map cone qs)]" by (auto dest: direct_decomp_appendD intro!: empty_not_in_map_cone) let ?ss = "[sum_list ` listset (map cone ps), sum_list ` listset (map cone qs)]" show ?thesis proof (intro cone_decompI direct_decompI) from 1 show "inj_on sum_list (listset (map cone ps))" by (simp only: direct_decomp_def bij_betw_def) next from assms(1) have "sum_list ` listset (map cone ps) \ ideal F \ P[X]" by (rule splits_wrt_image_sum_list_fst_subset) hence sub: "sum_list ` listset (map cone ps) \ ideal F" by simp show "sum_list ` listset (map cone ps) = T \ ideal F" proof (rule set_eqI) fix x show "x \ sum_list ` listset (map cone ps) \ x \ T \ ideal F" proof assume x_in: "x \ sum_list ` listset (map cone ps)" show "x \ T \ ideal F" proof (intro IntI) have "map (\_. 0) qs \ listset (map cone qs)" (is "?ys \ _") by (induct qs) (auto intro: listset_ConsI zero_in_cone simp del: listset.simps(2)) hence "sum_list ?ys \ sum_list ` listset (map cone qs)" by (rule imageI) hence "0 \ sum_list ` listset (map cone qs)" by simp with x_in have "[x, 0] \ listset ?ss" using refl by (rule listset_doubletonI) with 2 have "sum_list [x, 0] \ T" by (rule direct_decompD) thus "x \ T" by simp next from x_in sub show "x \ ideal F" .. qed next assume "x \ T \ ideal F" hence "x \ T" and "x \ ideal F" by simp_all from 2 this(1) obtain xs where "xs \ listset ?ss" and x: "x = sum_list xs" by (rule direct_decompE) from this(1) obtain p q where p: "p \ sum_list ` listset (map cone ps)" and q: "q \ sum_list ` listset (map cone qs)" and xs: "xs = [p, q]" by (rule listset_doubletonE) from \x \ ideal F\ have "p + q \ ideal F" by (simp add: x xs) moreover from p sub have "p \ ideal F" .. ultimately have "p + q - p \ ideal F" by (rule ideal.span_diff) hence "q \ ideal F" by simp have "q = 0" proof (rule ccontr) assume "q \ 0" hence "keys q \ {}" by simp then obtain t where "t \ keys q" by blast with assms(2) q obtain c h U where "(h, U) \ set qs" and "c \ 0" and "monomial c t \ cone (h, U)" by (rule monomial_decomp_sum_list_monomial_in_cone) moreover from assms(3) \q \ ideal F\ \t \ keys q\ have "monomial c t \ ideal F" by (rule punit.monomial_pmdl_field[simplified]) ultimately have "monomial c t \ cone (h, U) \ ideal F" by simp also from assms(1) \(h, U) \ set qs\ have "\ = {0}" by (rule splits_wrtD) finally have "c = 0" by (simp add: monomial_0_iff) with \c \ 0\ show False .. qed with p show "x \ sum_list ` listset (map cone ps)" by (simp add: x xs) qed qed qed qed text \Together, Theorems \splits_wrt_image_sum_list_fst_subset\ and \splits_wrt_cone_decomp_1\ imply that @{term ps} is also a cone decomposition of @{term "T \ ideal F \ P[X]"}.\ lemma splits_wrt_cone_decomp_2: assumes "finite X" and "splits_wrt (ps, qs) T F" and "monomial_decomp qs" and "is_monomial_set F" and "F \ P[X]" shows "cone_decomp (T \ normal_form F ` P[X]) qs" proof - from assms(2) have *: "cone_decomp T (ps @ qs)" by (rule splits_wrtD) hence "direct_decomp T (map cone ps @ map cone qs)" by (simp add: cone_decomp_def) hence 1: "direct_decomp (sum_list ` listset (map cone qs)) (map cone qs)" and 2: "direct_decomp T [sum_list ` listset (map cone ps), sum_list ` listset (map cone qs)]" by (auto dest: direct_decomp_appendD intro!: empty_not_in_map_cone) let ?ss = "[sum_list ` listset (map cone ps), sum_list ` listset (map cone qs)]" let ?G = "punit.reduced_GB F" from assms(1, 5) have "?G \ P[X]" and G_is_GB: "punit.is_Groebner_basis ?G" and ideal_G: "ideal ?G = ideal F" by (rule reduced_GB_Polys, rule reduced_GB_is_GB_Polys, rule reduced_GB_ideal_Polys) show ?thesis proof (intro cone_decompI direct_decompI) from 1 show "inj_on sum_list (listset (map cone qs))" by (simp only: direct_decomp_def bij_betw_def) next from assms(2) have "sum_list ` listset (map cone ps) \ ideal F \ P[X]" by (rule splits_wrt_image_sum_list_fst_subset) hence sub: "sum_list ` listset (map cone ps) \ ideal F" by simp show "sum_list ` listset (map cone qs) = T \ normal_form F ` P[X]" proof (rule set_eqI) fix x show "x \ sum_list ` listset (map cone qs) \ x \ T \ normal_form F ` P[X]" proof assume x_in: "x \ sum_list ` listset (map cone qs)" show "x \ T \ normal_form F ` P[X]" proof (intro IntI) have "map (\_. 0) ps \ listset (map cone ps)" (is "?ys \ _") by (induct ps) (auto intro: listset_ConsI zero_in_cone simp del: listset.simps(2)) hence "sum_list ?ys \ sum_list ` listset (map cone ps)" by (rule imageI) hence "0 \ sum_list ` listset (map cone ps)" by simp from this x_in have "[0, x] \ listset ?ss" using refl by (rule listset_doubletonI) with 2 have "sum_list [0, x] \ T" by (rule direct_decompD) thus "x \ T" by simp next from assms(2) have "sum_list ` listset (map cone qs) \ P[X]" by (rule splits_wrt_image_sum_list_snd_subset) with x_in have "x \ P[X]" .. moreover have "\ punit.is_red ?G x" proof assume "punit.is_red ?G x" then obtain g t where "g \ ?G" and "t \ keys x" and "g \ 0" and adds: "lpp g adds t" by (rule punit.is_red_addsE[simplified]) from assms(3) x_in this(2) obtain c h U where "(h, U) \ set qs" and "c \ 0" and "monomial c t \ cone (h, U)" by (rule monomial_decomp_sum_list_monomial_in_cone) note this(3) moreover have "monomial c t \ ideal ?G" proof (rule punit.is_red_monomial_monomial_set_in_pmdl[simplified]) from \c \ 0\ show "is_monomial (monomial c t)" by (rule monomial_is_monomial) next from assms(1, 5, 4) show "is_monomial_set ?G" by (rule reduced_GB_is_monomial_set_Polys) next from \c \ 0\ have "t \ keys (monomial c t)" by simp with \g \ ?G\ \g \ 0\ show "punit.is_red ?G (monomial c t)" using adds by (rule punit.is_red_addsI[simplified]) qed ultimately have "monomial c t \ cone (h, U) \ ideal F" by (simp add: ideal_G) also from assms(2) \(h, U) \ set qs\ have "\ = {0}" by (rule splits_wrtD) finally have "c = 0" by (simp add: monomial_0_iff) with \c \ 0\ show False .. qed ultimately show "x \ normal_form F ` P[X]" using assms(1, 5) by (simp add: image_normal_form_iff) qed next assume "x \ T \ normal_form F ` P[X]" hence "x \ T" and "x \ normal_form F ` P[X]" by simp_all from this(2) assms(1, 5) have "x \ P[X]" and irred: "\ punit.is_red ?G x" by (simp_all add: image_normal_form_iff) from 2 \x \ T\ obtain xs where "xs \ listset ?ss" and x: "x = sum_list xs" by (rule direct_decompE) from this(1) obtain p q where p: "p \ sum_list ` listset (map cone ps)" and q: "q \ sum_list ` listset (map cone qs)" and xs: "xs = [p, q]" by (rule listset_doubletonE) have "x = p + q" by (simp add: x xs) from p sub have "p \ ideal F" .. have "p = 0" proof (rule ccontr) assume "p \ 0" hence "keys p \ {}" by simp then obtain t where "t \ keys p" by blast from assms(4) \p \ ideal F\ \t \ keys p\ have 3: "monomial c t \ ideal F" for c by (rule punit.monomial_pmdl_field[simplified]) have "t \ keys q" proof assume "t \ keys q" with assms(3) q obtain c h U where "(h, U) \ set qs" and "c \ 0" and "monomial c t \ cone (h, U)" by (rule monomial_decomp_sum_list_monomial_in_cone) from this(3) 3 have "monomial c t \ cone (h, U) \ ideal F" by simp also from assms(2) \(h, U) \ set qs\ have "\ = {0}" by (rule splits_wrtD) finally have "c = 0" by (simp add: monomial_0_iff) with \c \ 0\ show False .. qed with \t \ keys p\ have "t \ keys x" unfolding \x = p + q\ by (rule in_keys_plusI1) have "punit.is_red ?G x" proof - note G_is_GB moreover from 3 have "monomial 1 t \ ideal ?G" by (simp only: ideal_G) moreover have "monomial (1::'a) t \ 0" by (simp add: monomial_0_iff) ultimately obtain g where "g \ ?G" and "g \ 0" and "lpp g adds lpp (monomial (1::'a) t)" by (rule punit.GB_adds_lt[simplified]) from this(3) have "lpp g adds t" by (simp add: punit.lt_monomial) with \g \ ?G\ \g \ 0\ \t \ keys x\ show ?thesis by (rule punit.is_red_addsI[simplified]) qed with irred show False .. qed with q show "x \ sum_list ` listset (map cone qs)" by (simp add: x xs) qed qed qed qed lemma quot_monomial_ideal_monomial: "ideal (monomial 1 ` S) \
monomial 1 (Poly_Mapping.single (x::'x) (1::nat)) = ideal (monomial (1::'a::comm_ring_1) ` (\s. s - Poly_Mapping.single x 1) ` S)" proof (rule set_eqI) let ?x = "Poly_Mapping.single x (1::nat)" fix a have "a \ ideal (monomial 1 ` S) \
monomial 1 ?x \ punit.monom_mult 1 ?x a \ ideal (monomial (1::'a) ` S)" by (simp only: quot_set_iff times_monomial_left) also have "\ \ a \ ideal (monomial 1 ` (\s. s - ?x) ` S)" proof (induct a rule: poly_mapping_plus_induct) case 1 show ?case by (simp add: ideal.span_zero) next case (2 a c t) let ?S = "monomial (1::'a) ` (\s. s - ?x) ` S" show ?case proof assume 0: "punit.monom_mult 1 ?x (monomial c t + a) \ ideal (monomial 1 ` S)" have "is_monomial_set (monomial (1::'a) ` S)" by (auto intro!: is_monomial_setI monomial_is_monomial) moreover from 0 have 1: "monomial c (?x + t) + punit.monom_mult 1 ?x a \ ideal (monomial 1 ` S)" by (simp add: punit.monom_mult_monomial punit.monom_mult_dist_right) moreover have "?x + t \ keys (monomial c (?x + t) + punit.monom_mult 1 ?x a)" proof (intro in_keys_plusI1 notI) from 2(1) show "?x + t \ keys (monomial c (?x + t))" by simp next assume "?x + t \ keys (punit.monom_mult 1 ?x a)" also have "\ \ (+) ?x ` keys a" by (rule punit.keys_monom_mult_subset[simplified]) finally obtain s where "s \ keys a" and "?x + t = ?x + s" .. from this(2) have "t = s" by simp with \s \ keys a\ 2(2) show False by simp qed ultimately obtain f where "f \ monomial (1::'a) ` S" and adds: "lpp f adds ?x + t" by (rule punit.keys_monomial_pmdl[simplified]) from this(1) obtain s where "s \ S" and f: "f = monomial 1 s" .. from adds have "s adds ?x + t" by (simp add: f punit.lt_monomial) hence "s - ?x adds t" by (metis (no_types, lifting) add_minus_2 adds_minus adds_triv_right plus_minus_assoc_pm_nat_1) then obtain s' where t: "t = (s - ?x) + s'" by (rule addsE) from \s \ S\ have "monomial 1 (s - ?x) \ ?S" by (intro imageI) also have "\ \ ideal ?S" by (rule ideal.span_superset) finally have "monomial c s' * monomial 1 (s - ?x) \ ideal ?S" by (rule ideal.span_scale) hence "monomial c t \ ideal ?S" by (simp add: times_monomial_monomial t add.commute) moreover have "a \ ideal ?S" proof - from \f \ monomial 1 ` S\ have "f \ ideal (monomial 1 ` S)" by (rule ideal.span_base) hence "punit.monom_mult c (?x + t - s) f \ ideal (monomial 1 ` S)" by (rule punit.pmdl_closed_monom_mult[simplified]) with \s adds ?x + t\ have "monomial c (?x + t) \ ideal (monomial 1 ` S)" by (simp add: f punit.monom_mult_monomial adds_minus) with 1 have "monomial c (?x + t) + punit.monom_mult 1 ?x a - monomial c (?x + t) \ ideal (monomial 1 ` S)" by (rule ideal.span_diff) thus ?thesis by (simp add: 2(3) del: One_nat_def) qed ultimately show "monomial c t + a \ ideal ?S" by (rule ideal.span_add) next have "is_monomial_set ?S" by (auto intro!: is_monomial_setI monomial_is_monomial) moreover assume 1: "monomial c t + a \ ideal ?S" moreover from _ 2(2) have "t \ keys (monomial c t + a)" proof (rule in_keys_plusI1) from 2(1) show "t \ keys (monomial c t)" by simp qed ultimately obtain f where "f \ ?S" and adds: "lpp f adds t" by (rule punit.keys_monomial_pmdl[simplified]) from this(1) obtain s where "s \ S" and f: "f = monomial 1 (s - ?x)" by blast from adds have "s - ?x adds t" by (simp add: f punit.lt_monomial) hence "s adds ?x + t" by (auto simp: adds_poly_mapping le_fun_def lookup_add lookup_minus lookup_single when_def split: if_splits) then obtain s' where t: "?x + t = s + s'" by (rule addsE) from \s \ S\ have "monomial 1 s \ monomial 1 ` S" by (rule imageI) also have "\ \ ideal (monomial 1 ` S)" by (rule ideal.span_superset) finally have "monomial c s' * monomial 1 s \ ideal (monomial 1 ` S)" by (rule ideal.span_scale) hence "monomial c (?x + t) \ ideal (monomial 1 ` S)" by (simp only: t) (simp add: times_monomial_monomial add.commute) moreover have "punit.monom_mult 1 ?x a \ ideal (monomial 1 ` S)" proof - from \f \ ?S\ have "f \ ideal ?S" by (rule ideal.span_base) hence "punit.monom_mult c (t - (s - ?x)) f \ ideal ?S" by (rule punit.pmdl_closed_monom_mult[simplified]) with \s - ?x adds t\ have "monomial c t \ ideal ?S" by (simp add: f punit.monom_mult_monomial adds_minus) with 1 have "monomial c t + a - monomial c t \ ideal ?S" by (rule ideal.span_diff) thus ?thesis by (simp add: 2(3) del: One_nat_def) qed ultimately have "monomial c (?x + t) + punit.monom_mult 1 ?x a \ ideal (monomial 1 ` S)" by (rule ideal.span_add) thus "punit.monom_mult 1 ?x (monomial c t + a) \ ideal (monomial 1 ` S)" by (simp add: punit.monom_mult_monomial punit.monom_mult_dist_right) qed qed finally show "a \ ideal (monomial 1 ` S) \
monomial 1 ?x \ a \ ideal (monomial 1 ` (\s. s - ?x) ` S)" . qed lemma lem_4_2_1: assumes "ideal F \
monomial 1 t = ideal (monomial (1::'a::comm_ring_1) ` S)" shows "cone (monomial 1 t, U) \ ideal F \ 0 \ S" proof have "monomial 1 t \ cone (monomial (1::'a) t, U)" by (rule tip_in_cone) also assume "cone (monomial 1 t, U) \ ideal F" finally have *: "monomial 1 t * 1 \ ideal F" by simp have "is_monomial_set (monomial (1::'a) ` S)" by (auto intro!: is_monomial_setI monomial_is_monomial) moreover from * have "1 \ ideal (monomial (1::'a) ` S)" by (simp only: quot_set_iff flip: assms) moreover have "0 \ keys (1::_ \\<^sub>0 'a)" by simp ultimately obtain g where "g \ monomial (1::'a) ` S" and adds: "lpp g adds 0" by (rule punit.keys_monomial_pmdl[simplified]) from this(1) obtain s where "s \ S" and g: "g = monomial 1 s" .. from adds have "s adds 0" by (simp add: g punit.lt_monomial flip: single_one) with \s \ S\ show "0 \ S" by (simp only: adds_zero) next assume "0 \ S" hence "monomial 1 0 \ monomial (1::'a) ` S" by (rule imageI) hence "1 \ ideal (monomial (1::'a) ` S)" unfolding single_one by (rule ideal.span_base) hence eq: "ideal F \
monomial 1 t = UNIV" (is "_ \
?t = _") by (simp only: assms ideal_eq_UNIV_iff_contains_one) show "cone (monomial 1 t, U) \ ideal F" proof fix a assume "a \ cone (?t, U)" then obtain q where a: "a = q * ?t" by (rule coneE) have "q \ ideal F \
?t" by (simp add: eq) thus "a \ ideal F" by (simp only: a quot_set_iff mult.commute) qed qed lemma lem_4_2_2: assumes "ideal F \
monomial 1 t = ideal (monomial (1::'a::comm_ring_1) ` S)" shows "cone (monomial 1 t, U) \ ideal F = {0} \ S \ .[U] = {}" proof let ?t = "monomial (1::'a) t" assume eq: "cone (?t, U) \ ideal F = {0}" { fix s assume "s \ S" hence "monomial 1 s \ monomial (1::'a) ` S" (is "?s \ _") by (rule imageI) hence "?s \ ideal (monomial 1 ` S)" by (rule ideal.span_base) also have "\ = ideal F \
?t" by (simp only: assms) finally have *: "?s * ?t \ ideal F" by (simp only: quot_set_iff mult.commute) assume "s \ .[U]" hence "?s \ P[U]" by (rule Polys_closed_monomial) with refl have "?s * ?t \ cone (?t, U)" by (rule coneI) with * have "?s * ?t \ cone (?t, U) \ ideal F" by simp hence False by (simp add: eq times_monomial_monomial monomial_0_iff) } thus "S \ .[U] = {}" by blast next let ?t = "monomial (1::'a) t" assume eq: "S \ .[U] = {}" { fix a assume "a \ cone (?t, U)" then obtain q where "q \ P[U]" and a: "a = q * ?t" by (rule coneE) assume "a \ ideal F" have "a = 0" proof (rule ccontr) assume "a \ 0" hence "q \ 0" by (auto simp: a) from \a \ ideal F\ have *: "q \ ideal F \
?t" by (simp only: quot_set_iff a mult.commute) have "is_monomial_set (monomial (1::'a) ` S)" by (auto intro!: is_monomial_setI monomial_is_monomial) moreover from * have q_in: "q \ ideal (monomial 1 ` S)" by (simp only: assms) moreover from \q \ 0\ have "lpp q \ keys q" by (rule punit.lt_in_keys) ultimately obtain g where "g \ monomial (1::'a) ` S" and adds: "lpp g adds lpp q" by (rule punit.keys_monomial_pmdl[simplified]) from this(1) obtain s where "s \ S" and g: "g = monomial 1 s" .. from \q \ 0\ have "lpp q \ keys q" by (rule punit.lt_in_keys) also from \q \ P[U]\ have "\ \ .[U]" by (rule PolysD) finally have "lpp q \ .[U]" . moreover from adds have "s adds lpp q" by (simp add: g punit.lt_monomial) ultimately have "s \ .[U]" by (rule PPs_closed_adds) with eq \s \ S\ show False by blast qed } thus "cone (?t, U) \ ideal F = {0}" using zero_in_cone ideal.span_zero by blast qed subsection \Function \split\\ definition max_subset :: "'a set \ ('a set \ bool) \ 'a set" where "max_subset A P = (ARG_MAX card B. B \ A \ P B)" lemma max_subset: assumes "finite A" and "B \ A" and "P B" shows "max_subset A P \ A" (is ?thesis1) and "P (max_subset A P)" (is ?thesis2) and "card B \ card (max_subset A P)" (is ?thesis3) proof - from assms(2, 3) have "B \ A \ P B" by simp moreover have "\C. C \ A \ P C \ card C < Suc (card A)" proof (intro allI impI, elim conjE) fix C assume "C \ A" with assms(1) have "card C \ card A" by (rule card_mono) thus "card C < Suc (card A)" by simp qed ultimately have "?thesis1 \ ?thesis2" and ?thesis3 unfolding max_subset_def by (rule arg_max_natI, rule arg_max_nat_le) thus ?thesis1 and ?thesis2 and ?thesis3 by simp_all qed function (domintros) split :: "('x \\<^sub>0 nat) \ 'x set \ ('x \\<^sub>0 nat) set \ ((((('x \\<^sub>0 nat) \\<^sub>0 'a) \ ('x set)) list) \ (((('x \\<^sub>0 nat) \\<^sub>0 'a::{zero,one}) \ ('x set)) list))" where "split t U S = (if 0 \ S then ([(monomial 1 t, U)], []) else if S \ .[U] = {} then ([], [(monomial 1 t, U)]) else let x = SOME x'. x' \ U - (max_subset U (\V. S \ .[V] = {})); (ps0, qs0) = split t (U - {x}) S; (ps1, qs1) = split (Poly_Mapping.single x 1 + t) U ((\f. f - Poly_Mapping.single x 1) ` S) in (ps0 @ ps1, qs0 @ qs1))" by auto text \Function @{const split} is not executable, because this is not necessary. With some effort, it could be made executable, though.\ lemma split_domI': assumes "finite X" and "fst (snd args) \ X" and "finite (snd (snd args))" shows "split_dom TYPE('a::{zero,one}) args" proof - let ?m = "\args'. card (fst (snd args')) + sum deg_pm (snd (snd args'))" from wf_measure[of ?m] assms(2, 3) show ?thesis proof (induct args) case (less args) obtain t U F where args: "args = (t, U, F)" using prod.exhaust by metis from less.prems have "U \ X" and "finite F" by (simp_all only: args fst_conv snd_conv) from this(1) assms(1) have "finite U" by (rule finite_subset) have IH: "split_dom TYPE('a) (t', U', F')" if "U' \ X" and "finite F'" and "card U' + sum deg_pm F' < card U + sum deg_pm F" for t' U' F' using less.hyps that by (simp add: args) define S where "S = max_subset U (\V. F \ .[V] = {})" define x where "x = (SOME x'. x' \ U \ x' \ S)" show ?case unfolding args proof (rule split.domintros, simp_all only: x_def[symmetric] S_def[symmetric]) fix f assume "0 \ F" and "f \ F" and "f \ .[U]" from this(1) have "F \ .[{}] = {}" by simp with \finite U\ empty_subsetI have "S \ U" and "F \ .[S] = {}" unfolding S_def by (rule max_subset)+ have "x \ U \ x \ S" unfolding x_def proof (rule someI_ex) from \f \ F\ \f \ .[U]\ \F \ .[S] = {}\ have "S \ U" by blast with \S \ U\ show "\y. y \ U \ y \ S" by blast qed hence "x \ U" and "x \ S" by simp_all { assume "\ split_dom TYPE('a) (t, U - {x}, F)" moreover from _ \finite F\ have "split_dom TYPE('a) (t, U - {x}, F)" proof (rule IH) from \U \ X\ show "U - {x} \ X" by blast next from \finite U\ \x \ U\ have "card (U - {x}) < card U" by (rule card_Diff1_less) thus "card (U - {x}) + sum deg_pm F < card U + sum deg_pm F" by simp qed ultimately show False .. } { let ?args = "(Poly_Mapping.single x (Suc 0) + t, U, (\f. f - Poly_Mapping.single x (Suc 0)) ` F)" assume "\ split_dom TYPE('a) ?args" moreover from \U \ X\ have "split_dom TYPE('a) ?args" proof (rule IH) from \finite F\ show "finite ((\f. f - Poly_Mapping.single x (Suc 0)) ` F)" by (rule finite_imageI) next have "sum deg_pm ((\f. f - Poly_Mapping.single x (Suc 0)) ` F) \ sum (deg_pm \ (\f. f - Poly_Mapping.single x (Suc 0))) F" using \finite F\ by (rule sum_image_le) simp also from \finite F\ have "\ < sum deg_pm F" proof (rule sum_strict_mono_ex1) show "\f\F. (deg_pm \ (\f. f - Poly_Mapping.single x (Suc 0))) f \ deg_pm f" by (simp add: deg_pm_minus_le) next show "\f\F. (deg_pm \ (\f. f - Poly_Mapping.single x (Suc 0))) f < deg_pm f" proof (rule ccontr) assume *: "\ (\f\F. (deg_pm \ (\f. f - Poly_Mapping.single x (Suc 0))) f < deg_pm f)" note \finite U\ moreover from \x \ U\ \S \ U\ have "insert x S \ U" by (rule insert_subsetI) moreover have "F \ .[insert x S] = {}" proof - { fix s assume "s \ F" with * have "\ deg_pm (s - Poly_Mapping.single x (Suc 0)) < deg_pm s" by simp with deg_pm_minus_le[of s "Poly_Mapping.single x (Suc 0)"] have "deg_pm (s - Poly_Mapping.single x (Suc 0)) = deg_pm s" by simp hence "keys s \ keys (Poly_Mapping.single x (Suc 0)) = {}" by (simp only: deg_pm_minus_id_iff) hence "x \ keys s" by simp moreover assume "s \ .[insert x S]" ultimately have "s \ .[S]" by (fastforce simp: PPs_def) with \s \ F\ \F \ .[S] = {}\ have False by blast } thus ?thesis by blast qed ultimately have "card (insert x S) \ card S" unfolding S_def by (rule max_subset) moreover from \S \ U\ \finite U\ have "finite S" by (rule finite_subset) ultimately show False using \x \ S\ by simp qed qed finally show "card U + sum deg_pm ((\f. f - monomial (Suc 0) x) ` F) < card U + sum deg_pm F" by simp qed ultimately show False .. } qed qed qed corollary split_domI: "finite X \ U \ X \ finite S \ split_dom TYPE('a::{zero,one}) (t, U, S)" using split_domI'[of "(t, U, S)"] by simp lemma split_empty: assumes "finite X" and "U \ X" shows "split t U {} = ([], [(monomial (1::'a::{zero,one}) t, U)])" proof - have "finite {}" .. with assms have "split_dom TYPE('a) (t, U, {})" by (rule split_domI) thus ?thesis by (simp add: split.psimps) qed lemma split_induct [consumes 3, case_names base1 base2 step]: fixes P :: "('x \\<^sub>0 nat) \ _" assumes "finite X" and "U \ X" and "finite S" assumes "\t U S. U \ X \ finite S \ 0 \ S \ P t U S ([(monomial (1::'a::{zero,one}) t, U)], [])" assumes "\t U S. U \ X \ finite S \ 0 \ S \ S \ .[U] = {} \ P t U S ([], [(monomial 1 t, U)])" assumes "\t U S V x ps0 ps1 qs0 qs1. U \ X \ finite S \ 0 \ S \ S \ .[U] \ {} \ V \ U \ S \ .[V] = {} \ (\V'. V' \ U \ S \ .[V'] = {} \ card V' \ card V) \ x \ U \ x \ V \ V = max_subset U (\V'. S \ .[V'] = {}) \ x = (SOME x'. x' \ U - V) \ (ps0, qs0) = split t (U - {x}) S \ (ps1, qs1) = split (Poly_Mapping.single x 1 + t) U ((\f. f - Poly_Mapping.single x 1) ` S) \ split t U S = (ps0 @ ps1, qs0 @ qs1) \ P t (U - {x}) S (ps0, qs0) \ P (Poly_Mapping.single x 1 + t) U ((\f. f - Poly_Mapping.single x 1) ` S) (ps1, qs1) \ P t U S (ps0 @ ps1, qs0 @ qs1)" shows "P t U S (split t U S)" proof - from assms(1-3) have "split_dom TYPE('a) (t, U, S)" by (rule split_domI) thus ?thesis using assms(2,3) proof (induct t U S rule: split.pinduct) case step: (1 t U F) from step(4) assms(1) have "finite U" by (rule finite_subset) define S where "S = max_subset U (\V. F \ .[V] = {})" define x where "x = (SOME x'. x' \ U \ x' \ S)" show ?case proof (simp add: split.psimps[OF step(1)] S_def[symmetric] x_def[symmetric] split: prod.split, intro allI conjI impI) assume "0 \ F" with step(4, 5) show "P t U F ([(monomial 1 t, U)], [])" by (rule assms(4)) thus "P t U F ([(monomial 1 t, U)], [])" . next assume "0 \ F" and "F \ .[U] = {}" with step(4, 5) show "P t U F ([], [(monomial 1 t, U)])" by (rule assms(5)) next fix ps0 qs0 ps1 qs1 :: "((_ \\<^sub>0 'a) \ _) list" assume "split (Poly_Mapping.single x (Suc 0) + t) U ((\f. f - Poly_Mapping.single x (Suc 0)) ` F) = (ps1, qs1)" hence PQ1[symmetric]: "split (Poly_Mapping.single x 1 + t) U ((\f. f - Poly_Mapping.single x 1) ` F) = (ps1, qs1)" by simp assume PQ0[symmetric]: "split t (U - {x}) F = (ps0, qs0)" assume "F \ .[U] \ {}" and "0 \ F" from this(2) have "F \ .[{}] = {}" by simp with \finite U\ empty_subsetI have "S \ U" and "F \ .[S] = {}" unfolding S_def by (rule max_subset)+ have S_max: "card S' \ card S" if "S' \ U" and "F \ .[S'] = {}" for S' using \finite U\ that unfolding S_def by (rule max_subset) have "x \ U \ x \ S" unfolding x_def proof (rule someI_ex) from \F \ .[U] \ {}\ \F \ .[S] = {}\ have "S \ U" by blast with \S \ U\ show "\y. y \ U \ y \ S" by blast qed hence "x \ U" and "x \ S" by simp_all from step(4, 5) \0 \ F\ \F \ .[U] \ {}\ \S \ U\ \F \ .[S] = {}\ S_max \x \ U\ \x \ S\ S_def _ PQ0 PQ1 show "P t U F (ps0 @ ps1, qs0 @ qs1)" proof (rule assms(6)) show "P t (U - {x}) F (ps0, qs0)" unfolding PQ0 using \0 \ F\ \F \ .[U] \ {}\ _ _ step(5) proof (rule step(2)) from \U \ X\ show "U - {x} \ X" by fastforce qed (simp add: x_def S_def) next show "P (Poly_Mapping.single x 1 + t) U ((\f. f - Poly_Mapping.single x 1) ` F) (ps1, qs1)" unfolding PQ1 using \0 \ F\ \F \ .[U] \ {}\ _ refl PQ0 \U \ X\ proof (rule step(3)) from \finite F\ show "finite ((\f. f - Poly_Mapping.single x 1) ` F)" by (rule finite_imageI) qed (simp add: x_def S_def) next show "split t U F = (ps0 @ ps1, qs0 @ qs1)" using \0 \ F\ \F \ .[U] \ {}\ by (simp add: split.psimps[OF step(1)] Let_def flip: S_def x_def PQ0 PQ1 del: One_nat_def) qed (assumption+, simp add: x_def S_def) qed qed qed lemma valid_decomp_split: assumes "finite X" and "U \ X" and "finite S" and "t \ .[X]" shows "valid_decomp X (fst ((split t U S)::(_ \ (((_ \\<^sub>0 'a::zero_neq_one) \ _) list))))" and "valid_decomp X (snd ((split t U S)::(_ \ (((_ \\<^sub>0 'a::zero_neq_one) \ _) list))))" (is "valid_decomp _ (snd ?s)") proof - from assms have "valid_decomp X (fst ?s) \ valid_decomp X (snd ?s)" proof (induct t U S rule: split_induct) case (base1 t U S) from base1(1, 4) show ?case by (simp add: valid_decomp_def monomial_0_iff Polys_closed_monomial) next case (base2 t U S) from base2(1, 5) show ?case by (simp add: valid_decomp_def monomial_0_iff Polys_closed_monomial) next case (step t U S V x ps0 ps1 qs0 qs1) from step.hyps(8, 1) have "x \ X" .. hence "Poly_Mapping.single x 1 \ .[X]" by (rule PPs_closed_single) hence "Poly_Mapping.single x 1 + t \ .[X]" using step.prems by (rule PPs_closed_plus) with step.hyps(15, 16) step.prems show ?case by (simp add: valid_decomp_append) qed thus "valid_decomp X (fst ?s)" and "valid_decomp X (snd ?s)" by simp_all qed lemma monomial_decomp_split: assumes "finite X" and "U \ X" and "finite S" shows "monomial_decomp (fst ((split t U S)::(_ \ (((_ \\<^sub>0 'a::zero_neq_one) \ _) list))))" and "monomial_decomp (snd ((split t U S)::(_ \ (((_ \\<^sub>0 'a::zero_neq_one) \ _) list))))" (is "monomial_decomp (snd ?s)") proof - from assms have "monomial_decomp (fst ?s) \ monomial_decomp (snd ?s)" proof (induct t U S rule: split_induct) case (base1 t U S) from base1(1) show ?case by (simp add: monomial_decomp_def monomial_is_monomial) next case (base2 t U S) from base2(1) show ?case by (simp add: monomial_decomp_def monomial_is_monomial) next case (step t U S V x ps0 ps1 qs0 qs1) from step.hyps(15, 16) show ?case by (auto simp: monomial_decomp_def) qed thus "monomial_decomp (fst ?s)" and "monomial_decomp (snd ?s)" by simp_all qed lemma split_splits_wrt: assumes "finite X" and "U \ X" and "finite S" and "t \ .[X]" and "ideal F \
monomial 1 t = ideal (monomial 1 ` S)" shows "splits_wrt (split t U S) (cone (monomial (1::'a::{comm_ring_1,ring_no_zero_divisors}) t, U)) F" using assms proof (induct t U S rule: split_induct) case (base1 t U S) from base1(3) have "cone (monomial 1 t, U) \ ideal F" by (simp only: lem_4_2_1 base1(5)) show ?case proof (rule splits_wrtI) fix h0 U0 assume "(h0, U0) \ set [(monomial (1::'a) t, U)]" hence h0: "h0 = monomial 1 t" and "U0 = U" by simp_all note this(1) also have "monomial 1 t \ cone (monomial (1::'a) t, U)" by (fact tip_in_cone) also have "\ \ ideal F" by fact finally show "h0 \ ideal F" . from base1(4) have "h0 \ P[X]" unfolding h0 by (rule Polys_closed_monomial) moreover from base1(1) have "U0 \ X" by (simp only: \U0 = U\) ultimately show "cone (h0, U0) \ P[X]" by (rule cone_subset_PolysI) qed (simp_all add: cone_decomp_singleton \U \ X\) next case (base2 t U S) from base2(4) have "cone (monomial 1 t, U) \ ideal F = {0}" by (simp only: lem_4_2_2 base2(6)) show ?case proof (rule splits_wrtI) fix h0 U0 assume "(h0, U0) \ set [(monomial (1::'a) t, U)]" hence h0: "h0 = monomial 1 t" and "U0 = U" by simp_all note this(1) also from base2(5) have "monomial 1 t \ P[X]" by (rule Polys_closed_monomial) finally have "h0 \ P[X]" . moreover from base2(1) have "U0 \ X" by (simp only: \U0 = U\) ultimately show "cone (h0, U0) \ P[X]" by (rule cone_subset_PolysI) next fix h0 U0 a assume "(h0, U0) \ set [(monomial (1::'a) t, U)]" and "a \ cone (h0, U0)" hence "a \ cone (monomial 1 t, U)" by simp moreover assume "a \ ideal F" ultimately have "a \ cone (monomial 1 t, U) \ ideal F" by (rule IntI) also have "\ = {0}" by fact finally show "a = 0" by simp qed (simp_all add: cone_decomp_singleton \U \ X\) next case (step t U S V x ps0 ps1 qs0 qs1) let ?x = "Poly_Mapping.single x 1" from step.prems have 0: "splits_wrt (ps0, qs0) (cone (monomial 1 t, U - {x})) F" by (rule step.hyps) have 1: "splits_wrt (ps1, qs1) (cone (monomial 1 (?x + t), U)) F" proof (rule step.hyps) from step.hyps(8, 1) have "x \ X" .. hence "?x \ .[X]" by (rule PPs_closed_single) thus "?x + t \ .[X]" using step.prems(1) by (rule PPs_closed_plus) next have "ideal F \
monomial 1 (?x + t) = ideal F \
monomial 1 t \
monomial 1 ?x" by (simp add: times_monomial_monomial add.commute) also have "\ = ideal (monomial 1 ` S) \
monomial 1 ?x" by (simp only: step.prems) finally show "ideal F \
monomial 1 (?x + t) = ideal (monomial 1 ` (\s. s - ?x) ` S)" by (simp only: quot_monomial_ideal_monomial) qed show ?case proof (rule splits_wrtI) from step.hyps(8) have U: "insert x U = U" by blast have "direct_decomp (cone (monomial (1::'a) t, insert x (U - {x}))) [cone (monomial 1 t, U - {x}), cone (monomial 1 (monomial (Suc 0) x) * monomial 1 t, insert x (U - {x}))]" by (rule direct_decomp_cone_insert) simp hence "direct_decomp (cone (monomial (1::'a) t, U)) [cone (monomial 1 t, U - {x}), cone (monomial 1 (?x + t), U)]" by (simp add: U times_monomial_monomial) moreover from 0 have "cone_decomp (cone (monomial 1 t, U - {x})) (ps0 @ qs0)" by (rule splits_wrtD) moreover from 1 have "cone_decomp (cone (monomial 1 (?x + t), U)) (ps1 @ qs1)" by (rule splits_wrtD) ultimately have "cone_decomp (cone (monomial 1 t, U)) ((ps0 @ qs0) @ (ps1 @ qs1))" by (rule cone_decomp_append) thus "cone_decomp (cone (monomial 1 t, U)) ((ps0 @ ps1) @ qs0 @ qs1)" - by (rule cone_decomp_perm) (metis append.assoc perm_append1 perm_append2 perm_append_swap) + by (rule cone_decomp_perm) simp next fix h0 U0 assume "(h0, U0) \ set (ps0 @ ps1)" hence "(h0, U0) \ set ps0 \ set ps1" by simp hence "cone (h0, U0) \ ideal F \ P[X]" proof assume "(h0, U0) \ set ps0" with 0 show ?thesis by (rule splits_wrtD) next assume "(h0, U0) \ set ps1" with 1 show ?thesis by (rule splits_wrtD) qed hence *: "cone (h0, U0) \ ideal F" and "cone (h0, U0) \ P[X]" by simp_all from this(2) show "cone (h0, U0) \ P[X]" . from tip_in_cone * show "h0 \ ideal F" .. next fix h0 U0 assume "(h0, U0) \ set (qs0 @ qs1)" hence "(h0, U0) \ set qs0 \ set qs1" by simp thus "cone (h0, U0) \ P[X]" proof assume "(h0, U0) \ set qs0" with 0 show ?thesis by (rule splits_wrtD) next assume "(h0, U0) \ set qs1" with 1 show ?thesis by (rule splits_wrtD) qed from \(h0, U0) \ set qs0 \ set qs1\ have "cone (h0, U0) \ ideal F = {0}" proof assume "(h0, U0) \ set qs0" with 0 show ?thesis by (rule splits_wrtD) next assume "(h0, U0) \ set qs1" with 1 show ?thesis by (rule splits_wrtD) qed thus "\a. a \ cone (h0, U0) \ a \ ideal F \ a = 0" by blast qed qed lemma lem_4_5: assumes "finite X" and "U \ X" and "t \ .[X]" and "F \ P[X]" and "ideal F \
monomial 1 t = ideal (monomial (1::'a) ` S)" and "cone (monomial (1::'a::field) t', V) \ cone (monomial 1 t, U) \ normal_form F ` P[X]" shows "V \ U" and "S \ .[V] = {}" proof - let ?t = "monomial (1::'a) t" let ?t' = "monomial (1::'a) t'" from assms(6) have 1: "cone (?t', V) \ cone (?t, U)" and 2: "cone (?t', V) \ normal_form F ` P[X]" by blast+ from this(1) show "V \ U" by (rule cone_subsetD) (simp add: monomial_0_iff) show "S \ .[V] = {}" proof let ?t = "monomial (1::'a) t" let ?t' = "monomial (1::'a) t'" show "S \ .[V] \ {}" proof fix s assume "s \ S \ .[V]" hence "s \ S" and "s \ .[V]" by simp_all from this(2) have "monomial (1::'a) s \ P[V]" (is "?s \ _") by (rule Polys_closed_monomial) with refl have "?s * ?t \ cone (?t, V)" by (rule coneI) from tip_in_cone 1 have "?t' \ cone (?t, U)" .. then obtain s' where "s' \ P[U]" and t': "?t' = s' * ?t" by (rule coneE) note this(1) also from assms(2) have "P[U] \ P[X]" by (rule Polys_mono) finally have "s' \ P[X]" . have "s' * ?s * ?t = ?s * ?t'" by (simp add: t') also from refl \?s \ P[V]\ have "\ \ cone (?t', V)" by (rule coneI) finally have "s' * ?s * ?t \ cone (?t', V)" . hence 1: "s' * ?s * ?t \ normal_form F ` P[X]" using 2 .. from \s \ S\ have "?s \ monomial 1 ` S" by (rule imageI) hence "?s \ ideal (monomial 1 ` S)" by (rule ideal.span_base) hence "s' * ?s \ ideal (monomial 1 ` S)" by (rule ideal.span_scale) hence "s' * ?s \ ideal F \
?t" by (simp only: assms(5)) hence "s' * ?s * ?t \ ideal F" by (simp only: quot_set_iff mult.commute) hence "s' * ?s * ?t \ ideal F \ normal_form F ` P[X]" using 1 by (rule IntI) also from assms(1, 4) have "\ \ {0}" by (auto simp: normal_form_normal_form simp flip: normal_form_zero_iff) finally have "?s * ?t' = 0" by (simp add: t' ac_simps) thus "s \ {}" by (simp add: times_monomial_monomial monomial_0_iff) qed qed (fact empty_subsetI) qed lemma lem_4_6: assumes "finite X" and "U \ X" and "finite S" and "t \ .[X]" and "F \ P[X]" and "ideal F \
monomial 1 t = ideal (monomial 1 ` S)" assumes "cone (monomial 1 t', V) \ cone (monomial 1 t, U) \ normal_form F ` P[X]" obtains V' where "(monomial 1 t, V') \ set (snd (split t U S))" and "card V \ card V'" proof - let ?t = "monomial (1::'a) t" let ?t' = "monomial (1::'a) t'" from assms(7) have "cone (?t', V) \ cone (?t, U)" and "cone (?t', V) \ normal_form F ` P[X]" by blast+ from assms(1, 2, 4, 5, 6, 7) have "V \ U" and "S \ .[V] = {}" by (rule lem_4_5)+ with assms(1, 2, 3) show ?thesis using that proof (induct t U S arbitrary: V thesis rule: split_induct) case (base1 t U S) from base1.hyps(3) have "0 \ S \ .[V]" using zero_in_PPs by (rule IntI) thus ?case by (simp add: base1.prems(2)) next case (base2 t U S) show ?case proof (rule base2.prems) from base2.hyps(1) assms(1) have "finite U" by (rule finite_subset) thus "card V \ card U" using base2.prems(1) by (rule card_mono) qed simp next case (step t U S V0 x ps0 ps1 qs0 qs1) from step.prems(1, 2) have 0: "card V \ card V0" by (rule step.hyps) from step.hyps(5, 9) have "V0 \ U - {x}" by blast then obtain V' where 1: "(monomial 1 t, V') \ set (snd (ps0, qs0))" and 2: "card V0 \ card V'" using step.hyps(6) by (rule step.hyps) show ?case proof (rule step.prems) from 1 show "(monomial 1 t, V') \ set (snd (ps0 @ ps1, qs0 @ qs1))" by simp next from 0 2 show "card V \ card V'" by (rule le_trans) qed qed qed lemma lem_4_7: assumes "finite X" and "S \ .[X]" and "g \ punit.reduced_GB (monomial (1::'a) ` S)" and "cone_decomp (P[X] \ ideal (monomial (1::'a::field) ` S)) ps" and "monomial_decomp ps" obtains U where "(g, U) \ set ps" proof - let ?S = "monomial (1::'a) ` S" let ?G = "punit.reduced_GB ?S" note assms(1) moreover from assms(2) have "?S \ P[X]" by (auto intro: Polys_closed_monomial) moreover have "is_monomial_set ?S" by (auto intro!: is_monomial_setI monomial_is_monomial) ultimately have "is_monomial_set ?G" by (rule reduced_GB_is_monomial_set_Polys) hence "is_monomial g" using assms(3) by (rule is_monomial_setD) hence "g \ 0" by (rule monomial_not_0) moreover from assms(1) \?S \ P[X]\ have "punit.is_monic_set ?G" by (rule reduced_GB_is_monic_set_Polys) ultimately have "punit.lc g = 1" using assms(3) by (simp add: punit.is_monic_set_def) moreover define t where "t = lpp g" moreover from \is_monomial g\ have "monomial (punit.lc g) (lpp g) = g" by (rule punit.monomial_eq_itself) ultimately have g: "g = monomial 1 t" by simp hence "t \ keys g" by simp from assms(3) have "g \ ideal ?G" by (rule ideal.span_base) also from assms(1) \?S \ P[X]\ have ideal_G: "\ = ideal ?S" by (rule reduced_GB_ideal_Polys) finally have "g \ ideal ?S" . moreover from assms(3) have "g \ P[X]" by rule (intro reduced_GB_Polys assms(1) \?S \ P[X]\) ultimately have "g \ P[X] \ ideal ?S" by simp with assms(4) have "g \ sum_list ` listset (map cone ps)" by (simp only: cone_decomp_def direct_decompD) with assms(5) obtain d h U where *: "(h, U) \ set ps" and "d \ 0" and "monomial d t \ cone (h, U)" using \t \ keys g\ by (rule monomial_decomp_sum_list_monomial_in_cone) from this(3) zero_in_PPs have "punit.monom_mult (1 / d) 0 (monomial d t) \ cone (h, U)" by (rule cone_closed_monom_mult) with \d \ 0\ have "g \ cone (h, U)" by (simp add: g punit.monom_mult_monomial) then obtain q where "q \ P[U]" and g': "g = q * h" by (rule coneE) from \g \ 0\ have "q \ 0" and "h \ 0" by (auto simp: g') hence lt_g': "lpp g = lpp q + lpp h" unfolding g' by (rule lp_times) hence adds1: "lpp h adds t" by (simp add: t_def) from assms(5) * have "is_monomial h" and "punit.lc h = 1" by (rule monomial_decompD)+ moreover from this(1) have "monomial (punit.lc h) (lpp h) = h" by (rule punit.monomial_eq_itself) moreover define s where "s = lpp h" ultimately have h: "h = monomial 1 s" by simp have "punit.lc q = punit.lc g" by (simp add: g' lc_times \punit.lc h = 1\) hence "punit.lc q = 1" by (simp only: \punit.lc g = 1\) note tip_in_cone also from assms(4) * have "cone (h, U) \ P[X] \ ideal ?S" by (rule cone_decomp_cone_subset) also have "\ \ ideal ?G" by (simp add: ideal_G) finally have "h \ ideal ?G" . from assms(1) \?S \ P[X]\ have "punit.is_Groebner_basis ?G" by (rule reduced_GB_is_GB_Polys) then obtain g' where "g' \ ?G" and "g' \ 0" and adds2: "lpp g' adds lpp h" using \h \ ideal ?G\ \h \ 0\ by (rule punit.GB_adds_lt[simplified]) from this(3) adds1 have "lpp g' adds t" by (rule adds_trans) with _ \g' \ 0\ \t \ keys g\ have "punit.is_red {g'} g" by (rule punit.is_red_addsI[simplified]) simp have "g' = g" proof (rule ccontr) assume "g' \ g" with \g' \ ?G\ have "{g'} \ ?G - {g}" by simp with \punit.is_red {g'} g\ have red: "punit.is_red (?G - {g}) g" by (rule punit.is_red_subset) from assms(1) \?S \ P[X]\ have "punit.is_auto_reduced ?G" by (rule reduced_GB_is_auto_reduced_Polys) hence "\ punit.is_red (?G - {g}) g" using assms(3) by (rule punit.is_auto_reducedD) thus False using red .. qed with adds2 have "t adds lpp h" by (simp only: t_def) with adds1 have "lpp h = t" by (rule adds_antisym) hence "lpp q = 0" using lt_g' by (simp add: t_def) hence "monomial (punit.lc q) 0 = q" by (rule punit.lt_eq_min_term_monomial[simplified]) hence "g = h" by (simp add: \punit.lc q = 1\ g') with * have "(g, U) \ set ps" by simp thus ?thesis .. qed lemma snd_splitI: assumes "finite X" and "U \ X" and "finite S" and "0 \ S" obtains V where "V \ U" and "(monomial 1 t, V) \ set (snd (split t U S))" using assms proof (induct t U S arbitrary: thesis rule: split_induct) case (base1 t U S) from base1.prems(2) base1.hyps(3) show ?case .. next case (base2 t U S) from subset_refl show ?case by (rule base2.prems) simp next case (step t U S V0 x ps0 ps1 qs0 qs1) from step.hyps(3) obtain V where 1: "V \ U - {x}" and 2: "(monomial 1 t, V) \ set (snd (ps0, qs0))" using step.hyps(15) by blast show ?case proof (rule step.prems) from 1 show "V \ U" by blast next from 2 show "(monomial 1 t, V) \ set (snd (ps0 @ ps1, qs0 @ qs1))" by fastforce qed qed lemma fst_splitE: assumes "finite X" and "U \ X" and "finite S" and "0 \ S" and "(monomial (1::'a) s, V) \ set (fst (split t U S))" obtains t' x where "t' \ .[X]" and "x \ X" and "V \ U" and "0 \ (\s. s - t') ` S" and "s = t' + t + Poly_Mapping.single x 1" and "(monomial (1::'a::zero_neq_one) s, V) \ set (fst (split (t' + t) V ((\s. s - t') ` S)))" and "set (snd (split (t' + t) V ((\s. s - t') ` S))) \ (set (snd (split t U S)) :: ((_ \\<^sub>0 'a) \ _) set)" using assms proof (induct t U S arbitrary: thesis rule: split_induct) case (base1 t U S) from base1.prems(2) base1.hyps(3) show ?case .. next case (base2 t U S) from base2.prems(3) show ?case by simp next case (step t U S V0 x ps0 ps1 qs0 qs1) from step.prems(3) have "(monomial 1 s, V) \ set ps0 \ set ps1" by simp thus ?case proof assume "(monomial 1 s, V) \ set ps0" hence "(monomial (1::'a) s, V) \ set (fst (ps0, qs0))" by (simp only: fst_conv) with step.hyps(3) obtain t' x' where "t' \ .[X]" and "x' \ X" and "V \ U - {x}" and "0 \ (\s. s - t') ` S" and "s = t' + t + Poly_Mapping.single x' 1" and "(monomial (1::'a) s, V) \ set (fst (split (t' + t) V ((\s. s - t') ` S)))" and "set (snd (split (t' + t) V ((\s. s - t') ` S))) \ set (snd (ps0, qs0))" using step.hyps(15) by blast note this(7) also have "set (snd (ps0, qs0)) \ set (snd (ps0 @ ps1, qs0 @ qs1))" by simp finally have "set (snd (split (t' + t) V ((\s. s - t') ` S))) \ set (snd (ps0 @ ps1, qs0 @ qs1))" . from \V \ U - {x}\ have "V \ U" by blast show ?thesis by (rule step.prems) fact+ next assume "(monomial 1 s, V) \ set ps1" show ?thesis proof (cases "0 \ (\f. f - Poly_Mapping.single x 1) ` S") case True from step.hyps(2) have fin: "finite ((\f. f - Poly_Mapping.single x 1) ` S)" by (rule finite_imageI) have "split (Poly_Mapping.single x 1 + t) U ((\f. f - Poly_Mapping.single x 1) ` S) = ([(monomial (1::'a) (Poly_Mapping.single x 1 + t), U)], [])" by (simp only: split.psimps[OF split_domI, OF assms(1) step.hyps(1) fin] True if_True) hence "ps1 = [(monomial 1 (Poly_Mapping.single x 1 + t), U)]" by (simp only: step.hyps(13)[symmetric] prod.inject) with \(monomial 1 s, V) \ set ps1\ have s: "s = Poly_Mapping.single x 1 + t" and "V = U" by (auto dest!: monomial_inj) show ?thesis proof (rule step.prems) show "0 \ .[X]" by (fact zero_in_PPs) next from step.hyps(8, 1) show "x \ X" .. next show "V \ U" by (simp add: \V = U\) next from step.hyps(3) show "0 \ (\s. s - 0) ` S" by simp next show "s = 0 + t + Poly_Mapping.single x 1" by (simp add: s add.commute) next show "(monomial (1::'a) s, V) \ set (fst (split (0 + t) V ((\s. s - 0) ` S)))" using \(monomial 1 s, V) \ set ps1\ by (simp add: step.hyps(14) \V = U\) next show "set (snd (split (0 + t) V ((\s. s - 0) ` S))) \ set (snd (ps0 @ ps1, qs0 @ qs1))" by (simp add: step.hyps(14) \V = U\) qed next case False moreover from \(monomial 1 s, V) \ set ps1\ have "(monomial 1 s, V) \ set (fst (ps1, qs1))" by (simp only: fst_conv) ultimately obtain t' x' where "t' \ .[X]" and "x' \ X" and "V \ U" and 1: "0 \ (\s. s - t') ` (\f. f - Poly_Mapping.single x 1) ` S" and s: "s = t' + (Poly_Mapping.single x 1 + t) + Poly_Mapping.single x' 1" and 2: "(monomial (1::'a) s, V) \ set (fst (split (t' + (Poly_Mapping.single x 1 + t)) V ((\s. s - t') ` (\f. f - Poly_Mapping.single x 1) ` S)))" and 3: "set (snd (split (t' + (Poly_Mapping.single x 1 + t)) V ((\s. s - t') ` (\f. f - monomial 1 x) ` S))) \ set (snd (ps1, qs1))" using step.hyps(16) by blast have eq: "(\s. s - t') ` (\f. f - Poly_Mapping.single x 1) ` S = (\s. s - (t' + Poly_Mapping.single x 1)) ` S" by (simp add: image_image add.commute diff_diff_eq) show ?thesis proof (rule step.prems) from step.hyps(8, 1) have "x \ X" .. hence "Poly_Mapping.single x 1 \ .[X]" by (rule PPs_closed_single) with \t' \ .[X]\ show "t' + Poly_Mapping.single x 1 \ .[X]" by (rule PPs_closed_plus) next from 1 show "0 \ (\s. s - (t' + Poly_Mapping.single x 1)) ` S" by (simp only: eq not_False_eq_True) next show "s = t' + Poly_Mapping.single x 1 + t + Poly_Mapping.single x' 1" by (simp only: s ac_simps) next show "(monomial (1::'a) s, V) \ set (fst (split (t' + Poly_Mapping.single x 1 + t) V ((\s. s - (t' + Poly_Mapping.single x 1)) ` S)))" using 2 by (simp only: eq add.assoc) next have "set (snd (split (t' + Poly_Mapping.single x 1 + t) V ((\s. s - (t' + Poly_Mapping.single x 1)) ` S))) \ set (snd (ps1, qs1))" (is "?x \ _") using 3 by (simp only: eq add.assoc) also have "\ \ set (snd (ps0 @ ps1, qs0 @ qs1))" by simp finally show "?x \ set (snd (ps0 @ ps1, qs0 @ qs1))" . qed fact+ qed qed qed lemma lem_4_8: assumes "finite X" and "finite S" and "S \ .[X]" and "0 \ S" and "g \ punit.reduced_GB (monomial (1::'a) ` S)" obtains t U where "U \ X" and "(monomial (1::'a::field) t, U) \ set (snd (split 0 X S))" and "poly_deg g = Suc (deg_pm t)" proof - let ?S = "monomial (1::'a) ` S" let ?G = "punit.reduced_GB ?S" have md1: "monomial_decomp (fst ((split 0 X S)::(_ \ (((_ \\<^sub>0 'a) \ _) list))))" and md2: "monomial_decomp (snd ((split 0 X S)::(_ \ (((_ \\<^sub>0 'a) \ _) list))))" using assms(1) subset_refl assms(2) by (rule monomial_decomp_split)+ from assms(3) have 0: "?S \ P[X]" by (auto intro: Polys_closed_monomial) with assms(1) have "punit.is_auto_reduced ?G" and "punit.is_monic_set ?G" and ideal_G: "ideal ?G = ideal ?S" and "0 \ ?G" by (rule reduced_GB_is_auto_reduced_Polys, rule reduced_GB_is_monic_set_Polys, rule reduced_GB_ideal_Polys, rule reduced_GB_nonzero_Polys) from this(2, 4) assms(5) have "punit.lc g = 1" by (auto simp: punit.is_monic_set_def) have "is_monomial_set ?S" by (auto intro!: is_monomial_setI monomial_is_monomial) with assms(1) 0 have "is_monomial_set ?G" by (rule reduced_GB_is_monomial_set_Polys) hence "is_monomial g" using assms(5) by (rule is_monomial_setD) moreover define s where "s = lpp g" ultimately have g: "g = monomial 1 s" using \punit.lc g = 1\ by (metis punit.monomial_eq_itself) note assms(1) subset_refl assms(2) zero_in_PPs moreover have "ideal ?G \
monomial 1 0 = ideal ?S" by (simp add: ideal_G) ultimately have "splits_wrt (split 0 X S) (cone (monomial (1::'a) 0, X)) ?G" by (rule split_splits_wrt) hence "splits_wrt (fst (split 0 X S), snd (split 0 X S)) P[X] ?G" by simp hence "cone_decomp (P[X] \ ideal ?G) (fst (split 0 X S))" using md2 \is_monomial_set ?G\ by (rule splits_wrt_cone_decomp_1) hence "cone_decomp (P[X] \ ideal ?S) (fst (split 0 X S))" by (simp only: ideal_G) with assms(1, 3, 5) obtain U where "(g, U) \ set (fst (split 0 X S))" using md1 by (rule lem_4_7) with assms(1) subset_refl assms(2, 4) obtain t' x where "t' \ .[X]" and "x \ X" and "U \ X" and "0 \ (\s. s - t') ` S" and s: "s = t' + 0 + Poly_Mapping.single x 1" and "(g, U) \ set (fst (split (t' + 0) U ((\s. s - t') ` S)))" and "set (snd (split (t' + 0) U ((\s. s - t') ` S))) \ (set (snd (split 0 X S)) :: ((_ \\<^sub>0 'a) \ _) set)" unfolding g by (rule fst_splitE) let ?S = "(\s. s - t') ` S" from assms(2) have "finite ?S" by (rule finite_imageI) with assms(1) \U \ X\ obtain V where "V \ U" and "(monomial (1::'a) (t' + 0), V) \ set (snd (split (t' + 0) U ?S))" using \0 \ ?S\ by (rule snd_splitI) note this(2) also have "\ \ set (snd (split 0 X S))" by fact finally have "(monomial (1::'a) t', V) \ set (snd (split 0 X S))" by simp have "poly_deg g = Suc (deg_pm t')" by (simp add: g s deg_pm_plus deg_pm_single poly_deg_monomial) from \V \ U\ \U \ X\ have "V \ X" by (rule subset_trans) show ?thesis by rule fact+ qed corollary cor_4_9: assumes "finite X" and "finite S" and "S \ .[X]" and "g \ punit.reduced_GB (monomial (1::'a::field) ` S)" shows "poly_deg g \ Suc (Max (poly_deg ` fst ` (set (snd (split 0 X S)) :: ((_ \\<^sub>0 'a) \ _) set)))" (is "_ \ Suc (Max (poly_deg ` fst ` ?S))") proof (cases "0 \ S") case True hence "1 \ monomial (1::'a) ` S" by (rule rev_image_eqI) (simp only: single_one) hence "1 \ ideal (monomial (1::'a) ` S)" by (rule ideal.span_base) hence "ideal (monomial (1::'a) ` S) = UNIV" by (simp only: ideal_eq_UNIV_iff_contains_one) moreover from assms(3) have "monomial (1::'a) ` S \ P[X]" by (auto intro: Polys_closed_monomial) ultimately have "punit.reduced_GB (monomial (1::'a) ` S) = {1}" using assms(1) by (simp only: ideal_eq_UNIV_iff_reduced_GB_eq_one_Polys) with assms(4) show ?thesis by simp next case False from finite_set have fin: "finite (poly_deg ` fst ` ?S)" by (intro finite_imageI) obtain t U where "(monomial 1 t, U) \ ?S" and g: "poly_deg g = Suc (deg_pm t)" using assms(1-3) False assms(4) by (rule lem_4_8) from this(1) have "poly_deg (fst (monomial (1::'a) t, U)) \ poly_deg ` fst ` ?S" by (intro imageI) hence "deg_pm t \ poly_deg ` fst ` ?S" by (simp add: poly_deg_monomial) with fin have "deg_pm t \ Max (poly_deg ` fst ` ?S)" by (rule Max_ge) thus "poly_deg g \ Suc (Max (poly_deg ` fst ` ?S))" by (simp add: g) qed lemma standard_decomp_snd_split: assumes "finite X" and "U \ X" and "finite S" and "S \ .[X]" and "t \ .[X]" shows "standard_decomp (deg_pm t) (snd (split t U S) :: ((_ \\<^sub>0 'a::field) \ _) list)" using assms proof (induct t U S rule: split_induct) case (base1 t U S) show ?case by (simp add: standard_decomp_Nil) next case (base2 t U S) have "deg_pm t = poly_deg (monomial (1::'a) t)" by (simp add: poly_deg_monomial) thus ?case by (simp add: standard_decomp_singleton) next case (step t U S V x ps0 ps1 qs0 qs1) from step.hyps(15) step.prems have qs0: "standard_decomp (deg_pm t) qs0" by (simp only: snd_conv) have "(\s. s - Poly_Mapping.single x 1) ` S \ .[X]" proof fix u assume "u \ (\s. s - Poly_Mapping.single x 1) ` S" then obtain s where "s \ S" and u: "u = s - Poly_Mapping.single x 1" .. from this(1) step.prems(1) have "s \ .[X]" .. thus "u \ .[X]" unfolding u by (rule PPs_closed_minus) qed moreover from _ step.prems(2) have "Poly_Mapping.single x 1 + t \ .[X]" proof (rule PPs_closed_plus) from step.hyps(8, 1) have "x \ X" .. thus "Poly_Mapping.single x 1 \ .[X]" by (rule PPs_closed_single) qed ultimately have qs1: "standard_decomp (Suc (deg_pm t)) qs1" using step.hyps(16) by (simp add: deg_pm_plus deg_pm_single) show ?case unfolding snd_conv proof (rule standard_decompI) fix h U0 assume "(h, U0) \ set ((qs0 @ qs1)\<^sub>+)" hence *: "(h, U0) \ set (qs0\<^sub>+) \ set (qs1\<^sub>+)" by (simp add: pos_decomp_append) thus "deg_pm t \ poly_deg h" proof assume "(h, U0) \ set (qs0\<^sub>+)" with qs0 show ?thesis by (rule standard_decompD) next assume "(h, U0) \ set (qs1\<^sub>+)" with qs1 have "Suc (deg_pm t) \ poly_deg h" by (rule standard_decompD) thus ?thesis by simp qed fix d assume d1: "deg_pm t \ d" and d2: "d \ poly_deg h" from * show "\t' U'. (t', U') \ set (qs0 @ qs1) \ poly_deg t' = d \ card U0 \ card U'" proof assume "(h, U0) \ set (qs0\<^sub>+)" with qs0 obtain h' U' where "(h', U') \ set qs0" and "poly_deg h' = d" and "card U0 \ card U'" using d1 d2 by (rule standard_decompE) moreover from this(1) have "(h', U') \ set (qs0 @ qs1)" by simp ultimately show ?thesis by blast next assume "(h, U0) \ set (qs1\<^sub>+)" hence "(h, U0) \ set qs1" by (simp add: pos_decomp_def) from assms(1) step.hyps(1, 2) have "monomial_decomp (snd (split t U S) :: ((_ \\<^sub>0 'a) \ _) list)" by (rule monomial_decomp_split) hence md: "monomial_decomp (qs0 @ qs1)" by (simp add: step.hyps(14)) moreover from \(h, U0) \ set qs1\ have "(h, U0) \ set (qs0 @ qs1)" by simp ultimately have "is_monomial h" and "punit.lc h = 1" by (rule monomial_decompD)+ moreover from this(1) have "monomial (punit.lc h) (lpp h) = h" by (rule punit.monomial_eq_itself) moreover define s where "s = lpp h" ultimately have h: "h = monomial 1 s" by simp from d1 have "deg_pm t = d \ Suc (deg_pm t) \ d" by auto thus ?thesis proof assume "deg_pm t = d" define F where "F = (*) (monomial 1 t) ` monomial (1::'a) ` S" have "F \ P[X]" proof fix f assume "f \ F" then obtain u where "u \ S" and f: "f = monomial 1 (t + u)" by (auto simp: F_def times_monomial_monomial) from this(1) step.prems(1) have "u \ .[X]" .. with step.prems(2) have "t + u \ .[X]" by (rule PPs_closed_plus) thus "f \ P[X]" unfolding f by (rule Polys_closed_monomial) qed have "ideal F = (*) (monomial 1 t) ` ideal (monomial 1 ` S)" by (simp only: ideal.span_image_scale_eq_image_scale F_def) moreover have "inj ((*) (monomial (1::'a) t))" by (auto intro!: injI simp: times_monomial_left monomial_0_iff dest!: punit.monom_mult_inj_3) ultimately have eq: "ideal F \
monomial 1 t = ideal (monomial 1 ` S)" by (simp only: quot_set_image_times) with assms(1) step.hyps(1, 2) step.prems(2) have "splits_wrt (split t U S) (cone (monomial (1::'a) t, U)) F" by (rule split_splits_wrt) hence "splits_wrt (ps0 @ ps1, qs0 @ qs1) (cone (monomial 1 t, U)) F" by (simp only: step.hyps(14)) with assms(1) have "cone_decomp (cone (monomial (1::'a) t, U) \ normal_form F ` P[X]) (qs0 @ qs1)" using md _ \F \ P[X]\ by (rule splits_wrt_cone_decomp_2) (auto intro!: is_monomial_setI monomial_is_monomial simp: F_def times_monomial_monomial) hence "cone (monomial 1 s, U0) \ cone (monomial (1::'a) t, U) \ normal_form F ` P[X]" using \(h, U0) \ set (qs0 @ qs1)\ unfolding h by (rule cone_decomp_cone_subset) with assms(1) step.hyps(1, 2) step.prems(2) \F \ P[X]\ eq obtain U' where "(monomial (1::'a) t, U') \ set (snd (split t U S))" and "card U0 \ card U'" by (rule lem_4_6) from this(1) have "(monomial 1 t, U') \ set (qs0 @ qs1)" by (simp add: step.hyps(14)) show ?thesis proof (intro exI conjI) show "poly_deg (monomial (1::'a) t) = d" by (simp add: poly_deg_monomial \deg_pm t = d\) qed fact+ next assume "Suc (deg_pm t) \ d" with qs1 \(h, U0) \ set (qs1\<^sub>+)\ obtain h' U' where "(h', U') \ set qs1" and "poly_deg h' = d" and "card U0 \ card U'" using d2 by (rule standard_decompE) moreover from this(1) have "(h', U') \ set (qs0 @ qs1)" by simp ultimately show ?thesis by blast qed qed qed qed theorem standard_cone_decomp_snd_split: fixes F defines "G \ punit.reduced_GB F" defines "ss \ (split 0 X (lpp ` G)) :: ((_ \\<^sub>0 'a::field) \ _) list \ _" defines "d \ Suc (Max (poly_deg ` fst ` set (snd ss)))" assumes "finite X" and "F \ P[X]" shows "standard_decomp 0 (snd ss)" (is ?thesis1) and "cone_decomp (normal_form F ` P[X]) (snd ss)" (is ?thesis2) and "(\f. f \ F \ homogeneous f) \ g \ G \ poly_deg g \ d" proof - have "ideal G = ideal F" and "punit.is_Groebner_basis G" and "finite G" and "0 \ G" and "G \ P[X]" and "punit.is_reduced_GB G" using assms(4, 5) unfolding G_def by (rule reduced_GB_ideal_Polys, rule reduced_GB_is_GB_Polys, rule finite_reduced_GB_Polys, rule reduced_GB_nonzero_Polys, rule reduced_GB_Polys, rule reduced_GB_is_reduced_GB_Polys) define S where "S = lpp ` G" note assms(4) subset_refl moreover from \finite G\ have "finite S" unfolding S_def by (rule finite_imageI) moreover from \G \ P[X]\ have "S \ .[X]" unfolding S_def by (rule PPs_closed_image_lpp) ultimately have "standard_decomp (deg_pm (0::'x \\<^sub>0 nat)) (snd ss)" using zero_in_PPs unfolding ss_def S_def by (rule standard_decomp_snd_split) thus ?thesis1 by simp let ?S = "monomial (1::'a) ` S" from \S \ .[X]\ have "?S \ P[X]" by (auto intro: Polys_closed_monomial) have "splits_wrt ss (cone (monomial 1 0, X)) ?S" using assms(4) subset_refl \finite S\ zero_in_PPs unfolding ss_def S_def by (rule split_splits_wrt) simp hence "splits_wrt (fst ss, snd ss) P[X] ?S" by simp with assms(4) have "cone_decomp (P[X] \ normal_form ?S ` P[X]) (snd ss)" using _ _ \?S \ P[X]\ proof (rule splits_wrt_cone_decomp_2) from assms(4) subset_refl \finite S\ show "monomial_decomp (snd ss)" unfolding ss_def S_def by (rule monomial_decomp_split) qed (auto intro!: is_monomial_setI monomial_is_monomial) moreover have "normal_form ?S ` P[X] = normal_form F ` P[X]" by (rule set_eqI) (simp add: image_normal_form_iff[OF assms(4)] assms(5) \?S \ P[X]\, simp add: S_def is_red_reduced_GB_monomial_lt_GB_Polys[OF assms(4)] \G \ P[X]\ \0 \ G\ flip: G_def) moreover from assms(4, 5) have "normal_form F ` P[X] \ P[X]" by (auto intro: Polys_closed_normal_form) ultimately show ?thesis2 by (simp only: Int_absorb1) assume "\f. f \ F \ homogeneous f" moreover note \punit.is_reduced_GB G\ \ideal G = ideal F\ moreover assume "g \ G" ultimately have "homogeneous g" by (rule is_reduced_GB_homogeneous) moreover have "lpp g \ keys g" proof (rule punit.lt_in_keys) from \g \ G\ \0 \ G\ show "g \ 0" by blast qed ultimately have deg_lt: "deg_pm (lpp g) = poly_deg g" by (rule homogeneousD_poly_deg) from \g \ G\ have "monomial 1 (lpp g) \ ?S" unfolding S_def by (intro imageI) also have "\ = punit.reduced_GB ?S" unfolding S_def G_def using assms(4, 5) by (rule reduced_GB_monomial_lt_reduced_GB_Polys[symmetric]) finally have "monomial 1 (lpp g) \ punit.reduced_GB ?S" . with assms(4) \finite S\ \S \ .[X]\ have "poly_deg (monomial (1::'a) (lpp g)) \ d" unfolding d_def ss_def S_def[symmetric] by (rule cor_4_9) thus "poly_deg g \ d" by (simp add: poly_deg_monomial deg_lt) qed subsection \Splitting Ideals\ qualified definition ideal_decomp_aux :: "(('x \\<^sub>0 nat) \\<^sub>0 'a) set \ (('x \\<^sub>0 nat) \\<^sub>0 'a) \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::field) set \ ((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list)" where "ideal_decomp_aux F f = (let J = ideal F; L = (J \
f) \ P[X]; L' = lpp ` punit.reduced_GB L in ((*) f ` normal_form L ` P[X], map (apfst ((*) f)) (snd (split 0 X L'))))" context assumes fin_X: "finite X" begin lemma ideal_decomp_aux: assumes "finite F" and "F \ P[X]" and "f \ P[X]" shows "fst (ideal_decomp_aux F f) \ ideal {f}" (is ?thesis1) and "ideal F \ fst (ideal_decomp_aux F f) = {0}" (is ?thesis2) and "direct_decomp (ideal (insert f F) \ P[X]) [fst (ideal_decomp_aux F f), ideal F \ P[X]]" (is ?thesis3) and "cone_decomp (fst (ideal_decomp_aux F f)) (snd (ideal_decomp_aux F f))" (is ?thesis4) and "f \ 0 \ valid_decomp X (snd (ideal_decomp_aux F f))" (is "_ \ ?thesis5") and "f \ 0 \ standard_decomp (poly_deg f) (snd (ideal_decomp_aux F f))" (is "_ \ ?thesis6") and "homogeneous f \ hom_decomp (snd (ideal_decomp_aux F f))" (is "_ \ ?thesis7") proof - define J where "J = ideal F" define L where "L = (J \
f) \ P[X]" define S where "S = (*) f ` normal_form L ` P[X]" define L' where "L' = lpp ` punit.reduced_GB L" have eq: "ideal_decomp_aux F f = (S, map (apfst ((*) f)) (snd (split 0 X L')))" by (simp add: J_def ideal_decomp_aux_def Let_def L_def L'_def S_def) have L_sub: "L \ P[X]" by (simp add: L_def) show ?thesis1 unfolding eq fst_conv proof fix s assume "s \ S" then obtain q where "s = normal_form L q * f" unfolding S_def by (elim imageE) auto also have "\ \ ideal {f}" by (intro ideal.span_scale ideal.span_base singletonI) finally show "s \ ideal {f}" . qed show ?thesis2 proof (rule set_eqI) fix h show "h \ ideal F \ fst (ideal_decomp_aux F f) \ h \ {0}" proof assume "h \ ideal F \ fst (ideal_decomp_aux F f)" hence "h \ J" and "h \ S" by (simp_all add: J_def S_def eq) from this(2) obtain q where "q \ P[X]" and h: "h = f * normal_form L q" by (auto simp: S_def) from fin_X L_sub this(1) have "normal_form L q \ P[X]" by (rule Polys_closed_normal_form) moreover from \h \ J\ have "f * normal_form L q \ J" by (simp add: h) ultimately have "normal_form L q \ L" by (simp add: L_def quot_set_iff) hence "normal_form L q \ ideal L" by (rule ideal.span_base) with normal_form_diff_in_ideal[OF fin_X L_sub] have "(q - normal_form L q) + normal_form L q \ ideal L" by (rule ideal.span_add) hence "normal_form L q = 0" using fin_X L_sub by (simp add: normal_form_zero_iff) thus "h \ {0}" by (simp add: h) next assume "h \ {0}" moreover have "0 \ (*) f ` normal_form L ` P[X]" proof (intro image_eqI) from fin_X L_sub show "0 = normal_form L 0" by (simp only: normal_form_zero) qed (simp_all add: zero_in_Polys) ultimately show "h \ ideal F \ fst (ideal_decomp_aux F f)" by (simp add: ideal.span_zero eq S_def) qed qed have "direct_decomp (ideal (insert f F) \ P[X]) [ideal F \ P[X], fst (ideal_decomp_aux F f)]" unfolding eq fst_conv S_def L_def J_def using fin_X assms(2, 3) by (rule direct_decomp_ideal_insert) - thus ?thesis3 using perm.swap by (rule direct_decomp_perm) + thus ?thesis3 by (rule direct_decomp_perm) simp have std: "standard_decomp 0 (snd (split 0 X L') :: ((_ \\<^sub>0 'a) \ _) list)" and "cone_decomp (normal_form L ` P[X]) (snd (split 0 X L'))" unfolding L'_def using fin_X \L \ P[X]\ by (rule standard_cone_decomp_snd_split)+ from this(2) show ?thesis4 unfolding eq fst_conv snd_conv S_def by (rule cone_decomp_map_times) from fin_X \L \ P[X]\ have "finite (punit.reduced_GB L)" by (rule finite_reduced_GB_Polys) hence "finite L'" unfolding L'_def by (rule finite_imageI) { have "monomial_decomp (snd (split 0 X L') :: ((_ \\<^sub>0 'a) \ _) list)" using fin_X subset_refl \finite L'\ by (rule monomial_decomp_split) hence "hom_decomp (snd (split 0 X L') :: ((_ \\<^sub>0 'a) \ _) list)" by (rule monomial_decomp_imp_hom_decomp) moreover assume "homogeneous f" ultimately show ?thesis7 unfolding eq snd_conv by (rule hom_decomp_map_times) } have vd: "valid_decomp X (snd (split 0 X L') :: ((_ \\<^sub>0 'a) \ _) list)" using fin_X subset_refl \finite L'\ zero_in_PPs by (rule valid_decomp_split) moreover note assms(3) moreover assume "f \ 0" ultimately show ?thesis5 unfolding eq snd_conv by (rule valid_decomp_map_times) from std vd \f \ 0\ have "standard_decomp (0 + poly_deg f) (map (apfst ((*) f)) (snd (split 0 X L')))" by (rule standard_decomp_map_times) thus ?thesis6 by (simp add: eq) qed lemma ideal_decompE: fixes f0 :: "_ \\<^sub>0 'a::field" assumes "finite F" and "F \ P[X]" and "f0 \ P[X]" and "\f. f \ F \ poly_deg f \ poly_deg f0" obtains T ps where "valid_decomp X ps" and "standard_decomp (poly_deg f0) ps" and "cone_decomp T ps" and "(\f. f \ F \ homogeneous f) \ hom_decomp ps" and "direct_decomp (ideal (insert f0 F) \ P[X]) [ideal {f0} \ P[X], T]" using assms(1, 2, 4) proof (induct F arbitrary: thesis) case empty show ?case proof (rule empty.prems) show "valid_decomp X []" by (rule valid_decompI) simp_all next show "standard_decomp (poly_deg f0) []" by (rule standard_decompI) simp_all next show "cone_decomp {0} []" by (rule cone_decompI) (simp add: direct_decomp_def bij_betw_def) next have "direct_decomp (ideal {f0} \ P[X]) [ideal {f0} \ P[X]]" by (fact direct_decomp_singleton) hence "direct_decomp (ideal {f0} \ P[X]) [{0}, ideal {f0} \ P[X]]" by (rule direct_decomp_Cons_zeroI) thus "direct_decomp (ideal {f0} \ P[X]) [ideal {f0} \ P[X], {0}]" - using perm.swap by (rule direct_decomp_perm) + by (rule direct_decomp_perm) simp qed (simp add: hom_decomp_def) next case (insert f F) from insert.prems(2) have "F \ P[X]" by simp moreover have "poly_deg f' \ poly_deg f0" if "f' \ F" for f' proof - from that have "f' \ insert f F" by simp thus ?thesis by (rule insert.prems) qed ultimately obtain T ps where valid_ps: "valid_decomp X ps" and std_ps: "standard_decomp (poly_deg f0) ps" and cn_ps: "cone_decomp T ps" and dd: "direct_decomp (ideal (insert f0 F) \ P[X]) [ideal {f0} \ P[X], T]" and hom_ps: "(\f. f \ F \ homogeneous f) \ hom_decomp ps" using insert.hyps(3) by metis show ?case proof (cases "f = 0") case True show ?thesis proof (rule insert.prems) from dd show "direct_decomp (ideal (insert f0 (insert f F)) \ P[X]) [ideal {f0} \ P[X], T]" by (simp only: insert_commute[of f0] True ideal.span_insert_zero) next assume "\f'. f' \ insert f F \ homogeneous f'" hence "\f. f \ F \ homogeneous f" by blast thus "hom_decomp ps" by (rule hom_ps) qed fact+ next case False let ?D = "ideal_decomp_aux (insert f0 F) f" from insert.hyps(1) have f0F_fin: "finite (insert f0 F)" by simp moreover from \F \ P[X]\ assms(3) have f0F_sub: "insert f0 F \ P[X]" by simp moreover from insert.prems(2) have "f \ P[X]" by simp ultimately have eq: "ideal (insert f0 F) \ fst ?D = {0}" and "valid_decomp X (snd ?D)" and cn_D: "cone_decomp (fst ?D) (snd ?D)" and "standard_decomp (poly_deg f) (snd ?D)" and dd': "direct_decomp (ideal (insert f (insert f0 F)) \ P[X]) [fst ?D, ideal (insert f0 F) \ P[X]]" and hom_D: "homogeneous f \ hom_decomp (snd ?D)" by (rule ideal_decomp_aux, auto intro: ideal_decomp_aux simp: False) note fin_X this(2-4) moreover have "poly_deg f \ poly_deg f0" by (rule insert.prems) simp ultimately obtain qs where valid_qs: "valid_decomp X qs" and cn_qs: "cone_decomp (fst ?D) qs" and std_qs: "standard_decomp (poly_deg f0) qs" and hom_qs: "hom_decomp (snd ?D) \ hom_decomp qs" by (rule standard_decomp_geE) blast let ?T = "sum_list ` listset [T, fst ?D]" let ?ps = "ps @ qs" show ?thesis proof (rule insert.prems) from valid_ps valid_qs show "valid_decomp X ?ps" by (rule valid_decomp_append) next from std_ps std_qs show "standard_decomp (poly_deg f0) ?ps" by (rule standard_decomp_append) next - from dd perm.swap have "direct_decomp (ideal (insert f0 F) \ P[X]) [T, ideal {f0} \ P[X]]" - by (rule direct_decomp_perm) + from dd have "direct_decomp (ideal (insert f0 F) \ P[X]) [T, ideal {f0} \ P[X]]" + by (rule direct_decomp_perm) simp hence "T \ ideal (insert f0 F) \ P[X]" by (rule direct_decomp_Cons_subsetI) (simp add: ideal.span_zero zero_in_Polys) hence "T \ fst ?D \ ideal (insert f0 F) \ fst ?D" by blast hence "T \ fst ?D \ {0}" by (simp only: eq) from refl have "direct_decomp ?T [T, fst ?D]" proof (intro direct_decompI inj_onI) fix xs ys assume "xs \ listset [T, fst ?D]" then obtain x1 x2 where "x1 \ T" and "x2 \ fst ?D" and xs: "xs = [x1, x2]" by (rule listset_doubletonE) assume "ys \ listset [T, fst ?D]" then obtain y1 y2 where "y1 \ T" and "y2 \ fst ?D" and ys: "ys = [y1, y2]" by (rule listset_doubletonE) assume "sum_list xs = sum_list ys" hence "x1 - y1 = y2 - x2" by (simp add: xs ys) (metis add_diff_cancel_left add_diff_cancel_right) moreover from cn_ps \x1 \ T\ \y1 \ T\ have "x1 - y1 \ T" by (rule cone_decomp_closed_minus) moreover from cn_D \y2 \ fst ?D\ \x2 \ fst ?D\ have "y2 - x2 \ fst ?D" by (rule cone_decomp_closed_minus) ultimately have "y2 - x2 \ T \ fst ?D" by simp also have "\ \ {0}" by fact finally have "x2 = y2" by simp with \x1 - y1 = y2 - x2\ show "xs = ys" by (simp add: xs ys) qed thus "cone_decomp ?T ?ps" using cn_ps cn_qs by (rule cone_decomp_append) next assume "\f'. f' \ insert f F \ homogeneous f'" hence "homogeneous f" and "\f'. f' \ F \ homogeneous f'" by blast+ from this(2) have "hom_decomp ps" by (rule hom_ps) moreover from \homogeneous f\ have "hom_decomp qs" by (intro hom_qs hom_D) ultimately show "hom_decomp (ps @ qs)" by (simp only: hom_decomp_append_iff) next from dd' have "direct_decomp (ideal (insert f0 (insert f F)) \ P[X]) [ideal (insert f0 F) \ P[X], fst ?D]" - by (simp add: insert_commute direct_decomp_perm perm.swap) + by (simp add: insert_commute direct_decomp_perm) hence "direct_decomp (ideal (insert f0 (insert f F)) \ P[X]) ([fst ?D] @ [ideal {f0} \ P[X], T])" using dd by (rule direct_decomp_direct_decomp) hence "direct_decomp (ideal (insert f0 (insert f F)) \ P[X]) ([ideal {f0} \ P[X]] @ [T, fst ?D])" by (rule direct_decomp_perm) auto hence "direct_decomp (ideal (insert f0 (insert f F)) \ P[X]) [sum_list ` listset [ideal {f0} \ P[X]], ?T]" by (rule direct_decomp_appendD) thus "direct_decomp (ideal (insert f0 (insert f F)) \ P[X]) [ideal {f0} \ P[X], ?T]" by (simp add: image_image) qed qed qed subsection \Exact Cone Decompositions\ definition exact_decomp :: "nat \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ 'x set) list \ bool" where "exact_decomp m ps \ (\(h, U)\set ps. h \ P[X] \ U \ X) \ (\(h, U)\set ps. \(h', U')\set ps. poly_deg h = poly_deg h' \ m < card U \ m < card U' \ (h, U) = (h', U'))" lemma exact_decompI: "(\h U. (h, U) \ set ps \ h \ P[X]) \ (\h U. (h, U) \ set ps \ U \ X) \ (\h h' U U'. (h, U) \ set ps \ (h', U') \ set ps \ poly_deg h = poly_deg h' \ m < card U \ m < card U' \ (h, U) = (h', U')) \ exact_decomp m ps" unfolding exact_decomp_def by fastforce lemma exact_decompD: assumes "exact_decomp m ps" and "(h, U) \ set ps" shows "h \ P[X]" and "U \ X" and "(h', U') \ set ps \ poly_deg h = poly_deg h' \ m < card U \ m < card U' \ (h, U) = (h', U')" using assms unfolding exact_decomp_def by fastforce+ lemma exact_decompI_zero: assumes "\h U. (h, U) \ set ps \ h \ P[X]" and "\h U. (h, U) \ set ps \ U \ X" and "\h h' U U'. (h, U) \ set (ps\<^sub>+) \ (h', U') \ set (ps\<^sub>+) \ poly_deg h = poly_deg h' \ (h, U) = (h', U')" shows "exact_decomp 0 ps" using assms(1, 2) proof (rule exact_decompI) fix h h' and U U' :: "'x set" assume "0 < card U" hence "U \ {}" by auto moreover assume "(h, U) \ set ps" ultimately have "(h, U) \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) assume "0 < card U'" hence "U' \ {}" by auto moreover assume "(h', U') \ set ps" ultimately have "(h', U') \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) assume "poly_deg h = poly_deg h'" with \(h, U) \ set (ps\<^sub>+)\ \(h', U') \ set (ps\<^sub>+)\ show "(h, U) = (h', U')" by (rule assms(3)) qed lemma exact_decompD_zero: assumes "exact_decomp 0 ps" and "(h, U) \ set (ps\<^sub>+)" and "(h', U') \ set (ps\<^sub>+)" and "poly_deg h = poly_deg h'" shows "(h, U) = (h', U')" proof - from assms(2) have "(h, U) \ set ps" and "U \ {}" by (simp_all add: pos_decomp_def) from assms(1) this(1) have "U \ X" by (rule exact_decompD) hence "finite U" using fin_X by (rule finite_subset) with \U \ {}\ have "0 < card U" by (simp add: card_gt_0_iff) from assms(3) have "(h', U') \ set ps" and "U' \ {}" by (simp_all add: pos_decomp_def) from assms(1) this(1) have "U' \ X" by (rule exact_decompD) hence "finite U'" using fin_X by (rule finite_subset) with \U' \ {}\ have "0 < card U'" by (simp add: card_gt_0_iff) show ?thesis by (rule exact_decompD) fact+ qed lemma exact_decomp_imp_valid_decomp: assumes "exact_decomp m ps" and "\h U. (h, U) \ set ps \ h \ 0" shows "valid_decomp X ps" proof (rule valid_decompI) fix h U assume *: "(h, U) \ set ps" with assms(1) show "h \ P[X]" and "U \ X" by (rule exact_decompD)+ from * show "h \ 0" by (rule assms(2)) qed lemma exact_decomp_card_X: assumes "valid_decomp X ps" and "card X \ m" shows "exact_decomp m ps" proof (rule exact_decompI) fix h U assume "(h, U) \ set ps" with assms(1) show "h \ P[X]" and "U \ X" by (rule valid_decompD)+ next fix h1 h2 U1 U2 assume "(h1, U1) \ set ps" with assms(1) have "U1 \ X" by (rule valid_decompD) with fin_X have "card U1 \ card X" by (rule card_mono) also have "\ \ m" by (fact assms(2)) also assume "m < card U1" finally show "(h1, U1) = (h2, U2)" by simp qed definition \ :: "((('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ 'x set) list \ nat" where "\ ps = (LEAST k. standard_decomp k ps)" definition \ :: "((('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ 'x set) list \ nat \ nat" where "\ ps i = (LEAST d. \ ps \ d \ (\(h, U)\set ps. i \ card U \ poly_deg h < d))" lemma \: "standard_decomp k ps \ standard_decomp (\ ps) ps" unfolding \_def by (rule LeastI) lemma \_Nil: assumes "ps\<^sub>+ = []" shows "\ ps = 0" proof - from assms have "standard_decomp 0 ps" by (rule standard_decomp_Nil) thus ?thesis unfolding \_def by (rule Least_eq_0) qed lemma \_nonempty: assumes "valid_decomp X ps" and "standard_decomp k ps" and "ps\<^sub>+ \ []" shows "\ ps = Min (poly_deg ` fst ` set (ps\<^sub>+))" using fin_X assms(1) _ assms(3) proof (rule standard_decomp_nonempty_unique) from assms(2) show "standard_decomp (\ ps) ps" by (rule \) qed lemma \_nonempty_unique: assumes "valid_decomp X ps" and "standard_decomp k ps" and "ps\<^sub>+ \ []" shows "\ ps = k" proof - from assms have "\ ps = Min (poly_deg ` fst ` set (ps\<^sub>+))" by (rule \_nonempty) moreover from fin_X assms have "k = Min (poly_deg ` fst ` set (ps\<^sub>+))" by (rule standard_decomp_nonempty_unique) ultimately show ?thesis by simp qed lemma \: shows "\ ps \ \ ps i" and "(h, U) \ set ps \ i \ card U \ poly_deg h < \ ps i" proof - let ?A = "poly_deg ` fst ` set ps" define A where "A = insert (\ ps) ?A" define m where "m = Suc (Max A)" from finite_set have "finite ?A" by (intro finite_imageI) hence "finite A" by (simp add: A_def) have "\ ps \ \ ps i \ (\(h', U')\set ps. i \ card U' \ poly_deg h' < \ ps i)" unfolding \_def proof (rule LeastI) have "\ ps \ A" by (simp add: A_def) with \finite A\ have "\ ps \ Max A" by (rule Max_ge) hence "\ ps \ m" by (simp add: m_def) moreover { fix h U assume "(h, U) \ set ps" hence "poly_deg (fst (h, U)) \ ?A" by (intro imageI) hence "poly_deg h \ A" by (simp add: A_def) with \finite A\ have "poly_deg h \ Max A" by (rule Max_ge) hence "poly_deg h < m" by (simp add: m_def) } ultimately show "\ ps \ m \ (\(h, U)\set ps. i \ card U \ poly_deg h < m)" by blast qed thus "\ ps \ \ ps i" and "(h, U) \ set ps \ i \ card U \ poly_deg h < \ ps i" by blast+ qed lemma \_le: "\ ps \ d \ (\h' U'. (h', U') \ set ps \ i \ card U' \ poly_deg h' < d) \ \ ps i \ d" unfolding \_def by (intro Least_le) blast lemma \_decreasing: assumes "i \ j" shows "\ ps j \ \ ps i" proof (rule \_le) fix h U assume "(h, U) \ set ps" assume "j \ card U" with assms(1) have "i \ card U" by (rule le_trans) with \(h, U) \ set ps\ show "poly_deg h < \ ps i" by (rule \) qed (fact \) lemma \_Nil: assumes "ps\<^sub>+ = []" and "Suc 0 \ i" shows "\ ps i = 0" unfolding \_def proof (rule Least_eq_0) from assms(1) have "\ ps = 0" by (rule \_Nil) moreover { fix h and U::"'x set" note assms(2) also assume "i \ card U" finally have "U \ {}" by auto moreover assume "(h, U) \ set ps" ultimately have "(h, U) \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) hence False by (simp add: assms) } ultimately show "\ ps \ 0 \ (\(h, U)\set ps. i \ card U \ poly_deg h < 0)" by blast qed lemma \_zero: assumes "ps \ []" shows "Suc (Max (poly_deg ` fst ` set ps)) \ \ ps 0" proof - from finite_set have "finite (poly_deg ` fst ` set ps)" by (intro finite_imageI) moreover from assms have "poly_deg ` fst ` set ps \ {}" by simp moreover have "\a\poly_deg ` fst ` set ps. a < \ ps 0" proof fix d assume "d \ poly_deg ` fst ` set ps" then obtain p where "p \ set ps" and "d = poly_deg (fst p)" by blast moreover obtain h U where "p = (h, U)" using prod.exhaust by blast ultimately have "(h, U) \ set ps" and d: "d = poly_deg h" by simp_all from this(1) le0 show "d < \ ps 0" unfolding d by (rule \) qed ultimately have "Max (poly_deg ` fst ` set ps) < \ ps 0" by simp thus ?thesis by simp qed corollary \_zero_gr: assumes "(h, U) \ set ps" shows "poly_deg h < \ ps 0" proof - have "poly_deg h \ Max (poly_deg ` fst ` set ps)" proof (rule Max_ge) from finite_set show "finite (poly_deg ` fst ` set ps)" by (intro finite_imageI) next from assms have "poly_deg (fst (h, U)) \ poly_deg ` fst ` set ps" by (intro imageI) thus "poly_deg h \ poly_deg ` fst ` set ps" by simp qed also have "\ < Suc \" by simp also have "\ \ \ ps 0" proof (rule \_zero) from assms show "ps \ []" by auto qed finally show ?thesis . qed lemma \_one: assumes "valid_decomp X ps" and "standard_decomp k ps" shows "\ ps (Suc 0) = (if ps\<^sub>+ = [] then 0 else Suc (Max (poly_deg ` fst ` set (ps\<^sub>+))))" proof (cases "ps\<^sub>+ = []") case True hence "\ ps (Suc 0) = 0" using le_refl by (rule \_Nil) with True show ?thesis by simp next case False with assms have aP: "\ ps = Min (poly_deg ` fst ` set (ps\<^sub>+))" (is "_ = Min ?A") by (rule \_nonempty) from pos_decomp_subset finite_set have "finite (set (ps\<^sub>+))" by (rule finite_subset) hence "finite ?A" by (intro finite_imageI) from False have "?A \ {}" by simp have "\ ps (Suc 0) = Suc (Max ?A)" unfolding \_def proof (rule Least_equality) from \finite ?A\ \?A \ {}\ have "\ ps \ ?A" unfolding aP by (rule Min_in) with \finite ?A\ have "\ ps \ Max ?A" by (rule Max_ge) hence "\ ps \ Suc (Max ?A)" by simp moreover { fix h U assume "(h, U) \ set ps" with fin_X assms(1) have "finite U" by (rule valid_decompD_finite) moreover assume "Suc 0 \ card U" ultimately have "U \ {}" by auto with \(h, U) \ set ps\ have "(h, U) \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) hence "poly_deg (fst (h, U)) \ ?A" by (intro imageI) hence "poly_deg h \ ?A" by (simp only: fst_conv) with \finite ?A\ have "poly_deg h \ Max ?A" by (rule Max_ge) hence "poly_deg h < Suc (Max ?A)" by simp } ultimately show "\ ps \ Suc (Max ?A) \ (\(h, U)\set ps. Suc 0 \ card U \ poly_deg h < Suc (Max ?A))" by blast next fix d assume "\ ps \ d \ (\(h, U)\set ps. Suc 0 \ card U \ poly_deg h < d)" hence rl: "poly_deg h < d" if "(h, U) \ set ps" and "0 < card U" for h U using that by auto have "Max ?A < d" unfolding Max_less_iff[OF \finite ?A\ \?A \ {}\] proof fix d0 assume "d0 \ poly_deg ` fst ` set (ps\<^sub>+)" then obtain h U where "(h, U) \ set (ps\<^sub>+)" and d0: "d0 = poly_deg h" by auto from this(1) have "(h, U) \ set ps" and "U \ {}" by (simp_all add: pos_decomp_def) from fin_X assms(1) this(1) have "finite U" by (rule valid_decompD_finite) with \U \ {}\ have "0 < card U" by (simp add: card_gt_0_iff) with \(h, U) \ set ps\ show "d0 < d" unfolding d0 by (rule rl) qed thus "Suc (Max ?A) \ d" by simp qed with False show ?thesis by simp qed corollary \_one_gr: assumes "valid_decomp X ps" and "standard_decomp k ps" and "(h, U) \ set (ps\<^sub>+)" shows "poly_deg h < \ ps (Suc 0)" proof - from assms(3) have "ps\<^sub>+ \ []" by auto with assms(1, 2) have eq: "\ ps (Suc 0) = Suc (Max (poly_deg ` fst ` set (ps\<^sub>+)))" by (simp add: \_one) have "poly_deg h \ Max (poly_deg ` fst ` set (ps\<^sub>+))" proof (rule Max_ge) from finite_set show "finite (poly_deg ` fst ` set (ps\<^sub>+))" by (intro finite_imageI) next from assms(3) have "poly_deg (fst (h, U)) \ poly_deg ` fst ` set (ps\<^sub>+)" by (intro imageI) thus "poly_deg h \ poly_deg ` fst ` set (ps\<^sub>+)" by simp qed also have "\ < \ ps (Suc 0)" by (simp add: eq) finally show ?thesis . qed lemma \_card_X: assumes "exact_decomp m ps" and "Suc (card X) \ i" shows "\ ps i = \ ps" unfolding \_def proof (rule Least_equality) { fix h U assume "(h, U) \ set ps" with assms(1) have "U \ X" by (rule exact_decompD) note assms(2) also assume "i \ card U" finally have "card X < card U" by simp with fin_X have "\ U \ X" by (auto dest: card_mono leD) hence False using \U \ X\ .. } thus "\ ps \ \ ps \ (\(h, U)\set ps. i \ card U \ poly_deg h < \ ps)" by blast qed simp lemma lem_6_1_1: assumes "standard_decomp k ps" and "exact_decomp m ps" and "Suc 0 \ i" and "i \ card X" and "\ ps (Suc i) \ d" and "d < \ ps i" obtains h U where "(h, U) \ set (ps\<^sub>+)" and "poly_deg h = d" and "card U = i" proof - have "ps\<^sub>+ \ []" proof assume "ps\<^sub>+ = []" hence "\ ps i = 0" using assms(3) by (rule \_Nil) with assms(6) show False by simp qed have eq1: "\ ps (Suc (card X)) = \ ps" using assms(2) le_refl by (rule \_card_X) from assms(1) have std: "standard_decomp (\ ps (Suc (card X))) ps" unfolding eq1 by (rule \) from assms(4) have "Suc i \ Suc (card X)" .. hence "\ ps (Suc (card X)) \ \ ps (Suc i)" by (rule \_decreasing) hence "\ ps \ \ ps (Suc i)" by (simp only: eq1) have "\h U. (h, U) \ set ps \ i \ card U \ \ ps i \ Suc (poly_deg h)" proof (rule ccontr) assume *: "\h U. (h, U) \ set ps \ i \ card U \ \ ps i \ Suc (poly_deg h)" note \\ ps \ \ ps (Suc i)\ also from assms(5, 6) have "\ ps (Suc i) < \ ps i" by (rule le_less_trans) finally have "\ ps < \ ps i" . hence "\ ps \ \ ps i - 1" by simp hence "\ ps i \ \ ps i - 1" proof (rule \_le) fix h U assume "(h, U) \ set ps" and "i \ card U" show "poly_deg h < \ ps i - 1" proof (rule ccontr) assume "\ poly_deg h < \ ps i - 1" hence "\ ps i \ Suc (poly_deg h)" by simp with * \(h, U) \ set ps\ \i \ card U\ show False by auto qed qed thus False using \\ ps < \ ps i\ by linarith qed then obtain h U where "(h, U) \ set ps" and "i \ card U" and "\ ps i \ Suc (poly_deg h)" by blast from assms(3) this(2) have "U \ {}" by auto with \(h, U) \ set ps\ have "(h, U) \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) note std this moreover have "\ ps (Suc (card X)) \ d" unfolding eq1 using \\ ps \ \ ps (Suc i)\ assms(5) by (rule le_trans) moreover have "d \ poly_deg h" proof - from assms(6) \\ ps i \ Suc (poly_deg h)\ have "d < Suc (poly_deg h)" by (rule less_le_trans) thus ?thesis by simp qed ultimately obtain h' U' where "(h', U') \ set ps" and d: "poly_deg h' = d" and "card U \ card U'" by (rule standard_decompE) from \i \ card U\ this(3) have "i \ card U'" by (rule le_trans) with assms(3) have "U' \ {}" by auto with \(h', U') \ set ps\ have "(h', U') \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) moreover note \poly_deg h' = d\ moreover have "card U' = i" proof (rule ccontr) assume "card U' \ i" with \i \ card U'\ have "Suc i \ card U'" by simp with \(h', U') \ set ps\ have "poly_deg h' < \ ps (Suc i)" by (rule \) with assms(5) show False by (simp add: d) qed ultimately show ?thesis .. qed corollary lem_6_1_2: assumes "standard_decomp k ps" and "exact_decomp 0 ps" and "Suc 0 \ i" and "i \ card X" and "\ ps (Suc i) \ d" and "d < \ ps i" obtains h U where "{(h', U') \ set (ps\<^sub>+). poly_deg h' = d} = {(h, U)}" and "card U = i" proof - from assms obtain h U where "(h, U) \ set (ps\<^sub>+)" and "poly_deg h = d" and "card U = i" by (rule lem_6_1_1) hence "{(h, U)} \ {(h', U') \ set (ps\<^sub>+). poly_deg h' = d}" (is "_ \ ?A") by simp moreover have "?A \ {(h, U)}" proof fix x assume "x \ ?A" then obtain h' U' where "(h', U') \ set (ps\<^sub>+)" and "poly_deg h' = d" and x: "x = (h', U')" by blast note assms(2) \(h, U) \ set (ps\<^sub>+)\ this(1) moreover have "poly_deg h = poly_deg h'" by (simp only: \poly_deg h = d\ \poly_deg h' = d\) ultimately have "(h, U) = (h', U')" by (rule exact_decompD_zero) thus "x \ {(h, U)}" by (simp add: x) qed ultimately have "{(h, U)} = ?A" .. hence "?A = {(h, U)}" by (rule sym) thus ?thesis using \card U = i\ .. qed corollary lem_6_1_2': assumes "standard_decomp k ps" and "exact_decomp 0 ps" and "Suc 0 \ i" and "i \ card X" and "\ ps (Suc i) \ d" and "d < \ ps i" shows "card {(h', U') \ set (ps\<^sub>+). poly_deg h' = d} = 1" (is "card ?A = _") and "{(h', U') \ set (ps\<^sub>+). poly_deg h' = d \ card U' = i} = {(h', U') \ set (ps\<^sub>+). poly_deg h' = d}" (is "?B = _") and "card {(h', U') \ set (ps\<^sub>+). poly_deg h' = d \ card U' = i} = 1" proof - from assms obtain h U where "?A = {(h, U)}" and "card U = i" by (rule lem_6_1_2) from this(1) show "card ?A = 1" by simp moreover show "?B = ?A" proof have "(h, U) \ ?A" by (simp add: \?A = {(h, U)}\) have "?A = {(h, U)}" by fact also from \(h, U) \ ?A\ \card U = i\ have "\ \ ?B" by simp finally show "?A \ ?B" . qed blast ultimately show "card ?B = 1" by simp qed corollary lem_6_1_3: assumes "standard_decomp k ps" and "exact_decomp 0 ps" and "Suc 0 \ i" and "i \ card X" and "(h, U) \ set (ps\<^sub>+)" and "card U = i" shows "\ ps (Suc i) \ poly_deg h" proof (rule ccontr) define j where "j = (LEAST j'. \ ps j' \ poly_deg h)" assume "\ \ ps (Suc i) \ poly_deg h" hence "poly_deg h < \ ps (Suc i)" by simp from assms(2) le_refl have "\ ps (Suc (card X)) = \ ps" by (rule \_card_X) also from _ assms(5) have "\ \ poly_deg h" proof (rule standard_decompD) from assms(1) show "standard_decomp (\ ps) ps" by (rule \) qed finally have "\ ps (Suc (card X)) \ poly_deg h" . hence 1: "\ ps j \ poly_deg h" unfolding j_def by (rule LeastI) have "Suc i < j" proof (rule ccontr) assume "\ Suc i < j" hence "j \ Suc i" by simp hence "\ ps (Suc i) \ \ ps j" by (rule \_decreasing) also have "\ \ poly_deg h" by fact finally show False using \poly_deg h < \ ps (Suc i)\ by simp qed hence eq: "Suc (j - 1) = j" by simp note assms(1, 2) moreover from assms(3) have "Suc 0 \ j - 1" proof (rule le_trans) from \Suc i < j\ show "i \ j - 1" by simp qed moreover have "j - 1 \ card X" proof - have "j \ Suc (card X)" unfolding j_def by (rule Least_le) fact thus ?thesis by simp qed moreover from 1 have "\ ps (Suc (j - 1)) \ poly_deg h" by (simp only: eq) moreover have "poly_deg h < \ ps (j - 1)" proof (rule ccontr) assume "\ poly_deg h < \ ps (j - 1)" hence "\ ps (j - 1) \ poly_deg h" by simp hence "j \ j - 1" unfolding j_def by (rule Least_le) also have "\ < Suc (j - 1)" by simp finally show False by (simp only: eq) qed ultimately obtain h0 U0 where eq1: "{(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = poly_deg h} = {(h0, U0)}" and "card U0 = j - 1" by (rule lem_6_1_2) from assms(5) have "(h, U) \ {(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = poly_deg h}" by simp hence "(h, U) \ {(h0, U0)}" by (simp only: eq1) hence "U = U0" by simp hence "card U = j - 1" by (simp only: \card U0 = j - 1\) hence "i = j - 1" by (simp only: assms(6)) hence "Suc i = j" by (simp only: eq) with \Suc i < j\ show False by simp qed qualified fun shift_list :: "((('x \\<^sub>0 nat) \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}) \ 'x set) \ 'x \ _ list \ _ list" where "shift_list (h, U) x ps = ((punit.monom_mult 1 (Poly_Mapping.single x 1) h, U) # (h, U - {x}) # removeAll (h, U) ps)" declare shift_list.simps[simp del] lemma monomial_decomp_shift_list: assumes "monomial_decomp ps" and "hU \ set ps" shows "monomial_decomp (shift_list hU x ps)" proof - let ?x = "Poly_Mapping.single x (1::nat)" obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast with assms(2) have "(h, U) \ set ps" by simp with assms(1) have 1: "is_monomial h" and 2: "lcf h = 1" by (rule monomial_decompD)+ from this(1) have "monomial (lcf h) (lpp h) = h" by (rule punit.monomial_eq_itself) moreover define t where "t = lpp h" ultimately have "h = monomial 1 t" by (simp only: 2) hence "is_monomial (punit.monom_mult 1 ?x h)" and "lcf (punit.monom_mult 1 ?x h) = 1" by (simp_all add: punit.monom_mult_monomial monomial_is_monomial) with assms(1) 1 2 show ?thesis by (simp add: shift_list.simps monomial_decomp_def hU) qed lemma hom_decomp_shift_list: assumes "hom_decomp ps" and "hU \ set ps" shows "hom_decomp (shift_list hU x ps)" proof - let ?x = "Poly_Mapping.single x (1::nat)" obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast with assms(2) have "(h, U) \ set ps" by simp with assms(1) have 1: "homogeneous h" by (rule hom_decompD) hence "homogeneous (punit.monom_mult 1 ?x h)" by (simp only: homogeneous_monom_mult) with assms(1) 1 show ?thesis by (simp add: shift_list.simps hom_decomp_def hU) qed lemma valid_decomp_shift_list: assumes "valid_decomp X ps" and "(h, U) \ set ps" and "x \ U" shows "valid_decomp X (shift_list (h, U) x ps)" proof - let ?x = "Poly_Mapping.single x (1::nat)" from assms(1, 2) have "h \ P[X]" and "h \ 0" and "U \ X" by (rule valid_decompD)+ moreover from this(1) have "punit.monom_mult 1 ?x h \ P[X]" proof (intro Polys_closed_monom_mult PPs_closed_single) from \x \ U\ \U \ X\ show "x \ X" .. qed moreover from \U \ X\ have "U - {x} \ X" by blast ultimately show ?thesis using assms(1) \h \ 0\ by (simp add: valid_decomp_def punit.monom_mult_eq_zero_iff shift_list.simps) qed lemma standard_decomp_shift_list: assumes "standard_decomp k ps" and "(h1, U1) \ set ps" and "(h2, U2) \ set ps" and "poly_deg h1 = poly_deg h2" and "card U2 \ card U1" and "(h1, U1) \ (h2, U2)" and "x \ U2" shows "standard_decomp k (shift_list (h2, U2) x ps)" proof (rule standard_decompI) let ?p1 = "(punit.monom_mult 1 (Poly_Mapping.single x 1) h2, U2)" let ?p2 = "(h2, U2 - {x})" let ?qs = "removeAll (h2, U2) ps" fix h U assume "(h, U) \ set ((shift_list (h2, U2) x ps)\<^sub>+)" hence disj: "(h, U) = ?p1 \ ((h, U) = ?p2 \ U2 - {x} \ {}) \ (h, U) \ set (ps\<^sub>+)" by (auto simp: pos_decomp_def shift_list.simps split: if_split_asm) from assms(7) have "U2 \ {}" by blast with assms(3) have "(h2, U2) \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) with assms(1) have k_le: "k \ poly_deg h2" by (rule standard_decompD) let ?x = "Poly_Mapping.single x 1" from disj show "k \ poly_deg h" proof (elim disjE) assume "(h, U) = ?p1" hence h: "h = punit.monom_mult (1::'a) ?x h2" by simp note k_le also have "poly_deg h2 \ poly_deg h" by (cases "h2 = 0") (simp_all add: h poly_deg_monom_mult) finally show ?thesis . next assume "(h, U) = ?p2 \ U2 - {x} \ {}" with k_le show ?thesis by simp next assume "(h, U) \ set (ps\<^sub>+)" with assms(1) show ?thesis by (rule standard_decompD) qed fix d assume "k \ d" and "d \ poly_deg h" from disj obtain h' U' where 1: "(h', U') \ set (?p1 # ps)" and "poly_deg h' = d" and "card U \ card U'" proof (elim disjE) assume "(h, U) = ?p1" hence h: "h = punit.monom_mult 1 ?x h2" and "U = U2" by simp_all from \d \ poly_deg h\ have "d \ poly_deg h2 \ poly_deg h = d" by (cases "h2 = 0") (auto simp: h poly_deg_monom_mult deg_pm_single) thus ?thesis proof assume "d \ poly_deg h2" with assms(1) \(h2, U2) \ set (ps\<^sub>+)\ \k \ d\ obtain h' U' where "(h', U') \ set ps" and "poly_deg h' = d" and "card U2 \ card U'" by (rule standard_decompE) from this(1) have "(h', U') \ set (?p1 # ps)" by simp moreover note \poly_deg h' = d\ moreover from \card U2 \ card U'\ have "card U \ card U'" by (simp only: \U = U2\) ultimately show ?thesis .. next have "(h, U) \ set (?p1 # ps)" by (simp add: \(h, U) = ?p1\) moreover assume "poly_deg h = d" ultimately show ?thesis using le_refl .. qed next assume "(h, U) = ?p2 \ U2 - {x} \ {}" hence "h = h2" and U: "U = U2 - {x}" by simp_all from \d \ poly_deg h\ this(1) have "d \ poly_deg h2" by simp with assms(1) \(h2, U2) \ set (ps\<^sub>+)\ \k \ d\ obtain h' U' where "(h', U') \ set ps" and "poly_deg h' = d" and "card U2 \ card U'" by (rule standard_decompE) from this(1) have "(h', U') \ set (?p1 # ps)" by simp moreover note \poly_deg h' = d\ moreover from _ \card U2 \ card U'\ have "card U \ card U'" unfolding U by (rule le_trans) (metis Diff_empty card_Diff1_le card.infinite finite_Diff_insert order_refl) ultimately show ?thesis .. next assume "(h, U) \ set (ps\<^sub>+)" from assms(1) this \k \ d\ \d \ poly_deg h\ obtain h' U' where "(h', U') \ set ps" and "poly_deg h' = d" and "card U \ card U'" by (rule standard_decompE) from this(1) have "(h', U') \ set (?p1 # ps)" by simp thus ?thesis using \poly_deg h' = d\ \card U \ card U'\ .. qed show "\h' U'. (h', U') \ set (shift_list (h2, U2) x ps) \ poly_deg h' = d \ card U \ card U'" proof (cases "(h', U') = (h2, U2)") case True hence "h' = h2" and "U' = U2" by simp_all from assms(2, 6) have "(h1, U1) \ set (shift_list (h2, U2) x ps)" by (simp add: shift_list.simps) moreover from \poly_deg h' = d\ have "poly_deg h1 = d" by (simp only: \h' = h2\ assms(4)) moreover from \card U \ card U'\ assms(5) have "card U \ card U1" by (simp add: \U' = U2\) ultimately show ?thesis by blast next case False with 1 have "(h', U') \ set (shift_list (h2, U2) x ps)" by (auto simp: shift_list.simps) thus ?thesis using \poly_deg h' = d\ \card U \ card U'\ by blast qed qed lemma cone_decomp_shift_list: assumes "valid_decomp X ps" and "cone_decomp T ps" and "(h, U) \ set ps" and "x \ U" shows "cone_decomp T (shift_list (h, U) x ps)" proof - let ?p1 = "(punit.monom_mult 1 (Poly_Mapping.single x 1) h, U)" let ?p2 = "(h, U - {x})" let ?qs = "removeAll (h, U) ps" from assms(3) obtain ps1 ps2 where ps: "ps = ps1 @ (h, U) # ps2" and *: "(h, U) \ set ps1" by (meson split_list_first) have "count_list ps2 (h, U) = 0" proof (rule ccontr) from assms(1, 3) have "h \ 0" by (rule valid_decompD) assume "count_list ps2 (h, U) \ 0" hence "1 < count_list ps (h, U)" by (simp add: ps count_list_append) also have "\ \ count_list (map cone ps) (cone (h, U))" by (fact count_list_map_ge) finally have "1 < count_list (map cone ps) (cone (h, U))" . with cone_decompD have "cone (h, U) = {0}" proof (rule direct_decomp_repeated_eq_zero) fix s assume "s \ set (map cone ps)" thus "0 \ s" by (auto intro: zero_in_cone) qed (fact assms(2)) with tip_in_cone[of h U] have "h = 0" by simp with \h \ 0\ show False .. qed hence **: "(h, U) \ set ps2" by (simp add: count_list_eq_0_iff) - have "perm ps ((h, U) # ps1 @ ps2)" (is "perm _ ?ps") - by (rule perm_sym) (simp only: perm_append_Cons ps) + have "mset ps = mset ((h, U) # ps1 @ ps2)" (is "mset _ = mset ?ps") + by (simp add: ps) with assms(2) have "cone_decomp T ?ps" by (rule cone_decomp_perm) hence "direct_decomp T (map cone ?ps)" by (rule cone_decompD) hence "direct_decomp T (cone (h, U) # map cone (ps1 @ ps2))" by simp hence "direct_decomp T ((map cone (ps1 @ ps2)) @ [cone ?p1, cone ?p2])" proof (rule direct_decomp_direct_decomp) let ?x = "Poly_Mapping.single x (Suc 0)" have "direct_decomp (cone (h, insert x (U - {x}))) [cone (h, U - {x}), cone (monomial (1::'a) ?x * h, insert x (U - {x}))]" by (rule direct_decomp_cone_insert) simp with assms(4) show "direct_decomp (cone (h, U)) [cone ?p1, cone ?p2]" - by (simp add: insert_absorb times_monomial_left direct_decomp_perm perm.swap) + by (simp add: insert_absorb times_monomial_left direct_decomp_perm) qed hence "direct_decomp T (map cone (ps1 @ ps2 @ [?p1, ?p2]))" by simp hence "cone_decomp T (ps1 @ ps2 @ [?p1, ?p2])" by (rule cone_decompI) - moreover have "perm (ps1 @ ps2 @ [?p1, ?p2]) (?p1 # ?p2 # (ps1 @ ps2))" - proof - - have "ps1 @ ps2 @ [?p1, ?p2] = (ps1 @ ps2) @ [?p1, ?p2]" by simp - also have "perm \ ([?p1, ?p2] @ (ps1 @ ps2))" by (rule perm_append_swap) - also have "\ = ?p1 # ?p2 # (ps1 @ ps2)" by simp - finally show ?thesis . - qed + moreover have "mset (ps1 @ ps2 @ [?p1, ?p2]) = mset (?p1 # ?p2 # (ps1 @ ps2))" + by simp ultimately have "cone_decomp T (?p1 # ?p2 # (ps1 @ ps2))" by (rule cone_decomp_perm) also from * ** have "ps1 @ ps2 = removeAll (h, U) ps" by (simp add: remove1_append ps) finally show ?thesis by (simp only: shift_list.simps) qed subsection \Functions \shift\ and \exact\\ context fixes k m :: nat begin context fixes d :: nat begin definition shift2_inv :: "((('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ 'x set) list \ bool" where "shift2_inv qs \ valid_decomp X qs \ standard_decomp k qs \ exact_decomp (Suc m) qs \ (\d0 set qs. poly_deg (fst q) = d0 \ m < card (snd q)} \ 1)" fun shift1_inv :: "(((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::zero) \ 'x set) set) \ bool" where "shift1_inv (qs, B) \ B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)} \ shift2_inv qs" lemma shift2_invI: "valid_decomp X qs \ standard_decomp k qs \ exact_decomp (Suc m) qs \ (\d0. d0 < d \ card {q \ set qs. poly_deg (fst q) = d0 \ m < card (snd q)} \ 1) \ shift2_inv qs" by (simp add: shift2_inv_def) lemma shift2_invD: assumes "shift2_inv qs" shows "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" and "d0 < d \ card {q \ set qs. poly_deg (fst q) = d0 \ m < card (snd q)} \ 1" using assms by (simp_all add: shift2_inv_def) lemma shift1_invI: "B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)} \ shift2_inv qs \ shift1_inv (qs, B)" by simp lemma shift1_invD: assumes "shift1_inv (qs, B)" shows "B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}" and "shift2_inv qs" using assms by simp_all declare shift1_inv.simps[simp del] lemma shift1_inv_finite_snd: assumes "shift1_inv (qs, B)" shows "finite B" proof (rule finite_subset) from assms have "B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}" by (rule shift1_invD) also have "\ \ set qs" by blast finally show "B \ set qs" . qed (fact finite_set) lemma shift1_inv_some_snd: assumes "shift1_inv (qs, B)" and "1 < card B" and "(h, U) = (SOME b. b \ B \ card (snd b) = Suc m)" shows "(h, U) \ B" and "(h, U) \ set qs" and "poly_deg h = d" and "card U = Suc m" proof - define A where "A = {q \ B. card (snd q) = Suc m}" define Y where "Y = {q \ set qs. poly_deg (fst q) = d \ Suc m < card (snd q)}" from assms(1) have B: "B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}" and inv2: "shift2_inv qs" by (rule shift1_invD)+ have B': "B = A \ Y" by (auto simp: B A_def Y_def) have "finite A" proof (rule finite_subset) show "A \ B" unfolding A_def by blast next from assms(1) show "finite B" by (rule shift1_inv_finite_snd) qed moreover have "finite Y" proof (rule finite_subset) show "Y \ set qs" unfolding Y_def by blast qed (fact finite_set) moreover have "A \ Y = {}" by (auto simp: A_def Y_def) ultimately have "card (A \ Y) = card A + card Y" by (rule card_Un_disjoint) with assms(2) have "1 < card A + card Y" by (simp only: B') thm card_le_Suc0_iff_eq[OF \finite Y\] moreover have "card Y \ 1" unfolding One_nat_def card_le_Suc0_iff_eq[OF \finite Y\] proof (intro ballI) fix q1 q2 :: "(('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set" obtain h1 U1 where q1: "q1 = (h1, U1)" using prod.exhaust by blast obtain h2 U2 where q2: "q2 = (h2, U2)" using prod.exhaust by blast assume "q1 \ Y" hence "(h1, U1) \ set qs" and "poly_deg h1 = d" and "Suc m < card U1" by (simp_all add: q1 Y_def) assume "q2 \ Y" hence "(h2, U2) \ set qs" and "poly_deg h2 = d" and "Suc m < card U2" by (simp_all add: q2 Y_def) from this(2) have "poly_deg h1 = poly_deg h2" by (simp only: \poly_deg h1 = d\) from inv2 have "exact_decomp (Suc m) qs" by (rule shift2_invD) thus "q1 = q2" unfolding q1 q2 by (rule exact_decompD) fact+ qed ultimately have "0 < card A" by simp hence "A \ {}" by auto then obtain a where "a \ A" by blast have "(h, U) \ B \ card (snd (h, U)) = Suc m" unfolding assms(3) proof (rule someI) from \a \ A\ show "a \ B \ card (snd a) = Suc m" by (simp add: A_def) qed thus "(h, U) \ B" and "card U = Suc m" by simp_all from this(1) show "(h, U) \ set qs" and "poly_deg h = d" by (simp_all add: B) qed lemma shift1_inv_preserved: assumes "shift1_inv (qs, B)" and "1 < card B" and "(h, U) = (SOME b. b \ B \ card (snd b) = Suc m)" and "x = (SOME y. y \ U)" shows "shift1_inv (shift_list (h, U) x qs, B - {(h, U)})" proof - let ?p1 = "(punit.monom_mult 1 (Poly_Mapping.single x 1) h, U)" let ?p2 = "(h, U - {x})" let ?qs = "removeAll (h, U) qs" let ?B = "B - {(h, U)}" from assms(1, 2, 3) have "(h, U) \ B" and "(h, U) \ set qs" and deg_h: "poly_deg h = d" and card_U: "card U = Suc m" by (rule shift1_inv_some_snd)+ from card_U have "U \ {}" by auto then obtain y where "y \ U" by blast hence "x \ U" unfolding assms(4) by (rule someI) with card_U have card_Ux: "card (U - {x}) = m" by (metis card_Diff_singleton card.infinite diff_Suc_1 nat.simps(3)) from assms(1) have B: "B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}" and inv2: "shift2_inv qs" by (rule shift1_invD)+ from inv2 have valid_qs: "valid_decomp X qs" by (rule shift2_invD) hence "h \ 0" using \(h, U) \ set qs\ by (rule valid_decompD) show ?thesis proof (intro shift1_invI shift2_invI) show "?B = {q \ set (shift_list (h, U) x qs). poly_deg (fst q) = d \ m < card (snd q)}" (is "_ = ?C") proof (rule Set.set_eqI) fix b show "b \ ?B \ b \ ?C" proof assume "b \ ?C" hence "b \ insert ?p1 (insert ?p2 (set ?qs))" and b1: "poly_deg (fst b) = d" and b2: "m < card (snd b)" by (simp_all add: shift_list.simps) from this(1) show "b \ ?B" proof (elim insertE) assume "b = ?p1" with \h \ 0\ have "poly_deg (fst b) = Suc d" by (simp add: poly_deg_monom_mult deg_pm_single deg_h) thus ?thesis by (simp add: b1) next assume "b = ?p2" hence "card (snd b) = m" by (simp add: card_Ux) with b2 show ?thesis by simp next assume "b \ set ?qs" with b1 b2 show ?thesis by (auto simp: B) qed qed (auto simp: B shift_list.simps) qed next from valid_qs \(h, U) \ set qs\ \x \ U\ show "valid_decomp X (shift_list (h, U) x qs)" by (rule valid_decomp_shift_list) next from inv2 have std: "standard_decomp k qs" by (rule shift2_invD) have "?B \ {}" proof assume "?B = {}" hence "B \ {(h, U)}" by simp with _ have "card B \ card {(h, U)}" by (rule card_mono) simp with assms(2) show False by simp qed then obtain h' U' where "(h', U') \ B" and "(h', U') \ (h, U)" by auto from this(1) have "(h', U') \ set qs" and "poly_deg h' = d" and "Suc m \ card U'" by (simp_all add: B) note std this(1) \(h, U) \ set qs\ moreover from \poly_deg h' = d\ have "poly_deg h' = poly_deg h" by (simp only: deg_h) moreover from \Suc m \ card U'\ have "card U \ card U'" by (simp only: card_U) ultimately show "standard_decomp k (shift_list (h, U) x qs)" by (rule standard_decomp_shift_list) fact+ next from inv2 have exct: "exact_decomp (Suc m) qs" by (rule shift2_invD) show "exact_decomp (Suc m) (shift_list (h, U) x qs)" proof (rule exact_decompI) fix h' U' assume "(h', U') \ set (shift_list (h, U) x qs)" hence *: "(h', U') \ insert ?p1 (insert ?p2 (set ?qs))" by (simp add: shift_list.simps) thus "h' \ P[X]" proof (elim insertE) assume "(h', U') = ?p1" hence h': "h' = punit.monom_mult 1 (Poly_Mapping.single x 1) h" by simp from exct \(h, U) \ set qs\ have "U \ X" by (rule exact_decompD) with \x \ U\ have "x \ X" .. hence "Poly_Mapping.single x 1 \ .[X]" by (rule PPs_closed_single) moreover from exct \(h, U) \ set qs\ have "h \ P[X]" by (rule exact_decompD) ultimately show ?thesis unfolding h' by (rule Polys_closed_monom_mult) next assume "(h', U') = ?p2" hence "h' = h" by simp also from exct \(h, U) \ set qs\ have "\ \ P[X]" by (rule exact_decompD) finally show ?thesis . next assume "(h', U') \ set ?qs" hence "(h', U') \ set qs" by simp with exct show ?thesis by (rule exact_decompD) qed from * show "U' \ X" proof (elim insertE) assume "(h', U') = ?p1" hence "U' = U" by simp also from exct \(h, U) \ set qs\ have "\ \ X" by (rule exact_decompD) finally show ?thesis . next assume "(h', U') = ?p2" hence "U' = U - {x}" by simp also have "\ \ U" by blast also from exct \(h, U) \ set qs\ have "\ \ X" by (rule exact_decompD) finally show ?thesis . next assume "(h', U') \ set ?qs" hence "(h', U') \ set qs" by simp with exct show ?thesis by (rule exact_decompD) qed next fix h1 h2 U1 U2 assume "(h1, U1) \ set (shift_list (h, U) x qs)" and "Suc m < card U1" hence "(h1, U1) \ set qs" using card_U card_Ux by (auto simp: shift_list.simps) assume "(h2, U2) \ set (shift_list (h, U) x qs)" and "Suc m < card U2" hence "(h2, U2) \ set qs" using card_U card_Ux by (auto simp: shift_list.simps) assume "poly_deg h1 = poly_deg h2" from exct show "(h1, U1) = (h2, U2)" by (rule exact_decompD) fact+ qed next fix d0 assume "d0 < d" have "finite {q \ set qs. poly_deg (fst q) = d0 \ m < card (snd q)}" (is "finite ?A") by auto moreover have "{q \ set (shift_list (h, U) x qs). poly_deg (fst q) = d0 \ m < card (snd q)} \ ?A" (is "?C \ _") proof fix q assume "q \ ?C" hence "q = ?p1 \ q = ?p2 \ q \ set ?qs" and 1: "poly_deg (fst q) = d0" and 2: "m < card (snd q)" by (simp_all add: shift_list.simps) from this(1) show "q \ ?A" proof (elim disjE) assume "q = ?p1" with \h \ 0\ have "d \ poly_deg (fst q)" by (simp add: poly_deg_monom_mult deg_h) with \d0 < d\ show ?thesis by (simp only: 1) next assume "q = ?p2" hence "d \ poly_deg (fst q)" by (simp add: deg_h) with \d0 < d\ show ?thesis by (simp only: 1) next assume "q \ set ?qs" with 1 2 show ?thesis by simp qed qed ultimately have "card ?C \ card ?A" by (rule card_mono) also from inv2 \d0 < d\ have "\ \ 1" by (rule shift2_invD) finally show "card ?C \ 1" . qed qed function (domintros) shift1 :: "(((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) set) \ (((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}) \ 'x set) set)" where "shift1 (qs, B) = (if 1 < card B then let (h, U) = SOME b. b \ B \ card (snd b) = Suc m; x = SOME y. y \ U in shift1 (shift_list (h, U) x qs, B - {(h, U)}) else (qs, B))" by auto lemma shift1_domI: assumes "shift1_inv args" shows "shift1_dom args" proof - from wf_measure[of "card \ snd"] show ?thesis using assms proof (induct) case (less args) obtain qs B where args: "args = (qs, B)" using prod.exhaust by blast have IH: "shift1_dom (qs0, B0)" if "card B0 < card B" and "shift1_inv (qs0, B0)" for qs0 and B0::"((_ \\<^sub>0 'a) \ _) set" using _ that(2) proof (rule less) from that(1) show "((qs0, B0), args) \ measure (card \ snd)" by (simp add: args) qed from less(2) have inv: "shift1_inv (qs, B)" by (simp only: args) show ?case unfolding args proof (rule shift1.domintros) fix h U assume hU: "(h, U) = (SOME b. b \ B \ card (snd b) = Suc m)" define x where "x = (SOME y. y \ U)" assume "Suc 0 < card B" hence "1 < card B" by simp have "shift1_dom (shift_list (h, U) x qs, B - {(h, U)})" proof (rule IH) from inv have "finite B" by (rule shift1_inv_finite_snd) moreover from inv \1 < card B\ hU have "(h, U) \ B" by (rule shift1_inv_some_snd) ultimately show "card (B - {(h, U)}) < card B" by (rule card_Diff1_less) next from inv \1 < card B\ hU x_def show "shift1_inv (shift_list (h, U) x qs, (B - {(h, U)}))" by (rule shift1_inv_preserved) qed thus "shift1_dom (shift_list (SOME b. b \ B \ card (snd b) = Suc m) (SOME y. y \ U) qs, B - {SOME b. b \ B \ card (snd b) = Suc m})" by (simp add: hU x_def) qed qed qed lemma shift1_induct [consumes 1, case_names base step]: assumes "shift1_inv args" assumes "\qs B. shift1_inv (qs, B) \ card B \ 1 \ P (qs, B) (qs, B)" assumes "\qs B h U x. shift1_inv (qs, B) \ 1 < card B \ (h, U) = (SOME b. b \ B \ card (snd b) = Suc m) \ x = (SOME y. y \ U) \ finite U \ x \ U \ card (U - {x}) = m \ P (shift_list (h, U) x qs, B - {(h, U)}) (shift1 (shift_list (h, U) x qs, B - {(h, U)})) \ P (qs, B) (shift1 (shift_list (h, U) x qs, B - {(h, U)}))" shows "P args (shift1 args)" proof - from assms(1) have "shift1_dom args" by (rule shift1_domI) thus ?thesis using assms(1) proof (induct args rule: shift1.pinduct) case step: (1 qs B) obtain h U where hU: "(h, U) = (SOME b. b \ B \ card (snd b) = Suc m)" by (smt prod.exhaust) define x where "x = (SOME y. y \ U)" show ?case proof (simp add: shift1.psimps[OF step.hyps(1)] flip: hU x_def del: One_nat_def, intro conjI impI) let ?args = "(shift_list (h, U) x qs, B - {(h, U)})" assume "1 < card B" with step.prems have card_U: "card U = Suc m" using hU by (rule shift1_inv_some_snd) from card_U have "finite U" using card.infinite by fastforce from card_U have "U \ {}" by auto then obtain y where "y \ U" by blast hence "x \ U" unfolding x_def by (rule someI) with step.prems \1 < card B\ hU x_def \finite U\ show "P (qs, B) (shift1 ?args)" proof (rule assms(3)) from \finite U\ \x \ U\ show "card (U - {x}) = m" by (simp add: card_U) next from \1 < card B\ refl hU x_def show "P ?args (shift1 ?args)" proof (rule step.hyps) from step.prems \1 < card B\ hU x_def show "shift1_inv ?args" by (rule shift1_inv_preserved) qed qed next assume "\ 1 < card B" hence "card B \ 1" by simp with step.prems show "P (qs, B) (qs, B)" by (rule assms(2)) qed qed qed lemma shift1_1: assumes "shift1_inv args" and "d0 \ d" shows "card {q \ set (fst (shift1 args)). poly_deg (fst q) = d0 \ m < card (snd q)} \ 1" using assms(1) proof (induct args rule: shift1_induct) case (base qs B) from assms(2) have "d0 < d \ d0 = d" by auto thus ?case proof from base(1) have "shift2_inv qs" by (rule shift1_invD) moreover assume "d0 < d" ultimately show ?thesis unfolding fst_conv by (rule shift2_invD) next assume "d0 = d" from base(1) have "B = {q \ set (fst (qs, B)). poly_deg (fst q) = d0 \ m < card (snd q)}" unfolding fst_conv \d0 = d\ by (rule shift1_invD) with base(2) show ?thesis by simp qed qed lemma shift1_2: "shift1_inv args \ card {q \ set (fst (shift1 args)). m < card (snd q)} \ card {q \ set (fst args). m < card (snd q)}" proof (induct args rule: shift1_induct) case (base qs B) show ?case .. next case (step qs B h U x) let ?x = "Poly_Mapping.single x (1::nat)" let ?p1 = "(punit.monom_mult 1 ?x h, U)" let ?A = "{q \ set qs. m < card (snd q)}" from step(1-3) have card_U: "card U = Suc m" and "(h, U) \ set qs" by (rule shift1_inv_some_snd)+ from step(1) have "shift2_inv qs" by (rule shift1_invD) hence "valid_decomp X qs" by (rule shift2_invD) hence "h \ 0" using \(h, U) \ set qs\ by (rule valid_decompD) have fin1: "finite ?A" by auto hence fin2: "finite (insert ?p1 ?A)" by simp from \(h, U) \ set qs\ have hU_in: "(h, U) \ insert ?p1 ?A" by (simp add: card_U) have "?p1 \ (h, U)" proof assume "?p1 = (h, U)" hence "lpp (punit.monom_mult 1 ?x h) = lpp h" by simp with \h \ 0\ show False by (simp add: punit.lt_monom_mult monomial_0_iff) qed let ?qs = "shift_list (h, U) x qs" have "{q \ set (fst (?qs, B - {(h, U)})). m < card (snd q)} = (insert ?p1 ?A) - {(h, U)}" using step(7) card_U \?p1 \ (h, U)\ by (fastforce simp: shift_list.simps) also from fin2 hU_in have "card \ = card (insert ?p1 ?A) - 1" by (simp add: card_Diff_singleton_if) also from fin1 have "\ \ Suc (card ?A) - 1" by (simp add: card_insert_if) also have "\ = card {q \ set (fst (qs, B)). m < card (snd q)}" by simp finally have "card {q \ set (fst (?qs, B - {(h, U)})). m < card (snd q)} \ card {q \ set (fst (qs, B)). m < card (snd q)}" . with step(8) show ?case by (rule le_trans) qed lemma shift1_3: "shift1_inv args \ cone_decomp T (fst args) \ cone_decomp T (fst (shift1 args))" proof (induct args rule: shift1_induct) case (base qs B) from base(3) show ?case . next case (step qs B h U x) from step.hyps(1) have "shift2_inv qs" by (rule shift1_invD) hence "valid_decomp X qs" by (rule shift2_invD) moreover from step.prems have "cone_decomp T qs" by (simp only: fst_conv) moreover from step.hyps(1-3) have "(h, U) \ set qs" by (rule shift1_inv_some_snd) ultimately have "cone_decomp T (fst (shift_list (h, U) x qs, B - {(h, U)}))" unfolding fst_conv using step.hyps(6) by (rule cone_decomp_shift_list) thus ?case by (rule step.hyps(8)) qed lemma shift1_4: "shift1_inv args \ Max (poly_deg ` fst ` set (fst args)) \ Max (poly_deg ` fst ` set (fst (shift1 args)))" proof (induct args rule: shift1_induct) case (base qs B) show ?case .. next case (step qs B h U x) let ?x = "Poly_Mapping.single x 1" let ?p1 = "(punit.monom_mult 1 ?x h, U)" let ?qs = "shift_list (h, U) x qs" from step(1) have "B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}" and inv2: "shift2_inv qs" by (rule shift1_invD)+ from this(1) have "B \ set qs" by auto with step(2) have "set qs \ {}" by auto from finite_set have fin: "finite (poly_deg ` fst ` set ?qs)" by (intro finite_imageI) have "Max (poly_deg ` fst ` set (fst (qs, B))) \ Max (poly_deg ` fst ` set (fst (?qs, B - {(h, U)})))" unfolding fst_conv proof (rule Max.boundedI) from finite_set show "finite (poly_deg ` fst ` set qs)" by (intro finite_imageI) next from \set qs \ {}\ show "poly_deg ` fst ` set qs \ {}" by simp next fix a assume "a \ poly_deg ` fst ` set qs" then obtain q where "q \ set qs" and a: "a = poly_deg (fst q)" by blast show "a \ Max (poly_deg ` fst ` set ?qs)" proof (cases "q = (h, U)") case True hence "a \ poly_deg (fst ?p1)" by (cases "h = 0") (simp_all add: a poly_deg_monom_mult) also from fin have "\ \ Max (poly_deg ` fst ` set ?qs)" proof (rule Max_ge) have "?p1 \ set ?qs" by (simp add: shift_list.simps) thus "poly_deg (fst ?p1) \ poly_deg ` fst ` set ?qs" by (intro imageI) qed finally show ?thesis . next case False with \q \ set qs\ have "q \ set ?qs" by (simp add: shift_list.simps) hence "a \ poly_deg ` fst ` set ?qs" unfolding a by (intro imageI) with fin show ?thesis by (rule Max_ge) qed qed thus ?case using step(8) by (rule le_trans) qed lemma shift1_5: "shift1_inv args \ fst (shift1 args) = [] \ fst args = []" proof (induct args rule: shift1_induct) case (base qs B) show ?case .. next case (step qs B h U x) let ?p1 = "(punit.monom_mult 1 (Poly_Mapping.single x 1) h, U)" let ?qs = "shift_list (h, U) x qs" from step(1) have "B = {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}" and inv2: "shift2_inv qs" by (rule shift1_invD)+ from this(1) have "B \ set qs" by auto with step(2) have "qs \ []" by auto moreover have "fst (shift1 (?qs, B - {(h, U)})) \ []" by (simp add: step.hyps(8) del: One_nat_def) (simp add: shift_list.simps) ultimately show ?case by simp qed lemma shift1_6: "shift1_inv args \ monomial_decomp (fst args) \ monomial_decomp (fst (shift1 args))" proof (induct args rule: shift1_induct) case (base qs B) from base(3) show ?case . next case (step qs B h U x) from step(1-3) have "(h, U) \ set qs" by (rule shift1_inv_some_snd) with step.prems have "monomial_decomp (fst (shift_list (h, U) x qs, B - {(h, U)}))" unfolding fst_conv by (rule monomial_decomp_shift_list) thus ?case by (rule step.hyps) qed lemma shift1_7: "shift1_inv args \ hom_decomp (fst args) \ hom_decomp (fst (shift1 args))" proof (induct args rule: shift1_induct) case (base qs B) from base(3) show ?case . next case (step qs B h U x) from step(1-3) have "(h, U) \ set qs" by (rule shift1_inv_some_snd) with step.prems have "hom_decomp (fst (shift_list (h, U) x qs, B - {(h, U)}))" unfolding fst_conv by (rule hom_decomp_shift_list) thus ?case by (rule step.hyps) qed end lemma shift2_inv_preserved: assumes "shift2_inv d qs" shows "shift2_inv (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})))" proof - define args where "args = (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})" from refl assms have inv1: "shift1_inv d args" unfolding args_def by (rule shift1_invI) hence "shift1_inv d (shift1 args)" by (induct args rule: shift1_induct) hence "shift1_inv d (fst (shift1 args), snd (shift1 args))" by simp hence "shift2_inv d (fst (shift1 args))" by (rule shift1_invD) hence "valid_decomp X (fst (shift1 args))" and "standard_decomp k (fst (shift1 args))" and "exact_decomp (Suc m) (fst (shift1 args))" by (rule shift2_invD)+ thus "shift2_inv (Suc d) (fst (shift1 args))" proof (rule shift2_invI) fix d0 assume "d0 < Suc d" hence "d0 \ d" by simp with inv1 show "card {q \ set (fst (shift1 args)). poly_deg (fst q) = d0 \ m < card (snd q)} \ 1" by (rule shift1_1) qed qed function shift2 :: "nat \ nat \ ((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}) \ 'x set) list" where "shift2 c d qs = (if c \ d then qs else shift2 c (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}))))" by auto termination proof show "wf (measure (\(c, d, _). c - d))" by (fact wf_measure) qed simp lemma shift2_1: "shift2_inv d qs \ shift2_inv c (shift2 c d qs)" proof (induct c d qs rule: shift2.induct) case IH: (1 c d qs) show ?case proof (subst shift2.simps, simp del: shift2.simps, intro conjI impI) assume "c \ d" show "shift2_inv c qs" proof (rule shift2_invI) from IH(2) show "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" by (rule shift2_invD)+ next fix d0 assume "d0 < c" hence "d0 < d" using \c \ d\ by (rule less_le_trans) with IH(2) show "card {q \ set qs. poly_deg (fst q) = d0 \ m < card (snd q)} \ 1" by (rule shift2_invD) qed next assume "\ c \ d" thus "shift2_inv c (shift2 c (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}))))" proof (rule IH) from IH(2) show "shift2_inv (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})))" by (rule shift2_inv_preserved) qed qed qed lemma shift2_2: "shift2_inv d qs \ card {q \ set (shift2 c d qs). m < card (snd q)} \ card {q \ set qs. m < card (snd q)}" proof (induct c d qs rule: shift2.induct) case IH: (1 c d qs) let ?A = "{q \ set (shift2 c (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})))). m < card (snd q)}" show ?case proof (subst shift2.simps, simp del: shift2.simps, intro impI) assume "\ c \ d" hence "card ?A \ card {q \ set (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}))). m < card (snd q)}" proof (rule IH) show "shift2_inv (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})))" using IH(2) by (rule shift2_inv_preserved) qed also have "\ \ card {q \ set (fst (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})). m < card (snd q)}" using refl IH(2) by (intro shift1_2 shift1_invI) finally show "card ?A \ card {q \ set qs. m < card (snd q)}" by (simp only: fst_conv) qed qed lemma shift2_3: "shift2_inv d qs \ cone_decomp T qs \ cone_decomp T (shift2 c d qs)" proof (induct c d qs rule: shift2.induct) case IH: (1 c d qs) have inv2: "shift2_inv (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})))" using IH(2) by (rule shift2_inv_preserved) show ?case proof (subst shift2.simps, simp add: IH.prems del: shift2.simps, intro impI) assume "\ c \ d" moreover note inv2 moreover have "cone_decomp T (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})))" proof (rule shift1_3) from refl IH(2) show "shift1_inv d (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})" by (rule shift1_invI) qed (simp add: IH.prems) ultimately show "cone_decomp T (shift2 c (Suc d) (fst (shift1 (qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)}))))" by (rule IH) qed qed lemma shift2_4: "shift2_inv d qs \ Max (poly_deg ` fst ` set qs) \ Max (poly_deg ` fst ` set (shift2 c d qs))" proof (induct c d qs rule: shift2.induct) case IH: (1 c d qs) let ?args = "(qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})" show ?case proof (subst shift2.simps, simp del: shift2.simps, intro impI) assume "\ c \ d" have "Max (poly_deg ` fst ` set (fst ?args)) \ Max (poly_deg ` fst ` set (fst (shift1 ?args)))" using refl IH(2) by (intro shift1_4 shift1_invI) also from \\ c \ d\ have "\ \ Max (poly_deg ` fst ` set (shift2 c (Suc d) (fst (shift1 ?args))))" proof (rule IH) from IH(2) show "shift2_inv (Suc d) (fst (shift1 ?args))" by (rule shift2_inv_preserved) qed finally show "Max (poly_deg ` fst ` set qs) \ Max (poly_deg ` fst ` set (shift2 c (Suc d) (fst (shift1 ?args))))" by (simp only: fst_conv) qed qed lemma shift2_5: "shift2_inv d qs \ shift2 c d qs = [] \ qs = []" proof (induct c d qs rule: shift2.induct) case IH: (1 c d qs) let ?args = "(qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})" show ?case proof (subst shift2.simps, simp del: shift2.simps, intro impI) assume "\ c \ d" hence "shift2 c (Suc d) (fst (shift1 ?args)) = [] \ fst (shift1 ?args) = []" proof (rule IH) from IH(2) show "shift2_inv (Suc d) (fst (shift1 ?args))" by (rule shift2_inv_preserved) qed also from refl IH(2) have "\ \ fst ?args = []" by (intro shift1_5 shift1_invI) finally show "shift2 c (Suc d) (fst (shift1 ?args)) = [] \ qs = []" by (simp only: fst_conv) qed qed lemma shift2_6: "shift2_inv d qs \ monomial_decomp qs \ monomial_decomp (shift2 c d qs)" proof (induct c d qs rule: shift2.induct) case IH: (1 c d qs) let ?args = "(qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})" show ?case proof (subst shift2.simps, simp del: shift2.simps, intro conjI impI IH) from IH(2) show "shift2_inv (Suc d) (fst (shift1 ?args))" by (rule shift2_inv_preserved) next from refl IH(2) have "shift1_inv d ?args" by (rule shift1_invI) moreover from IH(3) have "monomial_decomp (fst ?args)" by simp ultimately show "monomial_decomp (fst (shift1 ?args))" by (rule shift1_6) qed qed lemma shift2_7: "shift2_inv d qs \ hom_decomp qs \ hom_decomp (shift2 c d qs)" proof (induct c d qs rule: shift2.induct) case IH: (1 c d qs) let ?args = "(qs, {q \ set qs. poly_deg (fst q) = d \ m < card (snd q)})" show ?case proof (subst shift2.simps, simp del: shift2.simps, intro conjI impI IH) from IH(2) show "shift2_inv (Suc d) (fst (shift1 ?args))" by (rule shift2_inv_preserved) next from refl IH(2) have "shift1_inv d ?args" by (rule shift1_invI) moreover from IH(3) have "hom_decomp (fst ?args)" by simp ultimately show "hom_decomp (fst (shift1 ?args))" by (rule shift1_7) qed qed definition shift :: "((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}) \ 'x set) list" where "shift qs = shift2 (k + card {q \ set qs. m < card (snd q)}) k qs" lemma shift2_inv_init: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" shows "shift2_inv k qs" using assms proof (rule shift2_invI) fix d0 assume "d0 < k" have "{q \ set qs. poly_deg (fst q) = d0 \ m < card (snd q)} = {}" proof - { fix q assume "q \ set qs" obtain h U where q: "q = (h, U)" using prod.exhaust by blast assume "poly_deg (fst q) = d0" and "m < card (snd q)" hence "poly_deg h < k" and "m < card U" using \d0 < k\ by (simp_all add: q) from this(2) have "U \ {}" by auto with \q \ set qs\ have "(h, U) \ set (qs\<^sub>+)" by (simp add: q pos_decomp_def) with assms(2) have "k \ poly_deg h" by (rule standard_decompD) with \poly_deg h < k\ have False by simp } thus ?thesis by blast qed thus "card {q \ set qs. poly_deg (fst q) = d0 \ m < card (snd q)} \ 1" by (simp only: card.empty) qed lemma shift: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" shows "valid_decomp X (shift qs)" and "standard_decomp k (shift qs)" and "exact_decomp m (shift qs)" proof - define c where "c = card {q \ set qs. m < card (snd q)}" define A where "A = {q \ set (shift qs). m < card (snd q)}" from assms have "shift2_inv k qs" by (rule shift2_inv_init) hence inv2: "shift2_inv (k + c) (shift qs)" and "card A \ c" unfolding shift_def c_def A_def by (rule shift2_1, rule shift2_2) from inv2 have fin: "valid_decomp X (shift qs)" and std: "standard_decomp k (shift qs)" and exct: "exact_decomp (Suc m) (shift qs)" by (rule shift2_invD)+ show "valid_decomp X (shift qs)" and "standard_decomp k (shift qs)" by fact+ have "finite A" by (auto simp: A_def) show "exact_decomp m (shift qs)" proof (rule exact_decompI) fix h U assume "(h, U) \ set (shift qs)" with exct show "h \ P[X]" and "U \ X" by (rule exact_decompD)+ next fix h1 h2 U1 U2 assume 1: "(h1, U1) \ set (shift qs)" and 2: "(h2, U2) \ set (shift qs)" assume 3: "poly_deg h1 = poly_deg h2" and 4: "m < card U1" and 5: "m < card U2" from 5 have "U2 \ {}" by auto with 2 have "(h2, U2) \ set ((shift qs)\<^sub>+)" by (simp add: pos_decomp_def) let ?C = "{q \ set (shift qs). poly_deg (fst q) = poly_deg h2 \ m < card (snd q)}" define B where "B = {q \ A. k \ poly_deg (fst q) \ poly_deg (fst q) \ poly_deg h2}" have "Suc (poly_deg h2) - k \ card B" proof - have "B = (\d0\{k..poly_deg h2}. {q \ A. poly_deg (fst q) = d0})" by (auto simp: B_def) also have "card \ = (\d0=k..poly_deg h2. card {q \ A. poly_deg (fst q) = d0})" proof (intro card_UN_disjoint ballI impI) fix d0 from _ \finite A\ show "finite {q \ A. poly_deg (fst q) = d0}" by (rule finite_subset) blast next fix d0 d1 :: nat assume "d0 \ d1" thus "{q \ A. poly_deg (fst q) = d0} \ {q \ A. poly_deg (fst q) = d1} = {}" by blast qed (fact finite_atLeastAtMost) also have "\ \ (\d0=k..poly_deg h2. 1)" proof (rule sum_mono) fix d0 assume "d0 \ {k..poly_deg h2}" hence "k \ d0" and "d0 \ poly_deg h2" by simp_all with std \(h2, U2) \ set ((shift qs)\<^sub>+)\ obtain h' U' where "(h', U') \ set (shift qs)" and "poly_deg h' = d0" and "card U2 \ card U'" by (rule standard_decompE) from 5 this(3) have "m < card U'" by (rule less_le_trans) with \(h', U') \ set (shift qs)\ have "(h', U') \ {q \ A. poly_deg (fst q) = d0}" by (simp add: A_def \poly_deg h' = d0\) hence "{q \ A. poly_deg (fst q) = d0} \ {}" by blast moreover from _ \finite A\ have "finite {q \ A. poly_deg (fst q) = d0}" by (rule finite_subset) blast ultimately show "1 \ card {q \ A. poly_deg (fst q) = d0}" by (simp add: card_gt_0_iff Suc_le_eq) qed also have "(\d0=k..poly_deg h2. 1) = Suc (poly_deg h2) - k" by auto finally show ?thesis . qed also from \finite A\ _ have "\ \ card A" by (rule card_mono) (auto simp: B_def) also have "\ \ c" by fact finally have "poly_deg h2 < k + c" by simp with inv2 have "card ?C \ 1" by (rule shift2_invD) have "finite ?C" by auto moreover note \card ?C \ 1\ moreover from 1 3 4 have "(h1, U1) \ ?C" by simp moreover from 2 5 have "(h2, U2) \ ?C" by simp ultimately show "(h1, U1) = (h2, U2)" by (auto simp: card_le_Suc0_iff_eq) qed qed lemma monomial_decomp_shift: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" and "monomial_decomp qs" shows "monomial_decomp (shift qs)" proof - from assms(1, 2, 3) have "shift2_inv k qs" by (rule shift2_inv_init) thus ?thesis unfolding shift_def using assms(4) by (rule shift2_6) qed lemma hom_decomp_shift: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" and "hom_decomp qs" shows "hom_decomp (shift qs)" proof - from assms(1, 2, 3) have "shift2_inv k qs" by (rule shift2_inv_init) thus ?thesis unfolding shift_def using assms(4) by (rule shift2_7) qed lemma cone_decomp_shift: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" and "cone_decomp T qs" shows "cone_decomp T (shift qs)" proof - from assms(1, 2, 3) have "shift2_inv k qs" by (rule shift2_inv_init) thus ?thesis unfolding shift_def using assms(4) by (rule shift2_3) qed lemma Max_shift_ge: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" shows "Max (poly_deg ` fst ` set qs) \ Max (poly_deg ` fst ` set (shift qs))" proof - from assms(1-3) have "shift2_inv k qs" by (rule shift2_inv_init) thus ?thesis unfolding shift_def by (rule shift2_4) qed lemma shift_Nil_iff: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp (Suc m) qs" shows "shift qs = [] \ qs = []" proof - from assms(1-3) have "shift2_inv k qs" by (rule shift2_inv_init) thus ?thesis unfolding shift_def by (rule shift2_5) qed end primrec exact_aux :: "nat \ nat \ ((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}) \ 'x set) list" where "exact_aux k 0 qs = qs" | "exact_aux k (Suc m) qs = exact_aux k m (shift k m qs)" lemma exact_aux: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp m qs" shows "valid_decomp X (exact_aux k m qs)" (is ?thesis1) and "standard_decomp k (exact_aux k m qs)" (is ?thesis2) and "exact_decomp 0 (exact_aux k m qs)" (is ?thesis3) proof - from assms have "?thesis1 \ ?thesis2 \ ?thesis3" proof (induct m arbitrary: qs) case 0 thus ?case by simp next case (Suc m) let ?qs = "shift k m qs" have "valid_decomp X (exact_aux k m ?qs) \ standard_decomp k (exact_aux k m ?qs) \ exact_decomp 0 (exact_aux k m ?qs)" proof (rule Suc) from Suc.prems show "valid_decomp X ?qs" and "standard_decomp k ?qs" and "exact_decomp m ?qs" by (rule shift)+ qed thus ?case by simp qed thus ?thesis1 and ?thesis2 and ?thesis3 by simp_all qed lemma monomial_decomp_exact_aux: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp m qs" and "monomial_decomp qs" shows "monomial_decomp (exact_aux k m qs)" using assms proof (induct m arbitrary: qs) case 0 thus ?case by simp next case (Suc m) let ?qs = "shift k m qs" have "monomial_decomp (exact_aux k m ?qs)" proof (rule Suc) show "valid_decomp X ?qs" and "standard_decomp k ?qs" and "exact_decomp m ?qs" using Suc.prems(1, 2, 3) by (rule shift)+ next from Suc.prems show "monomial_decomp ?qs" by (rule monomial_decomp_shift) qed thus ?case by simp qed lemma hom_decomp_exact_aux: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp m qs" and "hom_decomp qs" shows "hom_decomp (exact_aux k m qs)" using assms proof (induct m arbitrary: qs) case 0 thus ?case by simp next case (Suc m) let ?qs = "shift k m qs" have "hom_decomp (exact_aux k m ?qs)" proof (rule Suc) show "valid_decomp X ?qs" and "standard_decomp k ?qs" and "exact_decomp m ?qs" using Suc.prems(1, 2, 3) by (rule shift)+ next from Suc.prems show "hom_decomp ?qs" by (rule hom_decomp_shift) qed thus ?case by simp qed lemma cone_decomp_exact_aux: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp m qs" and "cone_decomp T qs" shows "cone_decomp T (exact_aux k m qs)" using assms proof (induct m arbitrary: qs) case 0 thus ?case by simp next case (Suc m) let ?qs = "shift k m qs" have "cone_decomp T (exact_aux k m ?qs)" proof (rule Suc) show "valid_decomp X ?qs" and "standard_decomp k ?qs" and "exact_decomp m ?qs" using Suc.prems(1, 2, 3) by (rule shift)+ next from Suc.prems show "cone_decomp T ?qs" by (rule cone_decomp_shift) qed thus ?case by simp qed lemma Max_exact_aux_ge: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp m qs" shows "Max (poly_deg ` fst ` set qs) \ Max (poly_deg ` fst ` set (exact_aux k m qs))" using assms proof (induct m arbitrary: qs) case 0 thus ?case by simp next case (Suc m) let ?qs = "shift k m qs" from Suc.prems have "Max (poly_deg ` fst ` set qs) \ Max (poly_deg ` fst ` set ?qs)" by (rule Max_shift_ge) also have "\ \ Max (poly_deg ` fst ` set (exact_aux k m ?qs))" proof (rule Suc) from Suc.prems show "valid_decomp X ?qs" and "standard_decomp k ?qs" and "exact_decomp m ?qs" by (rule shift)+ qed finally show ?case by simp qed lemma exact_aux_Nil_iff: assumes "valid_decomp X qs" and "standard_decomp k qs" and "exact_decomp m qs" shows "exact_aux k m qs = [] \ qs = []" using assms proof (induct m arbitrary: qs) case 0 thus ?case by simp next case (Suc m) let ?qs = "shift k m qs" have "exact_aux k m ?qs = [] \ ?qs = []" proof (rule Suc) from Suc.prems show "valid_decomp X ?qs" and "standard_decomp k ?qs" and "exact_decomp m ?qs" by (rule shift)+ qed also from Suc.prems have "\ \ qs = []" by (rule shift_Nil_iff) finally show ?case by simp qed definition exact :: "nat \ ((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) list \ ((('x \\<^sub>0 nat) \\<^sub>0 'a::{comm_ring_1,ring_no_zero_divisors}) \ 'x set) list" where "exact k qs = exact_aux k (card X) qs" lemma exact: assumes "valid_decomp X qs" and "standard_decomp k qs" shows "valid_decomp X (exact k qs)" (is ?thesis1) and "standard_decomp k (exact k qs)" (is ?thesis2) and "exact_decomp 0 (exact k qs)" (is ?thesis3) proof - from assms(1) le_refl have "exact_decomp (card X) qs" by (rule exact_decomp_card_X) with assms show ?thesis1 and ?thesis2 and ?thesis3 unfolding exact_def by (rule exact_aux)+ qed lemma monomial_decomp_exact: assumes "valid_decomp X qs" and "standard_decomp k qs" and "monomial_decomp qs" shows "monomial_decomp (exact k qs)" proof - from assms(1) le_refl have "exact_decomp (card X) qs" by (rule exact_decomp_card_X) with assms(1, 2) show ?thesis unfolding exact_def using assms(3) by (rule monomial_decomp_exact_aux) qed lemma hom_decomp_exact: assumes "valid_decomp X qs" and "standard_decomp k qs" and "hom_decomp qs" shows "hom_decomp (exact k qs)" proof - from assms(1) le_refl have "exact_decomp (card X) qs" by (rule exact_decomp_card_X) with assms(1, 2) show ?thesis unfolding exact_def using assms(3) by (rule hom_decomp_exact_aux) qed lemma cone_decomp_exact: assumes "valid_decomp X qs" and "standard_decomp k qs" and "cone_decomp T qs" shows "cone_decomp T (exact k qs)" proof - from assms(1) le_refl have "exact_decomp (card X) qs" by (rule exact_decomp_card_X) with assms(1, 2) show ?thesis unfolding exact_def using assms(3) by (rule cone_decomp_exact_aux) qed lemma Max_exact_ge: assumes "valid_decomp X qs" and "standard_decomp k qs" shows "Max (poly_deg ` fst ` set qs) \ Max (poly_deg ` fst ` set (exact k qs))" proof - from assms(1) le_refl have "exact_decomp (card X) qs" by (rule exact_decomp_card_X) with assms(1, 2) show ?thesis unfolding exact_def by (rule Max_exact_aux_ge) qed lemma exact_Nil_iff: assumes "valid_decomp X qs" and "standard_decomp k qs" shows "exact k qs = [] \ qs = []" proof - from assms(1) le_refl have "exact_decomp (card X) qs" by (rule exact_decomp_card_X) with assms(1, 2) show ?thesis unfolding exact_def by (rule exact_aux_Nil_iff) qed corollary \_zero_exact: assumes "valid_decomp X qs" and "standard_decomp k qs" and "qs \ []" shows "Suc (Max (poly_deg ` fst ` set qs)) \ \ (exact k qs) 0" proof - from assms(1, 2) have "Max (poly_deg ` fst ` set qs) \ Max (poly_deg ` fst ` set (exact k qs))" by (rule Max_exact_ge) also have "Suc \ \ \ (exact k qs) 0" proof (rule \_zero) from assms show "exact k qs \ []" by (simp add: exact_Nil_iff) qed finally show ?thesis by simp qed lemma normal_form_exact_decompE: assumes "F \ P[X]" obtains qs where "valid_decomp X qs" and "standard_decomp 0 qs" and "monomial_decomp qs" and "cone_decomp (normal_form F ` P[X]) qs" and "exact_decomp 0 qs" and "\g. (\f. f \ F \ homogeneous f) \ g \ punit.reduced_GB F \ poly_deg g \ \ qs 0" proof - let ?G = "punit.reduced_GB F" let ?S = "lpp ` ?G" let ?N = "normal_form F ` P[X]" define qs::"((_ \\<^sub>0 'a) \ _) list" where "qs = snd (split 0 X ?S)" from fin_X assms have std: "standard_decomp 0 qs" and cn: "cone_decomp ?N qs" unfolding qs_def by (rule standard_cone_decomp_snd_split)+ from fin_X assms have "finite ?G" by (rule finite_reduced_GB_Polys) hence "finite ?S" by (rule finite_imageI) with fin_X subset_refl have valid: "valid_decomp X qs" unfolding qs_def using zero_in_PPs by (rule valid_decomp_split) from fin_X subset_refl \finite ?S\ have md: "monomial_decomp qs" unfolding qs_def by (rule monomial_decomp_split) let ?qs = "exact 0 qs" from valid std have "valid_decomp X ?qs" and "standard_decomp 0 ?qs" by (rule exact)+ moreover from valid std md have "monomial_decomp ?qs" by (rule monomial_decomp_exact) moreover from valid std cn have "cone_decomp ?N ?qs" by (rule cone_decomp_exact) moreover from valid std have "exact_decomp 0 ?qs" by (rule exact) moreover have "poly_deg g \ \ ?qs 0" if "\f. f \ F \ homogeneous f" and "g \ ?G" for g proof (cases "qs = []") case True from one_in_Polys have "normal_form F 1 \ ?N" by (rule imageI) also from True cn have "\ = {0}" by (simp add: cone_decomp_def direct_decomp_def bij_betw_def) finally have "?G = {1}" using fin_X assms by (simp add: normal_form_zero_iff ideal_eq_UNIV_iff_reduced_GB_eq_one_Polys flip: ideal_eq_UNIV_iff_contains_one) with that(2) show ?thesis by simp next case False from fin_X assms that have "poly_deg g \ Suc (Max (poly_deg ` fst ` set qs))" unfolding qs_def by (rule standard_cone_decomp_snd_split) also from valid std False have "\ \ \ ?qs 0" by (rule \_zero_exact) finally show ?thesis . qed ultimately show ?thesis .. qed end end end (* pm_powerprod *) end (* theory *) diff --git a/thys/Groebner_Macaulay/Dube_Bound.thy b/thys/Groebner_Macaulay/Dube_Bound.thy --- a/thys/Groebner_Macaulay/Dube_Bound.thy +++ b/thys/Groebner_Macaulay/Dube_Bound.thy @@ -1,1583 +1,1583 @@ (* Author: Alexander Maletzky *) section \Dub\'{e}'s Degree-Bound for Homogeneous Gr\"obner Bases\ theory Dube_Bound imports Poly_Fun Cone_Decomposition Degree_Bound_Utils begin context fixes n d :: nat begin function Dube_aux :: "nat \ nat" where "Dube_aux j = (if j + 2 < n then 2 + ((Dube_aux (j + 1)) choose 2) + (\i=j+3..n-1. (Dube_aux i) choose (Suc (i - j))) else if j + 2 = n then d\<^sup>2 + 2 * d else 2 * d)" by pat_completeness auto termination proof show "wf (measure ((-) n))" by (fact wf_measure) qed auto definition Dube :: nat where "Dube = (if n \ 1 \ d = 0 then d else Dube_aux 1)" lemma Dube_aux_ge_d: "d \ Dube_aux j" proof (induct j rule: Dube_aux.induct) case step: (1 j) have "j + 2 < n \ j + 2 = n \ n < j + 2" by auto show ?case proof (rule linorder_cases) assume *: "j + 2 < n" hence 1: "d \ Dube_aux (j + 1)" by (rule step.hyps)+ show ?thesis proof (cases "d \ 2") case True also from * have "2 \ Dube_aux j" by simp finally show ?thesis . next case False hence "2 < d" by simp hence "2 < Dube_aux (j + 1)" using 1 by (rule less_le_trans) with _ have "Dube_aux (j + 1) \ Dube_aux (j + 1) choose 2" by (rule upper_le_binomial) simp also from * have "\ \ Dube_aux j" by simp finally have "Dube_aux (j + 1) \ Dube_aux j" . with 1 show ?thesis by (rule le_trans) qed next assume "j + 2 = n" thus ?thesis by simp next assume "n < j + 2" thus ?thesis by simp qed qed corollary Dube_ge_d: "d \ Dube" by (simp add: Dube_def Dube_aux_ge_d del: Dube_aux.simps) text \Dub\'{e} in @{cite Dube1990} proves the following theorem, to obtain a short closed form for the degree bound. However, the proof he gives is wrong: In the last-but-one proof step of Lemma 8.1 the sum on the right-hand-side of the inequality can be greater than 1/2 (e.g. for @{prop "n = 7"}, @{prop "d = 2"} and @{prop "j = 1"}), rendering the value inside the big brackets negative. This is also true without the additional summand \2\ we had to introduce in function @{const Dube_aux} to correct another mistake found in @{cite Dube1990}. Nonetheless, experiments carried out in Mathematica still suggest that the short closed form is a valid upper bound for @{const Dube}, even with the additional summand \2\. So, with some effort it might be possible to prove the theorem below; but in fact function @{const Dube} gives typically much better (i.e. smaller) values for concrete values of \n\ and \d\, so it is better to stick to @{const Dube} instead of the closed form anyway. Asymptotically, as \n\ tends to infinity, @{const Dube} grows double exponentially, too.\ theorem "rat_of_nat Dube \ 2 * ((rat_of_nat d)\<^sup>2 / 2 + (rat_of_nat d)) ^ (2 ^ (n - 2))" oops end subsection \Hilbert Function and Hilbert Polynomial\ context pm_powerprod begin context fixes X :: "'x set" assumes fin_X: "finite X" begin lemma Hilbert_fun_cone_aux: assumes "h \ P[X]" and "h \ 0" and "U \ X" and "homogeneous (h::_ \\<^sub>0 'a::field)" shows "Hilbert_fun (cone (h, U)) z = card {t \ .[U]. deg_pm t + poly_deg h = z}" proof - from assms(2) have "lpp h \ keys h" by (rule punit.lt_in_keys) with assms(4) have deg_h[symmetric]: "deg_pm (lpp h) = poly_deg h" by (rule homogeneousD_poly_deg) from assms(1, 3) have "cone (h, U) \ P[X]" by (rule cone_subset_PolysI) with fin_X have "Hilbert_fun (cone (h, U)) z = card (lpp ` (hom_deg_set z (cone (h, U)) - {0}))" using subspace_cone[of "(h, U)"] by (simp only: Hilbert_fun_alt) also from assms(4) have "lpp ` (hom_deg_set z (cone (h, U)) - {0}) = {t \ lpp ` (cone (h, U) - {0}). deg_pm t = z}" by (intro image_lt_hom_deg_set homogeneous_set_coneI) also have "{t \ lpp ` (cone (h, U) - {0}). deg_pm t = z} = (\t. t + lpp h) ` {t \ .[U]. deg_pm t + poly_deg h = z}" (is "?A = ?B") proof show "?A \ ?B" proof fix t assume "t \ ?A" hence "t \ lpp ` (cone (h, U) - {0})" and "deg_pm t = z" by simp_all from this(1) obtain a where "a \ cone (h, U) - {0}" and 2: "t = lpp a" .. from this(1) have "a \ cone (h, U)" and "a \ 0" by simp_all from this(1) obtain q where "q \ P[U]" and a: "a = q * h" by (rule coneE) from \a \ 0\ have "q \ 0" by (auto simp: a) hence t: "t = lpp q + lpp h" using assms(2) unfolding 2 a by (rule lp_times) hence "deg_pm (lpp q) + poly_deg h = deg_pm t" by (simp add: deg_pm_plus deg_h) also have "\ = z" by fact finally have "deg_pm (lpp q) + poly_deg h = z" . moreover from \q \ P[U]\ have "lpp q \ .[U]" by (rule PPs_closed_lpp) ultimately have "lpp q \ {t \ .[U]. deg_pm t + poly_deg h = z}" by simp moreover have "t = lpp q + lpp h" by (simp only: t) ultimately show "t \ ?B" by (rule rev_image_eqI) qed next show "?B \ ?A" proof fix t assume "t \ ?B" then obtain s where "s \ {t \ .[U]. deg_pm t + poly_deg h = z}" and t1: "t = s + lpp h" .. from this(1) have "s \ .[U]" and 1: "deg_pm s + poly_deg h = z" by simp_all let ?q = "monomial (1::'a) s" have "?q \ 0" by (simp add: monomial_0_iff) hence "?q * h \ 0" and "lpp (?q * h) = lpp ?q + lpp h" using \h \ 0\ by (rule times_not_zero, rule lp_times) hence t: "t = lpp (?q * h)" by (simp add: t1 punit.lt_monomial) from \s \ .[U]\ have "?q \ P[U]" by (rule Polys_closed_monomial) with refl have "?q * h \ cone (h, U)" by (rule coneI) moreover from _ assms(2) have "?q * h \ 0" by (rule times_not_zero) (simp add: monomial_0_iff) ultimately have "?q * h \ cone (h, U) - {0}" by simp hence "t \ lpp ` (cone (h, U) - {0})" unfolding t by (rule imageI) moreover have "deg_pm t = int z" by (simp add: t1) (simp add: deg_pm_plus deg_h flip: 1) ultimately show "t \ ?A" by simp qed qed also have "card \ = card {t \ .[U]. deg_pm t + poly_deg h = z}" by (simp add: card_image) finally show ?thesis . qed lemma Hilbert_fun_cone_empty: assumes "h \ P[X]" and "h \ 0" and "homogeneous (h::_ \\<^sub>0 'a::field)" shows "Hilbert_fun (cone (h, {})) z = (if poly_deg h = z then 1 else 0)" proof - have "Hilbert_fun (cone (h, {})) z = card {t \ .[{}::'x set]. deg_pm t + poly_deg h = z}" using assms(1, 2) empty_subsetI assms(3) by (rule Hilbert_fun_cone_aux) also have "\ = (if poly_deg h = z then 1 else 0)" by simp finally show ?thesis . qed lemma Hilbert_fun_cone_nonempty: assumes "h \ P[X]" and "h \ 0" and "U \ X" and "homogeneous (h::_ \\<^sub>0 'a::field)" and "U \ {}" shows "Hilbert_fun (cone (h, U)) z = (if poly_deg h \ z then ((z - poly_deg h) + (card U - 1)) choose (card U - 1) else 0)" proof (cases "poly_deg h \ z") case True from assms(3) fin_X have "finite U" by (rule finite_subset) from assms(1-4) have "Hilbert_fun (cone (h, U)) z = card {t \ .[U]. deg_pm t + poly_deg h = z}" by (rule Hilbert_fun_cone_aux) also from True have "{t \ .[U]. deg_pm t + poly_deg h = z} = deg_sect U (z - poly_deg h)" by (auto simp: deg_sect_def) also from \finite U\ assms(5) have "card \ = (z - poly_deg h) + (card U - 1) choose (card U - 1)" by (rule card_deg_sect) finally show ?thesis by (simp add: True) next case False from assms(1-4) have "Hilbert_fun (cone (h, U)) z = card {t \ .[U]. deg_pm t + poly_deg h = z}" by (rule Hilbert_fun_cone_aux) also from False have "{t \ .[U]. deg_pm t + poly_deg h = z} = {}" by auto hence "card {t \ .[U]. deg_pm t + poly_deg h = z} = card ({}::('x \\<^sub>0 nat) set)" by (rule arg_cong) also have "\ = 0" by simp finally show ?thesis by (simp add: False) qed corollary Hilbert_fun_Polys: assumes "X \ {}" shows "Hilbert_fun (P[X]::(_ \\<^sub>0 'a::field) set) z = (z + (card X - 1)) choose (card X - 1)" proof - let ?one = "1::('x \\<^sub>0 nat) \\<^sub>0 'a" have "Hilbert_fun (P[X]::(_ \\<^sub>0 'a) set) z = Hilbert_fun (cone (?one, X)) z" by simp also have "\ = (if poly_deg ?one \ z then ((z - poly_deg ?one) + (card X - 1)) choose (card X - 1) else 0)" using one_in_Polys _ subset_refl _ assms by (rule Hilbert_fun_cone_nonempty) simp_all also have "\ = (z + (card X - 1)) choose (card X - 1)" by simp finally show ?thesis . qed lemma Hilbert_fun_cone_decomp: assumes "cone_decomp T ps" and "valid_decomp X ps" and "hom_decomp ps" shows "Hilbert_fun T z = (\hU\set ps. Hilbert_fun (cone hU) z)" proof - note fin_X moreover from assms(2, 1) have "T \ P[X]" by (rule valid_cone_decomp_subset_Polys) moreover from assms(1) have dd: "direct_decomp T (map cone ps)" by (rule cone_decompD) ultimately have "Hilbert_fun T z = (\s\set (map cone ps). Hilbert_fun s z)" proof (rule Hilbert_fun_direct_decomp) fix cn assume "cn \ set (map cone ps)" then obtain hU where "hU \ set ps" and cn: "cn = cone hU" unfolding set_map .. note this(1) moreover obtain h U where hU: "hU = (h, U)" using prod.exhaust by blast ultimately have "(h, U) \ set ps" by simp with assms(3) have "homogeneous h" by (rule hom_decompD) thus "homogeneous_set cn" unfolding cn hU by (rule homogeneous_set_coneI) show "phull.subspace cn" unfolding cn by (fact subspace_cone) qed also have "\ = (\hU\set ps. ((\s. Hilbert_fun s z) \ cone) hU)" unfolding set_map using finite_set proof (rule sum.reindex_nontrivial) fix hU1 hU2 assume "hU1 \ set ps" and "hU2 \ set ps" and "hU1 \ hU2" with dd have "cone hU1 \ cone hU2 = {0}" using zero_in_cone by (rule direct_decomp_map_Int_zero) moreover assume "cone hU1 = cone hU2" ultimately show "Hilbert_fun (cone hU1) z = 0" by simp qed finally show ?thesis by simp qed definition Hilbert_poly :: "(nat \ nat) \ int \ int" where "Hilbert_poly b = (\z::int. let n = card X in ((z - b (Suc n) + n) gchoose n) - 1 - (\i=1..n. (z - b i + i - 1) gchoose i))" lemma poly_fun_Hilbert_poly: "poly_fun (Hilbert_poly b)" by (simp add: Hilbert_poly_def Let_def) lemma Hilbert_fun_eq_Hilbert_poly_plus_card: assumes "X \ {}" and "valid_decomp X ps" and "hom_decomp ps" and "cone_decomp T ps" and "standard_decomp k ps" and "exact_decomp X 0 ps" and "\ ps (Suc 0) \ d" shows "int (Hilbert_fun T d) = card {h::_ \\<^sub>0 'a::field. (h, {}) \ set ps \ poly_deg h = d} + Hilbert_poly (\ ps) d" proof - define n where "n = card X" with assms(1) have "0 < n" using fin_X by (simp add: card_gt_0_iff) hence "1 \ n" and "Suc 0 \ n" by simp_all from pos_decomp_subset have eq0: "(set ps - set (ps\<^sub>+)) \ set (ps\<^sub>+) = set ps" by blast have "set ps - set (ps\<^sub>+) \ set ps" by blast hence fin2: "finite (set ps - set (ps\<^sub>+))" using finite_set by (rule finite_subset) have "(\hU\set ps - set (ps\<^sub>+). Hilbert_fun (cone hU) d) = (\(h, U)\set ps - set (ps\<^sub>+). if poly_deg h = d then 1 else 0)" using refl proof (rule sum.cong) fix x assume "x \ set ps - set (ps\<^sub>+)" moreover obtain h U where x: "x = (h, U)" using prod.exhaust by blast ultimately have "U = {}" and "(h, U) \ set ps" by (simp_all add: pos_decomp_def) from assms(2) this(2) have "h \ P[X]" and "h \ 0" by (rule valid_decompD)+ moreover from assms(3) \(h, U) \ set ps\ have "homogeneous h" by (rule hom_decompD) ultimately show "Hilbert_fun (cone x) d = (case x of (h, U) \ if poly_deg h = d then 1 else 0)" by (simp add: x \U = {}\ Hilbert_fun_cone_empty split del: if_split) qed also from fin2 have "\ = (\(h, U)\{(h', U') \ set ps - set (ps\<^sub>+). poly_deg h' = d}. 1)" by (rule sum.mono_neutral_cong_right) (auto split: if_splits) also have "\ = card {(h, U) \ set ps - set (ps\<^sub>+). poly_deg h = d}" by auto also have "\ = card {h. (h, {}) \ set ps \ poly_deg h = d}" by (fact card_Diff_pos_decomp) finally have eq1: "(\hU\set ps - set (ps\<^sub>+). Hilbert_fun (cone hU) d) = card {h. (h, {}) \ set ps \ poly_deg h = d}" . let ?f = "\a b. (int d) - a + b gchoose b" have "int (\hU\set (ps\<^sub>+). Hilbert_fun (cone hU) d) = (\hU\set (ps\<^sub>+). int (Hilbert_fun (cone hU) d))" by (simp add: int_sum prod.case_distrib) also have "\ = (\(h, U)\(\i\{1..n}. {(h, U) \ set (ps\<^sub>+). card U = i}). ?f (poly_deg h) (card U - 1))" proof (rule sum.cong) show "set (ps\<^sub>+) = (\i\{1..n}. {(h, U). (h, U) \ set (ps\<^sub>+) \ card U = i})" proof (rule Set.set_eqI, rule) fix x assume "x \ set (ps\<^sub>+)" moreover obtain h U where x: "x = (h, U)" using prod.exhaust by blast ultimately have "(h, U) \ set (ps\<^sub>+)" by simp hence "(h, U) \ set ps" and "U \ {}" by (simp_all add: pos_decomp_def) from fin_X assms(6) this(1) have "U \ X" by (rule exact_decompD) hence "finite U" using fin_X by (rule finite_subset) with \U \ {}\ have "0 < card U" by (simp add: card_gt_0_iff) moreover from fin_X \U \ X\ have "card U \ n" unfolding n_def by (rule card_mono) ultimately have "card U \ {1..n}" by simp moreover from \(h, U) \ set (ps\<^sub>+)\ have "(h, U) \ {(h', U'). (h', U') \ set (ps\<^sub>+) \ card U' = card U}" by simp ultimately show "x \ (\i\{1..n}. {(h, U). (h, U) \ set (ps\<^sub>+) \ card U = i})" by (simp add: x) qed blast next fix x assume "x \ (\i\{1..n}. {(h, U). (h, U) \ set (ps\<^sub>+) \ card U = i})" then obtain j where "j \ {1..n}" and "x \ {(h, U). (h, U) \ set (ps\<^sub>+) \ card U = j}" .. from this(2) obtain h U where "(h, U) \ set (ps\<^sub>+)" and "card U = j" and x: "x = (h, U)" by blast from fin_X assms(2, 5) this(1) have "poly_deg h < \ ps (Suc 0)" by (rule \_one_gr) also have "\ \ d" by fact finally have "poly_deg h < d" . hence int1: "int (d - poly_deg h) = int d - int (poly_deg h)" by simp from \card U = j\ \j \ {1..n}\ have "0 < card U" by simp hence int2: "int (card U - Suc 0) = int (card U) - 1" by simp from \(h, U) \ set (ps\<^sub>+)\ have "(h, U) \ set ps" using pos_decomp_subset .. with assms(2) have "h \ P[X]" and "h \ 0" and "U \ X" by (rule valid_decompD)+ moreover from assms(3) \(h, U) \ set ps\ have "homogeneous h" by (rule hom_decompD) moreover from \0 < card U\ have "U \ {}" by auto ultimately have "Hilbert_fun (cone (h, U)) d = (if poly_deg h \ d then (d - poly_deg h + (card U - 1)) choose (card U - 1) else 0)" by (rule Hilbert_fun_cone_nonempty) also from \poly_deg h < d\ have "\ = (d - poly_deg h + (card U - 1)) choose (card U - 1)" by simp finally have "int (Hilbert_fun (cone (h, U)) d) = (int d - int (poly_deg h) + (int (card U - 1))) gchoose (card U - 1)" by (simp add: int_binomial int1 int2) thus "int (Hilbert_fun (cone x) d) = (case x of (h, U) \ int d - int (poly_deg h) + (int (card U - 1)) gchoose (card U - 1))" by (simp add: x) qed also have "\ = (\j=1..n. \(h, U)\{(h', U') \ set (ps\<^sub>+). card U' = j}. ?f (poly_deg h) (card U - 1))" proof (intro sum.UNION_disjoint ballI) fix j have "{(h, U). (h, U) \ set (ps\<^sub>+) \ card U = j} \ set (ps\<^sub>+)" by blast thus "finite {(h, U). (h, U) \ set (ps\<^sub>+) \ card U = j}" using finite_set by (rule finite_subset) qed blast+ also from refl have "\ = (\j=1..n. ?f (\ ps (Suc j)) j - ?f (\ ps j) j)" proof (rule sum.cong) fix j assume "j \ {1..n}" hence "Suc 0 \ j" and "0 < j" and "j \ n" by simp_all from fin_X this(1) have "\ ps j \ \ ps (Suc 0)" by (rule \_decreasing) also have "\ \ d" by fact finally have "\ ps j \ d" . from fin_X have "\ ps (Suc j) \ \ ps j" by (rule \_decreasing) simp hence "\ ps (Suc j) \ d" using \\ ps j \ d\ by (rule le_trans) from \0 < j\ have int_j: "int (j - Suc 0) = int j - 1" by simp have "(\(h, U)\{(h', U'). (h', U') \ set (ps\<^sub>+) \ card U' = j}. ?f (poly_deg h) (card U - 1)) = (\(h, U)\(\d0\{\ ps (Suc j)..int (\ ps j) - 1}. {(h', U'). (h', U') \ set (ps\<^sub>+) \ int (poly_deg h') = d0 \ card U' = j}). ?f (poly_deg h) (card U - 1))" using _ refl proof (rule sum.cong) show "{(h', U'). (h', U') \ set (ps\<^sub>+) \ card U' = j} = (\d0\{\ ps (Suc j)..int (\ ps j) - 1}. {(h', U'). (h', U') \ set (ps\<^sub>+) \ int (poly_deg h') = d0 \ card U' = j})" proof (rule Set.set_eqI, rule) fix x assume "x \ {(h', U'). (h', U') \ set (ps\<^sub>+) \ card U' = j}" moreover obtain h U where x: "x = (h, U)" using prod.exhaust by blast ultimately have "(h, U) \ set (ps\<^sub>+)" and "card U = j" by simp_all with fin_X assms(5, 6) \Suc 0 \ j\ \j \ n\ have "\ ps (Suc j) \ poly_deg h" unfolding n_def by (rule lem_6_1_3) moreover from fin_X have "poly_deg h < \ ps j" proof (rule \) from \(h, U) \ set (ps\<^sub>+)\ show "(h, U) \ set ps" using pos_decomp_subset .. next show "j \ card U" by (simp add: \card U = j\) qed ultimately have "poly_deg h \ {\ ps (Suc j)..int (\ ps j) - 1}" by simp moreover have "(h, U) \ {(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = poly_deg h \ card U' = card U}" using \(h, U) \ set (ps\<^sub>+)\ by simp ultimately show "x \ (\d0\{\ ps (Suc j)..int (\ ps j) - 1}. {(h', U'). (h', U') \ set (ps\<^sub>+) \ int (poly_deg h') = d0 \ card U' = j})" by (simp add: x \card U = j\) qed blast qed also have "\ = (\d0=\ ps (Suc j)..int (\ ps j) - 1. \(h, U)\{(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = d0 \ card U' = j}. ?f (poly_deg h) (card U - 1))" proof (intro sum.UNION_disjoint ballI) fix d0::int have "{(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = d0 \ card U' = j} \ set (ps\<^sub>+)" by blast thus "finite {(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = d0 \ card U' = j}" using finite_set by (rule finite_subset) qed blast+ also from refl have "\ = (\d0=\ ps (Suc j)..int (\ ps j) - 1. ?f d0 (j - 1))" proof (rule sum.cong) fix d0 assume "d0 \ {\ ps (Suc j)..int (\ ps j) - 1}" hence "\ ps (Suc j) \ d0" and "d0 < int (\ ps j)" by simp_all hence "\ ps (Suc j) \ nat d0" and "nat d0 < \ ps j" by simp_all have "(\(h, U)\{(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = d0 \ card U' = j}. ?f (poly_deg h) (card U - 1)) = (\(h, U)\{(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = d0 \ card U' = j}. ?f d0 (j - 1))" using refl by (rule sum.cong) auto also have "\ = card {(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = nat d0 \ card U' = j} * ?f d0 (j - 1)" using \\ ps (Suc j) \ d0\ by (simp add: int_eq_iff) also have "\ = ?f d0 (j - 1)" using fin_X assms(5, 6) \Suc 0 \ j\ \j \ n\ \\ ps (Suc j) \ nat d0\ \nat d0 < \ ps j\ by (simp only: n_def lem_6_1_2'(3)) finally show "(\(h, U)\{(h', U'). (h', U') \ set (ps\<^sub>+) \ poly_deg h' = d0 \ card U' = j}. ?f (poly_deg h) (card U - 1)) = ?f d0 (j - 1)" . qed also have "\ = (\d0\(-) (int d) ` {\ ps (Suc j)..int (\ ps j) - 1}. d0 + int (j - 1) gchoose (j - 1))" proof - have "inj_on ((-) (int d)) {\ ps (Suc j)..int (\ ps j) - 1}" by (auto simp: inj_on_def) thus ?thesis by (simp only: sum.reindex o_def) qed also have "\ = (\d0\{0..int d - (\ ps (Suc j))}-{0..int d - \ ps j}. d0 + int (j - 1) gchoose (j - 1))" using _ refl proof (rule sum.cong) have "(-) (int d) ` {\ ps (Suc j)..int (\ ps j) - 1} = {int d - (int (\ ps j) - 1)..int d - int (\ ps (Suc j))}" by (simp only: image_diff_atLeastAtMost) also have "\ = {0..int d - int (\ ps (Suc j))} - {0..int d - int (\ ps j)}" proof - from \\ ps j \ d\ have "int (\ ps j) - 1 \ int d" by simp thus ?thesis by auto qed finally show "(-) (int d) ` {\ ps (Suc j)..int (\ ps j) - 1} = {0..int d - int (\ ps (Suc j))} - {0..int d - int (\ ps j)}" . qed also have "\ = (\d0=0..int d - (\ ps (Suc j)). d0 + int (j - 1) gchoose (j - 1)) - (\d0=0..int d - \ ps j. d0 + int (j - 1) gchoose (j - 1))" by (rule sum_diff) (auto simp: \\ ps (Suc j) \ \ ps j\) also from \\ ps (Suc j) \ d\ \\ ps j \ d\ have "\ = ?f (\ ps (Suc j)) j - ?f (\ ps j) j" by (simp add: gchoose_rising_sum, simp add: int_j ac_simps \0 < j\) finally show "(\(h, U)\{(h', U'). (h', U') \ set (ps\<^sub>+) \ card U' = j}. ?f (poly_deg h) (card U - 1)) = ?f (\ ps (Suc j)) j - ?f (\ ps j) j" . qed also have "\ = (\j=1..n. ?f (\ ps (Suc j)) j) - (\j=1..n. ?f (\ ps j) j)" by (fact sum_subtractf) also have "\ = ?f (\ ps (Suc n)) n + (\j=1..n-1. ?f (\ ps (Suc j)) j) - (\j=1..n. ?f (\ ps j) j)" by (simp only: sum_tail_nat[OF \0 < n\ \1 \ n\]) also have "\ = ?f (\ ps (Suc n)) n - ?f (\ ps 1) 1 + ((\j=1..n-1. ?f (\ ps (Suc j)) j) - (\j=1..n-1. ?f (\ ps (Suc j)) (Suc j)))" by (simp only: sum.atLeast_Suc_atMost[OF \1 \ n\] sum_atLeast_Suc_shift[OF \0 < n\ \1 \ n\]) also have "\ = ?f (\ ps (Suc n)) n - ?f (\ ps 1) 1 - (\j=1..n-1. ?f (\ ps (Suc j)) (Suc j) - ?f (\ ps (Suc j)) j)" by (simp only: sum_subtractf) also have "\ = ?f (\ ps (Suc n)) n - 1 - ((int d - \ ps (Suc 0)) gchoose (Suc 0)) - (\j=1..n-1. (int d - \ ps (Suc j) + j) gchoose (Suc j))" proof - have "?f (\ ps 1) 1 = 1 + ((int d - \ ps (Suc 0)) gchoose (Suc 0))" by (simp add: plus_Suc_gbinomial) moreover from refl have "(\j=1..n-1. ?f (\ ps (Suc j)) (Suc j) - ?f (\ ps (Suc j)) j) = (\j=1..n-1. (int d - \ ps (Suc j) + j) gchoose (Suc j))" by (rule sum.cong) (simp add: plus_Suc_gbinomial) ultimately show ?thesis by (simp only:) qed also have "\ = ?f (\ ps (Suc n)) n - 1 - (\j=0..n-1. (int d - \ ps (Suc j) + j) gchoose (Suc j))" by (simp only: sum.atLeast_Suc_atMost[OF le0], simp) also have "\ = ?f (\ ps (Suc n)) n - 1 - (\j=Suc 0..Suc (n-1). (int d - \ ps j + j - 1) gchoose j)" by (simp only: sum.shift_bounds_cl_Suc_ivl, simp add: ac_simps) also have "\ = Hilbert_poly (\ ps) d" using \0 < n\ by (simp add: Hilbert_poly_def Let_def n_def) finally have eq2: "int (\hU\set (ps\<^sub>+). Hilbert_fun (cone hU) d) = Hilbert_poly (\ ps) (int d)" . from assms(4, 2, 3) have "Hilbert_fun T d = (\hU\set ps. Hilbert_fun (cone hU) d)" by (rule Hilbert_fun_cone_decomp) also have "\ = (\hU\(set ps - set (ps\<^sub>+)) \ set (ps\<^sub>+). Hilbert_fun (cone hU) d)" by (simp only: eq0) also have "\ = (\hU\set ps - set (ps\<^sub>+). Hilbert_fun (cone hU) d) + (\hU\set (ps\<^sub>+). Hilbert_fun (cone hU) d)" using fin2 finite_set by (rule sum.union_disjoint) blast also have "\ = card {h. (h, {}) \ set ps \ poly_deg h = d} + (\hU\set (ps\<^sub>+). Hilbert_fun (cone hU) d)" by (simp only: eq1) also have "int \ = card {h. (h, {}) \ set ps \ poly_deg h = d} + Hilbert_poly (\ ps) d" by (simp only: eq2 int_plus) finally show ?thesis . qed corollary Hilbert_fun_eq_Hilbert_poly: assumes "X \ {}" and "valid_decomp X ps" and "hom_decomp ps" and "cone_decomp T ps" and "standard_decomp k ps" and "exact_decomp X 0 ps" and "\ ps 0 \ d" shows "int (Hilbert_fun (T::(_ \\<^sub>0 'a::field) set) d) = Hilbert_poly (\ ps) d" proof - from fin_X have "\ ps (Suc 0) \ \ ps 0" using le0 by (rule \_decreasing) also have "\ \ d" by fact finally have "\ ps (Suc 0) \ d" . with assms(1-6) have "int (Hilbert_fun T d) = int (card {h. (h, {}) \ set ps \ poly_deg h = d}) + Hilbert_poly (\ ps) (int d)" by (rule Hilbert_fun_eq_Hilbert_poly_plus_card) also have "\ = Hilbert_poly (\ ps) (int d)" proof - have eq: "{h. (h, {}) \ set ps \ poly_deg h = d} = {}" proof - { fix h assume "(h, {}) \ set ps" and "poly_deg h = d" from fin_X this(1) le0 have "poly_deg h < \ ps 0" by (rule \) with assms(7) have False by (simp add: \poly_deg h = d\) } thus ?thesis by blast qed show ?thesis by (simp add: eq) qed finally show ?thesis . qed subsection \Dub\'{e}'s Bound\ context fixes f :: "('x \\<^sub>0 nat) \\<^sub>0 'a::field" fixes F assumes n_gr_1: "1 < card X" and fin_F: "finite F" and F_sub: "F \ P[X]" and f_in: "f \ F" and hom_F: "\f'. f' \ F \ homogeneous f'" and f_max: "\f'. f' \ F \ poly_deg f' \ poly_deg f" and d_gr_0: "0 < poly_deg f" and ideal_f_neq: "ideal {f} \ ideal F" begin private abbreviation (input) "n \ card X" private abbreviation (input) "d \ poly_deg f" lemma f_in_Polys: "f \ P[X]" using f_in F_sub .. lemma hom_f: "homogeneous f" using f_in by (rule hom_F) lemma f_not_0: "f \ 0" using d_gr_0 by auto lemma X_not_empty: "X \ {}" using n_gr_1 by auto lemma n_gr_0: "0 < n" using \1 < n\ by simp corollary int_n_minus_1 [simp]: "int (n - Suc 0) = int n - 1" using n_gr_0 by simp lemma int_n_minus_2 [simp]: "int (n - Suc (Suc 0)) = int n - 2" using n_gr_1 by simp lemma cone_f_X_sub: "cone (f, X) \ P[X]" proof - have "cone (f, X) = cone (f * 1, X)" by simp also from f_in_Polys have "\ \ cone (1, X)" by (rule cone_mono_1) finally show ?thesis by simp qed lemma ideal_Int_Polys_eq_cone: "ideal {f} \ P[X] = cone (f, X)" proof (intro subset_antisym subsetI) fix p assume "p \ ideal {f} \ P[X]" hence "p \ ideal {f}" and "p \ P[X]" by simp_all have "finite {f}" by simp then obtain q where "p = (\f'\{f}. q f' * f')" using \p \ ideal {f}\ by (rule ideal.span_finiteE) hence p: "p = q f * f" by simp with \p \ P[X]\ have "f * q f \ P[X]" by (simp only: mult.commute) hence "q f \ P[X]" using f_in_Polys f_not_0 by (rule times_in_PolysD) with p show "p \ cone (f, X)" by (rule coneI) next fix p assume "p \ cone (f, X)" then obtain q where "q \ P[X]" and p: "p = q * f" by (rule coneE) have "f \ ideal {f}" by (rule ideal.span_base) simp with \q \ P[X]\ f_in_Polys show "p \ ideal {f} \ P[X]" unfolding p by (intro IntI ideal.span_scale Polys_closed_times) qed private definition P_ps where "P_ps = (SOME x. valid_decomp X (snd x) \ standard_decomp d (snd x) \ exact_decomp X 0 (snd x) \ cone_decomp (fst x) (snd x) \ hom_decomp (snd x) \ direct_decomp (ideal F \ P[X]) [ideal {f} \ P[X], fst x])" private definition P where "P = fst P_ps" private definition ps where "ps = snd P_ps" lemma shows valid_ps: "valid_decomp X ps" (is ?thesis1) and std_ps: "standard_decomp d ps" (is ?thesis2) and ext_ps: "exact_decomp X 0 ps" (is ?thesis3) and cn_ps: "cone_decomp P ps" (is ?thesis4) and hom_ps: "hom_decomp ps" (is ?thesis5) and decomp_F: "direct_decomp (ideal F \ P[X]) [ideal {f} \ P[X], P]" (is ?thesis6) proof - note fin_X moreover from fin_F have "finite (F - {f})" by simp moreover from F_sub have "F - {f} \ P[X]" by blast ultimately obtain P' ps' where 1: "valid_decomp X ps'" and 2: "standard_decomp d ps'" and 3: "cone_decomp P' ps'" and 40: "(\f'. f' \ F - {f} \ homogeneous f') \ hom_decomp ps'" and 50: "direct_decomp (ideal (insert f (F - {f})) \ P[X]) [ideal {f} \ P[X], P']" using f_in_Polys f_max by (rule ideal_decompE) blast+ have 4: "hom_decomp ps'" by (intro 40 hom_F) simp from 50 f_in have 5: "direct_decomp (ideal F \ P[X]) [ideal {f} \ P[X], P']" by (simp add: insert_absorb) let ?ps = "exact X (poly_deg f) ps'" from fin_X 1 2 have "valid_decomp X ?ps" and "standard_decomp d ?ps" and "exact_decomp X 0 ?ps" by (rule exact)+ moreover from fin_X 1 2 3 have "cone_decomp P' ?ps" by (rule cone_decomp_exact) moreover from fin_X 1 2 4 have "hom_decomp ?ps" by (rule hom_decomp_exact) ultimately have "valid_decomp X (snd (P', ?ps)) \ standard_decomp d (snd (P', ?ps)) \ exact_decomp X 0 (snd (P', ?ps)) \ cone_decomp (fst (P', ?ps)) (snd (P', ?ps)) \ hom_decomp (snd (P', ?ps)) \ direct_decomp (ideal F \ P[X]) [ideal {f} \ P[X], fst (P', ?ps)]" using 5 by simp hence "?thesis1 \ ?thesis2 \ ?thesis3 \ ?thesis4 \ ?thesis5 \ ?thesis6" unfolding P_def ps_def P_ps_def by (rule someI) thus ?thesis1 and ?thesis2 and ?thesis3 and ?thesis4 and ?thesis5 and ?thesis6 by simp_all qed lemma P_sub: "P \ P[X]" using valid_ps cn_ps by (rule valid_cone_decomp_subset_Polys) lemma ps_not_Nil: "ps\<^sub>+ \ []" proof assume "ps\<^sub>+ = []" have "Keys P \ (\hU\set ps. keys (fst hU))" (is "_ \ ?A") proof fix t assume "t \ Keys P" then obtain p where "p \ P" and "t \ keys p" by (rule in_KeysE) from cn_ps have "direct_decomp P (map cone ps)" by (rule cone_decompD) then obtain qs where qs: "qs \ listset (map cone ps)" and p: "p = sum_list qs" using \p \ P\ by (rule direct_decompE) from \t \ keys p\ keys_sum_list_subset have "t \ Keys (set qs)" unfolding p .. then obtain q where "q \ set qs" and "t \ keys q" by (rule in_KeysE) from this(1) obtain i where "i < length qs" and "q = qs ! i" by (metis in_set_conv_nth) with qs have "i < length ps" and "q \ (map cone ps) ! i" by (simp_all add: listsetD del: nth_map) hence "q \ cone (ps ! i)" by simp obtain h U where eq: "ps ! i = (h, U)" using prod.exhaust by blast from \i < length ps\ this[symmetric] have "(h, U) \ set ps" by simp have "U = {}" proof (rule ccontr) assume "U \ {}" with \(h, U) \ set ps\ have "(h, U) \ set (ps\<^sub>+)" by (simp add: pos_decomp_def) with \ps\<^sub>+ = []\ show False by simp qed with \q \ cone (ps ! i)\ have "q \ range (\c. c \ h)" by (simp only: eq cone_empty) then obtain c where "q = c \ h" .. also have "keys \ \ keys h" by (fact keys_map_scale_subset) finally have "t \ keys h" using \t \ keys q\ .. hence "t \ keys (fst (h, U))" by simp with \(h, U) \ set ps\ show "t \ ?A" .. qed moreover from finite_set finite_keys have "finite ?A" by (rule finite_UN_I) ultimately have "finite (Keys P)" by (rule finite_subset) have "\q\ideal F. q \ P[X] \ q \ 0 \ \ lpp f adds lpp q" proof (rule ccontr) assume "\ (\q\ideal F. q \ P[X] \ q \ 0 \ \ lpp f adds lpp q)" hence adds: "lpp f adds lpp q" if "q \ ideal F" and "q \ P[X]" and "q \ 0" for q using that by blast from fin_X _ F_sub have "ideal {f} = ideal F" proof (rule punit.pmdl_eqI_adds_lt_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]) from f_in_Polys show "{f} \ P[X]" by simp next from f_in have "{f} \ F" by simp thus "ideal {f} \ ideal F" by (rule ideal.span_mono) next fix q assume "q \ ideal F" and "q \ P[X]" and "q \ 0" hence "lpp f adds lpp q" by (rule adds) with f_not_0 show "\g\{f}. g \ 0 \ lpp g adds lpp q" by blast qed with ideal_f_neq show False .. qed then obtain q0 where "q0 \ ideal F" and "q0 \ P[X]" and "q0 \ 0" and nadds_q0: "\ lpp f adds lpp q0" by blast define q where "q = hom_component q0 (deg_pm (lpp q0))" from hom_F \q0 \ ideal F\ have "q \ ideal F" unfolding q_def by (rule homogeneous_ideal) from homogeneous_set_Polys \q0 \ P[X]\ have "q \ P[X]" unfolding q_def by (rule homogeneous_setD) from \q0 \ 0\ have "q \ 0" and "lpp q = lpp q0" unfolding q_def by (rule hom_component_lpp)+ from nadds_q0 this(2) have nadds_q: "\ lpp f adds lpp q" by simp have hom_q: "homogeneous q" by (simp only: q_def homogeneous_hom_component) from nadds_q obtain x where x: "\ lookup (lpp f) x \ lookup (lpp q) x" by (auto simp add: adds_poly_mapping le_fun_def) obtain y where "y \ X" and "y \ x" proof - from n_gr_1 have "2 \ n" by simp then obtain Y where "Y \ X" and "card Y = 2" by (rule card_geq_ex_subset) from this(2) obtain u v where "u \ v" and "Y = {u, v}" by (rule card_2_E) from this obtain y where "y \ Y" and "y \ x" by blast from this(1) \Y \ X\ have "y \ X" .. thus ?thesis using \y \ x\ .. qed define q' where "q' = (\k. punit.monom_mult 1 (Poly_Mapping.single y k) q)" have inj1: "inj q'" by (auto intro!: injI simp: q'_def \q \ 0\ dest: punit.monom_mult_inj_2 monomial_inj) have q'_in: "q' k \ ideal F \ P[X]" for k unfolding q'_def using \q \ ideal F\ \q \ P[X]\ \y \ X\ by (intro IntI punit.pmdl_closed_monom_mult[simplified] Polys_closed_monom_mult PPs_closed_single) have lpp_q': "lpp (q' k) = Poly_Mapping.single y k + lpp q" for k using \q \ 0\ by (simp add: q'_def punit.lt_monom_mult) have inj2: "inj_on (deg_pm \ lpp) (range q')" by (auto intro!: inj_onI simp: lpp_q' deg_pm_plus deg_pm_single dest: monomial_inj) have "(deg_pm \ lpp) ` range q' \ deg_pm ` Keys P" proof fix d assume "d \ (deg_pm \ lpp) ` range q'" then obtain k where d: "d = deg_pm (lpp (q' k))" (is "_ = deg_pm ?t") by auto from hom_q have hom_q': "homogeneous (q' k)" by (simp add: q'_def homogeneous_monom_mult) from \q \ 0\ have "q' k \ 0" by (simp add: q'_def punit.monom_mult_eq_zero_iff) hence "?t \ keys (q' k)" by (rule punit.lt_in_keys) with hom_q' have deg_q': "d = poly_deg (q' k)" unfolding d by (rule homogeneousD_poly_deg) from decomp_F q'_in obtain qs where "qs \ listset [ideal {f} \ P[X], P]" and "q' k = sum_list qs" by (rule direct_decompE) moreover from this(1) obtain f0 p0 where f0: "f0 \ ideal {f} \ P[X]" and p0: "p0 \ P" and "qs = [f0, p0]" by (rule listset_doubletonE) ultimately have q': "q' k = f0 + p0" by simp define f1 where "f1 = hom_component f0 d" define p1 where "p1 = hom_component p0 d" from hom_q have "homogeneous (q' k)" by (simp add: q'_def homogeneous_monom_mult) hence "q' k = hom_component (q' k) d" by (simp add: hom_component_of_homogeneous deg_q') also have "\ = f1 + p1" by (simp only: q' hom_component_plus f1_def p1_def) finally have "q' k = f1 + p1" . have "keys p1 \ {}" proof assume "keys p1 = {}" with \q' k = f1 + p1\ \q' k \ 0\ have t: "?t = lpp f1" and "f1 \ 0" by simp_all from f0 have "f0 \ ideal {f}" by simp with _ have "f1 \ ideal {f}" unfolding f1_def by (rule homogeneous_ideal) (simp add: hom_f) with punit.is_Groebner_basis_singleton obtain g where "g \ {f}" and "lpp g adds lpp f1" using \f1 \ 0\ by (rule punit.GB_adds_lt[simplified]) hence "lpp f adds ?t" by (simp add: t) hence "lookup (lpp f) x \ lookup ?t x" by (simp add: adds_poly_mapping le_fun_def) also have "\ = lookup (lpp q) x" by (simp add: lpp_q' lookup_add lookup_single \y \ x\) finally have "lookup (lpp f) x \ lookup (lpp q) x" . with x show False .. qed then obtain t where "t \ keys p1" by blast hence "d = deg_pm t" by (simp add: p1_def keys_hom_component) from cn_ps hom_ps have "homogeneous_set P" by (intro homogeneous_set_cone_decomp) hence "p1 \ P" using \p0 \ P\ unfolding p1_def by (rule homogeneous_setD) with \t \ keys p1\ have "t \ Keys P" by (rule in_KeysI) with \d = deg_pm t\ show "d \ deg_pm ` Keys P" by (rule image_eqI) qed moreover from inj1 inj2 have "infinite ((deg_pm \ lpp) ` range q')" by (simp add: finite_image_iff o_def) ultimately have "infinite (deg_pm ` Keys P)" by (rule infinite_super) hence "infinite (Keys P)" by blast thus False using \finite (Keys P)\ .. qed private definition N where "N = normal_form F ` P[X]" private definition qs where "qs = (SOME qs'. valid_decomp X qs' \ standard_decomp 0 qs' \ monomial_decomp qs' \ cone_decomp N qs' \ exact_decomp X 0 qs' \ (\g\punit.reduced_GB F. poly_deg g \ \ qs' 0))" private definition "aa \ \ ps" private definition "bb \ \ qs" private abbreviation (input) "cc \ (\i. aa i + bb i)" lemma shows valid_qs: "valid_decomp X qs" (is ?thesis1) and std_qs: "standard_decomp 0 qs" (is ?thesis2) and mon_qs: "monomial_decomp qs" (is ?thesis3) and hom_qs: "hom_decomp qs" (is ?thesis6) and cn_qs: "cone_decomp N qs" (is ?thesis4) and ext_qs: "exact_decomp X 0 qs" (is ?thesis5) and deg_RGB: "g \ punit.reduced_GB F \ poly_deg g \ bb 0" proof - from fin_X F_sub obtain qs' where 1: "valid_decomp X qs'" and 2: "standard_decomp 0 qs'" and 3: "monomial_decomp qs'" and 4: "cone_decomp (normal_form F ` P[X]) qs'" and 5: "exact_decomp X 0 qs'" and 60: "\g. (\f. f \ F \ homogeneous f) \ g \ punit.reduced_GB F \ poly_deg g \ \ qs' 0" by (rule normal_form_exact_decompE) blast from hom_F have "\g. g \ punit.reduced_GB F \ poly_deg g \ \ qs' 0" by (rule 60) with 1 2 3 4 5 have "valid_decomp X qs' \ standard_decomp 0 qs' \ monomial_decomp qs' \ cone_decomp N qs' \ exact_decomp X 0 qs' \ (\g\punit.reduced_GB F. poly_deg g \ \ qs' 0)" by (simp add: N_def) hence "?thesis1 \ ?thesis2 \ ?thesis3 \ ?thesis4 \ ?thesis5 \ (\g\punit.reduced_GB F. poly_deg g \ bb 0)" unfolding qs_def bb_def by (rule someI) thus ?thesis1 and ?thesis2 and ?thesis3 and ?thesis4 and ?thesis5 and "g \ punit.reduced_GB F \ poly_deg g \ bb 0" by simp_all from \?thesis3\ show ?thesis6 by (rule monomial_decomp_imp_hom_decomp) qed lemma N_sub: "N \ P[X]" using valid_qs cn_qs by (rule valid_cone_decomp_subset_Polys) lemma decomp_Polys: "direct_decomp P[X] [ideal {f} \ P[X], P, N]" proof - from fin_X F_sub have "direct_decomp P[X] [ideal F \ P[X], N]" unfolding N_def by (rule direct_decomp_ideal_normal_form) hence "direct_decomp P[X] ([N] @ [ideal {f} \ P[X], P])" using decomp_F by (rule direct_decomp_direct_decomp) - hence "direct_decomp P[X] ([ideal {f} \ P[X], P] @ [N])" using perm_append_swap - by (rule direct_decomp_perm) + hence "direct_decomp P[X] ([ideal {f} \ P[X], P] @ [N])" + by (rule direct_decomp_perm) simp thus ?thesis by simp qed lemma aa_Suc_n [simp]: "aa (Suc n) = d" proof - from fin_X ext_ps le_refl have "aa (Suc n) = \ ps" unfolding aa_def by (rule \_card_X) also from fin_X valid_ps std_ps ps_not_Nil have "\ = d" by (rule \_nonempty_unique) finally show ?thesis . qed lemma bb_Suc_n [simp]: "bb (Suc n) = 0" proof - from fin_X ext_qs le_refl have "bb (Suc n) = \ qs" unfolding bb_def by (rule \_card_X) also from std_qs have "\ = 0" unfolding \_def[OF fin_X] by (rule Least_eq_0) finally show ?thesis . qed lemma Hilbert_fun_X: assumes "d \ z" shows "Hilbert_fun (P[X]::(_ \\<^sub>0 'a) set) z = ((z - d) + (n - 1)) choose (n - 1) + Hilbert_fun P z + Hilbert_fun N z" proof - define ss where "ss = [ideal {f} \ P[X], P, N]" have "homogeneous_set A \ phull.subspace A" if "A \ set ss" for A proof - from that have "A = ideal {f} \ P[X] \ A = P \ A = N" by (simp add: ss_def) thus ?thesis proof (elim disjE) assume A: "A = ideal {f} \ P[X]" show ?thesis unfolding A by (intro conjI homogeneous_set_IntI phull.subspace_inter homogeneous_set_homogeneous_ideal homogeneous_set_Polys subspace_ideal subspace_Polys) (simp add: hom_f) next assume A: "A = P" from cn_ps hom_ps show ?thesis unfolding A by (intro conjI homogeneous_set_cone_decomp subspace_cone_decomp) next assume A: "A = N" from cn_qs hom_qs show ?thesis unfolding A by (intro conjI homogeneous_set_cone_decomp subspace_cone_decomp) qed qed hence 1: "\A. A \ set ss \ homogeneous_set A" and 2: "\A. A \ set ss \ phull.subspace A" by simp_all have "Hilbert_fun (P[X]::(_ \\<^sub>0 'a) set) z = (\p\set ss. Hilbert_fun p z)" using fin_X subset_refl decomp_Polys unfolding ss_def proof (rule Hilbert_fun_direct_decomp) fix A assume "A \ set [ideal {f} \ P[X], P, N]" hence "A \ set ss" by (simp only: ss_def) thus "homogeneous_set A" and "phull.subspace A" by (rule 1, rule 2) qed also have "\ = (\p\set ss. count_list ss p * Hilbert_fun p z)" using refl proof (rule sum.cong) fix p assume "p \ set ss" hence "count_list ss p \ 0" by (simp only: count_list_eq_0_iff not_not) hence "count_list ss p = 1 \ 1 < count_list ss p" by auto thus "Hilbert_fun p z = count_list ss p * Hilbert_fun p z" proof assume "1 < count_list ss p" with decomp_Polys have "p = {0}" unfolding ss_def[symmetric] using phull.subspace_0 by (rule direct_decomp_repeated_eq_zero) (rule 2) thus ?thesis by simp qed simp qed also have "\ = sum_list (map (\p. Hilbert_fun p z) ss)" by (rule sym) (rule sum_list_map_eq_sum_count) also have "\ = Hilbert_fun (cone (f, X)) z + Hilbert_fun P z + Hilbert_fun N z" by (simp add: ss_def ideal_Int_Polys_eq_cone) also have "Hilbert_fun (cone (f, X)) z = (z - d + (n - 1)) choose (n - 1)" using f_not_0 f_in_Polys fin_X hom_f X_not_empty by (simp add: Hilbert_fun_cone_nonempty assms) finally show ?thesis . qed lemma dube_eq_0: "(\z::int. (z + int n - 1) gchoose (n - 1)) = (\z::int. ((z - d + n - 1) gchoose (n - 1)) + Hilbert_poly aa z + Hilbert_poly bb z)" (is "?f = ?g") proof (rule poly_fun_eqI_ge) fix z::int let ?z = "nat z" assume "max (aa 0) (bb 0) \ z" hence "aa 0 \ nat z" and "bb 0 \ nat z" and "0 \ z" by simp_all from this(3) have int_z: "int ?z = z" by simp have "d \ aa 0" unfolding aa_Suc_n[symmetric] using fin_X le0 unfolding aa_def by (rule \_decreasing) hence "d \ ?z" using \aa 0 \ nat z\ by (rule le_trans) hence int_zd: "int (?z - d) = z - int d" using int_z by linarith from \d \ ?z\ have "Hilbert_fun (P[X]::(_ \\<^sub>0 'a) set) ?z = ((?z - d) + (n - 1)) choose (n - 1) + Hilbert_fun P ?z + Hilbert_fun N ?z" by (rule Hilbert_fun_X) also have "int \ = (z - d + (n - 1)) gchoose (n - 1) + Hilbert_poly aa z + Hilbert_poly bb z" using X_not_empty valid_ps hom_ps cn_ps std_ps ext_ps \aa 0 \ nat z\ valid_qs hom_qs cn_qs std_qs ext_qs \bb 0 \ nat z\ \0 \ z\ by (simp add: Hilbert_fun_eq_Hilbert_poly int_z aa_def bb_def int_binomial int_zd) finally show "?f z = ?g z" using fin_X X_not_empty \0 \ z\ by (simp add: Hilbert_fun_Polys int_binomial) smt qed (simp_all add: poly_fun_Hilbert_poly) corollary dube_eq_1: "(\z::int. (z + int n - 1) gchoose (n - 1)) = (\z::int. ((z - d + n - 1) gchoose (n - 1)) + ((z - d + n) gchoose n) + ((z + n) gchoose n) - 2 - (\i=1..n. ((z - aa i + i - 1) gchoose i) + ((z - bb i + i - 1) gchoose i)))" by (simp only: dube_eq_0) (auto simp: Hilbert_poly_def Let_def sum.distrib) lemma dube_eq_2: assumes "j < n" shows "(\z::int. (z + int n - int j - 1) gchoose (n - j - 1)) = (\z::int. ((z - d + n - int j - 1) gchoose (n - j - 1)) + ((z - d + n - j) gchoose (n - j)) + ((z + n - j) gchoose (n - j)) - 2 - (\i=Suc j..n. ((z - aa i + i - j - 1) gchoose (i - j)) + ((z - bb i + i - j - 1) gchoose (i - j))))" (is "?f = ?g") proof - let ?h = "\z i. ((z + (int i - aa i - 1)) gchoose i) + ((z + (int i - bb i - 1)) gchoose i)" let ?hj = "\z i. ((z + (int i - aa i - 1) - j) gchoose (i - j)) + ((z + (int i - bb i - 1) - j) gchoose (i - j))" from assms have 1: "j \ n - Suc 0" and 2: "j \ n" by simp_all have eq1: "(bw_diff ^^ j) (\z. \i=1..j. ?h z i) = (\_. if j = 0 then 0 else 2)" proof (cases j) case 0 thus ?thesis by simp next case (Suc j0) hence "j \ 0" by simp have "(\z::int. \i = 1..j. ?h z i) = (\z::int. (\i = 1..j0. ?h z i) + ?h z j)" by (simp add: \j = Suc j0\) moreover have "(bw_diff ^^ j) \ = (\z::int. (\i = 1..j0. (bw_diff ^^ j) (\z. ?h z i) z) + 2)" by (simp add: bw_diff_gbinomial_pow) moreover have "(\i = 1..j0. (bw_diff ^^ j) (\z. ?h z i) z) = (\i = 1..j0. 0)" for z::int using refl proof (rule sum.cong) fix i assume "i \ {1..j0}" hence "\ j \ i" by (simp add: \j = Suc j0\) thus "(bw_diff ^^ j) (\z. ?h z i) z = 0" by (simp add: bw_diff_gbinomial_pow) qed ultimately show ?thesis by (simp add: \j \ 0\) qed have eq2: "(bw_diff ^^ j) (\z. \i=Suc j..n. ?h z i) = (\z. (\i=Suc j..n. ?hj z i))" proof - have "(bw_diff ^^ j) (\z. \i=Suc j..n. ?h z i) = (\z. \i=Suc j..n. (bw_diff ^^ j) (\z. ?h z i) z)" by simp also have "\ = (\z. (\i=Suc j..n. ?hj z i))" proof (intro ext sum.cong) fix z i assume "i \ {Suc j..n}" hence "j \ i" by simp thus "(bw_diff ^^ j) (\z. ?h z i) z = ?hj z i" by (simp add: bw_diff_gbinomial_pow) qed (fact refl) finally show ?thesis . qed from 1 have "?f = (bw_diff ^^ j) (\z::int. (z + (int n - 1)) gchoose (n - 1))" by (simp add: bw_diff_gbinomial_pow) (simp only: algebra_simps) also have "\ = (bw_diff ^^ j) (\z::int. (z + int n - 1) gchoose (n - 1))" by (simp only: algebra_simps) also have "\ = (bw_diff ^^ j) (\z::int. ((z - d + n - 1) gchoose (n - 1)) + ((z - d + n) gchoose n) + ((z + n) gchoose n) - 2 - (\i=1..n. ((z - aa i + i - 1) gchoose i) + ((z - bb i + i - 1) gchoose i)))" by (simp only: dube_eq_1) also have "\ = (bw_diff ^^ j) (\z::int. ((z + (int n - d - 1)) gchoose (n - 1)) + ((z + (int n - d)) gchoose n) + ((z + n) gchoose n) - 2 - (\i=1..n. ?h z i))" by (simp only: algebra_simps) also have "\ = (\z::int. ((z + (int n - d - 1) - j) gchoose (n - 1 - j)) + ((z + (int n - d) - j) gchoose (n - j)) + ((z + n - j) gchoose (n - j)) - (if j = 0 then 2 else 0) - (bw_diff ^^ j) (\z. \i=1..n. ?h z i) z)" using 1 2 by (simp add: bw_diff_const_pow bw_diff_gbinomial_pow del: bw_diff_sum_pow) also from \j \ n\ have "(\z. \i=1..n. ?h z i) = (\z. (\i=1..j. ?h z i) + (\i=Suc j..n. ?h z i))" by (simp add: sum_split_nat_ivl) also have "(bw_diff ^^ j) \ = (\z. (bw_diff ^^ j) (\z. \i=1..j. ?h z i) z + (bw_diff ^^ j) (\z. \i=Suc j..n. ?h z i) z)" by (simp only: bw_diff_plus_pow) also have "\ = (\z. (if j = 0 then 0 else 2) + (\i=Suc j..n. ?hj z i))" by (simp only: eq1 eq2) finally show ?thesis by (simp add: algebra_simps) qed lemma dube_eq_3: assumes "j < n" shows "(1::int) = (- 1)^(n - Suc j) * ((int d - 1) gchoose (n - Suc j)) + (- 1)^(n - j) * ((int d - 1) gchoose (n - j)) - 1 - (\i=Suc j..n. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" proof - from assms have 1: "int (n - Suc j) = int n - j - 1" and 2: "int (n - j) = int n - j" by simp_all from assms have "int n - int j - 1 = int (n - j - 1)" by simp hence eq1: "int n - int j - 1 gchoose (n - Suc j) = 1" by simp from assms have "int n - int j = int (n - j)" by simp hence eq2: "int n - int j gchoose (n - j) = 1" by simp have eq3: "int n - d - j - 1 gchoose (n - Suc j) = (- 1)^(n - Suc j) * (int d - 1 gchoose (n - Suc j))" by (simp add: gbinomial_int_negated_upper[of "int n - d - j - 1"] 1) have eq4: "int n - d - j gchoose (n - j) = (- 1)^(n - j) * (int d - 1 gchoose (n - j))" by (simp add: gbinomial_int_negated_upper[of "int n - d - j"] 2) have eq5: "(\i = Suc j..n. int i - aa i - j - 1 gchoose (i - j) + (int i - bb i - j - 1 gchoose (i - j))) = (\i=Suc j..n. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" using refl proof (rule sum.cong) fix i assume "i \ {Suc j..n}" hence "j \ i" by simp hence 3: "int (i - j) = int i - j" by simp show "int i - aa i - j - 1 gchoose (i - j) + (int i - bb i - j - 1 gchoose (i - j)) = (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j)))" by (simp add: gbinomial_int_negated_upper[of "int i - aa i - j - 1"] gbinomial_int_negated_upper[of "int i - bb i - j - 1"] 3 distrib_left) qed from fun_cong[OF dube_eq_2, OF assms, of 0] show ?thesis by (simp add: eq1 eq2 eq3 eq4 eq5) qed lemma dube_aux_1: assumes "(h, {}) \ set ps \ set qs" shows "poly_deg h < max (aa 1) (bb 1)" proof (rule ccontr) define z where "z = poly_deg h" assume "\ z < max (aa 1) (bb 1)" let ?S = "\A. {h. (h, {}) \ A \ poly_deg h = z}" have fin: "finite (?S A)" if "finite A" for A::"((('x \\<^sub>0 nat) \\<^sub>0 'a) \ 'x set) set" proof - have "(\t. (t, {})) ` ?S A \ A" by blast hence "finite ((\t. (t, {}::'x set)) ` ?S A)" using that by (rule finite_subset) moreover have "inj_on (\t. (t, {}::'x set)) (?S A)" by (rule inj_onI) simp ultimately show ?thesis by (rule finite_imageD) qed from finite_set have 1: "finite (?S (set ps))" by (rule fin) from finite_set have 2: "finite (?S (set qs))" by (rule fin) from \\ z < max (aa 1) (bb 1)\ have "aa 1 \ z" and "bb 1 \ z" by simp_all have "d \ aa 1" unfolding aa_Suc_n[symmetric] aa_def using fin_X by (rule \_decreasing) simp hence "d \ z" using \aa 1 \ z\ by (rule le_trans) hence eq: "int (z - d) = int z - int d" by simp from \d \ z\ have "Hilbert_fun (P[X]::(_ \\<^sub>0 'a) set) z = ((z - d) + (n - 1)) choose (n - 1) + Hilbert_fun P z + Hilbert_fun N z" by (rule Hilbert_fun_X) also have "int \ = ((int z - d + (n - 1)) gchoose (n - 1) + Hilbert_poly aa z + Hilbert_poly bb z) + (int (card (?S (set ps))) + int (card (?S (set qs))))" using X_not_empty valid_ps hom_ps cn_ps std_ps ext_ps \aa 1 \ z\ valid_qs hom_qs cn_qs std_qs ext_qs \bb 1 \ z\ by (simp add: Hilbert_fun_eq_Hilbert_poly_plus_card aa_def bb_def int_binomial eq) finally have "((int z - d + n - 1) gchoose (n - 1) + Hilbert_poly aa z + Hilbert_poly bb z) + (int (card (?S (set ps))) + int (card (?S (set qs)))) = int z + n - 1 gchoose (n - 1)" using fin_X X_not_empty by (simp add: Hilbert_fun_Polys int_binomial algebra_simps) also have "\ = (int z - d + n - 1) gchoose (n - 1) + Hilbert_poly aa z + Hilbert_poly bb z" by (fact dube_eq_0[THEN fun_cong]) finally have "int (card (?S (set ps))) + int (card (?S (set qs))) = 0" by simp hence "card (?S (set ps)) = 0" and "card (?S (set qs)) = 0" by simp_all with 1 2 have "?S (set ps \ set qs) = {}" by auto moreover from assms have "h \ ?S (set ps \ set qs)" by (simp add: z_def) ultimately have "h \ {}" by (rule subst) thus False by simp qed lemma shows aa_n: "aa n = d" and bb_n: "bb n = 0" and bb_0: "bb 0 \ max (aa 1) (bb 1)" proof - let ?j = "n - Suc 0" from n_gr_0 have "?j < n" and eq1: "Suc ?j = n" and eq2: "n - ?j = 1" by simp_all from this(1) have "(1::int) = (- 1)^(n - Suc ?j) * ((int d - 1) gchoose (n - Suc ?j)) + (- 1)^(n - ?j) * ((int d - 1) gchoose (n - ?j)) - 1 - (\i=Suc ?j..n. (- 1)^(i - ?j) * ((int (aa i) gchoose (i - ?j)) + (int (bb i) gchoose (i - ?j))))" by (rule dube_eq_3) hence eq: "aa n + bb n = d" by (simp add: eq1 eq2) hence "aa n \ d" by simp moreover have "d \ aa n" unfolding aa_Suc_n[symmetric] aa_def using fin_X by (rule \_decreasing) simp ultimately show "aa n = d" by (rule antisym) with eq show "bb n = 0" by simp have "bb 0 = \ qs 0" by (simp only: bb_def) also from fin_X have "\ \ max (aa 1) (bb 1)" (is "_ \ ?m") proof (rule \_le) from fin_X ext_qs have "\ qs = bb (Suc n)" by (simp add: \_card_X bb_def) also have "\ \ bb 1" unfolding bb_def using fin_X by (rule \_decreasing) simp also have "\ \ ?m" by (rule max.cobounded2) finally show "\ qs \ ?m" . next fix h U assume "(h, U) \ set qs" show "poly_deg h < ?m" proof (cases "card U = 0") case True from fin_X valid_qs \(h, U) \ set qs\ have "finite U" by (rule valid_decompD_finite) with True have "U = {}" by simp with \(h, U) \ set qs\ have "(h, {}) \ set ps \ set qs" by simp thus ?thesis by (rule dube_aux_1) next case False hence "1 \ card U" by simp with fin_X \(h, U) \ set qs\ have "poly_deg h < bb 1" unfolding bb_def by (rule \) also have "\ \ ?m" by (rule max.cobounded2) finally show ?thesis . qed qed finally show "bb 0 \ ?m" . qed lemma dube_eq_4: assumes "j < n" shows "(1::int) = 2 * (- 1)^(n - Suc j) * ((int d - 1) gchoose (n - Suc j)) - 1 - (\i=Suc j..n-1. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" proof - from assms have "Suc j \ n" and "0 < n" and 1: "Suc (n - Suc j) = n - j" by simp_all have 2: "(- 1) ^ (n - Suc j) = - ((- (1::int)) ^ (n - j))" by (simp flip: 1) from assms have "(1::int) = (- 1)^(n - Suc j) * ((int d - 1) gchoose (n - Suc j)) + (- 1)^(n - j) * ((int d - 1) gchoose (n - j)) - 1 - (\i=Suc j..n. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" by (rule dube_eq_3) also have "\ = (- 1)^(n - Suc j) * ((int d - 1) gchoose (n - Suc j)) + (- 1)^(n - j) * ((int d - 1) gchoose (n - j)) - 1 - (- 1)^(n - j) * ((int (aa n) gchoose (n - j)) + (int (bb n) gchoose (n - j))) - (\i=Suc j..n-1. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" using \0 < n\ \Suc j \ n\ by (simp only: sum_tail_nat) also have "\ = (- 1)^(n - Suc j) * ((int d - 1) gchoose (n - Suc j)) + (- 1)^(n - j) * (((int d - 1) gchoose (n - j)) - (int d gchoose (n - j))) - 1 - (\i=Suc j..n-1. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" using assms by (simp add: aa_n bb_n gbinomial_0_left right_diff_distrib) also have "(- 1)^(n - j) * (((int d - 1) gchoose (n - j)) - (int d gchoose (n - j))) = (- 1)^(n - Suc j) * (((int d - 1 + 1) gchoose (Suc (n - Suc j))) - ((int d - 1) gchoose (Suc (n - Suc j))))" by (simp add: 1 2 flip: mult_minus_right) also have "\ = (- 1)^(n - Suc j) * ((int d - 1) gchoose (n - Suc j))" by (simp only: gbinomial_int_Suc_Suc, simp) finally show ?thesis by simp qed lemma cc_Suc: assumes "j < n - 1" shows "int (cc (Suc j)) = 2 + 2 * (- 1)^(n - j) * ((int d - 1) gchoose (n - Suc j)) + (\i=j+2..n-1. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" proof - from assms have "j < n" and "Suc j \ n - 1" by simp_all hence "n - j = Suc (n - Suc j)" by simp hence eq: "(- 1) ^ (n - Suc j) = - ((- (1::int)) ^ (n - j))" by simp from \j < n\ have "(1::int) = 2 * (- 1)^(n - Suc j) * ((int d - 1) gchoose (n - Suc j)) - 1 - (\i=Suc j..n-1. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" by (rule dube_eq_4) also have "\ = cc (Suc j) - 2 * (- 1)^(n - j) * ((int d - 1) gchoose (n - Suc j)) - 1 - (\i=j+2..n-1. (- 1)^(i - j) * ((int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))))" using \Suc j \ n - 1\ by (simp add: sum.atLeast_Suc_atMost eq) finally show ?thesis by simp qed lemma cc_n_minus_1: "cc (n - 1) = 2 * d" proof - let ?j = "n - 2" from n_gr_1 have 1: "Suc ?j = n - 1" and "?j < n - 1" and 2: "Suc (n - 1) = n" and 3: "n - (n - Suc 0) = Suc 0" and 4: "n - ?j = 2" by simp_all have "int (cc (n - 1)) = int (cc (Suc ?j))" by (simp only: 1) also from \?j < n - 1\ have "\ = 2 + 2 * (- 1) ^ (n - ?j) * (int d - 1 gchoose (n - Suc ?j)) + (\i = ?j+2..n-1. (- 1) ^ (i - ?j) * (int (aa i) gchoose (i - ?j) + (int (bb i) gchoose (i - ?j))))" by (rule cc_Suc) also have "\ = int (2 * d)" by (simp add: 1 2 3 4) finally show ?thesis by (simp only: int_int_eq) qed text \Since the case @{prop "n = 2"} is settled, we can concentrate on @{prop "2 < n"} now.\ context assumes n_gr_2: "2 < n" begin lemma cc_n_minus_2: "cc (n - 2) \ d\<^sup>2 + 2 * d" proof - let ?j = "n - 3" from n_gr_2 have 1: "Suc ?j = n - 2" and "?j < n - 1" and 2: "Suc (n - 2) = n - Suc 0" and 3: "n - (n - 2) = 2" and 4: "n - ?j = 3" by simp_all have "int (cc (n - 2)) = int (cc (Suc ?j))" by (simp only: 1) also from \?j < n - 1\ have "\ = 2 + 2 * (- 1) ^ (n - ?j) * (int d - 1 gchoose (n - Suc ?j)) + (\i = ?j+2..n-1. (- 1) ^ (i - ?j) * (int (aa i) gchoose (i - ?j) + (int (bb i) gchoose (i - ?j))))" by (rule cc_Suc) also have "\ = (2 - 2 * (int d - 1 gchoose 2)) + ((int (aa (n - 1)) gchoose 2) + (int (bb (n - 1)) gchoose 2))" by (simp add: 1 2 3 4) also have "\ \ (2 - 2 * (int d - 1 gchoose 2)) + (2 * int d gchoose 2)" proof (rule add_left_mono) have "int (aa (n - 1)) gchoose 2 + (int (bb (n - 1)) gchoose 2) \ int (aa (n - 1)) + int (bb (n - 1)) gchoose 2" by (rule gbinomial_int_plus_le) simp_all also have "\ = int (2 * d) gchoose 2" by (simp flip: cc_n_minus_1) also have "\ = 2 * int d gchoose 2" by (simp add: int_ops(7)) finally show "int (aa (n - 1)) gchoose 2 + (int (bb (n - 1)) gchoose 2) \ 2 * int d gchoose 2" . qed also have "\ = 2 - fact 2 * (int d - 1 gchoose 2) + (2 * int d gchoose 2)" by (simp only: fact_2) also have "\ = 2 - (int d - 1) * (int d - 2) + (2 * int d gchoose 2)" by (simp only: gbinomial_int_mult_fact) (simp add: numeral_2_eq_2 prod.atLeast0_lessThan_Suc) also have "\ = 2 - (int d - 1) * (int d - 2) + int d * (2 * int d - 1)" by (simp add: gbinomial_prod_rev numeral_2_eq_2 prod.atLeast0_lessThan_Suc) also have "\ = int (d\<^sup>2 + 2 * d)" by (simp add: power2_eq_square) (simp only: algebra_simps) finally show ?thesis by (simp only: int_int_eq) qed lemma cc_Suc_le: assumes "j < n - 3" shows "int (cc (Suc j)) \ 2 + (int (cc (j + 2)) gchoose 2) + (\i=j+4..n-1. int (cc i) gchoose (i - j))" \\Could be proved without coercing to @{typ int}, because everything is non-negative.\ proof - let ?f = "\i j. (int (aa i) gchoose (i - j)) + (int (bb i) gchoose (i - j))" let ?S = "\x y. (\i=j+x..n-y. (- 1)^(i - j) * ?f i j)" let ?S3 = "\x y. (\i=j+x..n-y. (int (cc i) gchoose (i - j)))" have ie1: "int (aa i) gchoose k + (int (bb i) gchoose k) \ int (cc i) gchoose k" if "0 < k" for i k proof - from that have "int (aa i) gchoose k + (int (bb i) gchoose k) \ int (aa i) + int (bb i) gchoose k" by (rule gbinomial_int_plus_le) simp_all also have "\ = int (cc i) gchoose k" by simp finally show ?thesis . qed from d_gr_0 have "0 \ int d - 1" by simp from assms have "0 < n - Suc j" by simp have f_nonneg: "0 \ ?f i j" for i by (simp add: gbinomial_int_nonneg) show ?thesis proof (cases "n = j + 4") case True hence j: "j = n - 4" by simp have 1: "n - Suc j = 3" and "j < n - 1" and 2: "Suc (n - 3) = Suc (Suc j)" and 3: "n - (n - 3) = 3" and 4: "n - j = 4" and 5: "n - Suc 0 = Suc (Suc (Suc j))" and 6: "n - 2 = Suc (Suc j)" by (simp_all add: True) from \j < n - 1\ have "int (cc (Suc j)) = 2 + 2 * (- 1) ^ (n - j) * (int d - 1 gchoose (n - Suc j)) + (\i = j+2..n-1. (- 1) ^ (i - j) * (int (aa i) gchoose (i - j) + (int (bb i) gchoose (i - j))))" by (rule cc_Suc) also have "\ = (2 + ((int (aa (n - 2)) gchoose 2) + (int (bb (n - 2)) gchoose 2))) + (2 * (int d - 1 gchoose 3) - ((int (aa (n - 1)) gchoose 3) + (int (bb (n - 1)) gchoose 3)))" by (simp add: 1 2 3 4 5 6) also have "\ \ (2 + ((int (aa (n - 2)) gchoose 2) + (int (bb (n - 2)) gchoose 2))) + 0" proof (rule add_left_mono) from cc_n_minus_1 have eq1: "int (aa (n - 1)) + int (bb (n - 1)) = 2 * int d" by simp hence ie2: "int (aa (n - 1)) \ 2 * int d" by simp from \0 \ int d - 1\ have "int d - 1 gchoose 3 \ int d gchoose 3" by (rule gbinomial_int_mono) simp hence "2 * (int d - 1 gchoose 3) \ 2 * (int d gchoose 3)" by simp also from _ ie2 have "\ \ int (aa (n - 1)) gchoose 3 + (2 * int d - int (aa (n - 1)) gchoose 3)" by (rule binomial_int_ineq_3) simp also have "\ = int (aa (n - 1)) gchoose 3 + (int (bb (n - 1)) gchoose 3)" by (simp flip: eq1) finally show "2 * (int d - 1 gchoose 3) - (int (aa (n - 1)) gchoose 3 + (int (bb (n - 1)) gchoose 3)) \ 0" by simp qed also have "\ = 2 + ((int (aa (n - 2)) gchoose 2) + (int (bb (n - 2)) gchoose 2))" by simp also from ie1 have "\ \ 2 + (int (cc (n - 2)) gchoose 2)" by (rule add_left_mono) simp also have "\ = 2 + (int (cc (j + 2)) gchoose 2) + ?S3 4 1" by (simp add: True) finally show ?thesis . next case False with assms have "j + 4 \ n - 1" by simp from n_gr_1 have "0 < n - 1" by simp from assms have "j + 2 \ n - 1" and "j + 2 \ n - 2" by simp_all hence "n - j = Suc (n - Suc j)" by simp hence 1: "(- 1) ^ (n - Suc j) = - ((- (1::int)) ^ (n - j))" by simp from assms have "j < n - 1" by simp hence "int (cc (Suc j)) = 2 + 2 * (- 1)^(n - j) * ((int d - 1) gchoose (n - Suc j)) + ?S 2 1" by (rule cc_Suc) also have "\ = 2 * (- 1)^(n - j) * ((int d - 1) gchoose (n - Suc j)) + (- 1)^(n - Suc j) * ((int (aa (n - 1)) gchoose (n - Suc j)) + (int (bb (n - 1)) gchoose (n - Suc j))) + (2 + ?S 2 2)" using \0 < n - 1\ \j + 2 \ n - 1\ by (simp only: sum_tail_nat) (simp flip: numeral_2_eq_2) also have "\ \ (int (cc (n - 1)) gchoose (n - Suc j)) + (2 + ?S 2 2)" proof (rule add_right_mono) have rl: "x - y \ x" if "0 \ y" for x y :: int using that by simp have "2 * (- 1)^(n - j) * ((int d - 1) gchoose (n - Suc j)) + (- 1)^(n - Suc j) * ((int (aa (n - 1)) gchoose (n - Suc j)) + (int (bb (n - 1)) gchoose (n - Suc j))) = (-1)^(n - j) * (2 * ((int d - 1) gchoose (n - Suc j)) - (int (aa (n - 1)) gchoose (n - Suc j)) - (int (bb (n - 1)) gchoose (n - Suc j)))" by (simp only: 1 algebra_simps) also have "\ \ (int (cc (n - 1))) gchoose (n - Suc j)" proof (cases "even (n - j)") case True hence "(- 1) ^ (n - j) * (2 * (int d - 1 gchoose (n - Suc j)) - (int (aa (n - 1)) gchoose (n - Suc j)) - (int (bb (n - 1)) gchoose (n - Suc j))) = 2 * (int d - 1 gchoose (n - Suc j)) - ((int (aa (n - 1)) gchoose (n - Suc j)) + (int (bb (n - 1)) gchoose (n - Suc j)))" by simp also have "\ \ 2 * (int d - 1 gchoose (n - Suc j))" by (rule rl) (simp add: gbinomial_int_nonneg) also have "\ = (int d - 1 gchoose (n - Suc j)) + (int d - 1 gchoose (n - Suc j))" by simp also have "\ \ (int d - 1) + (int d - 1) gchoose (n - Suc j)" using \0 < n - Suc j\ \0 \ int d - 1\ \0 \ int d - 1\ by (rule gbinomial_int_plus_le) also have "\ \ 2 * int d gchoose (n - Suc j)" proof (rule gbinomial_int_mono) from \0 \ int d - 1\ show "0 \ int d - 1 + (int d - 1)" by simp qed simp also have "\ = int (cc (n - 1)) gchoose (n - Suc j)" by (simp only: cc_n_minus_1) simp finally show ?thesis . next case False hence "(- 1) ^ (n - j) * (2 * (int d - 1 gchoose (n - Suc j)) - (int (aa (n - 1)) gchoose (n - Suc j)) - (int (bb (n - 1)) gchoose (n - Suc j))) = ((int (aa (n - 1)) gchoose (n - Suc j)) + (int (bb (n - 1)) gchoose (n - Suc j))) - 2 * (int d - 1 gchoose (n - Suc j))" by simp also have "\ \ (int (aa (n - 1)) gchoose (n - Suc j)) + (int (bb (n - 1)) gchoose (n - Suc j))" by (rule rl) (simp add: gbinomial_int_nonneg d_gr_0) also from \0 < n - Suc j\ have "\ \ int (cc (n - 1)) gchoose (n - Suc j)" by (rule ie1) finally show ?thesis . qed finally show "2 * (- 1)^(n - j) * ((int d - 1) gchoose (n - Suc j)) + (- 1)^(n - Suc j) * ((int (aa (n - 1)) gchoose (n - Suc j)) + (int (bb (n - 1)) gchoose (n - Suc j))) \ (int (cc (n - 1))) gchoose (n - Suc j)" . qed also have "\ = 2 + (int (cc (n - 1)) gchoose ((n - 1) - j)) + ((int (aa (j + 2)) gchoose 2) + (int (bb (j + 2)) gchoose 2)) + ?S 3 2" using \j + 2 \ n - 2\ by (simp add: sum.atLeast_Suc_atMost numeral_3_eq_3) also have "\ \ 2 + (int (cc (n - 1)) gchoose ((n - 1) - j)) + ((int (aa (j + 2)) gchoose 2) + (int (bb (j + 2)) gchoose 2)) + ?S3 4 2" proof (rule add_left_mono) from \j + 4 \ n - 1\ have "j + 3 \ n - 2" by simp hence "?S 3 2 = ?S 4 2 - ?f (j + 3) j" by (simp add: sum.atLeast_Suc_atMost add.commute) hence "?S 3 2 \ ?S 4 2" using f_nonneg[of "j + 3"] by simp also have "\ \ ?S3 4 2" proof (rule sum_mono) fix i assume "i \ {j + 4..n - 2}" hence "0 < i - j" by simp from f_nonneg[of i] have "(- 1)^(i - j) * ?f i j \ ?f i j" by (smt minus_one_mult_self mult_cancel_right1 pos_zmult_eq_1_iff_lemma zero_less_mult_iff) also from \0 < i - j\ have "\ \ int (cc i) gchoose (i - j)" by (rule ie1) finally show "(- 1)^(i - j) * ?f i j \ int (cc i) gchoose (i - j)" . qed finally show "?S 3 2 \ ?S3 4 2" . qed also have "\ = ((int (aa (j + 2)) gchoose 2) + (int (bb (j + 2)) gchoose 2)) + (2 + ?S3 4 1)" using \0 < n - 1\ \j + 4 \ n - 1\ by (simp only: sum_tail_nat) (simp flip: numeral_2_eq_2) also from ie1 have "\ \ int (cc (j + 2)) gchoose 2 + (2 + ?S3 4 1)" by (rule add_right_mono) simp also have "\ = 2 + (int (cc (j + 2)) gchoose 2) + ?S3 4 1" by (simp only: ac_simps) finally show ?thesis . qed qed corollary cc_le: assumes "0 < j" and "j < n - 2" shows "cc j \ 2 + (cc (j + 1) choose 2) + (\i=j+3..n-1. cc i choose (Suc (i - j)))" proof - define j0 where "j0 = j - 1" with assms have j: "j = Suc j0" and "j0 < n - 3" by simp_all have "int (cc j) = int (cc (Suc j0))" by (simp only: j) also have "\ \ 2 + (int (cc (j0 + 2)) gchoose 2) + (\i=j0+4..n-1. int (cc i) gchoose (i - j0))" using \j0 < n - 3\ by (rule cc_Suc_le) also have "\ = 2 + (int (cc (j + 1)) gchoose 2) + (\i=j0+4..n-1. int (cc i) gchoose (i - j0))" by (simp add: j) also have "(\i=j0+4..n-1. int (cc i) gchoose (i - j0)) = int (\i=j+3..n-1. cc i choose (Suc (i - j)))" unfolding int_sum proof (rule sum.cong) fix i assume "i \ {j + 3..n - 1}" hence "Suc j0 < i" by (simp add: j) hence "i - j0 = Suc (i - j)" by (simp add: j) thus "int (cc i) gchoose (i - j0) = int (cc i choose (Suc (i - j)))" by (simp add: int_binomial) qed (simp add: j) finally have "int (cc j) \ int (2 + (cc (j + 1) choose 2) + (\i = j + 3..n - 1. cc i choose (Suc (i - j))))" by (simp only: int_plus int_binomial) thus ?thesis by (simp only: zle_int) qed corollary cc_le_Dube_aux: "0 < j \ j + 1 \ n \ cc j \ Dube_aux n d j" proof (induct j rule: Dube_aux.induct[where n=n]) case step: (1 j) from step.prems(2) have "j + 2 < n \ j + 2 = n \ j + 1 = n" by auto thus ?case proof (elim disjE) assume *: "j + 2 < n" moreover have "0 < j + 1" by simp moreover from * have "j + 1 + 1 \ n" by simp ultimately have "cc (j + 1) \ Dube_aux n d (j + 1)" by (rule step.hyps) hence 1: "cc (j + 1) choose 2 \ Dube_aux n d (j + 1) choose 2" by (rule Binomial_Int.binomial_mono) have 2: "(\i = j + 3..n - 1. cc i choose Suc (i - j)) \ (\i = j + 3..n - 1. Dube_aux n d i choose Suc (i - j))" proof (rule sum_mono) fix i::nat note * moreover assume "i \ {j + 3..n - 1}" moreover from this \2 < n\ have "0 < i" and "i + 1 \ n" by auto ultimately have "cc i \ Dube_aux n d i" by (rule step.hyps) thus "cc i choose Suc (i - j) \ Dube_aux n d i choose Suc (i - j)" by (rule Binomial_Int.binomial_mono) qed from * have "j < n - 2" by simp with step.prems(1) have "cc j \ 2 + (cc (j + 1) choose 2) + (\i = j + 3..n - 1. cc i choose Suc (i - j))" by (rule cc_le) also from * 1 2 have "\ \ Dube_aux n d j" by simp finally show ?thesis . next assume "j + 2 = n" hence "j = n - 2" and "Dube_aux n d j = d\<^sup>2 + 2 * d" by simp_all thus ?thesis by (simp only: cc_n_minus_2) next assume "j + 1 = n" hence "j = n - 1" and "Dube_aux n d j = 2 * d" by simp_all thus ?thesis by (simp only: cc_n_minus_1) qed qed end lemma Dube_aux: assumes "g \ punit.reduced_GB F" shows "poly_deg g \ Dube_aux n d 1" proof (cases "n = 2") case True from assms have "poly_deg g \ bb 0" by (rule deg_RGB) also have "\ \ max (aa 1) (bb 1)" by (fact bb_0) also have "\ \ cc (n - 1)" by (simp add: True) also have "\ = 2 * d" by (fact cc_n_minus_1) also have "\ = Dube_aux n d 1" by (simp add: True) finally show ?thesis . next case False with \1 < n\ have "2 < n" and "1 + 1 \ n" by simp_all from assms have "poly_deg g \ bb 0" by (rule deg_RGB) also have "\ \ max (aa 1) (bb 1)" by (fact bb_0) also have "\ \ cc 1" by simp also from \2 < n\ _ \1 + 1 \ n\ have "\ \ Dube_aux n d 1" by (rule cc_le_Dube_aux) simp finally show ?thesis . qed end theorem Dube: assumes "finite F" and "F \ P[X]" and "\f. f \ F \ homogeneous f" and "g \ punit.reduced_GB F" shows "poly_deg g \ Dube (card X) (maxdeg F)" proof (cases "F \ {0}") case True hence "F = {} \ F = {0}" by blast with assms(4) show ?thesis by (auto simp: punit.reduced_GB_empty punit.reduced_GB_singleton) next case False hence "F - {0} \ {}" by simp hence "F \ {}" by blast hence "poly_deg ` F \ {}" by simp from assms(1) have fin1: "finite (poly_deg ` F)" by (rule finite_imageI) from assms(1) have "finite (F - {0})" by simp hence fin: "finite (poly_deg ` (F - {0}))" by (rule finite_imageI) moreover from \F - {0} \ {}\ have *: "poly_deg ` (F - {0}) \ {}" by simp ultimately have "maxdeg (F - {0}) \ poly_deg ` (F - {0})" unfolding maxdeg_def by (rule Max_in) then obtain f where "f \ F - {0}" and md1: "maxdeg (F - {0}) = poly_deg f" .. note this(2) moreover have "maxdeg (F - {0}) \ maxdeg F" unfolding maxdeg_def using image_mono * fin1 by (rule Max_mono) blast ultimately have "poly_deg f \ maxdeg F" by simp from \f \ F - {0}\ have "f \ F" and "f \ 0" by simp_all from this(1) assms(2) have "f \ P[X]" .. have f_max: "poly_deg f' \ poly_deg f" if "f' \ F" for f' proof (cases "f' = 0") case True thus ?thesis by simp next case False with that have "f' \ F - {0}" by simp hence "poly_deg f' \ poly_deg ` (F - {0})" by (rule imageI) with fin show "poly_deg f' \ poly_deg f" unfolding md1[symmetric] maxdeg_def by (rule Max_ge) qed have "maxdeg F \ poly_deg f" unfolding maxdeg_def using fin1 \poly_deg ` F \ {}\ proof (rule Max.boundedI) fix d assume "d \ poly_deg ` F" then obtain f' where "f' \ F" and "d = poly_deg f'" .. note this(2) also from \f' \ F\ have "poly_deg f' \ poly_deg f" by (rule f_max) finally show "d \ poly_deg f" . qed with \poly_deg f \ maxdeg F\ have md: "poly_deg f = maxdeg F" by (rule antisym) show ?thesis proof (cases "ideal {f} = ideal F") case True note assms(4) also have "punit.reduced_GB F = punit.reduced_GB {f}" using punit.finite_reduced_GB_finite punit.reduced_GB_is_reduced_GB_finite by (rule punit.reduced_GB_unique) (simp_all add: punit.reduced_GB_pmdl_finite[simplified] True) also have "\ \ {punit.monic f}" by (simp add: punit.reduced_GB_singleton) finally have "g \ {punit.monic f}" . hence "poly_deg g = poly_deg (punit.monic f)" by simp also from poly_deg_monom_mult_le[where c="1 / lcf f" and t=0 and p=f] have "\ \ poly_deg f" by (simp add: punit.monic_def) also have "\ = maxdeg F" by (fact md) also have "\ \ Dube (card X) (maxdeg F)" by (fact Dube_ge_d) finally show ?thesis . next case False show ?thesis proof (cases "poly_deg f = 0") case True hence "monomial (lookup f 0) 0 = f" by (rule poly_deg_zero_imp_monomial) moreover define c where "c = lookup f 0" ultimately have f: "f = monomial c 0" by simp with \f \ 0\ have "c \ 0" by (simp add: monomial_0_iff) from \f \ F\ have "f \ ideal F" by (rule ideal.span_base) hence "punit.monom_mult (1 / c) 0 f \ ideal F" by (rule punit.pmdl_closed_monom_mult[simplified]) with \c \ 0\ have "ideal F = UNIV" by (simp add: f punit.monom_mult_monomial ideal_eq_UNIV_iff_contains_one) with assms(1) have "punit.reduced_GB F = {1}" by (simp only: ideal_eq_UNIV_iff_reduced_GB_eq_one_finite) with assms(4) show ?thesis by simp next case False hence "0 < poly_deg f" by simp have "card X \ 1 \ 1 < card X" by auto thus ?thesis proof note fin_X moreover assume "card X \ 1" moreover note assms(2) moreover from \f \ F\ have "f \ ideal F" by (rule ideal.span_base) ultimately have "poly_deg g \ poly_deg f" using \f \ 0\ assms(4) by (rule deg_reduced_GB_univariate_le) also have "\ \ Dube (card X) (maxdeg F)" unfolding md by (fact Dube_ge_d) finally show ?thesis . next assume "1 < card X" hence "poly_deg g \ Dube_aux (card X) (poly_deg f) 1" using assms(1, 2) \f \ F\ assms(3) f_max \0 < poly_deg f\ \ideal {f} \ ideal F\ assms(4) by (rule Dube_aux) also from \1 < card X\ \0 < poly_deg f\ have "\ = Dube (card X) (maxdeg F)" by (simp add: Dube_def md) finally show ?thesis . qed qed qed qed corollary Dube_is_hom_GB_bound: "finite F \ F \ P[X] \ is_hom_GB_bound F (Dube (card X) (maxdeg F))" by (intro is_hom_GB_boundI Dube) end corollary Dube_indets: assumes "finite F" and "\f. f \ F \ homogeneous f" and "g \ punit.reduced_GB F" shows "poly_deg g \ Dube (card (\(indets ` F))) (maxdeg F)" using _ assms(1) _ assms(2, 3) proof (rule Dube) from assms show "finite (\(indets ` F))" by (simp add: finite_indets) next show "F \ P[\(indets ` F)]" by (auto simp: Polys_alt) qed corollary Dube_is_hom_GB_bound_indets: "finite F \ is_hom_GB_bound F (Dube (card (\(indets ` F))) (maxdeg F))" by (intro is_hom_GB_boundI Dube_indets) end (* pm_powerprod *) hide_const (open) pm_powerprod.\ pm_powerprod.\ context extended_ord_pm_powerprod begin lemma Dube_is_GB_cofactor_bound: assumes "finite X" and "finite F" and "F \ P[X]" shows "is_GB_cofactor_bound F (Dube (Suc (card X)) (maxdeg F))" using assms(1, 3) proof (rule hom_GB_bound_is_GB_cofactor_bound) let ?F = "homogenize None ` extend_indets ` F" let ?X = "insert None (Some ` X)" from assms(1) have "finite ?X" by simp moreover from assms(2) have "finite ?F" by (intro finite_imageI) moreover have "?F \ P[?X]" proof fix f' assume "f' \ ?F" then obtain f where "f \ F" and f': "f' = homogenize None (extend_indets f)" by blast from this(1) assms(3) have "f \ P[X]" .. hence "extend_indets f \ P[Some ` X]" by (auto simp: Polys_alt indets_extend_indets) thus "f' \ P[?X]" unfolding f' by (rule homogenize_in_Polys) qed ultimately have "extended_ord.is_hom_GB_bound ?F (Dube (card ?X) (maxdeg ?F))" by (rule extended_ord.Dube_is_hom_GB_bound) moreover have "maxdeg ?F = maxdeg F" proof - have "maxdeg ?F = maxdeg (extend_indets ` F)" by (auto simp: indets_extend_indets intro: maxdeg_homogenize) also have "\ = maxdeg F" by (simp add: maxdeg_def image_image) finally show "maxdeg ?F = maxdeg F" . qed moreover from assms(1) have "card ?X = card X + 1" by (simp add: card_image) ultimately show "extended_ord.is_hom_GB_bound ?F (Dube (Suc (card X)) (maxdeg F))" by simp qed lemma Dube_is_GB_cofactor_bound_explicit: assumes "finite X" and "finite F" and "F \ P[X]" obtains G where "punit.is_Groebner_basis G" and "ideal G = ideal F" and "G \ P[X]" and "\g. g \ G \ \q. g = (\f\F. q f * f) \ (\f. q f \ P[X] \ poly_deg (q f * f) \ Dube (Suc (card X)) (maxdeg F) \ (f \ F \ q f = 0))" proof - from assms have "is_GB_cofactor_bound F (Dube (Suc (card X)) (maxdeg F))" (is "is_GB_cofactor_bound _ ?b") by (rule Dube_is_GB_cofactor_bound) moreover note assms(3) ultimately obtain G where "punit.is_Groebner_basis G" and "ideal G = ideal F" and "G \ P[X]" and 1: "\g. g \ G \ \F' q. finite F' \ F' \ F \ g = (\f\F'. q f * f) \ (\f. q f \ P[X] \ poly_deg (q f * f) \ ?b \ (f \ F' \ q f = 0))" by (rule is_GB_cofactor_boundE_Polys) blast from this(1-3) show ?thesis proof fix g assume "g \ G" hence "\F' q. finite F' \ F' \ F \ g = (\f\F'. q f * f) \ (\f. q f \ P[X] \ poly_deg (q f * f) \ ?b \ (f \ F' \ q f = 0))" by (rule 1) then obtain F' q where "F' \ F" and g: "g = (\f\F'. q f * f)" and "\f. q f \ P[X]" and "\f. poly_deg (q f * f) \ ?b" and 2: "\f. f \ F' \ q f = 0" by blast show "\q. g = (\f\F. q f * f) \ (\f. q f \ P[X] \ poly_deg (q f * f) \ ?b \ (f \ F \ q f = 0))" proof (intro exI allI conjI impI) from assms(2) \F' \ F\ have "(\f\F'. q f * f) = (\f\F. q f * f)" proof (intro sum.mono_neutral_left ballI) fix f assume "f \ F - F'" hence "f \ F'" by simp hence "q f = 0" by (rule 2) thus "q f * f = 0" by simp qed thus "g = (\f\F. q f * f)" by (simp only: g) next fix f assume "f \ F" with \F' \ F\ have "f \ F'" by blast thus "q f = 0" by (rule 2) qed fact+ qed qed corollary Dube_is_GB_cofactor_bound_indets: assumes "finite F" shows "is_GB_cofactor_bound F (Dube (Suc (card (\(indets ` F)))) (maxdeg F))" using _ assms _ proof (rule Dube_is_GB_cofactor_bound) from assms show "finite (\(indets ` F))" by (simp add: finite_indets) next show "F \ P[\(indets ` F)]" by (auto simp: Polys_alt) qed end (* extended_ord_pm_powerprod *) end (* theory *) diff --git a/thys/Groebner_Macaulay/Hilbert_Function.thy b/thys/Groebner_Macaulay/Hilbert_Function.thy --- a/thys/Groebner_Macaulay/Hilbert_Function.thy +++ b/thys/Groebner_Macaulay/Hilbert_Function.thy @@ -1,1404 +1,1388 @@ (* Author: Alexander Maletzky *) section \Direct Decompositions and Hilbert Functions\ theory Hilbert_Function - imports Dube_Prelims Degree_Section "HOL-Library.List_Permutation" + imports Dube_Prelims Degree_Section "HOL-Library.Permutations" begin subsection \Direct Decompositions\ text \The main reason for defining \direct_decomp\ in terms of lists rather than sets is that lemma \direct_decomp_direct_decomp\ can be proved easier. At some point one could invest the time to re-define \direct_decomp\ in terms of sets (possibly adding a couple of further assumptions to \direct_decomp_direct_decomp\).\ definition direct_decomp :: "'a set \ 'a::comm_monoid_add set list \ bool" where "direct_decomp A ss \ bij_betw sum_list (listset ss) A" lemma direct_decompI: "inj_on sum_list (listset ss) \ sum_list ` listset ss = A \ direct_decomp A ss" by (simp add: direct_decomp_def bij_betw_def) lemma direct_decompI_alt: "(\qs. qs \ listset ss \ sum_list qs \ A) \ (\a. a \ A \ \!qs\listset ss. a = sum_list qs) \ direct_decomp A ss" by (auto simp: direct_decomp_def intro!: bij_betwI') blast lemma direct_decompD: assumes "direct_decomp A ss" shows "qs \ listset ss \ sum_list qs \ A" and "inj_on sum_list (listset ss)" and "sum_list ` listset ss = A" using assms by (auto simp: direct_decomp_def bij_betw_def) lemma direct_decompE: assumes "direct_decomp A ss" and "a \ A" obtains qs where "qs \ listset ss" and "a = sum_list qs" using assms by (auto simp: direct_decomp_def bij_betw_def) lemma direct_decomp_unique: "direct_decomp A ss \ qs \ listset ss \ qs' \ listset ss \ sum_list qs = sum_list qs' \ qs = qs'" by (auto dest: direct_decompD simp: inj_on_def) lemma direct_decomp_singleton: "direct_decomp A [A]" proof (rule direct_decompI_alt) fix qs assume "qs \ listset [A]" then obtain q where "q \ A" and "qs = [q]" by (rule listset_singletonE) thus "sum_list qs \ A" by simp next fix a assume "a \ A" show "\!qs\listset [A]. a = sum_list qs" proof (intro ex1I conjI allI impI) from \a \ A\ refl show "[a] \ listset [A]" by (rule listset_singletonI) next fix qs assume "qs \ listset [A] \ a = sum_list qs" hence a: "a = sum_list qs" and "qs \ listset [A]" by simp_all from this(2) obtain b where qs: "qs = [b]" by (rule listset_singletonE) with a show "qs = [a]" by simp qed simp_all qed (* TODO: Move. *) lemma mset_bij: assumes "bij_betw f {..i. i < length xs \ xs ! i = ys ! f i" shows "mset xs = mset ys" proof - from assms(1) have 1: "inj_on f {0..i < length xs\ have "\ = map ((!) ys \ f) [0.. f) [0.. = image_mset ((!) ys) (image_mset f (mset_set {0.. = image_mset ((!) ys) (mset_set {0.. = mset (map ((!) ys) [0..f. bij_betw f {.. (\ii. i < length ss2 \ ss1 ! i = ss2 ! f i" unfolding len_ss1 by blast + from assms(2) have len_ss1: "length ss1 = length ss2" + using mset_eq_length by blast + from assms(2) obtain f where \f permutes {.. + \permute_list f ss2 = ss1\ + by (rule mset_eq_permutation) + then have f_bij: "bij_betw f {..i. i < length ss2 \ ss1 ! i = ss2 ! f i" + by (auto simp add: permutes_imp_bij permute_list_nth) define g where "g = inv_into {.. f ` {.. {.. {.. g ` {.. = {.. listset ss2" then obtain qs1 where qs1_in: "qs1 \ listset ss1" and len_qs1: "length qs1 = length qs2" and *: "\i. i < length qs2 \ qs1 ! i = qs2 ! f i" using f_bij f by (rule listset_permE) blast+ from \qs2 \ listset ss2\ have "length qs2 = length ss2" by (rule listsetD) with f_bij have "bij_betw f {.. \ A" by (rule direct_decompD) finally show "sum_list qs2 \ A" . next fix a assume "a \ A" with assms(1) obtain qs where a: "a = sum_list qs" and qs_in: "qs \ listset ss1" by (rule direct_decompE) from qs_in obtain qs2 where qs2_in: "qs2 \ listset ss2" and len_qs2: "length qs2 = length qs" and 1: "\i. i < length qs \ qs2 ! i = qs ! g i" using g_bij g by (rule listset_permE) blast+ show "\!qs\listset ss2. a = sum_list qs" proof (intro ex1I conjI allI impI) from qs_in have len_qs: "length qs = length ss1" by (rule listsetD) with g_bij have g_bij2: "bij_betw g {.. listset ss2 \ a = sum_list qs'" hence qs'_in: "qs' \ listset ss2" and a': "a = sum_list qs'" by simp_all from this(1) obtain qs1 where qs1_in: "qs1 \ listset ss1" and len_qs1: "length qs1 = length qs'" and 2: "\i. i < length qs' \ qs1 ! i = qs' ! f i" using f_bij f by (rule listset_permE) blast+ from \qs' \ listset ss2\ have "length qs' = length ss2" by (rule listsetD) with f_bij have "bij_betw f {.. {.. g ` {.. = {..i < length qs\ have "qs2 ! i = qs ! g i" by (rule 1) also have "\ = qs1 ! g i" by (simp only: \qs1 = qs\) also from \g i < length qs'\ have "\ = qs' ! f (g i)" by (rule 2) also from \i < length ss1\ have "\ = qs' ! i" by (simp only: f_g) finally show "qs' ! i = qs2 ! i" by (rule sym) qed qed fact qed qed lemma direct_decomp_split_map: "direct_decomp A (map f ss) \ direct_decomp A (map f (filter P ss) @ map f (filter (- P) ss))" proof (rule direct_decomp_perm) - show "perm (map f ss) (map f (filter P ss) @ map f (filter (- P) ss))" - proof (induct ss) - case Nil - show ?case by simp - next - case (Cons s ss) - show ?case - proof (cases "P s") - case True - with Cons show ?thesis by simp - next - case False - have "map f (s # ss) = f s # map f ss" by simp - also from Cons have "perm (f s # map f ss) (f s # map f (filter P ss) @ map f (filter (- P) ss))" - by (rule perm.intros) - also have "perm \ (map f (filter P ss) @ map f (s # filter (- P) ss))" - by (simp add: perm_append_Cons) - also(trans) from False have "\ = map f (filter P (s # ss)) @ map f (filter (- P) (s # ss))" - by simp - finally show ?thesis . - qed - qed + show "mset (map f ss) = mset (map f (filter P ss) @ map f (filter (- P) ss))" + by simp (metis image_mset_union multiset_partition) qed lemmas direct_decomp_split = direct_decomp_split_map[where f=id, simplified] lemma direct_decomp_direct_decomp: assumes "direct_decomp A (s # ss)" and "direct_decomp s rs" shows "direct_decomp A (ss @ rs)" (is "direct_decomp A ?ss") proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain qs1 qs2 where qs1: "qs1 \ listset ss" and qs2: "qs2 \ listset rs" and qs: "qs = qs1 @ qs2" by (rule listset_appendE) have "sum_list qs = sum_list ((sum_list qs2) # qs1)" by (simp add: qs add.commute) also from assms(1) have "\ \ A" proof (rule direct_decompD) from assms(2) qs2 have "sum_list qs2 \ s" by (rule direct_decompD) thus "sum_list qs2 # qs1 \ listset (s # ss)" using qs1 refl by (rule listset_ConsI) qed finally show "sum_list qs \ A" . next fix a assume "a \ A" with assms(1) obtain qs1 where qs1_in: "qs1 \ listset (s # ss)" and a: "a = sum_list qs1" by (rule direct_decompE) from qs1_in obtain qs11 qs12 where "qs11 \ s" and qs12_in: "qs12 \ listset ss" and qs1: "qs1 = qs11 # qs12" by (rule listset_ConsE) from assms(2) this(1) obtain qs2 where qs2_in: "qs2 \ listset rs" and qs11: "qs11 = sum_list qs2" by (rule direct_decompE) let ?qs = "qs12 @ qs2" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI allI impI) from qs12_in qs2_in refl show "?qs \ listset ?ss" by (rule listset_appendI) show "a = sum_list ?qs" by (simp add: a qs1 qs11 add.commute) fix qs0 assume "qs0 \ listset ?ss \ a = sum_list qs0" hence qs0_in: "qs0 \ listset ?ss" and a2: "a = sum_list qs0" by simp_all from this(1) obtain qs01 qs02 where qs01_in: "qs01 \ listset ss" and qs02_in: "qs02 \ listset rs" and qs0: "qs0 = qs01 @ qs02" by (rule listset_appendE) note assms(1) moreover from _ qs01_in refl have "(sum_list qs02) # qs01 \ listset (s # ss)" (is "?qs' \ _") proof (rule listset_ConsI) from assms(2) qs02_in show "sum_list qs02 \ s" by (rule direct_decompD) qed moreover note qs1_in moreover from a2 have "sum_list ?qs' = sum_list qs1" by (simp add: qs0 a add.commute) ultimately have "?qs' = qs11 # qs12" unfolding qs1 by (rule direct_decomp_unique) hence "qs11 = sum_list qs02" and 1: "qs01 = qs12" by simp_all from this(1) have "sum_list qs02 = sum_list qs2" by (simp only: qs11) with assms(2) qs02_in qs2_in have "qs02 = qs2" by (rule direct_decomp_unique) thus "qs0 = qs12 @ qs2" by (simp only: 1 qs0) qed qed lemma sum_list_map_times: "sum_list (map ((*) x) xs) = (x::'a::semiring_0) * sum_list xs" by (induct xs) (simp_all add: algebra_simps) lemma direct_decomp_image_times: assumes "direct_decomp (A::'a::semiring_0 set) ss" and "\a b. x * a = x * b \ x \ 0 \ a = b" shows "direct_decomp ((*) x ` A) (map ((`) ((*) x)) ss)" (is "direct_decomp ?A ?ss") proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain qs0 where qs0_in: "qs0 \ listset ss" and qs: "qs = map ((*) x) qs0" by (rule listset_map_imageE) have "sum_list qs = x * sum_list qs0" by (simp only: qs sum_list_map_times) moreover from assms(1) qs0_in have "sum_list qs0 \ A" by (rule direct_decompD) ultimately show "sum_list qs \ (*) x ` A" by (rule image_eqI) next fix a assume "a \ ?A" then obtain a' where "a' \ A" and a: "a = x * a'" .. from assms(1) this(1) obtain qs' where qs'_in: "qs' \ listset ss" and a': "a' = sum_list qs'" by (rule direct_decompE) define qs where "qs = map ((*) x) qs'" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI allI impI) from qs'_in qs_def show "qs \ listset ?ss" by (rule listset_map_imageI) fix qs0 assume "qs0 \ listset ?ss \ a = sum_list qs0" hence "qs0 \ listset ?ss" and a0: "a = sum_list qs0" by simp_all from this(1) obtain qs1 where qs1_in: "qs1 \ listset ss" and qs0: "qs0 = map ((*) x) qs1" by (rule listset_map_imageE) show "qs0 = qs" proof (cases "x = 0") case True from qs1_in have "length qs1 = length ss" by (rule listsetD) moreover from qs'_in have "length qs' = length ss" by (rule listsetD) ultimately show ?thesis by (simp add: qs_def qs0 list_eq_iff_nth_eq True) next case False have "x * sum_list qs1 = a" by (simp only: a0 qs0 sum_list_map_times) also have "\ = x * sum_list qs'" by (simp only: a' a) finally have "sum_list qs1 = sum_list qs'" using False by (rule assms(2)) with assms(1) qs1_in qs'_in have "qs1 = qs'" by (rule direct_decomp_unique) thus ?thesis by (simp only: qs0 qs_def) qed qed (simp only: a a' qs_def sum_list_map_times) qed lemma direct_decomp_appendD: assumes "direct_decomp A (ss1 @ ss2)" shows "{} \ set ss2 \ direct_decomp (sum_list ` listset ss1) ss1" (is "_ \ ?thesis1") and "{} \ set ss1 \ direct_decomp (sum_list ` listset ss2) ss2" (is "_ \ ?thesis2") and "direct_decomp A [sum_list ` listset ss1, sum_list ` listset ss2]" (is "direct_decomp _ ?ss") proof - have rl: "direct_decomp (sum_list ` listset ts1) ts1" if "direct_decomp A (ts1 @ ts2)" and "{} \ set ts2" for ts1 ts2 proof (intro direct_decompI inj_onI refl) fix qs1 qs2 assume qs1: "qs1 \ listset ts1" and qs2: "qs2 \ listset ts1" assume eq: "sum_list qs1 = sum_list qs2" from that(2) have "listset ts2 \ {}" by (simp add: listset_empty_iff) then obtain qs3 where qs3: "qs3 \ listset ts2" by blast note that(1) moreover from qs1 qs3 refl have "qs1 @ qs3 \ listset (ts1 @ ts2)" by (rule listset_appendI) moreover from qs2 qs3 refl have "qs2 @ qs3 \ listset (ts1 @ ts2)" by (rule listset_appendI) moreover have "sum_list (qs1 @ qs3) = sum_list (qs2 @ qs3)" by (simp add: eq) ultimately have "qs1 @ qs3 = qs2 @ qs3" by (rule direct_decomp_unique) thus "qs1 = qs2" by simp qed { assume "{} \ set ss2" with assms show ?thesis1 by (rule rl) } { - from assms perm_append_swap have "direct_decomp A (ss2 @ ss1)" by (rule direct_decomp_perm) + from assms have "direct_decomp A (ss2 @ ss1)" + by (rule direct_decomp_perm) simp moreover assume "{} \ set ss1" ultimately show ?thesis2 by (rule rl) } show "direct_decomp A ?ss" proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain q1 q2 where q1: "q1 \ sum_list ` listset ss1" and q2: "q2 \ sum_list ` listset ss2" and qs: "qs = [q1, q2]" by (rule listset_doubletonE) from q1 obtain qs1 where qs1: "qs1 \ listset ss1" and q1: "q1 = sum_list qs1" .. from q2 obtain qs2 where qs2: "qs2 \ listset ss2" and q2: "q2 = sum_list qs2" .. from qs1 qs2 refl have "qs1 @ qs2 \ listset (ss1 @ ss2)" by (rule listset_appendI) with assms have "sum_list (qs1 @ qs2) \ A" by (rule direct_decompD) thus "sum_list qs \ A" by (simp add: qs q1 q2) next fix a assume "a \ A" with assms obtain qs0 where qs0_in: "qs0 \ listset (ss1 @ ss2)" and a: "a = sum_list qs0" by (rule direct_decompE) from this(1) obtain qs1 qs2 where qs1: "qs1 \ listset ss1" and qs2: "qs2 \ listset ss2" and qs0: "qs0 = qs1 @ qs2" by (rule listset_appendE) from qs1 have len_qs1: "length qs1 = length ss1" by (rule listsetD) define qs where "qs = [sum_list qs1, sum_list qs2]" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI) from qs1 have "sum_list qs1 \ sum_list ` listset ss1" by (rule imageI) moreover from qs2 have "sum_list qs2 \ sum_list ` listset ss2" by (rule imageI) ultimately show "qs \ listset ?ss" using qs_def by (rule listset_doubletonI) fix qs' assume "qs' \ listset ?ss \ a = sum_list qs'" hence "qs' \ listset ?ss" and a': "a = sum_list qs'" by simp_all from this(1) obtain q1 q2 where q1: "q1 \ sum_list ` listset ss1" and q2: "q2 \ sum_list ` listset ss2" and qs': "qs' = [q1, q2]" by (rule listset_doubletonE) from q1 obtain qs1' where qs1': "qs1' \ listset ss1" and q1: "q1 = sum_list qs1'" .. from q2 obtain qs2' where qs2': "qs2' \ listset ss2" and q2: "q2 = sum_list qs2'" .. from qs1' have len_qs1': "length qs1' = length ss1" by (rule listsetD) note assms moreover from qs1' qs2' refl have "qs1' @ qs2' \ listset (ss1 @ ss2)" by (rule listset_appendI) moreover note qs0_in moreover have "sum_list (qs1' @ qs2') = sum_list qs0" by (simp add: a' qs' flip: a q1 q2) ultimately have "qs1' @ qs2' = qs0" by (rule direct_decomp_unique) also have "\ = qs1 @ qs2" by fact finally show "qs' = qs" by (simp add: qs_def qs' q1 q2 len_qs1 len_qs1') qed (simp add: qs_def a qs0) qed qed lemma direct_decomp_Cons_zeroI: assumes "direct_decomp A ss" shows "direct_decomp A ({0} # ss)" proof (rule direct_decompI_alt) fix qs assume "qs \ listset ({0} # ss)" then obtain q qs' where "q \ {0}" and "qs' \ listset ss" and "qs = q # qs'" by (rule listset_ConsE) from this(1, 3) have "sum_list qs = sum_list qs'" by simp also from assms \qs' \ listset ss\ have "\ \ A" by (rule direct_decompD) finally show "sum_list qs \ A" . next fix a assume "a \ A" with assms obtain qs' where qs': "qs' \ listset ss" and a: "a = sum_list qs'" by (rule direct_decompE) define qs where "qs = 0 # qs'" show "\!qs. qs \ listset ({0} # ss) \ a = sum_list qs" proof (intro ex1I conjI) from _ qs' qs_def show "qs \ listset ({0} # ss)" by (rule listset_ConsI) simp next fix qs0 assume "qs0 \ listset ({0} # ss) \ a = sum_list qs0" hence "qs0 \ listset ({0} # ss)" and a0: "a = sum_list qs0" by simp_all from this(1) obtain q0 qs0' where "q0 \ {0}" and qs0': "qs0' \ listset ss" and qs0: "qs0 = q0 # qs0'" by (rule listset_ConsE) from this(1, 3) have "sum_list qs0' = sum_list qs'" by (simp add: a0 flip: a) with assms qs0' qs' have "qs0' = qs'" by (rule direct_decomp_unique) with \q0 \ {0}\ show "qs0 = qs" by (simp add: qs_def qs0) qed (simp add: qs_def a) qed lemma direct_decomp_Cons_zeroD: assumes "direct_decomp A ({0} # ss)" shows "direct_decomp A ss" proof - have "direct_decomp {0} []" by (simp add: direct_decomp_def bij_betw_def) with assms have "direct_decomp A (ss @ [])" by (rule direct_decomp_direct_decomp) thus ?thesis by simp qed lemma direct_decomp_Cons_subsetI: assumes "direct_decomp A (s # ss)" and "\s0. s0 \ set ss \ 0 \ s0" shows "s \ A" proof fix x assume "x \ s" moreover from assms(2) have "map (\_. 0) ss \ listset ss" by (induct ss, auto simp del: listset.simps(2) intro: listset_ConsI) ultimately have "x # (map (\_. 0) ss) \ listset (s # ss)" using refl by (rule listset_ConsI) with assms(1) have "sum_list (x # (map (\_. 0) ss)) \ A" by (rule direct_decompD) thus "x \ A" by simp qed lemma direct_decomp_Int_zero: assumes "direct_decomp A ss" and "i < j" and "j < length ss" and "\s. s \ set ss \ 0 \ s" shows "ss ! i \ ss ! j = {0}" proof - from assms(2, 3) have "i < length ss" by (rule less_trans) hence i_in: "ss ! i \ set ss" by simp from assms(3) have j_in: "ss ! j \ set ss" by simp show ?thesis proof show "ss ! i \ ss ! j \ {0}" proof fix x assume "x \ ss ! i \ ss ! j" hence x_i: "x \ ss ! i" and x_j: "x \ ss ! j" by simp_all have 1: "(map (\_. 0) ss)[k := y] \ listset ss" if "k < length ss" and "y \ ss ! k" for k y using assms(4) that proof (induct ss arbitrary: k) case Nil from Nil(2) show ?case by simp next case (Cons s ss) have *: "\s'. s' \ set ss \ 0 \ s'" by (rule Cons.prems) simp show ?case proof (cases k) case k: 0 with Cons.prems(3) have "y \ s" by simp moreover from * have "map (\_. 0) ss \ listset ss" by (induct ss) (auto simp del: listset.simps(2) intro: listset_ConsI) moreover have "(map (\_. 0) (s # ss))[k := y] = y # map (\_. 0) ss" by (simp add: k) ultimately show ?thesis by (rule listset_ConsI) next case k: (Suc k') have "0 \ s" by (rule Cons.prems) simp moreover from * have "(map (\_. 0) ss)[k' := y] \ listset ss" proof (rule Cons.hyps) from Cons.prems(2) show "k' < length ss" by (simp add: k) next from Cons.prems(3) show "y \ ss ! k'" by (simp add: k) qed moreover have "(map (\_. 0) (s # ss))[k := y] = 0 # (map (\_. 0) ss)[k' := y]" by (simp add: k) ultimately show ?thesis by (rule listset_ConsI) qed qed have 2: "sum_list ((map (\_. 0) ss)[k := y]) = y" if "k < length ss" for k and y::'a using that by (induct ss arbitrary: k) (auto simp: add_ac split: nat.split) define qs1 where "qs1 = (map (\_. 0) ss)[i := x]" define qs2 where "qs2 = (map (\_. 0) ss)[j := x]" note assms(1) moreover from \i < length ss\ x_i have "qs1 \ listset ss" unfolding qs1_def by (rule 1) moreover from assms(3) x_j have "qs2 \ listset ss" unfolding qs2_def by (rule 1) thm sum_list_update moreover from \i < length ss\ assms(3) have "sum_list qs1 = sum_list qs2" by (simp add: qs1_def qs2_def 2) ultimately have "qs1 = qs2" by (rule direct_decomp_unique) hence "qs1 ! i = qs2 ! i" by simp with \i < length ss\ assms(2, 3) show "x \ {0}" by (simp add: qs1_def qs2_def) qed next from i_in have "0 \ ss ! i" by (rule assms(4)) moreover from j_in have "0 \ ss ! j" by (rule assms(4)) ultimately show "{0} \ ss ! i \ ss ! j" by simp qed qed corollary direct_decomp_pairwise_zero: assumes "direct_decomp A ss" and "\s. s \ set ss \ 0 \ s" shows "pairwise (\s1 s2. s1 \ s2 = {0}) (set ss)" proof (rule pairwiseI) fix s1 s2 assume "s1 \ set ss" then obtain i where "i < length ss" and s1: "s1 = ss ! i" by (metis in_set_conv_nth) assume "s2 \ set ss" then obtain j where "j < length ss" and s2: "s2 = ss ! j" by (metis in_set_conv_nth) assume "s1 \ s2" hence "i < j \ j < i" by (auto simp: s1 s2) thus "s1 \ s2 = {0}" proof assume "i < j" with assms(1) show ?thesis unfolding s1 s2 using \j < length ss\ assms(2) by (rule direct_decomp_Int_zero) next assume "j < i" with assms(1) have "s2 \ s1 = {0}" unfolding s1 s2 using \i < length ss\ assms(2) by (rule direct_decomp_Int_zero) thus ?thesis by (simp only: Int_commute) qed qed corollary direct_decomp_repeated_eq_zero: assumes "direct_decomp A ss" and "1 < count_list ss X" and "\s. s \ set ss \ 0 \ s" shows "X = {0}" proof - from assms(2) obtain i j where "i < j" and "j < length ss" and 1: "ss ! i = X" and 2: "ss ! j = X" by (rule count_list_gr_1_E) from assms(1) this(1, 2) assms(3) have "ss ! i \ ss ! j = {0}" by (rule direct_decomp_Int_zero) thus ?thesis by (simp add: 1 2) qed corollary direct_decomp_map_Int_zero: assumes "direct_decomp A (map f ss)" and "s1 \ set ss" and "s2 \ set ss" and "s1 \ s2" and "\s. s \ set ss \ 0 \ f s" shows "f s1 \ f s2 = {0}" proof - from assms(2) obtain i where "i < length ss" and s1: "s1 = ss ! i" by (metis in_set_conv_nth) from this(1) have i: "i < length (map f ss)" by simp from assms(3) obtain j where "j < length ss" and s2: "s2 = ss ! j" by (metis in_set_conv_nth) from this(1) have j: "j < length (map f ss)" by simp have *: "0 \ s" if "s \ set (map f ss)" for s proof - from that obtain s' where "s' \ set ss" and s: "s = f s'" unfolding set_map .. from this(1) show "0 \ s" unfolding s by (rule assms(5)) qed show ?thesis proof (rule linorder_cases) assume "i < j" with assms(1) have "(map f ss) ! i \ (map f ss) ! j = {0}" using j * by (rule direct_decomp_Int_zero) with i j show ?thesis by (simp add: s1 s2) next assume "j < i" with assms(1) have "(map f ss) ! j \ (map f ss) ! i = {0}" using i * by (rule direct_decomp_Int_zero) with i j show ?thesis by (simp add: s1 s2 Int_commute) next assume "i = j" with assms(4) show ?thesis by (simp add: s1 s2) qed qed subsection \Direct Decompositions and Vector Spaces\ definition (in vector_space) is_basis :: "'b set \ 'b set \ bool" where "is_basis V B \ (B \ V \ independent B \ V \ span B \ card B = dim V)" definition (in vector_space) some_basis :: "'b set \ 'b set" where "some_basis V = Eps (local.is_basis V)" hide_const (open) real_vector.is_basis real_vector.some_basis context vector_space begin lemma dim_empty [simp]: "dim {} = 0" using dim_span_eq_card_independent independent_empty by fastforce lemma dim_zero [simp]: "dim {0} = 0" using dim_span_eq_card_independent independent_empty by fastforce lemma independent_UnI: assumes "independent A" and "independent B" and "span A \ span B = {0}" shows "independent (A \ B)" proof from span_superset have "A \ B \ span A \ span B" by blast hence "A \ B = {}" unfolding assms(3) using assms(1, 2) dependent_zero by blast assume "dependent (A \ B)" then obtain T u v where "finite T" and "T \ A \ B" and eq: "(\v\T. u v *s v) = 0" and "v \ T" and "u v \ 0" unfolding dependent_explicit by blast define TA where "TA = T \ A" define TB where "TB = T \ B" from \T \ A \ B\ have T: "T = TA \ TB" by (auto simp: TA_def TB_def) from \finite T\ have "finite TA" and "TA \ A" by (simp_all add: TA_def) from \finite T\ have "finite TB" and "TB \ B" by (simp_all add: TB_def) from \A \ B = {}\ \TA \ A\ this(2) have "TA \ TB = {}" by blast have "0 = (\v\TA \ TB. u v *s v)" by (simp only: eq flip: T) also have "\ = (\v\TA. u v *s v) + (\v\TB. u v *s v)" by (rule sum.union_disjoint) fact+ finally have "(\v\TA. u v *s v) = (\v\TB. (- u) v *s v)" (is "?x = ?y") by (simp add: sum_negf eq_neg_iff_add_eq_0) from \finite TB\ \TB \ B\ have "?y \ span B" by (auto simp: span_explicit simp del: uminus_apply) moreover from \finite TA\ \TA \ A\ have "?x \ span A" by (auto simp: span_explicit) ultimately have "?y \ span A \ span B" by (simp add: \?x = ?y\) hence "?x = 0" and "?y = 0" by (simp_all add: \?x = ?y\ assms(3)) from \v \ T\ have "v \ TA \ TB" by (simp only: T) hence "u v = 0" proof assume "v \ TA" with assms(1) \finite TA\ \TA \ A\ \?x = 0\ show "u v = 0" by (rule independentD) next assume "v \ TB" with assms(2) \finite TB\ \TB \ B\ \?y = 0\ have "(- u) v = 0" by (rule independentD) thus "u v = 0" by simp qed with \u v \ 0\ show False .. qed lemma subspace_direct_decomp: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "subspace A" proof (rule subspaceI) let ?qs = "map (\_. 0) ss" from assms(2) have "?qs \ listset ss" by (induct ss) (auto simp del: listset.simps(2) dest: subspace_0 intro: listset_ConsI) with assms(1) have "sum_list ?qs \ A" by (rule direct_decompD) thus "0 \ A" by simp next fix p q assume "p \ A" with assms(1) obtain ps where ps: "ps \ listset ss" and p: "p = sum_list ps" by (rule direct_decompE) assume "q \ A" with assms(1) obtain qs where qs: "qs \ listset ss" and q: "q = sum_list qs" by (rule direct_decompE) from ps qs have l: "length ps = length qs" by (simp only: listsetD) from ps qs have "map2 (+) ps qs \ listset ss" (is "?qs \ _") by (rule listset_closed_map2) (auto dest: assms(2) subspace_add) with assms(1) have "sum_list ?qs \ A" by (rule direct_decompD) thus "p + q \ A" using l by (simp only: p q sum_list_map2_plus) next fix c p assume "p \ A" with assms(1) obtain ps where "ps \ listset ss" and p: "p = sum_list ps" by (rule direct_decompE) from this(1) have "map ((*s) c) ps \ listset ss" (is "?qs \ _") by (rule listset_closed_map) (auto dest: assms(2) subspace_scale) with assms(1) have "sum_list ?qs \ A" by (rule direct_decompD) also have "sum_list ?qs = c *s sum_list ps" by (induct ps) (simp_all add: scale_right_distrib) finally show "c *s p \ A" by (simp only: p) qed lemma is_basis_alt: "subspace V \ is_basis V B \ (independent B \ span B = V)" by (metis (full_types) is_basis_def dim_eq_card span_eq span_eq_iff) lemma is_basis_finite: "is_basis V A \ is_basis V B \ finite A \ finite B" unfolding is_basis_def using independent_span_bound by auto lemma some_basis_is_basis: "is_basis V (some_basis V)" proof - obtain B where "B \ V" and "independent B" and "V \ span B" and "card B = dim V" by (rule basis_exists) hence "is_basis V B" by (simp add: is_basis_def) thus ?thesis unfolding some_basis_def by (rule someI) qed corollary shows some_basis_subset: "some_basis V \ V" and independent_some_basis: "independent (some_basis V)" and span_some_basis_supset: "V \ span (some_basis V)" and card_some_basis: "card (some_basis V) = dim V" using some_basis_is_basis[of V] by (simp_all add: is_basis_def) lemma some_basis_not_zero: "0 \ some_basis V" using independent_some_basis dependent_zero by blast lemma span_some_basis: "subspace V \ span (some_basis V) = V" by (simp add: span_subspace some_basis_subset span_some_basis_supset) lemma direct_decomp_some_basis_pairwise_disjnt: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "pairwise (\s1 s2. disjnt (some_basis s1) (some_basis s2)) (set ss)" proof (rule pairwiseI) fix s1 s2 assume "s1 \ set ss" and "s2 \ set ss" and "s1 \ s2" have "some_basis s1 \ some_basis s2 \ s1 \ s2" using some_basis_subset by blast also from direct_decomp_pairwise_zero have "\ = {0}" proof (rule pairwiseD) fix s assume "s \ set ss" hence "subspace s" by (rule assms(2)) thus "0 \ s" by (rule subspace_0) qed fact+ finally have "some_basis s1 \ some_basis s2 \ {0}" . with some_basis_not_zero show "disjnt (some_basis s1) (some_basis s2)" unfolding disjnt_def by blast qed lemma direct_decomp_span_some_basis: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "span (\(some_basis ` set ss)) = A" proof - from assms(1) have eq0[symmetric]: "sum_list ` listset ss = A" by (rule direct_decompD) show ?thesis unfolding eq0 using assms(2) proof (induct ss) case Nil show ?case by simp next case (Cons s ss) have "subspace s" by (rule Cons.prems) simp hence eq1: "span (some_basis s) = s" by (rule span_some_basis) have "\s'. s' \ set ss \ subspace s'" by (rule Cons.prems) simp hence eq2: "span (\ (some_basis ` set ss)) = sum_list ` listset ss" by (rule Cons.hyps) have "span (\ (some_basis ` set (s # ss))) = {x + y |x y. x \ s \ y \ sum_list ` listset ss}" by (simp add: span_Un eq1 eq2) also have "\ = sum_list ` listset (s # ss)" (is "?A = ?B") proof show "?A \ ?B" proof fix a assume "a \ ?A" then obtain x y where "x \ s" and "y \ sum_list ` listset ss" and a: "a = x + y" by blast from this(2) obtain qs where "qs \ listset ss" and y: "y = sum_list qs" .. from \x \ s\ this(1) refl have "x # qs \ listset (s # ss)" by (rule listset_ConsI) hence "sum_list (x # qs) \ ?B" by (rule imageI) also have "sum_list (x # qs) = a" by (simp add: a y) finally show "a \ ?B" . qed next show "?B \ ?A" proof fix a assume "a \ ?B" then obtain qs' where "qs' \ listset (s # ss)" and a: "a = sum_list qs'" .. from this(1) obtain x qs where "x \ s" and "qs \ listset ss" and qs': "qs' = x # qs" by (rule listset_ConsE) from this(2) have "sum_list qs \ sum_list ` listset ss" by (rule imageI) moreover have "a = x + sum_list qs" by (simp add: a qs') ultimately show "a \ ?A" using \x \ s\ by blast qed qed finally show ?case . qed qed lemma direct_decomp_independent_some_basis: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "independent (\(some_basis ` set ss))" using assms proof (induct ss arbitrary: A) case Nil from independent_empty show ?case by simp next case (Cons s ss) have 1: "\s'. s' \ set ss \ subspace s'" by (rule Cons.prems) simp have "subspace s" by (rule Cons.prems) simp hence "0 \ s" and eq1: "span (some_basis s) = s" by (rule subspace_0, rule span_some_basis) from Cons.prems(1) have *: "direct_decomp A ([s] @ ss)" by simp moreover from \0 \ s\ have "{} \ set [s]" by auto ultimately have 2: "direct_decomp (sum_list ` listset ss) ss" by (rule direct_decomp_appendD) hence eq2: "span (\ (some_basis ` set ss)) = sum_list ` listset ss" using 1 by (rule direct_decomp_span_some_basis) note independent_some_basis[of s] moreover from 2 1 have "independent (\ (some_basis ` set ss))" by (rule Cons.hyps) moreover have "span (some_basis s) \ span (\ (some_basis ` set ss)) = {0}" proof - from * have "direct_decomp A [sum_list ` listset [s], sum_list ` listset ss]" by (rule direct_decomp_appendD) hence "direct_decomp A [s, sum_list ` listset ss]" by (simp add: image_image) moreover have "0 < (1::nat)" by simp moreover have "1 < length [s, sum_list ` listset ss]" by simp ultimately have "[s, sum_list ` listset ss] ! 0 \ [s, sum_list ` listset ss] ! 1 = {0}" by (rule direct_decomp_Int_zero) (auto simp: \0 \ s\ eq2[symmetric] span_zero) thus ?thesis by (simp add: eq1 eq2) qed ultimately have "independent (some_basis s \ (\ (some_basis ` set ss)))" by (rule independent_UnI) thus ?case by simp qed corollary direct_decomp_is_basis: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "is_basis A (\(some_basis ` set ss))" proof - from assms have "subspace A" by (rule subspace_direct_decomp) moreover from assms have "span (\(some_basis ` set ss)) = A" by (rule direct_decomp_span_some_basis) moreover from assms have "independent (\(some_basis ` set ss))" by (rule direct_decomp_independent_some_basis) ultimately show ?thesis by (simp add: is_basis_alt) qed lemma dim_direct_decomp: assumes "direct_decomp A ss" and "finite B" and "A \ span B" and "\s. s \ set ss \ subspace s" shows "dim A = (\s\set ss. dim s)" proof - from assms(1, 4) have "is_basis A (\(some_basis ` set ss))" (is "is_basis A ?B") by (rule direct_decomp_is_basis) hence "dim A = card ?B" and "independent ?B" and "?B \ A" by (simp_all add: is_basis_def) from this(3) assms(3) have "?B \ span B" by (rule subset_trans) with assms(2) \independent ?B\ have "finite ?B" using independent_span_bound by blast note \dim A = card ?B\ also from finite_set have "card ?B = (\s\set ss. card (some_basis s))" proof (intro card_UN_disjoint ballI impI) fix s assume "s \ set ss" with \finite ?B\ show "finite (some_basis s)" by auto next fix s1 s2 have "pairwise (\s t. disjnt (some_basis s) (some_basis t)) (set ss)" using assms(1, 4) by (rule direct_decomp_some_basis_pairwise_disjnt) moreover assume "s1 \ set ss" and "s2 \ set ss" and "s1 \ s2" thm pairwiseD ultimately have "disjnt (some_basis s1) (some_basis s2)" by (rule pairwiseD) thus "some_basis s1 \ some_basis s2 = {}" by (simp only: disjnt_def) qed also from refl card_some_basis have "\ = (\s\set ss. dim s)" by (rule sum.cong) finally show ?thesis . qed end (* vector_space *) subsection \Homogeneous Sets of Polynomials with Fixed Degree\ lemma homogeneous_set_direct_decomp: assumes "direct_decomp A ss" and "\s. s \ set ss \ homogeneous_set s" shows "homogeneous_set A" proof (rule homogeneous_setI) fix a n assume "a \ A" with assms(1) obtain qs where "qs \ listset ss" and a: "a = sum_list qs" by (rule direct_decompE) have "hom_component a n = hom_component (sum_list qs) n" by (simp only: a) also have "\ = sum_list (map (\q. hom_component q n) qs)" by (induct qs) (simp_all add: hom_component_plus) also from assms(1) have "\ \ A" proof (rule direct_decompD) show "map (\q. hom_component q n) qs \ listset ss" proof (rule listset_closed_map) fix s q assume "s \ set ss" hence "homogeneous_set s" by (rule assms(2)) moreover assume "q \ s" ultimately show "hom_component q n \ s" by (rule homogeneous_setD) qed fact qed finally show "hom_component a n \ A" . qed definition hom_deg_set :: "nat \ (('x \\<^sub>0 nat) \\<^sub>0 'a) set \ (('x \\<^sub>0 nat) \\<^sub>0 'a::zero) set" where "hom_deg_set z A = (\a. hom_component a z) ` A" lemma hom_deg_setD: assumes "p \ hom_deg_set z A" shows "homogeneous p" and "p \ 0 \ poly_deg p = z" proof - from assms obtain a where "a \ A" and p: "p = hom_component a z" unfolding hom_deg_set_def .. show *: "homogeneous p" by (simp only: p homogeneous_hom_component) assume "p \ 0" hence "keys p \ {}" by simp then obtain t where "t \ keys p" by blast with * have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) moreover from \t \ keys p\ have "deg_pm t = z" unfolding p by (rule keys_hom_componentD) ultimately show "poly_deg p = z" by simp qed lemma zero_in_hom_deg_set: assumes "0 \ A" shows "0 \ hom_deg_set z A" proof - have "0 = hom_component 0 z" by simp also from assms have "\ \ hom_deg_set z A" unfolding hom_deg_set_def by (rule imageI) finally show ?thesis . qed lemma hom_deg_set_closed_uminus: assumes "\a. a \ A \ - a \ A" and "p \ hom_deg_set z A" shows "- p \ hom_deg_set z A" proof - from assms(2) obtain a where "a \ A" and p: "p = hom_component a z" unfolding hom_deg_set_def .. from this(1) have "- a \ A" by (rule assms(1)) moreover have "- p = hom_component (- a) z" by (simp add: p) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_plus: assumes "\a1 a2. a1 \ A \ a2 \ A \ a1 + a2 \ A" and "p \ hom_deg_set z A" and "q \ hom_deg_set z A" shows "p + q \ hom_deg_set z A" proof - from assms(2) obtain a1 where "a1 \ A" and p: "p = hom_component a1 z" unfolding hom_deg_set_def .. from assms(3) obtain a2 where "a2 \ A" and q: "q = hom_component a2 z" unfolding hom_deg_set_def .. from \a1 \ A\ this(1) have "a1 + a2 \ A" by (rule assms(1)) moreover have "p + q = hom_component (a1 + a2) z" by (simp only: p q hom_component_plus) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_minus: assumes "\a1 a2. a1 \ A \ a2 \ A \ a1 - a2 \ A" and "p \ hom_deg_set z A" and "q \ hom_deg_set z A" shows "p - q \ hom_deg_set z A" proof - from assms(2) obtain a1 where "a1 \ A" and p: "p = hom_component a1 z" unfolding hom_deg_set_def .. from assms(3) obtain a2 where "a2 \ A" and q: "q = hom_component a2 z" unfolding hom_deg_set_def .. from \a1 \ A\ this(1) have "a1 - a2 \ A" by (rule assms(1)) moreover have "p - q = hom_component (a1 - a2) z" by (simp only: p q hom_component_minus) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_scalar: assumes "\a. a \ A \ c \ a \ A" and "p \ hom_deg_set z A" shows "(c::'a::semiring_0) \ p \ hom_deg_set z A" proof - from assms(2) obtain a where "a \ A" and p: "p = hom_component a z" unfolding hom_deg_set_def .. from this(1) have "c \ a \ A" by (rule assms(1)) moreover have "c \ p = hom_component (c \ a) z" by (simp add: p punit.map_scale_eq_monom_mult hom_component_monom_mult) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_sum: assumes "0 \ A" and "\a1 a2. a1 \ A \ a2 \ A \ a1 + a2 \ A" and "\i. i \ I \ f i \ hom_deg_set z A" shows "sum f I \ hom_deg_set z A" using assms(3) proof (induct I rule: infinite_finite_induct) case (infinite I) with assms(1) show ?case by (simp add: zero_in_hom_deg_set) next case empty with assms(1) show ?case by (simp add: zero_in_hom_deg_set) next case (insert j I) from insert.hyps(1, 2) have "sum f (insert j I) = f j + sum f I" by simp also from assms(2) have "\ \ hom_deg_set z A" proof (intro hom_deg_set_closed_plus insert.hyps) show "f j \ hom_deg_set z A" by (rule insert.prems) simp next fix i assume "i \ I" hence "i \ insert j I" by simp thus "f i \ hom_deg_set z A" by (rule insert.prems) qed finally show ?case . qed lemma hom_deg_set_subset: "homogeneous_set A \ hom_deg_set z A \ A" by (auto dest: homogeneous_setD simp: hom_deg_set_def) lemma Polys_closed_hom_deg_set: assumes "A \ P[X]" shows "hom_deg_set z A \ P[X]" proof fix p assume "p \ hom_deg_set z A" then obtain p' where "p' \ A" and p: "p = hom_component p' z" unfolding hom_deg_set_def .. from this(1) assms have "p' \ P[X]" .. have "keys p \ keys p'" by (simp add: p keys_hom_component) also from \p' \ P[X]\ have "\ \ .[X]" by (rule PolysD) finally show "p \ P[X]" by (rule PolysI) qed lemma hom_deg_set_alt_homogeneous_set: assumes "homogeneous_set A" shows "hom_deg_set z A = {p \ A. homogeneous p \ (p = 0 \ poly_deg p = z)}" (is "?A = ?B") proof show "?A \ ?B" proof fix h assume "h \ ?A" also from assms have "\ \ A" by (rule hom_deg_set_subset) finally show "h \ ?B" using \h \ ?A\ by (auto dest: hom_deg_setD) qed next show "?B \ ?A" proof fix h assume "h \ ?B" hence "h \ A" and "homogeneous h" and "h = 0 \ poly_deg h = z" by simp_all from this(3) show "h \ ?A" proof assume "h = 0" with \h \ A\ have "0 \ A" by simp thus ?thesis unfolding \h = 0\ by (rule zero_in_hom_deg_set) next assume "poly_deg h = z" with \homogeneous h\ have "h = hom_component h z" by (simp add: hom_component_of_homogeneous) with \h \ A\ show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed qed qed lemma hom_deg_set_sum_list_listset: assumes "A = sum_list ` listset ss" shows "hom_deg_set z A = sum_list ` listset (map (hom_deg_set z) ss)" (is "?A = ?B") proof show "?A \ ?B" proof fix h assume "h \ ?A" then obtain a where "a \ A" and h: "h = hom_component a z" unfolding hom_deg_set_def .. from this(1) obtain qs where "qs \ listset ss" and a: "a = sum_list qs" unfolding assms .. have "h = hom_component (sum_list qs) z" by (simp only: a h) also have "\ = sum_list (map (\q. hom_component q z) qs)" by (induct qs) (simp_all add: hom_component_plus) also have "\ \ ?B" proof (rule imageI) show "map (\q. hom_component q z) qs \ listset (map (hom_deg_set z) ss)" unfolding hom_deg_set_def using \qs \ listset ss\ refl by (rule listset_map_imageI) qed finally show "h \ ?B" . qed next show "?B \ ?A" proof fix h assume "h \ ?B" then obtain qs where "qs \ listset (map (hom_deg_set z) ss)" and h: "h = sum_list qs" .. from this(1) obtain qs' where "qs' \ listset ss" and qs: "qs = map (\q. hom_component q z) qs'" unfolding hom_deg_set_def by (rule listset_map_imageE) have "h = sum_list (map (\q. hom_component q z) qs')" by (simp only: h qs) also have "\ = hom_component (sum_list qs') z" by (induct qs') (simp_all add: hom_component_plus) finally have "h = hom_component (sum_list qs') z" . moreover have "sum_list qs' \ A" unfolding assms using \qs' \ listset ss\ by (rule imageI) ultimately show "h \ ?A" unfolding hom_deg_set_def by (rule image_eqI) qed qed lemma direct_decomp_hom_deg_set: assumes "direct_decomp A ss" and "\s. s \ set ss \ homogeneous_set s" shows "direct_decomp (hom_deg_set z A) (map (hom_deg_set z) ss)" proof (rule direct_decompI) from assms(1) have "sum_list ` listset ss = A" by (rule direct_decompD) from this[symmetric] show "sum_list ` listset (map (hom_deg_set z) ss) = hom_deg_set z A" by (simp only: hom_deg_set_sum_list_listset) next from assms(1) have "inj_on sum_list (listset ss)" by (rule direct_decompD) moreover have "listset (map (hom_deg_set z) ss) \ listset ss" proof (rule listset_mono) fix i assume "i < length ss" hence "map (hom_deg_set z) ss ! i = hom_deg_set z (ss ! i)" by simp also from \i < length ss\ have "\ \ ss ! i" by (intro hom_deg_set_subset assms(2) nth_mem) finally show "map (hom_deg_set z) ss ! i \ ss ! i" . qed simp ultimately show "inj_on sum_list (listset (map (hom_deg_set z) ss))" by (rule inj_on_subset) qed subsection \Interpreting Polynomial Rings as Vector Spaces over the Coefficient Field\ text \There is no need to set up any further interpretation, since interpretation \phull\ is exactly what we need.\ lemma subspace_ideal: "phull.subspace (ideal (F::('b::comm_powerprod \\<^sub>0 'a::field) set))" using ideal.span_zero ideal.span_add proof (rule phull.subspaceI) fix c p assume "p \ ideal F" thus "c \ p \ ideal F" unfolding map_scale_eq_times by (rule ideal.span_scale) qed lemma subspace_Polys: "phull.subspace (P[X]::(('x \\<^sub>0 nat) \\<^sub>0 'a::field) set)" using zero_in_Polys Polys_closed_plus Polys_closed_map_scale by (rule phull.subspaceI) lemma subspace_hom_deg_set: assumes "phull.subspace A" shows "phull.subspace (hom_deg_set z A)" (is "phull.subspace ?A") proof (rule phull.subspaceI) from assms have "0 \ A" by (rule phull.subspace_0) thus "0 \ ?A" by (rule zero_in_hom_deg_set) next fix p q assume "p \ ?A" and "q \ ?A" with phull.subspace_add show "p + q \ ?A" by (rule hom_deg_set_closed_plus) (rule assms) next fix c p assume "p \ ?A" with phull.subspace_scale show "c \ p \ ?A" by (rule hom_deg_set_closed_scalar) (rule assms) qed lemma hom_deg_set_Polys_eq_span: "hom_deg_set z P[X] = phull.span (monomial (1::'a::field) ` deg_sect X z)" (is "?A = ?B") proof show "?A \ ?B" proof fix p assume "p \ ?A" also from this have "\ = {p \ P[X]. homogeneous p \ (p = 0 \ poly_deg p = z)}" by (simp only: hom_deg_set_alt_homogeneous_set[OF homogeneous_set_Polys]) finally have "p \ P[X]" and "homogeneous p" and "p \ 0 \ poly_deg p = z" by simp_all thus "p \ ?B" proof (induct p rule: poly_mapping_plus_induct) case 1 from phull.span_zero show ?case . next case (2 p c t) let ?m = "monomial c t" from 2(1) have "t \ keys ?m" by simp hence "t \ keys (?m + p)" using 2(2) by (rule in_keys_plusI1) hence "?m + p \ 0" by auto hence "poly_deg (monomial c t + p) = z" by (rule 2) from 2(4) have "keys (?m + p) \ .[X]" by (rule PolysD) with \t \ keys (?m + p)\ have "t \ .[X]" .. hence "?m \ P[X]" by (rule Polys_closed_monomial) have "t \ deg_sect X z" proof (rule deg_sectI) from 2(5) \t \ keys (?m + p)\ have "deg_pm t = poly_deg (?m + p)" by (rule homogeneousD_poly_deg) also have "\ = z" by fact finally show "deg_pm t = z" . qed fact hence "monomial 1 t \ monomial 1 ` deg_sect X z" by (rule imageI) hence "monomial 1 t \ ?B" by (rule phull.span_base) hence "c \ monomial 1 t \ ?B" by (rule phull.span_scale) hence "?m \ ?B" by simp moreover have "p \ ?B" proof (rule 2) from 2(4) \?m \ P[X]\ have "(?m + p) - ?m \ P[X]" by (rule Polys_closed_minus) thus "p \ P[X]" by simp next have 1: "deg_pm s = z" if "s \ keys p" for s proof - from that 2(2) have "s \ t" by blast hence "s \ keys ?m" by simp with that have "s \ keys (?m + p)" by (rule in_keys_plusI2) with 2(5) have "deg_pm s = poly_deg (?m + p)" by (rule homogeneousD_poly_deg) also have "\ = z" by fact finally show ?thesis . qed show "homogeneous p" by (rule homogeneousI) (simp add: 1) assume "p \ 0" show "poly_deg p = z" proof (rule antisym) show "poly_deg p \ z" by (rule poly_deg_leI) (simp add: 1) next from \p \ 0\ have "keys p \ {}" by simp then obtain s where "s \ keys p" by blast hence "z = deg_pm s" by (simp only: 1) also from \s \ keys p\ have "\ \ poly_deg p" by (rule poly_deg_max_keys) finally show "z \ poly_deg p" . qed qed ultimately show ?case by (rule phull.span_add) qed qed next show "?B \ ?A" proof fix p assume "p \ ?B" then obtain M u where "M \ monomial 1 ` deg_sect X z" and "finite M" and p: "p = (\m\M. u m \ m)" by (auto simp: phull.span_explicit) from this(1) obtain T where "T \ deg_sect X z" and M: "M = monomial 1 ` T" and inj: "inj_on (monomial (1::'a)) T" by (rule subset_imageE_inj) define c where "c = (\t. u (monomial 1 t))" from inj have "p = (\t\T. monomial (c t) t)" by (simp add: p M sum.reindex c_def) also have "\ \ ?A" proof (intro hom_deg_set_closed_sum zero_in_Polys Polys_closed_plus) fix t assume "t \ T" hence "t \ deg_sect X z" using \T \ deg_sect X z\ .. hence "t \ .[X]" and eq: "deg_pm t = z" by (rule deg_sectD)+ from this(1) have "monomial (c t) t \ P[X]" (is "?m \ _") by (rule Polys_closed_monomial) thus "?m \ ?A" by (simp add: hom_deg_set_alt_homogeneous_set[OF homogeneous_set_Polys] poly_deg_monomial monomial_0_iff eq) qed finally show "p \ ?A" . qed qed subsection \(Projective) Hilbert Function\ interpretation phull: vector_space map_scale apply standard subgoal by (fact map_scale_distrib_left) subgoal by (fact map_scale_distrib_right) subgoal by (fact map_scale_assoc) subgoal by (fact map_scale_one_left) done definition Hilbert_fun :: "(('x \\<^sub>0 nat) \\<^sub>0 'a::field) set \ nat \ nat" where "Hilbert_fun A z = phull.dim (hom_deg_set z A)" lemma Hilbert_fun_empty [simp]: "Hilbert_fun {} = 0" by (rule ext) (simp add: Hilbert_fun_def hom_deg_set_def) lemma Hilbert_fun_zero [simp]: "Hilbert_fun {0} = 0" by (rule ext) (simp add: Hilbert_fun_def hom_deg_set_def) lemma Hilbert_fun_direct_decomp: assumes "finite X" and "A \ P[X]" and "direct_decomp (A::(('x::countable \\<^sub>0 nat) \\<^sub>0 'a::field) set) ps" and "\s. s \ set ps \ homogeneous_set s" and "\s. s \ set ps \ phull.subspace s" shows "Hilbert_fun A z = (\p\set ps. Hilbert_fun p z)" proof - from assms(3, 4) have dd: "direct_decomp (hom_deg_set z A) (map (hom_deg_set z) ps)" by (rule direct_decomp_hom_deg_set) have "Hilbert_fun A z = phull.dim (hom_deg_set z A)" by (fact Hilbert_fun_def) also from dd have "\ = sum phull.dim (set (map (hom_deg_set z) ps))" proof (rule phull.dim_direct_decomp) from assms(1) have "finite (deg_sect X z)" by (rule finite_deg_sect) thus "finite (monomial (1::'a) ` deg_sect X z)" by (rule finite_imageI) next from assms(2) have "hom_deg_set z A \ hom_deg_set z P[X]" unfolding hom_deg_set_def by (rule image_mono) thus "hom_deg_set z A \ phull.span (monomial 1 ` deg_sect X z)" by (simp only: hom_deg_set_Polys_eq_span) next fix s assume "s \ set (map (hom_deg_set z) ps)" then obtain s' where "s' \ set ps" and s: "s = hom_deg_set z s'" unfolding set_map .. from this(1) have "phull.subspace s'" by (rule assms(5)) thus "phull.subspace s" unfolding s by (rule subspace_hom_deg_set) qed also have "\ = sum (phull.dim \ hom_deg_set z) (set ps)" unfolding set_map using finite_set proof (rule sum.reindex_nontrivial) fix s1 s2 note dd moreover assume "s1 \ set ps" and "s2 \ set ps" and "s1 \ s2" moreover have "0 \ hom_deg_set z s" if "s \ set ps" for s proof (rule zero_in_hom_deg_set) from that have "phull.subspace s" by (rule assms(5)) thus "0 \ s" by (rule phull.subspace_0) qed ultimately have "hom_deg_set z s1 \ hom_deg_set z s2 = {0}" by (rule direct_decomp_map_Int_zero) moreover assume "hom_deg_set z s1 = hom_deg_set z s2" ultimately show "phull.dim (hom_deg_set z s1) = 0" by simp qed also have "\ = (\p\set ps. Hilbert_fun p z)" by (simp only: o_def Hilbert_fun_def) finally show ?thesis . qed context pm_powerprod begin lemma image_lt_hom_deg_set: assumes "homogeneous_set A" shows "lpp ` (hom_deg_set z A - {0}) = {t \ lpp ` (A - {0}). deg_pm t = z}" (is "?B = ?A") proof (intro set_eqI iffI) fix t assume "t \ ?A" hence "t \ lpp ` (A - {0})" and deg_t[symmetric]: "deg_pm t = z" by simp_all from this(1) obtain p where "p \ A - {0}" and t: "t = lpp p" .. from this(1) have "p \ A" and "p \ 0" by simp_all from this(1) have 1: "hom_component p z \ hom_deg_set z A" (is "?p \ _") unfolding hom_deg_set_def by (rule imageI) from \p \ 0\ have "?p \ 0" and "lpp ?p = t" unfolding t deg_t by (rule hom_component_lpp)+ note this(2)[symmetric] moreover from 1 \?p \ 0\ have "?p \ hom_deg_set z A - {0}" by simp ultimately show "t \ ?B" by (rule image_eqI) next fix t assume "t \ ?B" then obtain p where "p \ hom_deg_set z A - {0}" and t: "t = lpp p" .. from this(1) have "p \ hom_deg_set z A" and "p \ 0" by simp_all with assms have "p \ A" and "homogeneous p" and "poly_deg p = z" by (simp_all add: hom_deg_set_alt_homogeneous_set) from this(1) \p \ 0\ have "p \ A - {0}" by simp hence 1: "t \ lpp ` (A - {0})" using t by (rule rev_image_eqI) from \p \ 0\ have "t \ keys p" unfolding t by (rule punit.lt_in_keys) with \homogeneous p\ have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) with 1 show "t \ ?A" by (simp add: \poly_deg p = z\) qed lemma Hilbert_fun_alt: assumes "finite X" and "A \ P[X]" and "phull.subspace A" shows "Hilbert_fun A z = card (lpp ` (hom_deg_set z A - {0}))" (is "_ = card ?A") proof - have "?A \ lpp ` (hom_deg_set z A - {0})" by simp then obtain B where sub: "B \ hom_deg_set z A - {0}" and eq1: "?A = lpp ` B" and inj: "inj_on lpp B" by (rule subset_imageE_inj) have "Hilbert_fun A z = phull.dim (hom_deg_set z A)" by (fact Hilbert_fun_def) also have "\ = card B" proof (rule phull.dim_eq_card) show "phull.span B = phull.span (hom_deg_set z A)" proof from sub have "B \ hom_deg_set z A" by blast thus "phull.span B \ phull.span (hom_deg_set z A)" by (rule phull.span_mono) next from assms(3) have "phull.subspace (hom_deg_set z A)" by (rule subspace_hom_deg_set) hence "phull.span (hom_deg_set z A) = hom_deg_set z A" by (simp only: phull.span_eq_iff) also have "\ \ phull.span B" proof (rule ccontr) assume "\ hom_deg_set z A \ phull.span B" then obtain p0 where "p0 \ hom_deg_set z A - phull.span B" (is "_ \ ?B") by blast note assms(1) this moreover have "?B \ P[X]" proof (rule subset_trans) from assms(2) show "hom_deg_set z A \ P[X]" by (rule Polys_closed_hom_deg_set) qed blast ultimately obtain p where "p \ ?B" and p_min: "\q. punit.ord_strict_p q p \ q \ ?B" by (rule punit.ord_p_minimum_dgrad_p_set[OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]) blast from this(1) have "p \ hom_deg_set z A" and "p \ phull.span B" by simp_all from phull.span_zero this(2) have "p \ 0" by blast with \p \ hom_deg_set z A\ have "p \ hom_deg_set z A - {0}" by simp hence "lpp p \ lpp ` (hom_deg_set z A - {0})" by (rule imageI) also have "\ = lpp ` B" by (simp only: eq1) finally obtain b where "b \ B" and eq2: "lpp p = lpp b" .. from this(1) sub have "b \ hom_deg_set z A - {0}" .. hence "b \ hom_deg_set z A" and "b \ 0" by simp_all from this(2) have lcb: "punit.lc b \ 0" by (rule punit.lc_not_0) from \p \ 0\ have lcp: "punit.lc p \ 0" by (rule punit.lc_not_0) from \b \ B\ have "b \ phull.span B" by (rule phull.span_base) hence "(punit.lc p / punit.lc b) \ b \ phull.span B" (is "?b \ _") by (rule phull.span_scale) with \p \ phull.span B\ have "p - ?b \ 0" by auto moreover from lcb lcp \b \ 0\ have "lpp ?b = lpp p" by (simp add: punit.map_scale_eq_monom_mult punit.lt_monom_mult eq2) moreover from lcb have "punit.lc ?b = punit.lc p" by (simp add: punit.map_scale_eq_monom_mult) ultimately have "lpp (p - ?b) \ lpp p" by (rule punit.lt_minus_lessI) hence "punit.ord_strict_p (p - ?b) p" by (rule punit.lt_ord_p) hence "p - ?b \ ?B" by (rule p_min) hence "p - ?b \ hom_deg_set z A \ p - ?b \ phull.span B" by simp thus False proof assume *: "p - ?b \ hom_deg_set z A" from phull.subspace_scale have "?b \ hom_deg_set z A" proof (rule hom_deg_set_closed_scalar) show "phull.subspace A" by fact next show "b \ hom_deg_set z A" by fact qed with phull.subspace_diff \p \ hom_deg_set z A\ have "p - ?b \ hom_deg_set z A" by (rule hom_deg_set_closed_minus) (rule assms(3)) with * show ?thesis .. next assume "p - ?b \ phull.span B" hence "p - ?b + ?b \ phull.span B" using \?b \ phull.span B\ by (rule phull.span_add) hence "p \ phull.span B" by simp with \p \ phull.span B\ show ?thesis .. qed qed finally show "phull.span (hom_deg_set z A) \ phull.span B" . qed next show "phull.independent B" proof assume "phull.dependent B" then obtain B' u b' where "finite B'" and "B' \ B" and "(\b\B'. u b \ b) = 0" and "b' \ B'" and "u b' \ 0" unfolding phull.dependent_explicit by blast define B0 where "B0 = {b \ B'. u b \ 0}" have "B0 \ B'" by (simp add: B0_def) with \finite B'\ have "(\b\B0. u b \ b) = (\b\B'. u b \ b)" by (rule sum.mono_neutral_left) (simp add: B0_def) also have "\ = 0" by fact finally have eq: "(\b\B0. u b \ b) = 0" . define t where "t = ordered_powerprod_lin.Max (lpp ` B0)" from \b' \ B'\ \u b' \ 0\ have "b' \ B0" by (simp add: B0_def) hence "lpp b' \ lpp ` B0" by (rule imageI) hence "lpp ` B0 \ {}" by blast from \B0 \ B'\ \finite B'\ have "finite B0" by (rule finite_subset) hence "finite (lpp ` B0)" by (rule finite_imageI) hence "t \ lpp ` B0" unfolding t_def using \lpp ` B0 \ {}\ by (rule ordered_powerprod_lin.Max_in) then obtain b0 where "b0 \ B0" and t: "t = lpp b0" .. note this(1) moreover from \B0 \ B'\ \B' \ B\ have "B0 \ B" by (rule subset_trans) also have "\ \ hom_deg_set z A - {0}" by fact finally have "b0 \ hom_deg_set z A - {0}" . hence "b0 \ 0" by simp hence "t \ keys b0" unfolding t by (rule punit.lt_in_keys) have "lookup (\b\B0. u b \ b) t = (\b\B0. u b * lookup b t)" by (simp add: lookup_sum) also from \finite B0\ have "\ = (\b\{b0}. u b * lookup b t)" proof (rule sum.mono_neutral_right) from \b0 \ B0\ show "{b0} \ B0" by simp next show "\b\B0 - {b0}. u b * lookup b t = 0" proof fix b assume "b \ B0 - {b0}" hence "b \ B0" and "b \ b0" by simp_all from this(1) have "lpp b \ lpp ` B0" by (rule imageI) with \finite (lpp ` B0)\ have "lpp b \ t" unfolding t_def by (rule ordered_powerprod_lin.Max_ge) have "t \ keys b" proof assume "t \ keys b" hence "t \ lpp b" by (rule punit.lt_max_keys) with \lpp b \ t\ have "lpp b = lpp b0" unfolding t by (rule ordered_powerprod_lin.antisym) from inj \B0 \ B\ have "inj_on lpp B0" by (rule inj_on_subset) hence "b = b0" using \lpp b = lpp b0\ \b \ B0\ \b0 \ B0\ by (rule inj_onD) with \b \ b0\ show False .. qed thus "u b * lookup b t = 0" by (simp add: in_keys_iff) qed qed also from \t \ keys b0\ \b0 \ B0\ have "\ \ 0" by (simp add: B0_def in_keys_iff) finally show False by (simp add: eq) qed qed also have "\ = card ?A" unfolding eq1 using inj by (rule card_image[symmetric]) finally show ?thesis . qed end (* pm_powerprod *) end (* theory *) diff --git a/thys/Program-Conflict-Analysis/Interleave.thy b/thys/Program-Conflict-Analysis/Interleave.thy --- a/thys/Program-Conflict-Analysis/Interleave.thy +++ b/thys/Program-Conflict-Analysis/Interleave.thy @@ -1,417 +1,417 @@ (* Title: Conflict analysis/List Interleaving Operator Author: Peter Lammich Maintainer: Peter Lammich *) section "List Interleaving Operator" theory Interleave -imports Main "HOL-Library.List_Permutation" Misc +imports Main Misc begin text_raw \\label{thy:Interleave}\ text \ This theory defines an operator to return the set of all possible interleavings of two lists. \ (* Is \-operator needed ? Should we better do a reformulation of \ on sets of lists, to get an associative, commutative operator with identity (ACI ?) ? *) subsection "Definitions" subsubsection "Prepend and concatenate lifted to sets" definition list_set_cons :: "'a \ 'a list set \ 'a list set" (infixr "\" 65) where [simp]: "a\S = (#) a ` S" definition list_set_append :: "'a list \ 'a list set \ 'a list set" (infixr "\" 65) where [simp]: "a\S = (@) a ` S" subsubsection "The interleaving operator" function interleave :: "'a list \ 'a list \ 'a list set" (infixr "\" 64) where "[]\b = {b}" | "a\[] = {a}" | "a#l \ b#r = ((a\(l \ b#r)) \ (b\(a#l \ r)))" by pat_completeness auto termination by lexicographic_order subsection "Properties" subsubsection "Lifted prepend and concatenate operators" lemma cons_set_cons_eq: "a#l \ b\S = (a=b & l\S)" by auto lemma append_set_append_eq: "\length a = length b\ \ a@l \ b\S = (a=b & l\S)" by auto lemma list_set_cons_cases[cases set]: "\w\a\S; !!w'. \ w=a#w'; w'\S \ \ P\ \ P" by auto lemma list_set_append_cases[cases set]: "\w\a\S; !!w'. \ w=a@w'; w'\S \ \ P\ \ P" by auto subsubsection "Standard simplification-, introduction-, elimination-, and induction rules" lemma interleave_cons1: "l \ a\b \ x#l \ x#a \ b" apply(case_tac b) apply(auto) done lemma interleave_cons2: "l \ a\b \ x#l \ a \ x#b" apply(case_tac a) apply(auto) done lemmas interleave_cons = interleave_cons1 interleave_cons2 lemma interleave_empty[simp]: "[]\a\b = (a=[] & b=[])" apply(case_tac a) apply(case_tac b) apply(simp) apply(simp) apply(case_tac b) apply(auto) done (* TODO: Is this wise as default simp ?*) lemma interleave_length[rule_format, simp]: "ALL x : a\b . length x = length a + length b" apply(induct rule: interleave.induct) apply(auto) done lemma interleave_same[simp]: "y\l\y = (l=[])" apply (rule iffI) apply (subgoal_tac "length y = length l + length y") apply (simp del: interleave_length) apply (erule interleave_length) apply simp done lemma interleave_comm[simp]: "a\b = b\a" (is "?P f a b") apply(induct rule: interleave.induct) apply(auto) done lemma interleave_cont_conc[rule_format, simp]: "ALL b . a@b \ a\b" apply(induct_tac a) apply(auto simp add: interleave_cons) done lemma interleave_cont_rev_conc[simp]: "b@a \ a\b" apply (subgoal_tac "b@a \ b\a") apply(simp) apply(simp only: interleave_cont_conc) done lemma interleave_not_empty: "a\b ~= {}" apply(induct rule: interleave.induct) apply(auto) done lemma cons_interleave_split: "\a#l \ l1\l2\ \ (EX lh . l1=a#lh & l \ lh\l2 | l2=a#lh & l \ l1\lh )" apply(case_tac l1) apply(auto) apply(case_tac l2) apply(auto) done lemma cons_interleave_cases[cases set, case_names left right]: "\a#l \ l1\l2; !!lh . \l1=a#lh; l \ lh\l2\ \ P; !!lh. \l2=a#lh; l \ l1\lh\ \ P\ \ P" by (blast dest: cons_interleave_split) lemma interleave_cases[cases set, case_names empty left right]: "\l\l1\l2; \ l=[]; l1=[]; l2=[] \ \ P; !!a l' l1'. \l=a#l'; l1=a#l1'; l'\l1'\l2\ \ P; !!a l' l2'. \l=a#l'; l2=a#l2'; l'\l1\l2'\ \ P \ \ P" apply (cases l) apply (simp) apply simp apply (erule cons_interleave_cases) apply simp_all done lemma interleave_elem_induct[induct set, case_names empty left right]: "!!w1 w2. \ w\w1\w2; P [] [] []; !!e w w1 w2. \ P w w1 w2; w\w1\w2 \ \ P (e#w) (e#w1) w2; !!e w w1 w2. \ P w w1 w2; w\w1\w2 \ \ P (e#w) w1 (e#w2) \ \ P w w1 w2" by (induct w) (auto elim!: cons_interleave_cases intro!: interleave_cons) lemma interleave_rev_cons1[rule_format]: "ALL a b . l \ a\b \ l@[x] \ a@[x] \ b" (is "?P l") proof (induct l) case Nil show ?case by (simp) case (Cons e l) assume IH[rule_format]: "?P l" show "?P (e#l)" proof (intro allI impI) fix a b assume "e#l \ a\b" then obtain c where SPLIT: "a=e#c & l\c\b | b=e#c & l\a\c" by (blast dest: cons_interleave_split) with IH have "a=e#c & l@[x]\c@[x]\b | b=e#c & l@[x]\a@[x]\c" by auto hence "a=e#c & e#l@[x]\e#c@[x]\b | b=e#c & e#l@[x]\a@[x]\e#c" by (auto simp add: interleave_cons) thus "(e#l)@[x]\a@[x]\b" by auto qed qed corollary interleave_rev_cons2: "l \ a\b \ l@[x] \ a \ b@[x]" proof - assume "l \ a\b" hence "l\b\a" by auto hence "l@[x] \ b@[x] \ a" by (blast dest: interleave_rev_cons1) thus ?thesis by auto qed lemmas interleave_rev_cons = interleave_rev_cons1 interleave_rev_cons2 subsubsection "Interleaving and list concatenation" lemma interleave_append1: "\l \ a\b\ \ x@l \ x@a \ b" proof - have "ALL l a b . l \ a\b \ x@l \ x@a \ b" (is "?P x") proof (induct x) show "?P []" by simp next fix e fix x::"'a list" assume IH: "?P x" show "?P (e#x)" proof (intro impI allI) fix l::"'a list" fix a b assume "l \ a \ b" with IH have "x@l \ x@a \ b" by auto with interleave_cons1 show "(e # x) @ l \ (e # x) @ a \ b" by force qed qed moreover assume "l \ a \ b" ultimately show ?thesis by blast qed lemma interleave_append2: "\l \ a\b\ \ x@l \ a \ x@b" proof - have "ALL l a b . l \ a\b \ x@l \ a \ x@b" (is "?P x") proof (induct x) show "?P []" by simp next fix a fix x::"'a list" assume IH: "\l a b. l \ a \ b \ x @ l \ a \ x @ b" show "\l aa b. l \ aa \ b \ (a # x) @ l \ aa \ (a # x) @ b" proof (intro impI allI) fix l::"'a list" fix aa b assume "l \ aa \ b" with IH have "x@l \ aa \ x@b" by auto with interleave_cons2 show "(a # x) @ l \ aa \ (a # x) @ b" by force qed qed moreover assume "l \ a \ b" ultimately show ?thesis by blast qed lemmas interleave_append = interleave_append1 interleave_append2 lemma interleave_rev_append1: "!!a b w. w\a\b \ w@w' \ a@w' \ b" proof (induct w' rule: rev_induct) case Nil thus ?case by simp next case (snoc e w') note IHP=this hence "w@w'\a@w'\b" by auto thus ?case using interleave_rev_cons1[of "w@w'" "a@w'" "b"] by (simp) qed lemma interleave_rev_append2: "w\a\b \ w@w' \ a \ b@w'" by (simp only: interleave_comm[of a b] interleave_comm[of a "b@w'"]) (blast dest: interleave_rev_append1) lemmas interleave_rev_append = interleave_rev_append1 interleave_rev_append2 lemma interleave_rev1[rule_format]: "ALL w1 w2 . (w\w1\w2) \ (rev w \ rev w1 \ rev w2)" (is "?P w") proof (induct w) case Nil show ?case by (simp) case (Cons e w) assume IH[rule_format]: "?P w" show ?case proof (clarsimp) fix w1 w2 assume "e # w \ w1 \ w2" then obtain w' where "w1=e#w' & w\w'\w2 | w2=e#w' & w\w1\w'" by (blast dest: cons_interleave_split) with IH have "w1=e#w' & rev w\rev w' \ rev w2 | w2=e#w' & rev w \ rev w1 \ rev w'" by auto thus "rev w @ [e]\rev w1 \ rev w2" by (auto simp add: interleave_rev_cons) qed qed corollary interleave_rev2: "(rev w \ rev w1 \ rev w2) \ (w\w1\w2)" apply (subgoal_tac "(rev w\rev w1\rev w2) = (rev (rev w) \ rev (rev w1) \ rev (rev w2))") apply(simp) apply (blast dest: interleave_rev1) done lemmas interleave_rev = interleave_rev1 interleave_rev2 lemma rev_cons_interleave_split: "\l@[a] \ l1\l2\ \ (EX lh . l1=lh@[a] & l \ lh\l2 | l2=lh@[a] & l \ l1\lh )" proof - assume "l@[a] \ l1\l2" hence "a#rev l \ rev l1 \ rev l2" by (auto dest: interleave_rev) then obtain lh where "rev l1 = a#lh & rev l \ lh\rev l2 | rev l2 = a#lh & rev l \ rev l1 \ lh" by (blast dest: cons_interleave_split) hence "rev l1 = a#lh & l \ rev lh \ l2 | rev l2 = a#lh & l \ l1 \ rev lh" by (auto dest: interleave_rev) hence "l1 = rev lh @ [a] & l \ rev lh \ l2 | l2 = rev lh @ [a] & l \ l1 \ rev lh" by (simp add: rev_swap) thus ?thesis by blast qed subsubsection "Associativity" lemma interleave_s_assoc1[rule_format]: "ALL w1 ws w2 w3 . (w\w1\ws & ws\w2\w3 \ (EX ws' : w1\w2 . w \ ws'\w3))" proof (induct w) case Nil show ?case by (auto) case (Cons e w) assume IH: "ALL w1 ws w2 w3 . w\w1\ws & ws\w2\w3 \ (EX ws' : w1\w2 . w \ ws'\w3)" show ?case proof (intro impI allI) fix w1 ws w2 w3 assume A: "e#w \ w1 \ ws \ ws \ w2 \ w3" then obtain wh where CASES: "w1=e#wh & w\wh\ws | ws=e#wh & w\w1\wh" by (blast dest!: cons_interleave_split) moreover { assume CASE: "w1=e#wh & w\wh\ws" with A IH obtain ws' where "ws'\wh\w2 & w\ws'\w3" by (blast) hence "e#ws'\ (e#wh)\w2 & e#w \ (e#ws')\w3" by (auto simp add: interleave_cons) with CASE have "\ws'\w1 \ w2. e # w \ ws' \ w3" by blast } moreover { assume CASE: "ws=e#wh & w\w1\wh" with A obtain whh where "w2=e#whh & wh\whh\w3 | w3=e#whh & wh\w2\whh" by (blast dest!: cons_interleave_split) moreover { assume SCASE: "w2=e#whh & wh\whh\w3" with CASE IH obtain ws' where "ws'\w1\whh & w\ws'\w3" by blast with SCASE have "e#ws'\w1\w2 & e#w \ (e#ws')\w3" by (auto simp add: interleave_cons) hence "\ws'\w1 \ w2. e # w \ ws' \ w3" by blast } moreover { assume SCASE: "w3=e#whh & wh\w2\whh" with CASE IH obtain ws' where "ws'\w1\w2 & w\ws'\whh" by blast with SCASE have "ws'\w1\w2 & e#w \ ws'\w3" by (auto simp add: interleave_cons) hence "\ws'\w1 \ w2. e # w \ ws' \ w3" by blast } ultimately have "\ws'\w1 \ w2. e # w \ ws' \ w3" by blast } ultimately show "\ws'\w1 \ w2. e # w \ ws' \ w3" by blast qed qed lemma interleave_s_assoc2: "\w\ws\w3; ws\w1\w2\ \ EX ws' : w2\w3 . w \ w1\ws'" proof - assume "w \ ws \ w3" "ws \ w1 \ w2" hence "w \ w3 \ ws & ws \ w2 \ w1" by simp hence "EX ws':w3\w2 . w\ws'\w1" by (simp only: interleave_s_assoc1) thus ?thesis by simp qed lemmas interleave_s_assoc = interleave_s_assoc1 interleave_s_assoc2 subsubsection "Relation to other standard list operations" lemma interleave_set: "w\w1\w2 \ set w = set w1 \ set w2" by (induct rule: interleave_elem_induct) auto lemma interleave_map: "w\w1\w2 \ map f w \ map f w1 \ map f w2" by (induct rule: interleave_elem_induct) (auto intro!: interleave_cons) lemma interleave_filter: "w\w1\w2 \ filter f w \ filter f w1 \ filter f w2" by (induct rule: interleave_elem_induct) (auto intro!: interleave_cons) subsubsection "Relation to sublist order" lemma ileq_interleave_alt: "l'\l == \lb. l\lb \ l'" proof (rule eq_reflection) {fix l' l have "l'\l \ \lb. l\lb \ l'" by (induct rule: less_eq_list_induct) (simp, (blast intro: interleave_cons)+)} moreover {fix l have "!!lb l'. l\lb\l' \ l'\l" by (induct l) (auto intro: less_eq_list_drop elim!: cons_interleave_cases)} ultimately show "(l'\l) = (\lb. l\lb \ l')" by blast qed lemma ileq_interleave: "w\w1\w2 \ w1\w & w2\w" by (unfold ileq_interleave_alt, auto) lemma ilt_ex_notempty: "x < y \ (\xs. xs \ [] \ y \ xs \ x)" apply (auto simp add: order_less_le ileq_interleave_alt) apply (case_tac lb) apply auto done lemma ilt_interleave1: "\w\w1\w2; w1~=[]\ \ w2w\w1\w2; w2~=[]\ \ w1 \Recover structure of @{text w} wrt. to structure of @{text "w1"}\ lemma interleave_recover1[rule_format]: "ALL w1a w1b w2 . w\(w1a@w1b)\w2 \ (EX wa wb w2a w2b . w=wa@wb & w2=w2a@w2b & wa\w1a\w2a & wb\w1b\w2b)" (is "?P w" is "ALL w1a w1b w2 . ?PRE w w1a w1b w2 \ ?CONS w w1a w1b w2") proof (induct w) case Nil show ?case by (auto) case (Cons e w) assume IH[rule_format]: "?P w" show "?P (e#w)" proof (intro allI impI) fix w1a w1b w2 assume PRE: "e # w \ w1a @ w1b \ w2" { assume CASE: "w1a=[]" with PRE have "e#w=[]@(e#w) & w2=[]@w2 & []\w1a\[] & e#w\w1b\w2" by (auto) hence "?CONS (e#w) w1a w1b w2" by blast } moreover { assume CASE: "w1a~=[]" with PRE obtain wh where SCASES: "w1a@w1b=e#wh & w\wh\w2 | w2=e#wh & w\w1a@w1b\wh" by (blast dest: cons_interleave_split) moreover { assume SCASE: "w1a@w1b=e#wh & w\wh\w2" with CASE obtain w1ah where W1AFMT: "w1a=e#w1ah & w1ah@w1b=wh" by (cases w1a, auto) with SCASE have "w\w1ah@w1b\w2" by auto with IH[of w1ah w1b w2] obtain wa wb w2a w2b where "w=wa@wb & w2=w2a@w2b & wa\w1ah\w2a & wb\w1b\w2b" by (blast) with W1AFMT have "e#w=(e#wa)@wb & e#w\e#wa\wb & w2=w2a@w2b & e#wa\w1a\w2a & wb\w1b\w2b" by (auto simp add: interleave_cons) hence "?CONS (e#w) w1a w1b w2" by blast } moreover { assume SCASE: "w2=e#wh & w\w1a@w1b\wh" with IH[of w1a w1b wh] obtain wa wb w2a w2b where "w=wa@wb & wh=w2a@w2b & wa\w1a\w2a & wb\w1b\w2b" by blast with SCASE have "e#w=(e#wa)@wb & w2=(e#w2a)@w2b & e#wa\w1a\e#w2a & wb\w1b\w2b" by (auto simp add: interleave_cons) hence "?CONS (e#w) w1a w1b w2" by blast } ultimately have "?CONS (e#w) w1a w1b w2" by blast } ultimately show "?CONS (e#w) w1a w1b w2" by blast qed qed lemma interleave_recover2: "w\w1\(w2a@w2b) \ EX wa wb w1a w1b . w=wa@wb & w1=w1a@w1b & wa\w1a\w2a & wb\w1b\w2b" proof - assume "w\w1\(w2a@w2b)" hence "w\(w2a@w2b)\w1" by auto hence "EX wa wb w1a w1b . w=wa@wb & w1=w1a@w1b & wa\w2a\w1a & wb\w2b\w1b" by (blast dest: interleave_recover1) thus ?thesis by auto qed lemmas interleave_recover = interleave_recover1 interleave_recover2 \ \Split operands according to element of result\ lemma interleave_unconc: "!! l2 w1 w2 . l1@l2 \ w1\w2 \ \ w11 w12 w21 w22 . w1=w11@w12 \ w2=w21@w22 \ l1\w11\w21 \ l2\w12\w22" proof (induct l1) case Nil hence "w1=[]@w1 & w2=[]@w2 & []\[]\[] & l2\w1\w2" by auto thus ?case by fast next case (Cons e l1') note IHP=this hence "e#(l1'@l2)\w1\w2" by auto with cons_interleave_split obtain w' where "(w1=e#w' & l1'@l2\w'\w2) | (w2=e#w' & l1'@l2\w1\w')" (is "?C1 | ?C2") by (fast) moreover { assume CASE: ?C1 moreover then obtain w11' w12' w21 w22 where "w'=w11'@w12' \ w2=w21@w22 \ l1'\w11'\w21 \ l2\w12'\w22" by (fast dest: IHP(1)) moreover hence "e#w'=(e#w11')@w12' & e#l1'\(e#w11')\w21" by (auto dest: interleave_cons) ultimately have ?case by fast } moreover { assume CASE: ?C2 moreover then obtain w11 w12 w21' w22' where "w1=w11@w12 \ w'=w21'@w22' \ l1'\w11\w21' \ l2\w12\w22'" by (fast dest: IHP(1)) moreover hence "e#w'=(e#w21')@w22' & e#l1'\w11\(e#w21')" by (auto dest: interleave_cons) ultimately have ?case by fast } ultimately show ?case by fast qed \ \Reverse direction of @{thm [source] "interleave_unconc"}\ lemma interleave_reconc: "!!w11 w21 l2 w12 w22 . \l1\w11\w21;l2\w12\w22\ \ l1@l2\(w11@w12)\(w21@w22)" proof (induct l1) case Nil thus ?case by (auto) next case (Cons e l1') note IHP=this then obtain w' where "(w11=e#w' & l1'\w'\w21) | (w21=e#w' & l1'\w11\w')" (is "?C1 | ?C2") by (fast dest: cons_interleave_split) moreover { assume CASE: ?C1 moreover with IHP have "l1'@l2\(w'@w12)\(w21@w22)" by auto ultimately have ?case by (auto dest: interleave_cons) } moreover { assume CASE: ?C2 moreover with IHP have "l1'@l2\(w11@w12)\(w'@w22)" by auto ultimately have ?case by (auto dest: interleave_cons) } ultimately show ?case by fast qed (* interleave_unconc and interleave_reconc as equivalence *) lemma interleave_unconc_eq: "!! l2 w1 w2 . l1@l2 \ w1\w2 = (\ w11 w12 w21 w22 . w1=w11@w12 \ w2=w21@w22 \ l1\w11\w21 \ l2\w12\w22)" by (fast dest: interleave_reconc interleave_unconc) end diff --git a/thys/Program-Conflict-Analysis/MainResult.thy b/thys/Program-Conflict-Analysis/MainResult.thy --- a/thys/Program-Conflict-Analysis/MainResult.thy +++ b/thys/Program-Conflict-Analysis/MainResult.thy @@ -1,48 +1,47 @@ (* Title: Conflict analysis/Main Result Author: Peter Lammich Maintainer: Peter Lammich *) section "Main Result" theory MainResult imports ConstraintSystems begin text_raw \\label{thy:MainResult}\ text \At this point everything is available to prove the main result of this project: {\em The constraint system @{term RUV_cs} precisely characterizes simultaneously reachable control nodes w.r.t. to our semantic reference point.} The ,,trusted base'' of this proof, that are all definitions a reader that trusts the Isabelle prover must additionally trust, is the following: \begin{itemize} \item The flowgraph and the assumptions made on it in the \flowgraph\- and \eflowgraph\-locales. Note that we show in Section~\ref{sec:Flowgraph:ex_flowgraph} that there is at least one non-trivial model of \eflowgraph\. \item The reference point semantics (@{term refpoint}) and the transitive closure operator (@{term trcl}). \item The definition of @{term atUV}. \item All dependencies of the above definitions in the Isabelle standard libraries. \end{itemize} \ theorem (in eflowgraph) RUV_is_sim_reach: "(\w c'. ({#[entry fg (main fg)]#},w,c')\trcl (refpoint fg) \ atUV U V c') \ (\Ml Me. (entry fg (main fg),Ml,Me)\RUV_cs fg U V)" \ \The proof uses the soundness and precision theorems wrt. to normalized paths (@{thm [source] flowgraph.RUV_sound}, @{thm [source] flowgraph.RUV_precise}) as well as the normalization result, i.e. that every reachable configuration is also reachable using a normalized path (@{thm [source] eflowgraph.normalize}) and, vice versa, that every normalized path is also a usual path (@{thm [source] ntr_is_tr}). Finally the conversion between our working semantics and the semantic reference point is exploited (@{thm [source] flowgraph.refpoint_eq}).\ (is "?lhs \ ?rhs") proof assume ?lhs then obtain w c' where C: "({#[entry fg (main fg)]#}, w, c') \ trcl (tr fg)" "atUV U V c'" by (auto simp add: refpoint_eq) from normalize[OF C(1), of "main fg", simplified] obtain ww where "({#[entry fg (main fg)]#}, ww, c') \ trcl (ntr fg)" by blast from ntrs.gtr2gtrp[where c="{#}", simplified, OF this] obtain s' ce' wwl where 1: "c'=add_mset s' ce'" "ww = map le_rem_s wwl" "(([entry fg (main fg)], {#}), wwl, s', ce') \ trcl (ntrp fg)" by blast with C(2) have 2: "atUV U V ({#s'#}+ce')" by auto from RUV_sound[OF 1(3) 2] show ?rhs by blast next assume ?rhs then obtain Ml Me where C: "(entry fg (main fg), Ml, Me) \ RUV_cs fg U V" by blast from RUV_precise[OF C] obtain wwl s' c' where P: "(([entry fg (main fg)], {#}), wwl, s', c') \ trcl (ntrp fg)" "atUV U V ({#s'#} + c')" by blast from gtrp2gtr[OF P(1)] have "({# [entry fg (main fg)] #}, map le_rem_s wwl, {#s'#}+c') \ trcl (ntr fg)" by (auto) from ntr_is_tr[OF this] P(2) have "\w c'. ({#[entry fg (main fg)]#}, w, c') \ trcl (tr fg) \ atUV U V c'" by blast thus ?lhs by (simp add: refpoint_eq) qed - end